added pip to a new repository
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Thu, 06 Dec 2012 16:32:03 +0000
changeset 373 0679a84b11ad
parent 372 2c56b20032a7
child 374 01d223421ba0
added pip to a new repository
prio/Attic/Ext.thy
prio/Attic/ExtGG_1.thy
prio/Attic/ExtS.thy
prio/Attic/ExtSG.thy
prio/Attic/Happen_within.thy
prio/Attic/Lsp.thy
prio/Attic/Prio.thy
prio/CpsG.thy
prio/ExtGG.thy
prio/IsaMakefile
prio/Moment.thy
prio/Paper/Paper.thy
prio/Paper/PrioGDef.tex
prio/Paper/ROOT.ML
prio/Paper/document/llncs.cls
prio/Paper/document/mathpartir.sty
prio/Paper/document/root.bib
prio/Paper/document/root.tex
prio/Paper/tt.thy
prio/Precedence_ord.thy
prio/PrioG.thy
prio/PrioGDef.thy
prio/README
prio/ROOT.ML
prio/Slides/ROOT1.ML
prio/Slides/Slides1.thy
prio/Slides/document/beamerthemeplaincu.sty
prio/Slides/document/mathpartir.sty
prio/Slides/document/root.beamer.tex
prio/Slides/document/root.notes.tex
prio/Slides/document/root.tex
prio/Slides/slides.pdf
prio/document/beamerthemeplaincu.sty
prio/document/llncs.cls
prio/document/root.bib
prio/document/root.tex
prio/paper.pdf
--- a/prio/Attic/Ext.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1057 +0,0 @@
-theory Ext
-imports Prio
-begin
-
-locale highest_create =
-  fixes s' th prio fixes s 
-  defines s_def : "s \<equiv> (Create th prio#s')"
-  assumes vt_s: "vt step s"
-  and highest: "cp s th = Max ((cp s)`threads s)"
-
-context highest_create
-begin
-
-lemma threads_s: "threads s = threads s' \<union> {th}"
-  by (unfold s_def, simp)
-
-lemma vt_s': "vt step s'"
-  by (insert vt_s, unfold s_def, drule_tac step_back_vt, simp)
-
-lemma step_create: "step s' (Create th prio)"
-  by (insert vt_s, unfold s_def, drule_tac step_back_step, simp)
-
-lemma step_create_elim: 
-  "\<lbrakk>\<And>max_prio. \<lbrakk>prio \<le> max_prio; th \<notin> threads s'\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-  by (insert step_create, ind_cases "step s' (Create th prio)", auto)
-
-lemma eq_cp_s: 
-  assumes th'_in: "th' \<in> threads s'"
-  shows "cp s th' = cp s' th'"
-proof(unfold cp_eq_cpreced cpreced_def cs_dependents_def s_def 
-    eq_depend depend_create_unchanged)
-  show "Max ((\<lambda>tha. preced tha (Create th prio # s')) `
-         ({th'} \<union> {th'a. (Th th'a, Th th') \<in> (depend s')\<^sup>+})) =
-        Max ((\<lambda>th. preced th s') ` ({th'} \<union> {th'a. (Th th'a, Th th') \<in> (depend s')\<^sup>+}))"
-    (is "Max (?f ` ?A) = Max (?g ` ?A)")
-  proof -
-    have "?f ` ?A = ?g ` ?A"
-    proof(rule f_image_eq)
-      fix a
-      assume a_in: "a \<in> ?A"
-      thus "?f a = ?g a" 
-      proof -
-        from a_in
-        have "a = th' \<or> (Th a, Th th') \<in> (depend s')\<^sup>+" by auto 
-        hence "a \<noteq> th"
-        proof
-          assume "a = th'"
-          moreover have "th' \<noteq> th"
-          proof(rule step_create_elim)
-            assume th_not_in: "th \<notin> threads s'" with th'_in
-            show ?thesis by auto
-          qed
-          ultimately show ?thesis by auto
-        next
-          assume "(Th a, Th th') \<in> (depend s')\<^sup>+"
-          hence "Th a \<in> Domain \<dots>"
-            by (auto simp:Domain_def)
-          hence "Th a \<in> Domain (depend s')"
-            by (simp add:trancl_domain)
-          from dm_depend_threads[OF vt_s' this]
-          have h: "a \<in> threads s'" .
-          show ?thesis
-          proof(rule step_create_elim)
-            assume "th \<notin> threads s'" with h
-            show ?thesis by auto
-          qed
-        qed
-        thus ?thesis 
-          by (unfold preced_def, auto)
-      qed
-    qed
-    thus ?thesis by auto
-  qed
-qed
-
-lemma same_depend: "depend s = depend s'"
-  by (insert depend_create_unchanged, unfold s_def, simp)
-
-lemma same_dependents:
-  "dependents (wq s) th = dependents (wq s') th"
-  apply (unfold cs_dependents_def)
-  by (unfold eq_depend same_depend, simp)
-
-lemma nil_dependents_s': "dependents (wq s') th = {}"
-proof -
-  { assume ne: "dependents (wq s') th \<noteq> {}"
-    then obtain th' where "th' \<in>  dependents (wq s') th"
-      by (unfold cs_dependents_def, auto)
-    hence "(Th th', Th th) \<in> (depend (wq s'))^+"
-      by (unfold cs_dependents_def, auto)
-    hence "(Th th', Th th) \<in> (depend s')^+"
-      by (simp add:eq_depend)
-    hence "Th th \<in> Range ((depend s')^+)" by (auto simp:Range_def Domain_def)
-    hence "Th th \<in> Range (depend s')" by (simp add:trancl_range)
-    from range_in [OF vt_s' this]
-    have h: "th \<in> threads s'" .
-    have "False"
-    proof(rule step_create_elim)
-      assume "th \<notin> threads s'" with h show ?thesis by auto
-    qed
-  } thus ?thesis by auto
-qed
-
-lemma nil_dependents: "dependents (wq s) th = {}"
-proof -
-  have "wq s' = wq s"
-    by (unfold wq_def s_def, auto simp:Let_def)
-  with nil_dependents_s' show ?thesis by auto
-qed
-
-lemma eq_cp_s_th: "cp s th = preced th s"
-  by (unfold cp_eq_cpreced cpreced_def nil_dependents, auto)
-
-lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold max_cp_eq[OF vt_s], unfold highest, simp)
-
-lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
-
-lemma is_ready: "th \<in> readys s"
-proof -
-  { assume "th \<notin> readys s"
-    with threads_s obtain cs where 
-      "waiting s th cs"
-      by (unfold readys_def, auto)
-    hence "(Th th, Cs cs) \<in> depend s"
-      by (unfold s_depend_def, unfold eq_waiting, simp)
-    hence "Th th \<in> Domain (depend s')"
-      by (unfold same_depend, auto simp:Domain_def)
-    from dm_depend_threads [OF vt_s' this] 
-    have h: "th \<in> threads s'" .
-    have "False"
-    proof (rule_tac step_create_elim)
-      assume "th \<notin> threads s'" with h show ?thesis by simp
-    qed
-  } thus ?thesis by auto
-qed
-
-lemma is_runing: "th \<in> runing s"
-proof -
-  have "Max (cp s ` threads s) = Max (cp s ` readys s)"
-  proof -
-    have " Max (cp s ` readys s) = cp s th"
-    proof(rule Max_eqI)
-      from finite_threads[OF vt_s] readys_threads finite_subset
-      have "finite (readys s)" by blast
-      thus "finite (cp s ` readys s)" by auto
-    next
-      from is_ready show "cp s th \<in> cp s ` readys s" by auto
-    next
-      fix y
-      assume h: "y \<in> cp s ` readys s"
-      have "y \<le> Max (cp s ` readys s)"
-      proof(rule Max_ge [OF _ h])
-        from finite_threads[OF vt_s] readys_threads finite_subset
-        have "finite (readys s)" by blast
-        thus "finite (cp s ` readys s)" by auto
-      qed
-      moreover have "\<dots> \<le> Max (cp s ` threads s)"
-      proof(rule Max_mono)
-        from readys_threads 
-        show "cp s ` readys s \<subseteq> cp s ` threads s" by auto
-      next
-        from is_ready show "cp s ` readys s \<noteq> {}" by auto
-      next
-        from finite_threads [OF vt_s]
-        show "finite (cp s ` threads s)" by auto
-      qed
-      moreover note highest
-      ultimately show "y \<le> cp s th" by auto
-    qed
-    with highest show ?thesis by auto
-  qed
-  thus ?thesis
-    by (unfold runing_def, insert highest is_ready, auto)
-qed
-
-end
-
-locale extend_highest = highest_create + 
-  fixes t 
-  assumes vt_t: "vt step (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"
-
-lemma step_back_vt_app: 
-  assumes vt_ts: "vt cs (t@s)" 
-  shows "vt cs 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 cs (t @ s) \<Longrightarrow> vt cs s"
-      and vt_et: "vt cs ((e # t) @ s)"
-    show ?case
-    proof(rule ih)
-      show "vt cs (t @ s)"
-      proof(rule step_back_vt)
-        from vt_et show "vt cs (e # t @ s)" by simp
-      qed
-    qed
-  qed
-qed
-
-context extend_highest
-begin
-
-lemma red_moment:
-  "extend_highest s' th prio (moment i t)"
-  apply (insert extend_highest_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
-  apply (unfold extend_highest_def extend_highest_axioms_def, clarsimp)
-  by (unfold highest_create_def, auto dest:step_back_vt_app)
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes 
-    h0: "R []"
-  and h2: "\<And> e t. \<lbrakk>vt step (t@s); step (t@s) e; 
-                    extend_highest s' th prio t; 
-                    extend_highest s' th prio (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
-  shows "R t"
-proof -
-  from vt_t extend_highest_axioms show ?thesis
-  proof(induct t)
-    from h0 show "R []" .
-  next
-    case (Cons e t')
-    assume ih: "\<lbrakk>vt step (t' @ s); extend_highest s' th prio t'\<rbrakk> \<Longrightarrow> R t'"
-      and vt_e: "vt step ((e # t') @ s)"
-      and et: "extend_highest s' th prio (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 step (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 s' th prio t'"
-          by (unfold extend_highest_def extend_highest_axioms_def, auto dest:step_back_vt)
-      next
-        from vt_ts show "vt step (t' @ s)" .
-      qed
-    next
-      from et show "extend_highest s' th prio (e # t')" .
-    next
-      from et show ext': "extend_highest s' th prio t'"
-          by (unfold extend_highest_def extend_highest_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 "th \<in> threads ([] @ s) \<and> preced th ([] @ s) = preced th s"
-      by auto
-  next
-    case (Cons e t)
-    show ?case
-    proof(cases e)
-      case (Create thread prio)
-      assume eq_e: " e = Create thread prio"
-      show ?thesis
-      proof -
-        from Cons and eq_e have "step (t@s) (Create thread prio)" by auto
-        hence "th \<noteq> thread"
-        proof(cases)
-          assume "thread \<notin> threads (t @ s)"
-          with Cons show ?thesis by auto
-        qed
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    next
-      case (Exit thread)
-      assume eq_e: "e = Exit thread"
-      from Cons have "extend_highest s' th prio (e # t)" by auto
-      from extend_highest.exit_diff [OF this] and eq_e
-      have neq_th: "thread \<noteq> th" by auto
-      with Cons
-      show ?thesis
-        by (unfold eq_e, auto simp:preced_def)
-    next
-      case (P thread cs)
-      assume eq_e: "e = P thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (V thread cs)
-      assume eq_e: "e = V thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (Set thread prio')
-      assume eq_e: " e = Set thread prio'"
-      show ?thesis
-      proof -
-        from Cons have "extend_highest s' th prio (e # t)" by auto
-        from extend_highest.set_diff_low[OF this] and eq_e
-        have "th \<noteq> thread" by auto
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    qed
-  qed
-qed
-
-lemma max_kept: "Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s"
-proof(induct rule:ind)
-  case Nil
-  from highest_preced_thread
-  show "Max ((\<lambda>th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s"
-    by simp
-next
-  case (Cons e t)
-  show ?case
-  proof(cases e)
-    case (Create thread prio')
-    assume eq_e: " e = Create thread prio'"
-    from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto
-    hence neq_thread: "thread \<noteq> th"
-    proof(cases)
-      assume "thread \<notin> threads (t @ s)"
-      moreover have "th \<in> threads (t@s)"
-      proof -
-        from Cons have "extend_highest s' th prio t" by auto
-        from extend_highest.th_kept[OF this] show ?thesis by (simp add:s_def)
-      qed
-      ultimately show ?thesis by auto
-    qed
-    from Cons have "extend_highest s' th prio t" by auto
-    from extend_highest.th_kept[OF this]
-    have h': " th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" 
-      by (auto simp:s_def)
-    from stp
-    have thread_ts: "thread \<notin> threads (t @ s)"
-      by (cases, auto)
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))"
-        by (unfold eq_e, simp)
-      moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
-      proof(rule Max_insert)
-        from Cons have "vt step (t @ s)" by auto
-        from finite_threads[OF this]
-        show "finite (?f ` (threads (t@s)))" by simp
-      next
-        from h' show "(?f ` (threads (t@s))) \<noteq> {}" by auto
-      qed
-      moreover have "(Max (?f ` (threads (t@s)))) = ?t"
-      proof -
-        have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
-          (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B")
-        proof -
-          { fix th' 
-            assume "th' \<in> ?B"
-            with thread_ts eq_e
-            have "?f1 th' = ?f2 th'" by (auto simp:preced_def)
-          } thus ?thesis 
-            apply (auto simp:Image_def)
-          proof -
-            fix th'
-            assume h: "\<And>th'. th' \<in> threads (t @ s) \<Longrightarrow> 
-              preced th' (e # t @ s) = preced th' (t @ s)"
-              and h1: "th' \<in> threads (t @ s)"
-            show "preced th' (t @ s) \<in> (\<lambda>th'. preced th' (e # t @ s)) ` threads (t @ s)"
-            proof -
-              from h1 have "?f1 th' \<in> ?f1 ` ?B" by auto
-              moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp
-              ultimately show ?thesis by simp
-            qed
-          qed
-        qed
-        with Cons show ?thesis by auto
-      qed
-      moreover have "?f thread < ?t"
-      proof -
-        from Cons have " extend_highest s' th prio (e # t)" by auto
-        from extend_highest.create_low[OF this] and eq_e
-        have "prio' \<le> prio" by auto
-        thus ?thesis
-        by (unfold eq_e, auto simp:preced_def s_def precedence_less_def)
-    qed
-    ultimately show ?thesis by (auto simp:max_def)
-  qed
-next
-    case (Exit thread)
-    assume eq_e: "e = Exit thread"
-    from Cons have vt_e: "vt step (e#(t @ s))" by auto
-    from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto
-    from stp have thread_ts: "thread \<in> threads (t @ s)"
-      by(cases, unfold runing_def readys_def, auto)
-    from Cons have "extend_highest s' th prio (e # t)" by auto
-    from extend_highest.exit_diff[OF this] and eq_e
-    have neq_thread: "thread \<noteq> th" by auto
-    from Cons have "extend_highest s' th prio t" by auto
-    from extend_highest.th_kept[OF this, folded s_def]
-    have h': "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "threads (t@s) = insert thread ?A"
-        by (insert stp thread_ts, unfold eq_e, auto)
-      hence "Max (?f ` (threads (t@s))) = Max (?f ` \<dots>)" by simp
-      also from this have "\<dots> = Max (insert (?f thread) (?f ` ?A))" by simp
-      also have "\<dots> = max (?f thread) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        from finite_threads [OF vt_e]
-        show "finite (?f ` ?A)" by simp
-      next
-        from Cons have "extend_highest s' th prio (e # t)" by auto
-        from extend_highest.th_kept[OF this]
-        show "?f ` ?A \<noteq> {}" by  (auto simp:s_def)
-      qed
-      finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" .
-      moreover have "Max (?f ` (threads (t@s))) = ?t"
-      proof -
-        from Cons show ?thesis
-          by (unfold eq_e, auto simp:preced_def)
-      qed
-      ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp
-      moreover have "?f thread < ?t"
-      proof(unfold eq_e, simp add:preced_def, fold preced_def)
-        show "preced thread (t @ s) < ?t"
-        proof -
-          have "preced thread (t @ s) \<le> ?t" 
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" 
-              (is "?t = Max (?g ` ?B)") by simp
-            moreover have "?g thread \<le> \<dots>"
-            proof(rule Max_ge)
-              have "vt step (t@s)" by fact
-              from finite_threads [OF this]
-              show "finite (?g ` ?B)" by simp
-            next
-              from thread_ts
-              show "?g thread \<in> (?g ` ?B)" by auto
-            qed
-            ultimately show ?thesis by auto
-          qed
-          moreover have "preced thread (t @ s) \<noteq> ?t"
-          proof
-            assume "preced thread (t @ s) = preced th s"
-            with h' have "preced thread (t @ s) = preced th (t@s)" by simp
-            from preced_unique [OF this] have "thread = th"
-            proof
-              from h' show "th \<in> threads (t @ s)" by simp
-            next
-              from thread_ts show "thread \<in> threads (t @ s)" .
-            qed(simp)
-            with neq_thread show "False" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      qed
-      ultimately show ?thesis 
-        by (auto simp:max_def split:if_splits)
-    qed
-  next
-    case (P thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (V thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (Set thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      let ?B = "threads (t@s)"
-      from Cons have "extend_highest s' th prio (e # t)" by auto
-      from extend_highest.set_diff_low[OF this] and Set
-      have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
-      from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
-      also have "\<dots> = ?t"
-      proof(rule Max_eqI)
-        fix y
-        assume y_in: "y \<in> ?f ` ?B"
-        then obtain th1 where 
-          th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
-        show "y \<le> ?t"
-        proof(cases "th1 = thread")
-          case True
-          with neq_thread le_p eq_y s_def Set
-          show ?thesis
-            by (auto simp:preced_def precedence_le_def)
-        next
-          case False
-          with Set eq_y
-          have "y  = preced th1 (t@s)"
-            by (simp add:preced_def)
-          moreover have "\<dots> \<le> ?t"
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
-              by auto
-            moreover have "preced th1 (t@s) \<le> \<dots>"
-            proof(rule Max_ge)
-              from th1_in 
-              show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
-                by simp
-            next
-              show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-              proof -
-                from Cons have "vt step (t @ s)" by auto
-                from finite_threads[OF this] show ?thesis by auto
-              qed
-            qed
-            ultimately show ?thesis by auto
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from Cons and finite_threads
-        show "finite (?f ` ?B)" by auto
-      next
-        from Cons have "extend_highest s' th prio t" by auto
-        from extend_highest.th_kept [OF this, folded s_def]
-        have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-        show "?t \<in> (?f ` ?B)" 
-        proof -
-          from neq_thread Set h
-          have "?t = ?f th" by (auto simp:preced_def)
-          with h show ?thesis by auto
-        qed
-      qed
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
-  by (insert th_kept max_kept, auto)
-
-lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
-  (is "?L = ?R")
-proof -
-  have "?L = cpreced (t@s) (wq (t@s)) th" 
-    by (unfold cp_eq_cpreced, simp)
-  also have "\<dots> = ?R"
-  proof(unfold cpreced_def)
-    show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependents (wq (t @ s)) th)) =
-          Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-      (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
-    proof(cases "?A = {}")
-      case False
-      have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
-      moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        show "finite (?f ` ?A)"
-        proof -
-          from dependents_threads[OF vt_t]
-          have "?A \<subseteq> threads (t@s)" .
-          moreover from finite_threads[OF vt_t] have "finite \<dots>" .
-          ultimately show ?thesis 
-            by (auto simp:finite_subset)
-        qed
-      next
-        from False show "(?f ` ?A) \<noteq> {}" by simp
-      qed
-      moreover have "\<dots> = Max (?f ` ?B)"
-      proof -
-        from max_preced have "?f th = Max (?f ` ?B)" .
-        moreover have "Max (?f ` ?A) \<le> \<dots>" 
-        proof(rule Max_mono)
-          from False show "(?f ` ?A) \<noteq> {}" by simp
-        next
-          show "?f ` ?A \<subseteq> ?f ` ?B" 
-          proof -
-            have "?A \<subseteq> ?B" by (rule dependents_threads[OF vt_t])
-            thus ?thesis by auto
-          qed
-        next
-          from finite_threads[OF vt_t] 
-          show "finite (?f ` ?B)" by simp
-        qed
-        ultimately show ?thesis
-          by (auto simp:max_def)
-      qed
-      ultimately show ?thesis by auto
-    next
-      case True
-      with max_preced show ?thesis by auto
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
-  by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)
-
-lemma th_cp_preced: "cp (t@s) th = preced th s"
-  by (fold max_kept, unfold th_cp_max_preced, simp)
-
-lemma preced_less':
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  shows "preced th' s < preced th s"
-proof -
-  have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
-  proof(rule Max_ge)
-    from finite_threads [OF vt_s]
-    show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
-  next
-    from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
-      by simp
-  qed
-  moreover have "preced th' s \<noteq> preced th s"
-  proof
-    assume "preced th' s = preced th s"
-    from preced_unique[OF this th'_in] neq_th' is_ready
-    show "False" by  (auto simp:readys_def)
-  qed
-  ultimately show ?thesis using highest_preced_thread
-    by auto
-qed
-
-lemma pv_blocked:
-  fixes th'
-  assumes th'_in: "th' \<in> threads (t@s)"
-  and neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
-  shows "th' \<notin> runing (t@s)"
-proof
-  assume "th' \<in> runing (t@s)"
-  hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
-    by (auto simp:runing_def)
-  with max_cp_readys_threads [OF vt_t]
-  have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
-    by auto
-  moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
-  ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
-  moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
-    by simp
-  finally have h: "cp (t @ s) th' = preced th (t @ s)" .
-  show False
-  proof -
-    have "dependents (wq (t @ s)) th' = {}" 
-      by (rule count_eq_dependents [OF vt_t eq_pv])
-    moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
-    proof
-      assume "preced th' (t @ s) = preced th (t @ s)"
-      hence "th' = th"
-      proof(rule preced_unique)
-        from th_kept show "th \<in> threads (t @ s)" by simp
-      next
-        from th'_in show "th' \<in> threads (t @ s)" by simp
-      qed
-      with assms show False by simp
-    qed
-    ultimately show ?thesis
-      by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
-  qed
-qed
-
-lemma runing_precond_pre:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<in> threads (t@s) \<and>
-         cntP (t@s) th' = cntV (t@s) th'"
-proof -
-  show ?thesis
-  proof(induct rule:ind)
-    case (Cons e t)
-    from Cons
-    have in_thread: "th' \<in> threads (t @ s)"
-      and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    have "extend_highest s' th prio t" by fact
-    from extend_highest.pv_blocked 
-    [OF this, folded s_def, OF in_thread neq_th' not_holding]
-    have not_runing: "th' \<notin> runing (t @ s)" .
-    show ?case
-    proof(cases e)
-      case (V thread cs)
-      from Cons and V have vt_v: "vt step (V thread cs#(t@s))" by auto
-
-      show ?thesis
-      proof -
-        from Cons and V have "step (t@s) (V thread cs)" by auto
-        hence neq_th': "thread \<noteq> th'"
-        proof(cases)
-          assume "thread \<in> runing (t@s)"
-          moreover have "th' \<notin> runing (t@s)" by fact
-          ultimately show ?thesis by auto
-        qed
-        with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
-          by (unfold V, simp add:cntP_def cntV_def count_def)
-        moreover from in_thread
-        have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
-        ultimately show ?thesis by auto
-      qed
-    next
-      case (P thread cs)
-      from Cons and P have "step (t@s) (P thread cs)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t@s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Create thread prio')
-      from Cons and Create have "step (t@s) (Create thread prio')" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<notin> threads (t @ s)"
-        moreover have "th' \<in> threads (t@s)" by fact
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Create 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Create 
-      have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
-      ultimately show ?thesis by auto
-    next
-      case (Exit thread)
-      from Cons and Exit have "step (t@s) (Exit thread)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t @ s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Exit 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Exit and neq_th' 
-      have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio')
-      with Cons
-      show ?thesis 
-        by (auto simp:cntP_def cntV_def count_def)
-    qed
-  next
-    case Nil
-    with assms
-    show ?case by auto
-  qed
-qed
-
-(*
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<notin> runing (t@s)"
-proof -
-  from runing_precond_pre[OF th'_in eq_pv neq_th']
-  have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-  from pv_blocked[OF h1 neq_th' h2] 
-  show ?thesis .
-qed
-*)
-
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  and is_runing: "th' \<in> runing (t@s)"
-  shows "cntP s th' > cntV s th'"
-proof -
-  have "cntP s th' \<noteq> cntV s th'"
-  proof
-    assume eq_pv: "cntP s th' = cntV s th'"
-    from runing_precond_pre[OF th'_in eq_pv neq_th']
-    have h1: "th' \<in> threads (t @ s)"  
-      and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    from pv_blocked[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
-    with is_runing show "False" by simp
-  qed
-  moreover from cnp_cnv_cncs[OF vt_s, of th'] 
-  have "cntV s th' \<le> cntP s th'" by auto
-  ultimately show ?thesis by auto
-qed
-
-lemma moment_blocked_pre:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
-         th' \<in> threads ((moment (i+j) t)@s)"
-proof(induct j)
-  case (Suc k)
-  show ?case
-  proof -
-    { assume True: "Suc (i+k) \<le> length t"
-      from moment_head [OF this] 
-      obtain e where
-        eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
-        by blast
-      from red_moment[of "Suc(i+k)"]
-      and eq_me have "extend_highest s' th prio (e # moment (i + k) t)" by simp
-      hence vt_e: "vt step (e#(moment (i + k) t)@s)"
-        by (unfold extend_highest_def extend_highest_axioms_def 
-          highest_create_def s_def, auto)
-      have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
-      proof(unfold s_def)
-        show "th' \<notin> runing (moment (i + k) t @ Create th prio # s')"
-        proof(rule extend_highest.pv_blocked)
-          from Suc show "th' \<in> threads (moment (i + k) t @ Create th prio # s')"
-            by (simp add:s_def)
-        next
-          from neq_th' show "th' \<noteq> th" .
-        next
-          from red_moment show "extend_highest s' th prio (moment (i + k) t)" .
-        next
-          from Suc show "cntP (moment (i + k) t @ Create th prio # s') th' =
-            cntV (moment (i + k) t @ Create th prio # s') th'"
-            by (auto simp:s_def)
-        qed
-      qed
-      from step_back_step[OF vt_e]
-      have "step ((moment (i + k) t)@s) e" .
-      hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
-        th' \<in> threads (e#(moment (i + k) t)@s)
-        "
-      proof(cases)
-        case (thread_create thread prio)
-        with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_exit thread)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_P thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_V thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_set thread prio')
-        with Suc show ?thesis
-          by (auto simp:cntP_def cntV_def count_def)
-      qed
-      with eq_me have ?thesis using eq_me by auto 
-    } note h = this
-    show ?thesis
-    proof(cases "Suc (i+k) \<le> length t")
-      case True
-      from h [OF this] show ?thesis .
-    next
-      case False
-      with moment_ge
-      have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
-      with Suc show ?thesis by auto
-    qed
-  qed
-next
-  case 0
-  from assms show ?case by auto
-qed
-
-lemma moment_blocked:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  and le_ij: "i \<le> j"
-  shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
-         th' \<in> threads ((moment j t)@s) \<and>
-         th' \<notin> runing ((moment j t)@s)"
-proof -
-  from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
-  have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
-    and h2: "th' \<in> threads ((moment j t)@s)" by auto
-  with extend_highest.pv_blocked [OF  red_moment [of j], folded s_def, OF h2 neq_th' h1]
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_1:
-  assumes neq_th': "th' \<noteq> th"
-  and runing': "th' \<in> runing (t@s)"
-  shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
-proof(cases "th' \<in> threads s")
-  case True
-  with runing_precond [OF this neq_th' runing'] show ?thesis by simp
-next
-  case False
-  let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
-  let ?q = "moment 0 t"
-  from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
-  from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
-  from p_split_gen [of ?Q, OF this not_thread]
-  obtain i where lt_its: "i < length t"
-    and le_i: "0 \<le> i"
-    and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
-    and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
-  from lt_its have "Suc i \<le> length t" by auto
-  from moment_head[OF this] obtain e where 
-   eq_me: "moment (Suc i) t = e # moment i t" by blast
-  from red_moment[of "Suc i"] and eq_me
-  have "extend_highest s' th prio (e # moment i t)" by simp
-  hence vt_e: "vt step (e#(moment i t)@s)"
-    by (unfold extend_highest_def extend_highest_axioms_def 
-      highest_create_def s_def, auto)
-  from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
-  from post[rule_format, of "Suc i"] and eq_me 
-  have not_in': "th' \<in> threads (e # moment i t@s)" by auto
-  from create_pre[OF stp_i pre this] 
-  obtain prio where eq_e: "e = Create th' prio" .
-  have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
-  proof(rule cnp_cnv_eq)
-    from step_back_vt [OF vt_e] 
-    show "vt step (moment i t @ s)" .
-  next
-    from eq_e and stp_i 
-    have "step (moment i t @ s) (Create th' prio)" by simp
-    thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
-  qed
-  with eq_e
-  have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
-    by (simp add:cntP_def cntV_def count_def)
-  with eq_me[symmetric]
-  have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
-    by simp
-  from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
-  with eq_me [symmetric]
-  have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
-  from moment_blocked [OF neq_th' h2 h1, of "length t"] and lt_its
-  and moment_ge
-  have "th' \<notin> runing (t @ s)" by auto
-  with runing'
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_2:
-  assumes runing': "th' \<in> runing (t@s)"
-  shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
-proof -
-  from runing_inversion_1[OF _ runing']
-  show ?thesis by auto
-qed
-
-lemma live: "runing (t@s) \<noteq> {}"
-proof(cases "th \<in> runing (t@s)")
-  case True thus ?thesis by auto
-next
-  case False
-  then have not_ready: "th \<notin> readys (t@s)"
-    apply (unfold runing_def, 
-            insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
-    by auto
-  from th_kept have "th \<in> threads (t@s)" by auto
-  from th_chain_to_ready[OF vt_t this] and not_ready
-  obtain th' where th'_in: "th' \<in> readys (t@s)"
-    and dp: "(Th th, Th th') \<in> (depend (t @ s))\<^sup>+" by auto
-  have "th' \<in> runing (t@s)"
-  proof -
-    have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
-    proof -
-      have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')) = 
-               preced th (t@s)"
-      proof(rule Max_eqI)
-        fix y
-        assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-        then obtain th1 where
-          h1: "th1 = th' \<or> th1 \<in>  dependents (wq (t @ s)) th'"
-          and eq_y: "y = preced th1 (t@s)" by auto
-        show "y \<le> preced th (t @ s)"
-        proof -
-          from max_preced
-          have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
-          moreover have "y \<le> \<dots>"
-          proof(rule Max_ge)
-            from h1
-            have "th1 \<in> threads (t@s)"
-            proof
-              assume "th1 = th'"
-              with th'_in show ?thesis by (simp add:readys_def)
-            next
-              assume "th1 \<in> dependents (wq (t @ s)) th'"
-              with dependents_threads [OF vt_t]
-              show "th1 \<in> threads (t @ s)" by auto
-            qed
-            with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
-          next
-            from finite_threads[OF vt_t]
-            show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th']
-        show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th'))"
-          by (auto intro:finite_subset)
-      next
-        from dp
-        have "th \<in> dependents (wq (t @ s)) th'" 
-          by (unfold cs_dependents_def, auto simp:eq_depend)
-        thus "preced th (t @ s) \<in> 
-                (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-          by auto
-      qed
-      moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
-      proof -
-        from max_preced and max_cp_eq[OF vt_t, symmetric]
-        have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
-        with max_cp_readys_threads[OF vt_t] show ?thesis by simp
-      qed
-      ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
-    qed
-    with th'_in show ?thesis by (auto simp:runing_def)
-  qed
-  thus ?thesis by auto
-qed
-
-end
-
-end
-
--- a/prio/Attic/ExtGG_1.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,973 +0,0 @@
-theory ExtGG
-imports PrioG
-begin
-
-lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
-  apply (induct s, simp)
-proof -
-  fix a s
-  assume ih: "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
-    and eq_as: "a # s \<noteq> []"
-  show "birthtime th (a # s) < length (a # s)"
-  proof(cases "s \<noteq> []")
-    case False
-    from False show ?thesis
-      by (cases a, auto simp:birthtime.simps)
-  next
-    case True
-    from ih [OF True] show ?thesis
-      by (cases a, auto simp:birthtime.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)
-
-locale highest_gen =
-  fixes s' th s e' prio tm
-  defines s_def : "s \<equiv> (e'#s')"
-  assumes vt_s: "vt step s"
-  and threads_s: "th \<in> threads s"
-  and highest: "preced th s = Max ((cp s)`threads s)"
-  and nh: "preced th s' \<noteq> Max ((cp s)`threads s')"
-  and preced_th: "preced th s = Prc prio tm"
-
-context highest_gen
-begin
-
-lemma lt_tm: "tm < length s"
-  by (insert preced_tm_lt[OF threads_s preced_th], simp)
-
-lemma vt_s': "vt step s'"
-  by (insert vt_s, unfold s_def, drule_tac step_back_vt, simp)
-
-lemma eq_cp_s_th: "cp s th = preced th s"
-proof -
-  from highest and max_cp_eq[OF vt_s]
-  have is_max: "preced th s = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-  have sbs: "({th} \<union> dependents (wq s) th) \<subseteq> threads s"
-  proof -
-    from threads_s and dependents_threads[OF vt_s, of th]
-    show ?thesis by auto
-  qed
-  show ?thesis
-  proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-    show "preced th s \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)" by simp
-  next
-    fix y 
-    assume "y \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)"
-    then obtain th1 where th1_in: "th1 \<in> ({th} \<union> dependents (wq s) th)"
-      and eq_y: "y = preced th1 s" by auto
-    show "y \<le> preced th s"
-    proof(unfold is_max, rule Max_ge)
-      from finite_threads[OF vt_s] 
-      show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-    next
-      from sbs th1_in and eq_y 
-      show "y \<in> (\<lambda>th. preced th s) ` threads s" by auto
-    qed
-  next
-    from sbs and finite_threads[OF vt_s]
-    show "finite ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))"
-      by (auto intro:finite_subset)
-  qed
-qed
-
-lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp)
-
-lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
-
-lemma highest': "cp s th = Max (cp s ` threads s)"
-proof -
-  from highest_cp_preced max_cp_eq[OF vt_s, symmetric]
-  show ?thesis by simp
-qed
-
-end
-
-locale extend_highest_gen = highest_gen + 
-  fixes t 
-  assumes vt_t: "vt step (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"
-
-lemma step_back_vt_app: 
-  assumes vt_ts: "vt cs (t@s)" 
-  shows "vt cs 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 cs (t @ s) \<Longrightarrow> vt cs s"
-      and vt_et: "vt cs ((e # t) @ s)"
-    show ?case
-    proof(rule ih)
-      show "vt cs (t @ s)"
-      proof(rule step_back_vt)
-        from vt_et show "vt cs (e # t @ s)" by simp
-      qed
-    qed
-  qed
-qed
-
-context extend_highest_gen
-begin
-
-lemma red_moment:
-  "extend_highest_gen s' th e' 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)
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes 
-    h0: "R []"
-  and h2: "\<And> e t. \<lbrakk>vt step (t@s); step (t@s) e; 
-                    extend_highest_gen s' th e' prio tm t; 
-                    extend_highest_gen s' th e' 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 step (t' @ s); extend_highest_gen s' th e' prio tm t'\<rbrakk> \<Longrightarrow> R t'"
-      and vt_e: "vt step ((e # t') @ s)"
-      and et: "extend_highest_gen s' th e' 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 step (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 e' 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 step (t' @ s)" .
-      qed
-    next
-      from et show "extend_highest_gen s' th e' prio tm (e # t')" .
-    next
-      from et show ext': "extend_highest_gen s' th e' 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 "th \<in> threads ([] @ s) \<and> preced th ([] @ s) = preced th s"
-      by auto
-  next
-    case (Cons e t)
-    show ?case
-    proof(cases e)
-      case (Create thread prio)
-      assume eq_e: " e = Create thread prio"
-      show ?thesis
-      proof -
-        from Cons and eq_e have "step (t@s) (Create thread prio)" by auto
-        hence "th \<noteq> thread"
-        proof(cases)
-          assume "thread \<notin> threads (t @ s)"
-          with Cons show ?thesis by auto
-        qed
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    next
-      case (Exit thread)
-      assume eq_e: "e = Exit thread"
-      from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
-      from extend_highest_gen.exit_diff [OF this] and eq_e
-      have neq_th: "thread \<noteq> th" by auto
-      with Cons
-      show ?thesis
-        by (unfold eq_e, auto simp:preced_def)
-    next
-      case (P thread cs)
-      assume eq_e: "e = P thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (V thread cs)
-      assume eq_e: "e = V thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (Set thread prio')
-      assume eq_e: " e = Set thread prio'"
-      show ?thesis
-      proof -
-        from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
-        from extend_highest_gen.set_diff_low[OF this] and eq_e
-        have "th \<noteq> thread" by auto
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    qed
-  qed
-qed
-
-lemma max_kept: "Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s"
-proof(induct rule:ind)
-  case Nil
-  from highest_preced_thread
-  show "Max ((\<lambda>th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s"
-    by simp
-next
-  case (Cons e t)
-  show ?case
-  proof(cases e)
-    case (Create thread prio')
-    assume eq_e: " e = Create thread prio'"
-    from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto
-    hence neq_thread: "thread \<noteq> th"
-    proof(cases)
-      assume "thread \<notin> threads (t @ s)"
-      moreover have "th \<in> threads (t@s)"
-      proof -
-        from Cons have "extend_highest_gen s' th e' prio tm t" by auto
-        from extend_highest_gen.th_kept[OF this] show ?thesis by (simp add:s_def)
-      qed
-      ultimately show ?thesis by auto
-    qed
-    from Cons have "extend_highest_gen s' th e' prio tm t" by auto
-    from extend_highest_gen.th_kept[OF this]
-    have h': " th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" 
-      by (auto simp:s_def)
-    from stp
-    have thread_ts: "thread \<notin> threads (t @ s)"
-      by (cases, auto)
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))"
-        by (unfold eq_e, simp)
-      moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
-      proof(rule Max_insert)
-        from Cons have "vt step (t @ s)" by auto
-        from finite_threads[OF this]
-        show "finite (?f ` (threads (t@s)))" by simp
-      next
-        from h' show "(?f ` (threads (t@s))) \<noteq> {}" by auto
-      qed
-      moreover have "(Max (?f ` (threads (t@s)))) = ?t"
-      proof -
-        have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
-          (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B")
-        proof -
-          { fix th' 
-            assume "th' \<in> ?B"
-            with thread_ts eq_e
-            have "?f1 th' = ?f2 th'" by (auto simp:preced_def)
-          } thus ?thesis 
-            apply (auto simp:Image_def)
-          proof -
-            fix th'
-            assume h: "\<And>th'. th' \<in> threads (t @ s) \<Longrightarrow> 
-              preced th' (e # t @ s) = preced th' (t @ s)"
-              and h1: "th' \<in> threads (t @ s)"
-            show "preced th' (t @ s) \<in> (\<lambda>th'. preced th' (e # t @ s)) ` threads (t @ s)"
-            proof -
-              from h1 have "?f1 th' \<in> ?f1 ` ?B" by auto
-              moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp
-              ultimately show ?thesis by simp
-            qed
-          qed
-        qed
-        with Cons show ?thesis by auto
-      qed
-      moreover have "?f thread < ?t"
-      proof -
-        from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
-        from extend_highest_gen.create_low[OF this] and eq_e
-        have "prio' \<le> prio" by auto
-        thus ?thesis
-        by (unfold preced_th, unfold eq_e, insert lt_tm, 
-          auto simp:preced_def s_def precedence_less_def preced_th)
-    qed
-    ultimately show ?thesis by (auto simp:max_def)
-  qed
-next
-    case (Exit thread)
-    assume eq_e: "e = Exit thread"
-    from Cons have vt_e: "vt step (e#(t @ s))" by auto
-    from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto
-    from stp have thread_ts: "thread \<in> threads (t @ s)"
-      by(cases, unfold runing_def readys_def, auto)
-    from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
-    from extend_highest_gen.exit_diff[OF this] and eq_e
-    have neq_thread: "thread \<noteq> th" by auto
-    from Cons have "extend_highest_gen s' th e' prio tm t" by auto
-    from extend_highest_gen.th_kept[OF this, folded s_def]
-    have h': "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "threads (t@s) = insert thread ?A"
-        by (insert stp thread_ts, unfold eq_e, auto)
-      hence "Max (?f ` (threads (t@s))) = Max (?f ` \<dots>)" by simp
-      also from this have "\<dots> = Max (insert (?f thread) (?f ` ?A))" by simp
-      also have "\<dots> = max (?f thread) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        from finite_threads [OF vt_e]
-        show "finite (?f ` ?A)" by simp
-      next
-        from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
-        from extend_highest_gen.th_kept[OF this]
-        show "?f ` ?A \<noteq> {}" by  (auto simp:s_def)
-      qed
-      finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" .
-      moreover have "Max (?f ` (threads (t@s))) = ?t"
-      proof -
-        from Cons show ?thesis
-          by (unfold eq_e, auto simp:preced_def)
-      qed
-      ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp
-      moreover have "?f thread < ?t"
-      proof(unfold eq_e, simp add:preced_def, fold preced_def)
-        show "preced thread (t @ s) < ?t"
-        proof -
-          have "preced thread (t @ s) \<le> ?t" 
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" 
-              (is "?t = Max (?g ` ?B)") by simp
-            moreover have "?g thread \<le> \<dots>"
-            proof(rule Max_ge)
-              have "vt step (t@s)" by fact
-              from finite_threads [OF this]
-              show "finite (?g ` ?B)" by simp
-            next
-              from thread_ts
-              show "?g thread \<in> (?g ` ?B)" by auto
-            qed
-            ultimately show ?thesis by auto
-          qed
-          moreover have "preced thread (t @ s) \<noteq> ?t"
-          proof
-            assume "preced thread (t @ s) = preced th s"
-            with h' have "preced thread (t @ s) = preced th (t@s)" by simp
-            from preced_unique [OF this] have "thread = th"
-            proof
-              from h' show "th \<in> threads (t @ s)" by simp
-            next
-              from thread_ts show "thread \<in> threads (t @ s)" .
-            qed(simp)
-            with neq_thread show "False" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      qed
-      ultimately show ?thesis 
-        by (auto simp:max_def split:if_splits)
-    qed
-  next
-    case (P thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (V thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (Set thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      let ?B = "threads (t@s)"
-      from Cons have "extend_highest_gen s' th e' prio tm (e # t)" by auto
-      from extend_highest_gen.set_diff_low[OF this] and Set
-      have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
-      from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
-      also have "\<dots> = ?t"
-      proof(rule Max_eqI)
-        fix y
-        assume y_in: "y \<in> ?f ` ?B"
-        then obtain th1 where 
-          th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
-        show "y \<le> ?t"
-        proof(cases "th1 = thread")
-          case True
-          with neq_thread le_p eq_y s_def Set
-          show ?thesis
-            apply (subst preced_th, insert lt_tm)
-            by (auto simp:preced_def precedence_le_def)
-        next
-          case False
-          with Set eq_y
-          have "y  = preced th1 (t@s)"
-            by (simp add:preced_def)
-          moreover have "\<dots> \<le> ?t"
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
-              by auto
-            moreover have "preced th1 (t@s) \<le> \<dots>"
-            proof(rule Max_ge)
-              from th1_in 
-              show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
-                by simp
-            next
-              show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-              proof -
-                from Cons have "vt step (t @ s)" by auto
-                from finite_threads[OF this] show ?thesis by auto
-              qed
-            qed
-            ultimately show ?thesis by auto
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from Cons and finite_threads
-        show "finite (?f ` ?B)" by auto
-      next
-        from Cons have "extend_highest_gen s' th e' prio tm t" by auto
-        from extend_highest_gen.th_kept [OF this, folded s_def]
-        have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-        show "?t \<in> (?f ` ?B)" 
-        proof -
-          from neq_thread Set h
-          have "?t = ?f th" by (auto simp:preced_def)
-          with h show ?thesis by auto
-        qed
-      qed
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
-  by (insert th_kept max_kept, auto)
-
-lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
-  (is "?L = ?R")
-proof -
-  have "?L = cpreced (t@s) (wq (t@s)) th" 
-    by (unfold cp_eq_cpreced, simp)
-  also have "\<dots> = ?R"
-  proof(unfold cpreced_def)
-    show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependents (wq (t @ s)) th)) =
-          Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-      (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
-    proof(cases "?A = {}")
-      case False
-      have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
-      moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        show "finite (?f ` ?A)"
-        proof -
-          from dependents_threads[OF vt_t]
-          have "?A \<subseteq> threads (t@s)" .
-          moreover from finite_threads[OF vt_t] have "finite \<dots>" .
-          ultimately show ?thesis 
-            by (auto simp:finite_subset)
-        qed
-      next
-        from False show "(?f ` ?A) \<noteq> {}" by simp
-      qed
-      moreover have "\<dots> = Max (?f ` ?B)"
-      proof -
-        from max_preced have "?f th = Max (?f ` ?B)" .
-        moreover have "Max (?f ` ?A) \<le> \<dots>" 
-        proof(rule Max_mono)
-          from False show "(?f ` ?A) \<noteq> {}" by simp
-        next
-          show "?f ` ?A \<subseteq> ?f ` ?B" 
-          proof -
-            have "?A \<subseteq> ?B" by (rule dependents_threads[OF vt_t])
-            thus ?thesis by auto
-          qed
-        next
-          from finite_threads[OF vt_t] 
-          show "finite (?f ` ?B)" by simp
-        qed
-        ultimately show ?thesis
-          by (auto simp:max_def)
-      qed
-      ultimately show ?thesis by auto
-    next
-      case True
-      with max_preced show ?thesis by auto
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
-  by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)
-
-lemma th_cp_preced: "cp (t@s) th = preced th s"
-  by (fold max_kept, unfold th_cp_max_preced, simp)
-
-lemma preced_less':
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  shows "preced th' s < preced th s"
-proof -
-  have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
-  proof(rule Max_ge)
-    from finite_threads [OF vt_s]
-    show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
-  next
-    from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
-      by simp
-  qed
-  moreover have "preced th' s \<noteq> preced th s"
-  proof
-    assume "preced th' s = preced th s"
-    from preced_unique[OF this th'_in] neq_th' threads_s
-    show "False" by  (auto simp:readys_def)
-  qed
-  ultimately show ?thesis using highest_preced_thread
-    by auto
-qed
-
-lemma pv_blocked:
-  fixes th'
-  assumes th'_in: "th' \<in> threads (t@s)"
-  and neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
-  shows "th' \<notin> runing (t@s)"
-proof
-  assume "th' \<in> runing (t@s)"
-  hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
-    by (auto simp:runing_def)
-  with max_cp_readys_threads [OF vt_t]
-  have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
-    by auto
-  moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
-  ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
-  moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
-    by simp
-  finally have h: "cp (t @ s) th' = preced th (t @ s)" .
-  show False
-  proof -
-    have "dependents (wq (t @ s)) th' = {}" 
-      by (rule count_eq_dependents [OF vt_t eq_pv])
-    moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
-    proof
-      assume "preced th' (t @ s) = preced th (t @ s)"
-      hence "th' = th"
-      proof(rule preced_unique)
-        from th_kept show "th \<in> threads (t @ s)" by simp
-      next
-        from th'_in show "th' \<in> threads (t @ s)" by simp
-      qed
-      with assms show False by simp
-    qed
-    ultimately show ?thesis
-      by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
-  qed
-qed
-
-lemma runing_precond_pre:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<in> threads (t@s) \<and>
-         cntP (t@s) th' = cntV (t@s) th'"
-proof -
-  show ?thesis
-  proof(induct rule:ind)
-    case (Cons e t)
-    from Cons
-    have in_thread: "th' \<in> threads (t @ s)"
-      and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    from Cons have "extend_highest_gen s' th e' prio tm t" by auto
-    from extend_highest_gen.pv_blocked 
-    [OF this, folded s_def, OF in_thread neq_th' not_holding]
-    have not_runing: "th' \<notin> runing (t @ s)" .
-    show ?case
-    proof(cases e)
-      case (V thread cs)
-      from Cons and V have vt_v: "vt step (V thread cs#(t@s))" by auto
-
-      show ?thesis
-      proof -
-        from Cons and V have "step (t@s) (V thread cs)" by auto
-        hence neq_th': "thread \<noteq> th'"
-        proof(cases)
-          assume "thread \<in> runing (t@s)"
-          moreover have "th' \<notin> runing (t@s)" by fact
-          ultimately show ?thesis by auto
-        qed
-        with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
-          by (unfold V, simp add:cntP_def cntV_def count_def)
-        moreover from in_thread
-        have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
-        ultimately show ?thesis by auto
-      qed
-    next
-      case (P thread cs)
-      from Cons and P have "step (t@s) (P thread cs)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t@s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Create thread prio')
-      from Cons and Create have "step (t@s) (Create thread prio')" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<notin> threads (t @ s)"
-        moreover have "th' \<in> threads (t@s)" by fact
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Create 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Create 
-      have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
-      ultimately show ?thesis by auto
-    next
-      case (Exit thread)
-      from Cons and Exit have "step (t@s) (Exit thread)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t @ s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Exit 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Exit and neq_th' 
-      have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio')
-      with Cons
-      show ?thesis 
-        by (auto simp:cntP_def cntV_def count_def)
-    qed
-  next
-    case Nil
-    with assms
-    show ?case by auto
-  qed
-qed
-
-(*
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<notin> runing (t@s)"
-proof -
-  from runing_precond_pre[OF th'_in eq_pv neq_th']
-  have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-  from pv_blocked[OF h1 neq_th' h2] 
-  show ?thesis .
-qed
-*)
-
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  and is_runing: "th' \<in> runing (t@s)"
-  shows "cntP s th' > cntV s th'"
-proof -
-  have "cntP s th' \<noteq> cntV s th'"
-  proof
-    assume eq_pv: "cntP s th' = cntV s th'"
-    from runing_precond_pre[OF th'_in eq_pv neq_th']
-    have h1: "th' \<in> threads (t @ s)"  
-      and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    from pv_blocked[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
-    with is_runing show "False" by simp
-  qed
-  moreover from cnp_cnv_cncs[OF vt_s, of th'] 
-  have "cntV s th' \<le> cntP s th'" by auto
-  ultimately show ?thesis by auto
-qed
-
-lemma moment_blocked_pre:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
-         th' \<in> threads ((moment (i+j) t)@s)"
-proof(induct j)
-  case (Suc k)
-  show ?case
-  proof -
-    { assume True: "Suc (i+k) \<le> length t"
-      from moment_head [OF this] 
-      obtain e where
-        eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
-        by blast
-      from red_moment[of "Suc(i+k)"]
-      and eq_me have "extend_highest_gen s' th e' prio tm (e # moment (i + k) t)" by simp
-      hence vt_e: "vt step (e#(moment (i + k) t)@s)"
-        by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
-                          highest_gen_def s_def, auto)
-      have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
-      proof(unfold s_def)
-        show "th' \<notin> runing (moment (i + k) t @ e' # s')"
-        proof(rule extend_highest_gen.pv_blocked)
-          from Suc show "th' \<in> threads (moment (i + k) t @ e' # s')"
-            by (simp add:s_def)
-        next
-          from neq_th' show "th' \<noteq> th" .
-        next
-          from red_moment show "extend_highest_gen s' th e' prio tm (moment (i + k) t)" .
-        next
-          from Suc show "cntP (moment (i + k) t @ e' # s') th' = cntV (moment (i + k) t @ e' # s') th'"
-            by (auto simp:s_def)
-        qed
-      qed
-      from step_back_step[OF vt_e]
-      have "step ((moment (i + k) t)@s) e" .
-      hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
-        th' \<in> threads (e#(moment (i + k) t)@s)
-        "
-      proof(cases)
-        case (thread_create thread prio)
-        with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_exit thread)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_P thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_V thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_set thread prio')
-        with Suc show ?thesis
-          by (auto simp:cntP_def cntV_def count_def)
-      qed
-      with eq_me have ?thesis using eq_me by auto 
-    } note h = this
-    show ?thesis
-    proof(cases "Suc (i+k) \<le> length t")
-      case True
-      from h [OF this] show ?thesis .
-    next
-      case False
-      with moment_ge
-      have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
-      with Suc show ?thesis by auto
-    qed
-  qed
-next
-  case 0
-  from assms show ?case by auto
-qed
-
-lemma moment_blocked:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  and le_ij: "i \<le> j"
-  shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
-         th' \<in> threads ((moment j t)@s) \<and>
-         th' \<notin> runing ((moment j t)@s)"
-proof -
-  from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
-  have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
-    and h2: "th' \<in> threads ((moment j t)@s)" by auto
-  with extend_highest_gen.pv_blocked [OF  red_moment [of j], folded s_def, OF h2 neq_th' h1]
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_1:
-  assumes neq_th': "th' \<noteq> th"
-  and runing': "th' \<in> runing (t@s)"
-  shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
-proof(cases "th' \<in> threads s")
-  case True
-  with runing_precond [OF this neq_th' runing'] show ?thesis by simp
-next
-  case False
-  let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
-  let ?q = "moment 0 t"
-  from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
-  from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
-  from p_split_gen [of ?Q, OF this not_thread]
-  obtain i where lt_its: "i < length t"
-    and le_i: "0 \<le> i"
-    and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
-    and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
-  from lt_its have "Suc i \<le> length t" by auto
-  from moment_head[OF this] obtain e where 
-   eq_me: "moment (Suc i) t = e # moment i t" by blast
-  from red_moment[of "Suc i"] and eq_me
-  have "extend_highest_gen s' th e' prio tm (e # moment i t)" by simp
-  hence vt_e: "vt step (e#(moment i t)@s)"
-    by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
-      highest_gen_def s_def, auto)
-  from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
-  from post[rule_format, of "Suc i"] and eq_me 
-  have not_in': "th' \<in> threads (e # moment i t@s)" by auto
-  from create_pre[OF stp_i pre this] 
-  obtain prio where eq_e: "e = Create th' prio" .
-  have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
-  proof(rule cnp_cnv_eq)
-    from step_back_vt [OF vt_e] 
-    show "vt step (moment i t @ s)" .
-  next
-    from eq_e and stp_i 
-    have "step (moment i t @ s) (Create th' prio)" by simp
-    thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
-  qed
-  with eq_e
-  have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
-    by (simp add:cntP_def cntV_def count_def)
-  with eq_me[symmetric]
-  have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
-    by simp
-  from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
-  with eq_me [symmetric]
-  have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
-  from moment_blocked [OF neq_th' h2 h1, of "length t"] and lt_its
-  and moment_ge
-  have "th' \<notin> runing (t @ s)" by auto
-  with runing'
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_2:
-  assumes runing': "th' \<in> runing (t@s)"
-  shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
-proof -
-  from runing_inversion_1[OF _ runing']
-  show ?thesis by auto
-qed
-
-lemma live: "runing (t@s) \<noteq> {}"
-proof(cases "th \<in> runing (t@s)")
-  case True thus ?thesis by auto
-next
-  case False
-  then have not_ready: "th \<notin> readys (t@s)"
-    apply (unfold runing_def, 
-            insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
-    by auto
-  from th_kept have "th \<in> threads (t@s)" by auto
-  from th_chain_to_ready[OF vt_t this] and not_ready
-  obtain th' where th'_in: "th' \<in> readys (t@s)"
-    and dp: "(Th th, Th th') \<in> (depend (t @ s))\<^sup>+" by auto
-  have "th' \<in> runing (t@s)"
-  proof -
-    have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
-    proof -
-      have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')) = 
-               preced th (t@s)"
-      proof(rule Max_eqI)
-        fix y
-        assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-        then obtain th1 where
-          h1: "th1 = th' \<or> th1 \<in>  dependents (wq (t @ s)) th'"
-          and eq_y: "y = preced th1 (t@s)" by auto
-        show "y \<le> preced th (t @ s)"
-        proof -
-          from max_preced
-          have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
-          moreover have "y \<le> \<dots>"
-          proof(rule Max_ge)
-            from h1
-            have "th1 \<in> threads (t@s)"
-            proof
-              assume "th1 = th'"
-              with th'_in show ?thesis by (simp add:readys_def)
-            next
-              assume "th1 \<in> dependents (wq (t @ s)) th'"
-              with dependents_threads [OF vt_t]
-              show "th1 \<in> threads (t @ s)" by auto
-            qed
-            with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
-          next
-            from finite_threads[OF vt_t]
-            show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th']
-        show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th'))"
-          by (auto intro:finite_subset)
-      next
-        from dp
-        have "th \<in> dependents (wq (t @ s)) th'" 
-          by (unfold cs_dependents_def, auto simp:eq_depend)
-        thus "preced th (t @ s) \<in> 
-                (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-          by auto
-      qed
-      moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
-      proof -
-        from max_preced and max_cp_eq[OF vt_t, symmetric]
-        have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
-        with max_cp_readys_threads[OF vt_t] show ?thesis by simp
-      qed
-      ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
-    qed
-    with th'_in show ?thesis by (auto simp:runing_def)
-  qed
-  thus ?thesis by auto
-qed
-
-end
-
-end
-
-
--- a/prio/Attic/ExtS.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1019 +0,0 @@
-theory ExtS
-imports Prio
-begin
-
-locale highest_set =
-  fixes s' th prio fixes s 
-  defines s_def : "s \<equiv> (Set th prio#s')"
-  assumes vt_s: "vt step s"
-  and highest: "preced th s = Max ((cp s)`threads s)"
-
-context highest_set
-begin
-
-
-lemma vt_s': "vt step s'"
-  by (insert vt_s, unfold s_def, drule_tac step_back_vt, simp)
-
-lemma step_set: "step s' (Set th prio)"
-  by (insert vt_s, unfold s_def, drule_tac step_back_step, simp)
-
-lemma step_set_elim: 
-  "\<lbrakk>\<lbrakk>th \<in> runing s'\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-  by (insert step_set, ind_cases "step s' (Set th prio)", auto)
-
-
-lemma threads_s: "th \<in> threads s"
-  by (rule step_set_elim, unfold runing_def readys_def, auto simp:s_def)
-
-lemma same_depend: "depend s = depend s'"
-  by (insert depend_set_unchanged, unfold s_def, simp)
-
-lemma same_dependents:
-  "dependents (wq s) th = dependents (wq s') th"
-  apply (unfold cs_dependents_def)
-  by (unfold eq_depend same_depend, simp)
-
-lemma eq_cp_s_th: "cp s th = preced th s"
-proof -
-  from highest and max_cp_eq[OF vt_s]
-  have is_max: "preced th s = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-  have sbs: "({th} \<union> dependents (wq s) th) \<subseteq> threads s"
-  proof -
-    from threads_s and dependents_threads[OF vt_s, of th]
-    show ?thesis by auto
-  qed
-  show ?thesis
-  proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-    show "preced th s \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)" by simp
-  next
-    fix y 
-    assume "y \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)"
-    then obtain th1 where th1_in: "th1 \<in> ({th} \<union> dependents (wq s) th)"
-      and eq_y: "y = preced th1 s" by auto
-    show "y \<le> preced th s"
-    proof(unfold is_max, rule Max_ge)
-      from finite_threads[OF vt_s] 
-      show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-    next
-      from sbs th1_in and eq_y 
-      show "y \<in> (\<lambda>th. preced th s) ` threads s" by auto
-    qed
-  next
-    from sbs and finite_threads[OF vt_s]
-    show "finite ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))"
-      by (auto intro:finite_subset)
-  qed
-qed
-
-lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp)
-
-lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
-
-lemma is_ready: "th \<in> readys s"
-proof -
-  have "\<forall>cs. \<not> waiting s th cs"
-    apply (rule step_set_elim, unfold s_def, insert depend_set_unchanged[of th prio s'])
-    apply (unfold s_depend_def, unfold runing_def readys_def)
-    apply (auto, fold s_def)
-    apply (erule_tac x = cs in allE, auto simp:waiting_eq)
-  proof -
-    fix cs
-    assume h: 
-      "{(Th t, Cs c) |t c. waiting (wq s) t c} \<union> {(Cs c, Th t) |c t. holding (wq s) t c} =
-          {(Th t, Cs c) |t c. waiting (wq s') t c} \<union> {(Cs c, Th t) |c t. holding (wq s') t c}"
-            (is "?L = ?R")
-    and wt: "waiting (wq s) th cs" and nwt: "\<not> waiting (wq s') th cs"
-    from wt have "(Th th, Cs cs) \<in> ?L" by auto
-    with h have "(Th th, Cs cs) \<in> ?R" by simp
-    hence "waiting (wq s') th cs" by auto with nwt
-    show False by auto
-  qed    
-  with threads_s show ?thesis 
-    by (unfold readys_def, auto)
-qed
-
-lemma highest': "cp s th = Max (cp s ` threads s)"
-proof -
-  from highest_cp_preced max_cp_eq[OF vt_s, symmetric]
-  show ?thesis by simp
-qed
-
-lemma is_runing: "th \<in> runing s"
-proof -
-  have "Max (cp s ` threads s) = Max (cp s ` readys s)"
-  proof -
-    have " Max (cp s ` readys s) = cp s th"
-    proof(rule Max_eqI)
-      from finite_threads[OF vt_s] readys_threads finite_subset
-      have "finite (readys s)" by blast
-      thus "finite (cp s ` readys s)" by auto
-    next
-      from is_ready show "cp s th \<in> cp s ` readys s" by auto
-    next
-      fix y
-      assume "y \<in> cp s ` readys s"
-      then obtain th1 where 
-        eq_y: "y = cp s th1" and th1_in: "th1 \<in> readys s" by auto
-      show  "y \<le> cp s th" 
-      proof -
-        have "y \<le> Max (cp s ` threads s)"
-        proof(rule Max_ge)
-          from eq_y and th1_in
-          show "y \<in> cp s ` threads s"
-            by (auto simp:readys_def)
-        next
-          from finite_threads[OF vt_s]
-          show "finite (cp s ` threads s)" by auto
-        qed
-        with highest' show ?thesis by auto
-      qed
-    qed
-    with highest' show ?thesis by auto
-  qed
-  thus ?thesis
-    by (unfold runing_def, insert highest' is_ready, auto)
-qed
-
-end
-
-locale extend_highest_set = highest_set + 
-  fixes t 
-  assumes vt_t: "vt step (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"
-
-lemma step_back_vt_app: 
-  assumes vt_ts: "vt cs (t@s)" 
-  shows "vt cs 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 cs (t @ s) \<Longrightarrow> vt cs s"
-      and vt_et: "vt cs ((e # t) @ s)"
-    show ?case
-    proof(rule ih)
-      show "vt cs (t @ s)"
-      proof(rule step_back_vt)
-        from vt_et show "vt cs (e # t @ s)" by simp
-      qed
-    qed
-  qed
-qed
-
-context extend_highest_set
-begin
-
-lemma red_moment:
-  "extend_highest_set s' th prio (moment i t)"
-  apply (insert extend_highest_set_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
-  apply (unfold extend_highest_set_def extend_highest_set_axioms_def, clarsimp)
-  by (unfold highest_set_def, auto dest:step_back_vt_app)
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes 
-    h0: "R []"
-  and h2: "\<And> e t. \<lbrakk>vt step (t@s); step (t@s) e; 
-                    extend_highest_set s' th prio t; 
-                    extend_highest_set s' th prio (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
-  shows "R t"
-proof -
-  from vt_t extend_highest_set_axioms show ?thesis
-  proof(induct t)
-    from h0 show "R []" .
-  next
-    case (Cons e t')
-    assume ih: "\<lbrakk>vt step (t' @ s); extend_highest_set s' th prio t'\<rbrakk> \<Longrightarrow> R t'"
-      and vt_e: "vt step ((e # t') @ s)"
-      and et: "extend_highest_set s' th prio (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 step (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_set s' th prio t'"
-          by (unfold extend_highest_set_def extend_highest_set_axioms_def, auto dest:step_back_vt)
-      next
-        from vt_ts show "vt step (t' @ s)" .
-      qed
-    next
-      from et show "extend_highest_set s' th prio (e # t')" .
-    next
-      from et show ext': "extend_highest_set s' th prio t'"
-          by (unfold extend_highest_set_def extend_highest_set_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 "th \<in> threads ([] @ s) \<and> preced th ([] @ s) = preced th s"
-      by auto
-  next
-    case (Cons e t)
-    show ?case
-    proof(cases e)
-      case (Create thread prio)
-      assume eq_e: " e = Create thread prio"
-      show ?thesis
-      proof -
-        from Cons and eq_e have "step (t@s) (Create thread prio)" by auto
-        hence "th \<noteq> thread"
-        proof(cases)
-          assume "thread \<notin> threads (t @ s)"
-          with Cons show ?thesis by auto
-        qed
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    next
-      case (Exit thread)
-      assume eq_e: "e = Exit thread"
-      from Cons have "extend_highest_set s' th prio (e # t)" by auto
-      from extend_highest_set.exit_diff [OF this] and eq_e
-      have neq_th: "thread \<noteq> th" by auto
-      with Cons
-      show ?thesis
-        by (unfold eq_e, auto simp:preced_def)
-    next
-      case (P thread cs)
-      assume eq_e: "e = P thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (V thread cs)
-      assume eq_e: "e = V thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (Set thread prio')
-      assume eq_e: " e = Set thread prio'"
-      show ?thesis
-      proof -
-        from Cons have "extend_highest_set s' th prio (e # t)" by auto
-        from extend_highest_set.set_diff_low[OF this] and eq_e
-        have "th \<noteq> thread" by auto
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    qed
-  qed
-qed
-
-lemma max_kept: "Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s"
-proof(induct rule:ind)
-  case Nil
-  from highest_preced_thread
-  show "Max ((\<lambda>th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s"
-    by simp
-next
-  case (Cons e t)
-  show ?case
-  proof(cases e)
-    case (Create thread prio')
-    assume eq_e: " e = Create thread prio'"
-    from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto
-    hence neq_thread: "thread \<noteq> th"
-    proof(cases)
-      assume "thread \<notin> threads (t @ s)"
-      moreover have "th \<in> threads (t@s)"
-      proof -
-        from Cons have "extend_highest_set s' th prio t" by auto
-        from extend_highest_set.th_kept[OF this] show ?thesis by (simp add:s_def)
-      qed
-      ultimately show ?thesis by auto
-    qed
-    from Cons have "extend_highest_set s' th prio t" by auto
-    from extend_highest_set.th_kept[OF this]
-    have h': " th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" 
-      by (auto simp:s_def)
-    from stp
-    have thread_ts: "thread \<notin> threads (t @ s)"
-      by (cases, auto)
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))"
-        by (unfold eq_e, simp)
-      moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
-      proof(rule Max_insert)
-        from Cons have "vt step (t @ s)" by auto
-        from finite_threads[OF this]
-        show "finite (?f ` (threads (t@s)))" by simp
-      next
-        from h' show "(?f ` (threads (t@s))) \<noteq> {}" by auto
-      qed
-      moreover have "(Max (?f ` (threads (t@s)))) = ?t"
-      proof -
-        have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
-          (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B")
-        proof -
-          { fix th' 
-            assume "th' \<in> ?B"
-            with thread_ts eq_e
-            have "?f1 th' = ?f2 th'" by (auto simp:preced_def)
-          } thus ?thesis 
-            apply (auto simp:Image_def)
-          proof -
-            fix th'
-            assume h: "\<And>th'. th' \<in> threads (t @ s) \<Longrightarrow> 
-              preced th' (e # t @ s) = preced th' (t @ s)"
-              and h1: "th' \<in> threads (t @ s)"
-            show "preced th' (t @ s) \<in> (\<lambda>th'. preced th' (e # t @ s)) ` threads (t @ s)"
-            proof -
-              from h1 have "?f1 th' \<in> ?f1 ` ?B" by auto
-              moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp
-              ultimately show ?thesis by simp
-            qed
-          qed
-        qed
-        with Cons show ?thesis by auto
-      qed
-      moreover have "?f thread < ?t"
-      proof -
-        from Cons have " extend_highest_set s' th prio (e # t)" by auto
-        from extend_highest_set.create_low[OF this] and eq_e
-        have "prio' \<le> prio" by auto
-        thus ?thesis
-        by (unfold eq_e, auto simp:preced_def s_def precedence_less_def)
-    qed
-    ultimately show ?thesis by (auto simp:max_def)
-  qed
-next
-    case (Exit thread)
-    assume eq_e: "e = Exit thread"
-    from Cons have vt_e: "vt step (e#(t @ s))" by auto
-    from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto
-    from stp have thread_ts: "thread \<in> threads (t @ s)"
-      by(cases, unfold runing_def readys_def, auto)
-    from Cons have "extend_highest_set s' th prio (e # t)" by auto
-    from extend_highest_set.exit_diff[OF this] and eq_e
-    have neq_thread: "thread \<noteq> th" by auto
-    from Cons have "extend_highest_set s' th prio t" by auto
-    from extend_highest_set.th_kept[OF this, folded s_def]
-    have h': "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "threads (t@s) = insert thread ?A"
-        by (insert stp thread_ts, unfold eq_e, auto)
-      hence "Max (?f ` (threads (t@s))) = Max (?f ` \<dots>)" by simp
-      also from this have "\<dots> = Max (insert (?f thread) (?f ` ?A))" by simp
-      also have "\<dots> = max (?f thread) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        from finite_threads [OF vt_e]
-        show "finite (?f ` ?A)" by simp
-      next
-        from Cons have "extend_highest_set s' th prio (e # t)" by auto
-        from extend_highest_set.th_kept[OF this]
-        show "?f ` ?A \<noteq> {}" by  (auto simp:s_def)
-      qed
-      finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" .
-      moreover have "Max (?f ` (threads (t@s))) = ?t"
-      proof -
-        from Cons show ?thesis
-          by (unfold eq_e, auto simp:preced_def)
-      qed
-      ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp
-      moreover have "?f thread < ?t"
-      proof(unfold eq_e, simp add:preced_def, fold preced_def)
-        show "preced thread (t @ s) < ?t"
-        proof -
-          have "preced thread (t @ s) \<le> ?t" 
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" 
-              (is "?t = Max (?g ` ?B)") by simp
-            moreover have "?g thread \<le> \<dots>"
-            proof(rule Max_ge)
-              have "vt step (t@s)" by fact
-              from finite_threads [OF this]
-              show "finite (?g ` ?B)" by simp
-            next
-              from thread_ts
-              show "?g thread \<in> (?g ` ?B)" by auto
-            qed
-            ultimately show ?thesis by auto
-          qed
-          moreover have "preced thread (t @ s) \<noteq> ?t"
-          proof
-            assume "preced thread (t @ s) = preced th s"
-            with h' have "preced thread (t @ s) = preced th (t@s)" by simp
-            from preced_unique [OF this] have "thread = th"
-            proof
-              from h' show "th \<in> threads (t @ s)" by simp
-            next
-              from thread_ts show "thread \<in> threads (t @ s)" .
-            qed(simp)
-            with neq_thread show "False" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      qed
-      ultimately show ?thesis 
-        by (auto simp:max_def split:if_splits)
-    qed
-  next
-    case (P thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (V thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (Set thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      let ?B = "threads (t@s)"
-      from Cons have "extend_highest_set s' th prio (e # t)" by auto
-      from extend_highest_set.set_diff_low[OF this] and Set
-      have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
-      from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
-      also have "\<dots> = ?t"
-      proof(rule Max_eqI)
-        fix y
-        assume y_in: "y \<in> ?f ` ?B"
-        then obtain th1 where 
-          th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
-        show "y \<le> ?t"
-        proof(cases "th1 = thread")
-          case True
-          with neq_thread le_p eq_y s_def Set
-          show ?thesis
-            by (auto simp:preced_def precedence_le_def)
-        next
-          case False
-          with Set eq_y
-          have "y  = preced th1 (t@s)"
-            by (simp add:preced_def)
-          moreover have "\<dots> \<le> ?t"
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
-              by auto
-            moreover have "preced th1 (t@s) \<le> \<dots>"
-            proof(rule Max_ge)
-              from th1_in 
-              show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
-                by simp
-            next
-              show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-              proof -
-                from Cons have "vt step (t @ s)" by auto
-                from finite_threads[OF this] show ?thesis by auto
-              qed
-            qed
-            ultimately show ?thesis by auto
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from Cons and finite_threads
-        show "finite (?f ` ?B)" by auto
-      next
-        from Cons have "extend_highest_set s' th prio t" by auto
-        from extend_highest_set.th_kept [OF this, folded s_def]
-        have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-        show "?t \<in> (?f ` ?B)" 
-        proof -
-          from neq_thread Set h
-          have "?t = ?f th" by (auto simp:preced_def)
-          with h show ?thesis by auto
-        qed
-      qed
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
-  by (insert th_kept max_kept, auto)
-
-lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
-  (is "?L = ?R")
-proof -
-  have "?L = cpreced (t@s) (wq (t@s)) th" 
-    by (unfold cp_eq_cpreced, simp)
-  also have "\<dots> = ?R"
-  proof(unfold cpreced_def)
-    show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependents (wq (t @ s)) th)) =
-          Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-      (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
-    proof(cases "?A = {}")
-      case False
-      have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
-      moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        show "finite (?f ` ?A)"
-        proof -
-          from dependents_threads[OF vt_t]
-          have "?A \<subseteq> threads (t@s)" .
-          moreover from finite_threads[OF vt_t] have "finite \<dots>" .
-          ultimately show ?thesis 
-            by (auto simp:finite_subset)
-        qed
-      next
-        from False show "(?f ` ?A) \<noteq> {}" by simp
-      qed
-      moreover have "\<dots> = Max (?f ` ?B)"
-      proof -
-        from max_preced have "?f th = Max (?f ` ?B)" .
-        moreover have "Max (?f ` ?A) \<le> \<dots>" 
-        proof(rule Max_mono)
-          from False show "(?f ` ?A) \<noteq> {}" by simp
-        next
-          show "?f ` ?A \<subseteq> ?f ` ?B" 
-          proof -
-            have "?A \<subseteq> ?B" by (rule dependents_threads[OF vt_t])
-            thus ?thesis by auto
-          qed
-        next
-          from finite_threads[OF vt_t] 
-          show "finite (?f ` ?B)" by simp
-        qed
-        ultimately show ?thesis
-          by (auto simp:max_def)
-      qed
-      ultimately show ?thesis by auto
-    next
-      case True
-      with max_preced show ?thesis by auto
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
-  by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)
-
-lemma th_cp_preced: "cp (t@s) th = preced th s"
-  by (fold max_kept, unfold th_cp_max_preced, simp)
-
-lemma preced_less':
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  shows "preced th' s < preced th s"
-proof -
-  have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
-  proof(rule Max_ge)
-    from finite_threads [OF vt_s]
-    show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
-  next
-    from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
-      by simp
-  qed
-  moreover have "preced th' s \<noteq> preced th s"
-  proof
-    assume "preced th' s = preced th s"
-    from preced_unique[OF this th'_in] neq_th' is_ready
-    show "False" by  (auto simp:readys_def)
-  qed
-  ultimately show ?thesis using highest_preced_thread
-    by auto
-qed
-
-lemma pv_blocked:
-  fixes th'
-  assumes th'_in: "th' \<in> threads (t@s)"
-  and neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
-  shows "th' \<notin> runing (t@s)"
-proof
-  assume "th' \<in> runing (t@s)"
-  hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
-    by (auto simp:runing_def)
-  with max_cp_readys_threads [OF vt_t]
-  have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
-    by auto
-  moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
-  ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
-  moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
-    by simp
-  finally have h: "cp (t @ s) th' = preced th (t @ s)" .
-  show False
-  proof -
-    have "dependents (wq (t @ s)) th' = {}" 
-      by (rule count_eq_dependents [OF vt_t eq_pv])
-    moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
-    proof
-      assume "preced th' (t @ s) = preced th (t @ s)"
-      hence "th' = th"
-      proof(rule preced_unique)
-        from th_kept show "th \<in> threads (t @ s)" by simp
-      next
-        from th'_in show "th' \<in> threads (t @ s)" by simp
-      qed
-      with assms show False by simp
-    qed
-    ultimately show ?thesis
-      by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
-  qed
-qed
-
-lemma runing_precond_pre:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<in> threads (t@s) \<and>
-         cntP (t@s) th' = cntV (t@s) th'"
-proof -
-  show ?thesis
-  proof(induct rule:ind)
-    case (Cons e t)
-    from Cons
-    have in_thread: "th' \<in> threads (t @ s)"
-      and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    have "extend_highest_set s' th prio t" by fact
-    from extend_highest_set.pv_blocked 
-    [OF this, folded s_def, OF in_thread neq_th' not_holding]
-    have not_runing: "th' \<notin> runing (t @ s)" .
-    show ?case
-    proof(cases e)
-      case (V thread cs)
-      from Cons and V have vt_v: "vt step (V thread cs#(t@s))" by auto
-
-      show ?thesis
-      proof -
-        from Cons and V have "step (t@s) (V thread cs)" by auto
-        hence neq_th': "thread \<noteq> th'"
-        proof(cases)
-          assume "thread \<in> runing (t@s)"
-          moreover have "th' \<notin> runing (t@s)" by fact
-          ultimately show ?thesis by auto
-        qed
-        with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
-          by (unfold V, simp add:cntP_def cntV_def count_def)
-        moreover from in_thread
-        have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
-        ultimately show ?thesis by auto
-      qed
-    next
-      case (P thread cs)
-      from Cons and P have "step (t@s) (P thread cs)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t@s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Create thread prio')
-      from Cons and Create have "step (t@s) (Create thread prio')" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<notin> threads (t @ s)"
-        moreover have "th' \<in> threads (t@s)" by fact
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Create 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Create 
-      have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
-      ultimately show ?thesis by auto
-    next
-      case (Exit thread)
-      from Cons and Exit have "step (t@s) (Exit thread)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t @ s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Exit 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Exit and neq_th' 
-      have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio')
-      with Cons
-      show ?thesis 
-        by (auto simp:cntP_def cntV_def count_def)
-    qed
-  next
-    case Nil
-    with assms
-    show ?case by auto
-  qed
-qed
-
-(*
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<notin> runing (t@s)"
-proof -
-  from runing_precond_pre[OF th'_in eq_pv neq_th']
-  have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-  from pv_blocked[OF h1 neq_th' h2] 
-  show ?thesis .
-qed
-*)
-
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  and is_runing: "th' \<in> runing (t@s)"
-  shows "cntP s th' > cntV s th'"
-proof -
-  have "cntP s th' \<noteq> cntV s th'"
-  proof
-    assume eq_pv: "cntP s th' = cntV s th'"
-    from runing_precond_pre[OF th'_in eq_pv neq_th']
-    have h1: "th' \<in> threads (t @ s)"  
-      and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    from pv_blocked[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
-    with is_runing show "False" by simp
-  qed
-  moreover from cnp_cnv_cncs[OF vt_s, of th'] 
-  have "cntV s th' \<le> cntP s th'" by auto
-  ultimately show ?thesis by auto
-qed
-
-lemma moment_blocked_pre:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
-         th' \<in> threads ((moment (i+j) t)@s)"
-proof(induct j)
-  case (Suc k)
-  show ?case
-  proof -
-    { assume True: "Suc (i+k) \<le> length t"
-      from moment_head [OF this] 
-      obtain e where
-        eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
-        by blast
-      from red_moment[of "Suc(i+k)"]
-      and eq_me have "extend_highest_set s' th prio (e # moment (i + k) t)" by simp
-      hence vt_e: "vt step (e#(moment (i + k) t)@s)"
-        by (unfold extend_highest_set_def extend_highest_set_axioms_def 
-          highest_set_def s_def, auto)
-      have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
-      proof(unfold s_def)
-        show "th' \<notin> runing (moment (i + k) t @ Set th prio # s')"
-        proof(rule extend_highest_set.pv_blocked)
-          from Suc show "th' \<in> threads (moment (i + k) t @ Set th prio # s')"
-            by (simp add:s_def)
-        next
-          from neq_th' show "th' \<noteq> th" .
-        next
-          from red_moment show "extend_highest_set s' th prio (moment (i + k) t)" .
-        next
-          from Suc show "cntP (moment (i + k) t @ Set th prio # s') th' =
-            cntV (moment (i + k) t @ Set th prio # s') th'"
-            by (auto simp:s_def)
-        qed
-      qed
-      from step_back_step[OF vt_e]
-      have "step ((moment (i + k) t)@s) e" .
-      hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
-        th' \<in> threads (e#(moment (i + k) t)@s)
-        "
-      proof(cases)
-        case (thread_create thread prio)
-        with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_exit thread)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_P thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_V thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_set thread prio')
-        with Suc show ?thesis
-          by (auto simp:cntP_def cntV_def count_def)
-      qed
-      with eq_me have ?thesis using eq_me by auto 
-    } note h = this
-    show ?thesis
-    proof(cases "Suc (i+k) \<le> length t")
-      case True
-      from h [OF this] show ?thesis .
-    next
-      case False
-      with moment_ge
-      have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
-      with Suc show ?thesis by auto
-    qed
-  qed
-next
-  case 0
-  from assms show ?case by auto
-qed
-
-lemma moment_blocked:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  and le_ij: "i \<le> j"
-  shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
-         th' \<in> threads ((moment j t)@s) \<and>
-         th' \<notin> runing ((moment j t)@s)"
-proof -
-  from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
-  have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
-    and h2: "th' \<in> threads ((moment j t)@s)" by auto
-  with extend_highest_set.pv_blocked [OF  red_moment [of j], folded s_def, OF h2 neq_th' h1]
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_1:
-  assumes neq_th': "th' \<noteq> th"
-  and runing': "th' \<in> runing (t@s)"
-  shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
-proof(cases "th' \<in> threads s")
-  case True
-  with runing_precond [OF this neq_th' runing'] show ?thesis by simp
-next
-  case False
-  let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
-  let ?q = "moment 0 t"
-  from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
-  from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
-  from p_split_gen [of ?Q, OF this not_thread]
-  obtain i where lt_its: "i < length t"
-    and le_i: "0 \<le> i"
-    and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
-    and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
-  from lt_its have "Suc i \<le> length t" by auto
-  from moment_head[OF this] obtain e where 
-   eq_me: "moment (Suc i) t = e # moment i t" by blast
-  from red_moment[of "Suc i"] and eq_me
-  have "extend_highest_set s' th prio (e # moment i t)" by simp
-  hence vt_e: "vt step (e#(moment i t)@s)"
-    by (unfold extend_highest_set_def extend_highest_set_axioms_def 
-      highest_set_def s_def, auto)
-  from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
-  from post[rule_format, of "Suc i"] and eq_me 
-  have not_in': "th' \<in> threads (e # moment i t@s)" by auto
-  from create_pre[OF stp_i pre this] 
-  obtain prio where eq_e: "e = Create th' prio" .
-  have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
-  proof(rule cnp_cnv_eq)
-    from step_back_vt [OF vt_e] 
-    show "vt step (moment i t @ s)" .
-  next
-    from eq_e and stp_i 
-    have "step (moment i t @ s) (Create th' prio)" by simp
-    thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
-  qed
-  with eq_e
-  have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
-    by (simp add:cntP_def cntV_def count_def)
-  with eq_me[symmetric]
-  have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
-    by simp
-  from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
-  with eq_me [symmetric]
-  have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
-  from moment_blocked [OF neq_th' h2 h1, of "length t"] and lt_its
-  and moment_ge
-  have "th' \<notin> runing (t @ s)" by auto
-  with runing'
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_2:
-  assumes runing': "th' \<in> runing (t@s)"
-  shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
-proof -
-  from runing_inversion_1[OF _ runing']
-  show ?thesis by auto
-qed
-
-lemma live: "runing (t@s) \<noteq> {}"
-proof(cases "th \<in> runing (t@s)")
-  case True thus ?thesis by auto
-next
-  case False
-  then have not_ready: "th \<notin> readys (t@s)"
-    apply (unfold runing_def, 
-            insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
-    by auto
-  from th_kept have "th \<in> threads (t@s)" by auto
-  from th_chain_to_ready[OF vt_t this] and not_ready
-  obtain th' where th'_in: "th' \<in> readys (t@s)"
-    and dp: "(Th th, Th th') \<in> (depend (t @ s))\<^sup>+" by auto
-  have "th' \<in> runing (t@s)"
-  proof -
-    have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
-    proof -
-      have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')) = 
-               preced th (t@s)"
-      proof(rule Max_eqI)
-        fix y
-        assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-        then obtain th1 where
-          h1: "th1 = th' \<or> th1 \<in>  dependents (wq (t @ s)) th'"
-          and eq_y: "y = preced th1 (t@s)" by auto
-        show "y \<le> preced th (t @ s)"
-        proof -
-          from max_preced
-          have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
-          moreover have "y \<le> \<dots>"
-          proof(rule Max_ge)
-            from h1
-            have "th1 \<in> threads (t@s)"
-            proof
-              assume "th1 = th'"
-              with th'_in show ?thesis by (simp add:readys_def)
-            next
-              assume "th1 \<in> dependents (wq (t @ s)) th'"
-              with dependents_threads [OF vt_t]
-              show "th1 \<in> threads (t @ s)" by auto
-            qed
-            with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
-          next
-            from finite_threads[OF vt_t]
-            show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th']
-        show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th'))"
-          by (auto intro:finite_subset)
-      next
-        from dp
-        have "th \<in> dependents (wq (t @ s)) th'" 
-          by (unfold cs_dependents_def, auto simp:eq_depend)
-        thus "preced th (t @ s) \<in> 
-                (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-          by auto
-      qed
-      moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
-      proof -
-        from max_preced and max_cp_eq[OF vt_t, symmetric]
-        have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
-        with max_cp_readys_threads[OF vt_t] show ?thesis by simp
-      qed
-      ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
-    qed
-    with th'_in show ?thesis by (auto simp:runing_def)
-  qed
-  thus ?thesis by auto
-qed
-
-end
-
-end
-
--- a/prio/Attic/ExtSG.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1019 +0,0 @@
-theory ExtSG
-imports PrioG
-begin
-
-locale highest_set =
-  fixes s' th prio fixes s 
-  defines s_def : "s \<equiv> (Set th prio#s')"
-  assumes vt_s: "vt step s"
-  and highest: "preced th s = Max ((cp s)`threads s)"
-
-context highest_set
-begin
-
-lemma vt_s': "vt step s'"
-  by (insert vt_s, unfold s_def, drule_tac step_back_vt, simp)
-
-lemma step_set: "step s' (Set th prio)"
-  by (insert vt_s, unfold s_def, drule_tac step_back_step, simp)
-
-lemma step_set_elim: 
-  "\<lbrakk>\<lbrakk>th \<in> runing s'\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-  by (insert step_set, ind_cases "step s' (Set th prio)", auto)
-
-
-lemma threads_s: "th \<in> threads s"
-  by (rule step_set_elim, unfold runing_def readys_def, auto simp:s_def)
-
-lemma same_depend: "depend s = depend s'"
-  by (insert depend_set_unchanged, unfold s_def, simp)
-
-lemma same_dependents:
-  "dependents (wq s) th = dependents (wq s') th"
-  apply (unfold cs_dependents_def)
-  by (unfold eq_depend same_depend, simp)
-
-lemma eq_cp_s_th: "cp s th = preced th s"
-proof -
-  from highest and max_cp_eq[OF vt_s]
-  have is_max: "preced th s = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-  have sbs: "({th} \<union> dependents (wq s) th) \<subseteq> threads s"
-  proof -
-    from threads_s and dependents_threads[OF vt_s, of th]
-    show ?thesis by auto
-  qed
-  show ?thesis
-  proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-    show "preced th s \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)" by simp
-  next
-    fix y 
-    assume "y \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)"
-    then obtain th1 where th1_in: "th1 \<in> ({th} \<union> dependents (wq s) th)"
-      and eq_y: "y = preced th1 s" by auto
-    show "y \<le> preced th s"
-    proof(unfold is_max, rule Max_ge)
-      from finite_threads[OF vt_s] 
-      show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-    next
-      from sbs th1_in and eq_y 
-      show "y \<in> (\<lambda>th. preced th s) ` threads s" by auto
-    qed
-  next
-    from sbs and finite_threads[OF vt_s]
-    show "finite ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))"
-      by (auto intro:finite_subset)
-  qed
-qed
-
-lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp)
-
-lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
-
-lemma is_ready: "th \<in> readys s"
-proof -
-  have "\<forall>cs. \<not> waiting s th cs"
-    apply (rule step_set_elim, unfold s_def, insert depend_set_unchanged[of th prio s'])
-    apply (unfold s_depend_def, unfold runing_def readys_def)
-    apply (auto, fold s_def)
-    apply (erule_tac x = cs in allE, auto simp:waiting_eq)
-  proof -
-    fix cs
-    assume h: 
-      "{(Th t, Cs c) |t c. waiting (wq s) t c} \<union> {(Cs c, Th t) |c t. holding (wq s) t c} =
-          {(Th t, Cs c) |t c. waiting (wq s') t c} \<union> {(Cs c, Th t) |c t. holding (wq s') t c}"
-            (is "?L = ?R")
-    and wt: "waiting (wq s) th cs" and nwt: "\<not> waiting (wq s') th cs"
-    from wt have "(Th th, Cs cs) \<in> ?L" by auto
-    with h have "(Th th, Cs cs) \<in> ?R" by simp
-    hence "waiting (wq s') th cs" by auto with nwt
-    show False by auto
-  qed    
-  with threads_s show ?thesis 
-    by (unfold readys_def, auto)
-qed
-
-lemma highest': "cp s th = Max (cp s ` threads s)"
-proof -
-  from highest_cp_preced max_cp_eq[OF vt_s, symmetric]
-  show ?thesis by simp
-qed
-
-lemma is_runing: "th \<in> runing s"
-proof -
-  have "Max (cp s ` threads s) = Max (cp s ` readys s)"
-  proof -
-    have " Max (cp s ` readys s) = cp s th"
-    proof(rule Max_eqI)
-      from finite_threads[OF vt_s] readys_threads finite_subset
-      have "finite (readys s)" by blast
-      thus "finite (cp s ` readys s)" by auto
-    next
-      from is_ready show "cp s th \<in> cp s ` readys s" by auto
-    next
-      fix y
-      assume "y \<in> cp s ` readys s"
-      then obtain th1 where 
-        eq_y: "y = cp s th1" and th1_in: "th1 \<in> readys s" by auto
-      show  "y \<le> cp s th" 
-      proof -
-        have "y \<le> Max (cp s ` threads s)"
-        proof(rule Max_ge)
-          from eq_y and th1_in
-          show "y \<in> cp s ` threads s"
-            by (auto simp:readys_def)
-        next
-          from finite_threads[OF vt_s]
-          show "finite (cp s ` threads s)" by auto
-        qed
-        with highest' show ?thesis by auto
-      qed
-    qed
-    with highest' show ?thesis by auto
-  qed
-  thus ?thesis
-    by (unfold runing_def, insert highest' is_ready, auto)
-qed
-
-end
-
-locale extend_highest_set = highest_set + 
-  fixes t 
-  assumes vt_t: "vt step (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"
-
-lemma step_back_vt_app: 
-  assumes vt_ts: "vt cs (t@s)" 
-  shows "vt cs 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 cs (t @ s) \<Longrightarrow> vt cs s"
-      and vt_et: "vt cs ((e # t) @ s)"
-    show ?case
-    proof(rule ih)
-      show "vt cs (t @ s)"
-      proof(rule step_back_vt)
-        from vt_et show "vt cs (e # t @ s)" by simp
-      qed
-    qed
-  qed
-qed
-
-context extend_highest_set
-begin
-
-lemma red_moment:
-  "extend_highest_set s' th prio (moment i t)"
-  apply (insert extend_highest_set_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
-  apply (unfold extend_highest_set_def extend_highest_set_axioms_def, clarsimp)
-  by (unfold highest_set_def, auto dest:step_back_vt_app)
-
-lemma ind [consumes 0, case_names Nil Cons, induct type]:
-  assumes 
-    h0: "R []"
-  and h2: "\<And> e t. \<lbrakk>vt step (t@s); step (t@s) e; 
-                    extend_highest_set s' th prio t; 
-                    extend_highest_set s' th prio (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
-  shows "R t"
-proof -
-  from vt_t extend_highest_set_axioms show ?thesis
-  proof(induct t)
-    from h0 show "R []" .
-  next
-    case (Cons e t')
-    assume ih: "\<lbrakk>vt step (t' @ s); extend_highest_set s' th prio t'\<rbrakk> \<Longrightarrow> R t'"
-      and vt_e: "vt step ((e # t') @ s)"
-      and et: "extend_highest_set s' th prio (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 step (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_set s' th prio t'"
-          by (unfold extend_highest_set_def extend_highest_set_axioms_def, auto dest:step_back_vt)
-      next
-        from vt_ts show "vt step (t' @ s)" .
-      qed
-    next
-      from et show "extend_highest_set s' th prio (e # t')" .
-    next
-      from et show ext': "extend_highest_set s' th prio t'"
-          by (unfold extend_highest_set_def extend_highest_set_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 "th \<in> threads ([] @ s) \<and> preced th ([] @ s) = preced th s"
-      by auto
-  next
-    case (Cons e t)
-    show ?case
-    proof(cases e)
-      case (Create thread prio)
-      assume eq_e: " e = Create thread prio"
-      show ?thesis
-      proof -
-        from Cons and eq_e have "step (t@s) (Create thread prio)" by auto
-        hence "th \<noteq> thread"
-        proof(cases)
-          assume "thread \<notin> threads (t @ s)"
-          with Cons show ?thesis by auto
-        qed
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    next
-      case (Exit thread)
-      assume eq_e: "e = Exit thread"
-      from Cons have "extend_highest_set s' th prio (e # t)" by auto
-      from extend_highest_set.exit_diff [OF this] and eq_e
-      have neq_th: "thread \<noteq> th" by auto
-      with Cons
-      show ?thesis
-        by (unfold eq_e, auto simp:preced_def)
-    next
-      case (P thread cs)
-      assume eq_e: "e = P thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (V thread cs)
-      assume eq_e: "e = V thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (Set thread prio')
-      assume eq_e: " e = Set thread prio'"
-      show ?thesis
-      proof -
-        from Cons have "extend_highest_set s' th prio (e # t)" by auto
-        from extend_highest_set.set_diff_low[OF this] and eq_e
-        have "th \<noteq> thread" by auto
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    qed
-  qed
-qed
-
-lemma max_kept: "Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s"
-proof(induct rule:ind)
-  case Nil
-  from highest_preced_thread
-  show "Max ((\<lambda>th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s"
-    by simp
-next
-  case (Cons e t)
-  show ?case
-  proof(cases e)
-    case (Create thread prio')
-    assume eq_e: " e = Create thread prio'"
-    from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto
-    hence neq_thread: "thread \<noteq> th"
-    proof(cases)
-      assume "thread \<notin> threads (t @ s)"
-      moreover have "th \<in> threads (t@s)"
-      proof -
-        from Cons have "extend_highest_set s' th prio t" by auto
-        from extend_highest_set.th_kept[OF this] show ?thesis by (simp add:s_def)
-      qed
-      ultimately show ?thesis by auto
-    qed
-    from Cons have "extend_highest_set s' th prio t" by auto
-    from extend_highest_set.th_kept[OF this]
-    have h': " th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" 
-      by (auto simp:s_def)
-    from stp
-    have thread_ts: "thread \<notin> threads (t @ s)"
-      by (cases, auto)
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))"
-        by (unfold eq_e, simp)
-      moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
-      proof(rule Max_insert)
-        from Cons have "vt step (t @ s)" by auto
-        from finite_threads[OF this]
-        show "finite (?f ` (threads (t@s)))" by simp
-      next
-        from h' show "(?f ` (threads (t@s))) \<noteq> {}" by auto
-      qed
-      moreover have "(Max (?f ` (threads (t@s)))) = ?t"
-      proof -
-        have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
-          (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B")
-        proof -
-          { fix th' 
-            assume "th' \<in> ?B"
-            with thread_ts eq_e
-            have "?f1 th' = ?f2 th'" by (auto simp:preced_def)
-          } thus ?thesis 
-            apply (auto simp:Image_def)
-          proof -
-            fix th'
-            assume h: "\<And>th'. th' \<in> threads (t @ s) \<Longrightarrow> 
-              preced th' (e # t @ s) = preced th' (t @ s)"
-              and h1: "th' \<in> threads (t @ s)"
-            show "preced th' (t @ s) \<in> (\<lambda>th'. preced th' (e # t @ s)) ` threads (t @ s)"
-            proof -
-              from h1 have "?f1 th' \<in> ?f1 ` ?B" by auto
-              moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp
-              ultimately show ?thesis by simp
-            qed
-          qed
-        qed
-        with Cons show ?thesis by auto
-      qed
-      moreover have "?f thread < ?t"
-      proof -
-        from Cons have " extend_highest_set s' th prio (e # t)" by auto
-        from extend_highest_set.create_low[OF this] and eq_e
-        have "prio' \<le> prio" by auto
-        thus ?thesis
-        by (unfold eq_e, auto simp:preced_def s_def precedence_less_def)
-    qed
-    ultimately show ?thesis by (auto simp:max_def)
-  qed
-next
-    case (Exit thread)
-    assume eq_e: "e = Exit thread"
-    from Cons have vt_e: "vt step (e#(t @ s))" by auto
-    from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto
-    from stp have thread_ts: "thread \<in> threads (t @ s)"
-      by(cases, unfold runing_def readys_def, auto)
-    from Cons have "extend_highest_set s' th prio (e # t)" by auto
-    from extend_highest_set.exit_diff[OF this] and eq_e
-    have neq_thread: "thread \<noteq> th" by auto
-    from Cons have "extend_highest_set s' th prio t" by auto
-    from extend_highest_set.th_kept[OF this, folded s_def]
-    have h': "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "threads (t@s) = insert thread ?A"
-        by (insert stp thread_ts, unfold eq_e, auto)
-      hence "Max (?f ` (threads (t@s))) = Max (?f ` \<dots>)" by simp
-      also from this have "\<dots> = Max (insert (?f thread) (?f ` ?A))" by simp
-      also have "\<dots> = max (?f thread) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        from finite_threads [OF vt_e]
-        show "finite (?f ` ?A)" by simp
-      next
-        from Cons have "extend_highest_set s' th prio (e # t)" by auto
-        from extend_highest_set.th_kept[OF this]
-        show "?f ` ?A \<noteq> {}" by  (auto simp:s_def)
-      qed
-      finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" .
-      moreover have "Max (?f ` (threads (t@s))) = ?t"
-      proof -
-        from Cons show ?thesis
-          by (unfold eq_e, auto simp:preced_def)
-      qed
-      ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp
-      moreover have "?f thread < ?t"
-      proof(unfold eq_e, simp add:preced_def, fold preced_def)
-        show "preced thread (t @ s) < ?t"
-        proof -
-          have "preced thread (t @ s) \<le> ?t" 
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" 
-              (is "?t = Max (?g ` ?B)") by simp
-            moreover have "?g thread \<le> \<dots>"
-            proof(rule Max_ge)
-              have "vt step (t@s)" by fact
-              from finite_threads [OF this]
-              show "finite (?g ` ?B)" by simp
-            next
-              from thread_ts
-              show "?g thread \<in> (?g ` ?B)" by auto
-            qed
-            ultimately show ?thesis by auto
-          qed
-          moreover have "preced thread (t @ s) \<noteq> ?t"
-          proof
-            assume "preced thread (t @ s) = preced th s"
-            with h' have "preced thread (t @ s) = preced th (t@s)" by simp
-            from preced_unique [OF this] have "thread = th"
-            proof
-              from h' show "th \<in> threads (t @ s)" by simp
-            next
-              from thread_ts show "thread \<in> threads (t @ s)" .
-            qed(simp)
-            with neq_thread show "False" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      qed
-      ultimately show ?thesis 
-        by (auto simp:max_def split:if_splits)
-    qed
-  next
-    case (P thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (V thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (Set thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      let ?B = "threads (t@s)"
-      from Cons have "extend_highest_set s' th prio (e # t)" by auto
-      from extend_highest_set.set_diff_low[OF this] and Set
-      have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
-      from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
-      also have "\<dots> = ?t"
-      proof(rule Max_eqI)
-        fix y
-        assume y_in: "y \<in> ?f ` ?B"
-        then obtain th1 where 
-          th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
-        show "y \<le> ?t"
-        proof(cases "th1 = thread")
-          case True
-          with neq_thread le_p eq_y s_def Set
-          show ?thesis
-            by (auto simp:preced_def precedence_le_def)
-        next
-          case False
-          with Set eq_y
-          have "y  = preced th1 (t@s)"
-            by (simp add:preced_def)
-          moreover have "\<dots> \<le> ?t"
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
-              by auto
-            moreover have "preced th1 (t@s) \<le> \<dots>"
-            proof(rule Max_ge)
-              from th1_in 
-              show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
-                by simp
-            next
-              show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-              proof -
-                from Cons have "vt step (t @ s)" by auto
-                from finite_threads[OF this] show ?thesis by auto
-              qed
-            qed
-            ultimately show ?thesis by auto
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from Cons and finite_threads
-        show "finite (?f ` ?B)" by auto
-      next
-        from Cons have "extend_highest_set s' th prio t" by auto
-        from extend_highest_set.th_kept [OF this, folded s_def]
-        have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-        show "?t \<in> (?f ` ?B)" 
-        proof -
-          from neq_thread Set h
-          have "?t = ?f th" by (auto simp:preced_def)
-          with h show ?thesis by auto
-        qed
-      qed
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
-  by (insert th_kept max_kept, auto)
-
-lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
-  (is "?L = ?R")
-proof -
-  have "?L = cpreced (t@s) (wq (t@s)) th" 
-    by (unfold cp_eq_cpreced, simp)
-  also have "\<dots> = ?R"
-  proof(unfold cpreced_def)
-    show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependents (wq (t @ s)) th)) =
-          Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-      (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
-    proof(cases "?A = {}")
-      case False
-      have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
-      moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        show "finite (?f ` ?A)"
-        proof -
-          from dependents_threads[OF vt_t]
-          have "?A \<subseteq> threads (t@s)" .
-          moreover from finite_threads[OF vt_t] have "finite \<dots>" .
-          ultimately show ?thesis 
-            by (auto simp:finite_subset)
-        qed
-      next
-        from False show "(?f ` ?A) \<noteq> {}" by simp
-      qed
-      moreover have "\<dots> = Max (?f ` ?B)"
-      proof -
-        from max_preced have "?f th = Max (?f ` ?B)" .
-        moreover have "Max (?f ` ?A) \<le> \<dots>" 
-        proof(rule Max_mono)
-          from False show "(?f ` ?A) \<noteq> {}" by simp
-        next
-          show "?f ` ?A \<subseteq> ?f ` ?B" 
-          proof -
-            have "?A \<subseteq> ?B" by (rule dependents_threads[OF vt_t])
-            thus ?thesis by auto
-          qed
-        next
-          from finite_threads[OF vt_t] 
-          show "finite (?f ` ?B)" by simp
-        qed
-        ultimately show ?thesis
-          by (auto simp:max_def)
-      qed
-      ultimately show ?thesis by auto
-    next
-      case True
-      with max_preced show ?thesis by auto
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
-  by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)
-
-lemma th_cp_preced: "cp (t@s) th = preced th s"
-  by (fold max_kept, unfold th_cp_max_preced, simp)
-
-lemma preced_less':
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  shows "preced th' s < preced th s"
-proof -
-  have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
-  proof(rule Max_ge)
-    from finite_threads [OF vt_s]
-    show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
-  next
-    from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
-      by simp
-  qed
-  moreover have "preced th' s \<noteq> preced th s"
-  proof
-    assume "preced th' s = preced th s"
-    from preced_unique[OF this th'_in] neq_th' is_ready
-    show "False" by  (auto simp:readys_def)
-  qed
-  ultimately show ?thesis using highest_preced_thread
-    by auto
-qed
-
-lemma pv_blocked:
-  fixes th'
-  assumes th'_in: "th' \<in> threads (t@s)"
-  and neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
-  shows "th' \<notin> runing (t@s)"
-proof
-  assume "th' \<in> runing (t@s)"
-  hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
-    by (auto simp:runing_def)
-  with max_cp_readys_threads [OF vt_t]
-  have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
-    by auto
-  moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
-  ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
-  moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
-    by simp
-  finally have h: "cp (t @ s) th' = preced th (t @ s)" .
-  show False
-  proof -
-    have "dependents (wq (t @ s)) th' = {}" 
-      by (rule count_eq_dependents [OF vt_t eq_pv])
-    moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
-    proof
-      assume "preced th' (t @ s) = preced th (t @ s)"
-      hence "th' = th"
-      proof(rule preced_unique)
-        from th_kept show "th \<in> threads (t @ s)" by simp
-      next
-        from th'_in show "th' \<in> threads (t @ s)" by simp
-      qed
-      with assms show False by simp
-    qed
-    ultimately show ?thesis
-      by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
-  qed
-qed
-
-lemma runing_precond_pre:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<in> threads (t@s) \<and>
-         cntP (t@s) th' = cntV (t@s) th'"
-proof -
-  show ?thesis
-  proof(induct rule:ind)
-    case (Cons e t)
-    from Cons
-    have in_thread: "th' \<in> threads (t @ s)"
-      and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    have "extend_highest_set s' th prio t" by fact
-    from extend_highest_set.pv_blocked 
-    [OF this, folded s_def, OF in_thread neq_th' not_holding]
-    have not_runing: "th' \<notin> runing (t @ s)" .
-    show ?case
-    proof(cases e)
-      case (V thread cs)
-      from Cons and V have vt_v: "vt step (V thread cs#(t@s))" by auto
-
-      show ?thesis
-      proof -
-        from Cons and V have "step (t@s) (V thread cs)" by auto
-        hence neq_th': "thread \<noteq> th'"
-        proof(cases)
-          assume "thread \<in> runing (t@s)"
-          moreover have "th' \<notin> runing (t@s)" by fact
-          ultimately show ?thesis by auto
-        qed
-        with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
-          by (unfold V, simp add:cntP_def cntV_def count_def)
-        moreover from in_thread
-        have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
-        ultimately show ?thesis by auto
-      qed
-    next
-      case (P thread cs)
-      from Cons and P have "step (t@s) (P thread cs)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t@s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Create thread prio')
-      from Cons and Create have "step (t@s) (Create thread prio')" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<notin> threads (t @ s)"
-        moreover have "th' \<in> threads (t@s)" by fact
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Create 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Create 
-      have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
-      ultimately show ?thesis by auto
-    next
-      case (Exit thread)
-      from Cons and Exit have "step (t@s) (Exit thread)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t @ s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Exit 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Exit and neq_th' 
-      have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio')
-      with Cons
-      show ?thesis 
-        by (auto simp:cntP_def cntV_def count_def)
-    qed
-  next
-    case Nil
-    with assms
-    show ?case by auto
-  qed
-qed
-
-(*
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<notin> runing (t@s)"
-proof -
-  from runing_precond_pre[OF th'_in eq_pv neq_th']
-  have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-  from pv_blocked[OF h1 neq_th' h2] 
-  show ?thesis .
-qed
-*)
-
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  and is_runing: "th' \<in> runing (t@s)"
-  shows "cntP s th' > cntV s th'"
-proof -
-  have "cntP s th' \<noteq> cntV s th'"
-  proof
-    assume eq_pv: "cntP s th' = cntV s th'"
-    from runing_precond_pre[OF th'_in eq_pv neq_th']
-    have h1: "th' \<in> threads (t @ s)"  
-      and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    from pv_blocked[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
-    with is_runing show "False" by simp
-  qed
-  moreover from cnp_cnv_cncs[OF vt_s, of th'] 
-  have "cntV s th' \<le> cntP s th'" by auto
-  ultimately show ?thesis by auto
-qed
-
-lemma moment_blocked_pre:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
-         th' \<in> threads ((moment (i+j) t)@s)"
-proof(induct j)
-  case (Suc k)
-  show ?case
-  proof -
-    { assume True: "Suc (i+k) \<le> length t"
-      from moment_head [OF this] 
-      obtain e where
-        eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
-        by blast
-      from red_moment[of "Suc(i+k)"]
-      and eq_me have "extend_highest_set s' th prio (e # moment (i + k) t)" by simp
-      hence vt_e: "vt step (e#(moment (i + k) t)@s)"
-        by (unfold extend_highest_set_def extend_highest_set_axioms_def 
-                          highest_set_def s_def, auto)
-      have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
-      proof(unfold s_def)
-        show "th' \<notin> runing (moment (i + k) t @ Set th prio # s')"
-        proof(rule extend_highest_set.pv_blocked)
-          from Suc show "th' \<in> threads (moment (i + k) t @ Set th prio # s')"
-            by (simp add:s_def)
-        next
-          from neq_th' show "th' \<noteq> th" .
-        next
-          from red_moment show "extend_highest_set s' th prio (moment (i + k) t)" .
-        next
-          from Suc show "cntP (moment (i + k) t @ Set th prio # s') th' =
-            cntV (moment (i + k) t @ Set th prio # s') th'"
-            by (auto simp:s_def)
-        qed
-      qed
-      from step_back_step[OF vt_e]
-      have "step ((moment (i + k) t)@s) e" .
-      hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
-        th' \<in> threads (e#(moment (i + k) t)@s)
-        "
-      proof(cases)
-        case (thread_create thread prio)
-        with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_exit thread)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_P thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_V thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_set thread prio')
-        with Suc show ?thesis
-          by (auto simp:cntP_def cntV_def count_def)
-      qed
-      with eq_me have ?thesis using eq_me by auto 
-    } note h = this
-    show ?thesis
-    proof(cases "Suc (i+k) \<le> length t")
-      case True
-      from h [OF this] show ?thesis .
-    next
-      case False
-      with moment_ge
-      have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
-      with Suc show ?thesis by auto
-    qed
-  qed
-next
-  case 0
-  from assms show ?case by auto
-qed
-
-lemma moment_blocked:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  and le_ij: "i \<le> j"
-  shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
-         th' \<in> threads ((moment j t)@s) \<and>
-         th' \<notin> runing ((moment j t)@s)"
-proof -
-  from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
-  have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
-    and h2: "th' \<in> threads ((moment j t)@s)" by auto
-  with extend_highest_set.pv_blocked [OF  red_moment [of j], folded s_def, OF h2 neq_th' h1]
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_1:
-  assumes neq_th': "th' \<noteq> th"
-  and runing': "th' \<in> runing (t@s)"
-  shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
-proof(cases "th' \<in> threads s")
-  case True
-  with runing_precond [OF this neq_th' runing'] show ?thesis by simp
-next
-  case False
-  let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
-  let ?q = "moment 0 t"
-  from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
-  from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
-  from p_split_gen [of ?Q, OF this not_thread]
-  obtain i where lt_its: "i < length t"
-    and le_i: "0 \<le> i"
-    and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
-    and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
-  from lt_its have "Suc i \<le> length t" by auto
-  from moment_head[OF this] obtain e where 
-   eq_me: "moment (Suc i) t = e # moment i t" by blast
-  from red_moment[of "Suc i"] and eq_me
-  have "extend_highest_set s' th prio (e # moment i t)" by simp
-  hence vt_e: "vt step (e#(moment i t)@s)"
-    by (unfold extend_highest_set_def extend_highest_set_axioms_def 
-      highest_set_def s_def, auto)
-  from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
-  from post[rule_format, of "Suc i"] and eq_me 
-  have not_in': "th' \<in> threads (e # moment i t@s)" by auto
-  from create_pre[OF stp_i pre this] 
-  obtain prio where eq_e: "e = Create th' prio" .
-  have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
-  proof(rule cnp_cnv_eq)
-    from step_back_vt [OF vt_e] 
-    show "vt step (moment i t @ s)" .
-  next
-    from eq_e and stp_i 
-    have "step (moment i t @ s) (Create th' prio)" by simp
-    thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
-  qed
-  with eq_e
-  have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
-    by (simp add:cntP_def cntV_def count_def)
-  with eq_me[symmetric]
-  have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
-    by simp
-  from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
-  with eq_me [symmetric]
-  have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
-  from moment_blocked [OF neq_th' h2 h1, of "length t"] and lt_its
-  and moment_ge
-  have "th' \<notin> runing (t @ s)" by auto
-  with runing'
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_2:
-  assumes runing': "th' \<in> runing (t@s)"
-  shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
-proof -
-  from runing_inversion_1[OF _ runing']
-  show ?thesis by auto
-qed
-
-lemma live: "runing (t@s) \<noteq> {}"
-proof(cases "th \<in> runing (t@s)")
-  case True thus ?thesis by auto
-next
-  case False
-  then have not_ready: "th \<notin> readys (t@s)"
-    apply (unfold runing_def, 
-            insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
-    by auto
-  from th_kept have "th \<in> threads (t@s)" by auto
-  from th_chain_to_ready[OF vt_t this] and not_ready
-  obtain th' where th'_in: "th' \<in> readys (t@s)"
-    and dp: "(Th th, Th th') \<in> (depend (t @ s))\<^sup>+" by auto
-  have "th' \<in> runing (t@s)"
-  proof -
-    have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
-    proof -
-      have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')) = 
-               preced th (t@s)"
-      proof(rule Max_eqI)
-        fix y
-        assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-        then obtain th1 where
-          h1: "th1 = th' \<or> th1 \<in>  dependents (wq (t @ s)) th'"
-          and eq_y: "y = preced th1 (t@s)" by auto
-        show "y \<le> preced th (t @ s)"
-        proof -
-          from max_preced
-          have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
-          moreover have "y \<le> \<dots>"
-          proof(rule Max_ge)
-            from h1
-            have "th1 \<in> threads (t@s)"
-            proof
-              assume "th1 = th'"
-              with th'_in show ?thesis by (simp add:readys_def)
-            next
-              assume "th1 \<in> dependents (wq (t @ s)) th'"
-              with dependents_threads [OF vt_t]
-              show "th1 \<in> threads (t @ s)" by auto
-            qed
-            with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
-          next
-            from finite_threads[OF vt_t]
-            show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th']
-        show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th'))"
-          by (auto intro:finite_subset)
-      next
-        from dp
-        have "th \<in> dependents (wq (t @ s)) th'" 
-          by (unfold cs_dependents_def, auto simp:eq_depend)
-        thus "preced th (t @ s) \<in> 
-                (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-          by auto
-      qed
-      moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
-      proof -
-        from max_preced and max_cp_eq[OF vt_t, symmetric]
-        have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
-        with max_cp_readys_threads[OF vt_t] show ?thesis by simp
-      qed
-      ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
-    qed
-    with th'_in show ?thesis by (auto simp:runing_def)
-  qed
-  thus ?thesis by auto
-qed
-
-end
-
-end
-
-
--- a/prio/Attic/Happen_within.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,126 +0,0 @@
-theory Happen_within
-imports Main Moment
-begin
-
-(* 
-  lemma 
-  fixes P :: "('a list) \<Rightarrow> bool"
-  and Q :: "('a list) \<Rightarrow> bool"
-  and k :: nat
-  and f :: "('a list) \<Rightarrow> nat"
-  assumes "\<And> s t. \<lbrakk>P s; \<not> Q s; P (t@s); k < length t\<rbrakk> \<Longrightarrow> f (t@s) < f s"
-  shows "\<And> s t. \<lbrakk> P s;  P(t @ s); f(s) * k < length t\<rbrakk> \<Longrightarrow> Q (t@s)"
-  sorry
-*)
-
-text {* 
-  The following two notions are introduced to improve the situation.
-  *}
-
-definition all_future :: "(('a list) \<Rightarrow> bool) \<Rightarrow> (('a list) \<Rightarrow> bool) \<Rightarrow> ('a list) \<Rightarrow> bool"
-where "all_future G R s = (\<forall> t. G (t@s) \<longrightarrow> R t)"
-
-definition happen_within :: "(('a list) \<Rightarrow> bool) \<Rightarrow> (('a list) \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> ('a list) \<Rightarrow> bool"
-where "happen_within G R k s = all_future G (\<lambda> t. k < length t \<longrightarrow> 
-                                  (\<exists> i \<le> k. R (moment i t @ s) \<and> G (moment i t @ s))) s"
-
-lemma happen_within_intro:
-  fixes P :: "('a list) \<Rightarrow> bool"
-  and Q :: "('a list) \<Rightarrow> bool"
-  and k :: nat
-  and f :: "('a list) \<Rightarrow> nat"
-  assumes 
-  lt_k: "0 < k"
-  and step: "\<And> s. \<lbrakk>P s; \<not> Q s\<rbrakk> \<Longrightarrow> happen_within P (\<lambda> s'. f s' < f s) k s"
-  shows "\<And> s. P s \<Longrightarrow> happen_within P Q ((f s + 1) * k) s"
-proof -
-  fix s
-  assume "P s"
-  thus "happen_within P Q ((f s + 1) * k) s"
-  proof(induct n == "f s + 1" arbitrary:s rule:nat_less_induct)
-    fix s
-    assume ih [rule_format]: "\<forall>m<f s + 1. \<forall>x. m = f x + 1 \<longrightarrow> P x 
-                                 \<longrightarrow> happen_within P Q ((f x + 1) * k) x"
-      and ps: "P s"
-    show "happen_within P Q ((f s + 1) * k) s"
-    proof(cases "Q s")
-      case True
-      show ?thesis 
-      proof -
-        { fix t
-          from True and ps have "0 \<le> ((f s + 1)*k) \<and> Q (moment 0 t @ s) \<and> P (moment 0 t @ s)" by auto
-          hence "\<exists>i\<le>(f s + 1) * k. Q (moment i t @ s) \<and> P (moment i t @ s)" by auto
-        } thus ?thesis by (auto simp: happen_within_def all_future_def)
-      qed
-    next
-      case False
-      from step [OF ps False] have kk: "happen_within P (\<lambda>s'. f s' < f s) k s" .
-      show ?thesis
-      proof -
-        { fix t
-          assume pts: "P (t @ s)" and ltk: "(f s + 1) * k < length t"
-          from ltk have lt_k_lt: "k < length t" by auto
-          with kk pts obtain i 
-            where le_ik: "i \<le> k" 
-            and lt_f: "f (moment i t @ s) < f s" 
-            and p_m: "P (moment i t @ s)"
-            by (auto simp:happen_within_def all_future_def)
-          from ih [of "f (moment i t @ s) + 1" "(moment i t @ s)", OF _ _ p_m] and lt_f
-          have hw: "happen_within P Q ((f (moment i t @ s) + 1) * k) (moment i t @ s)" by auto
-          have "(\<exists>j\<le>(f s + 1) * k. Q (moment j t @ s) \<and>  P (moment j t @ s))" (is "\<exists> j. ?T j")
-          proof -
-            let ?t = "restm i t"
-            have eq_t: "t = ?t @ moment i t" by (simp add:moment_restm_s) 
-            have h1: "P (restm i t @ moment i t @ s)" 
-            proof -
-              from pts and eq_t have "P ((restm i t @ moment i t) @ s)" by simp
-              thus ?thesis by simp
-            qed
-            moreover have h2: "(f (moment i t @ s) + 1) * k < length (restm i t)"
-            proof -
-              have h: "\<And> x y z. (x::nat) \<le> y \<Longrightarrow> x * z \<le> y * z" by simp
-              from lt_f have "(f (moment i t @ s) + 1) \<le> f s " by simp
-              from h [OF this, of k]
-              have "(f (moment i t @ s) + 1) * k \<le> f s * k" .
-              moreover from le_ik have "\<dots> \<le> ((f s) * k + k - i)" by simp
-              moreover from le_ik lt_k_lt and ltk have "(f s) * k + k - i < length t - i" by simp
-              moreover have "length (restm i t) = length t - i" using length_restm by metis
-              ultimately show ?thesis by simp
-            qed
-            from hw [unfolded happen_within_def all_future_def, rule_format, OF h1 h2]
-            obtain m where le_m: "m \<le> (f (moment i t @ s) + 1) * k"
-              and q_m: "Q (moment m ?t @ moment i t @ s)" 
-              and p_m: "P (moment m ?t @ moment i t @ s)" by auto
-            have eq_mm: "moment m ?t @ moment i t @ s = (moment (m+i) t)@s"
-            proof -
-              have "moment m (restm i t) @ moment i t = moment (m + i) t"
-                using moment_plus_split by metis
-              thus ?thesis by simp
-            qed
-            let ?j = "m + i"
-            have "?T ?j"
-            proof -
-              have "m + i \<le> (f s + 1) * k"
-              proof -
-                have h: "\<And> x y z. (x::nat) \<le> y \<Longrightarrow> x * z \<le> y * z" by simp
-                from lt_f have "(f (moment i t @ s) + 1) \<le> f s " by simp
-                from h [OF this, of k]
-                have "(f (moment i t @ s) + 1) * k \<le> f s * k" .
-                with le_m have "m \<le> f s * k" by simp
-                hence "m + i \<le> f s * k + i" by simp
-                with le_ik show ?thesis by simp
-              qed
-              moreover from eq_mm q_m have " Q (moment (m + i) t @ s)" by metis
-              moreover from eq_mm p_m have " P (moment (m + i) t @ s)" by metis
-              ultimately show ?thesis by blast
-            qed
-            thus ?thesis by blast
-          qed
-        } thus ?thesis by  (simp add:happen_within_def all_future_def firstn.simps)
-      qed
-    qed
-  qed
-qed
-
-end
-
--- a/prio/Attic/Lsp.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,323 +0,0 @@
-theory Lsp
-imports Main
-begin
-
-fun lsp :: "('a \<Rightarrow> ('b::linorder)) \<Rightarrow> 'a list \<Rightarrow> ('a list \<times> 'a list \<times> 'a list)"
-where 
-   "lsp f [] = ([], [], [])" |
-   "lsp f [x] = ([], [x], [])" |
-   "lsp f (x#xs) = (case (lsp f xs) of
-                     (l, [], r) \<Rightarrow> ([], [x], []) |
-                     (l, y#ys, r) \<Rightarrow> if f x \<ge> f y then ([], [x], xs) else (x#l, y#ys, r))"
-
-inductive lsp_p :: "('a \<Rightarrow> ('b::linorder)) \<Rightarrow> 'a list \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
-for f :: "('a \<Rightarrow> ('b::linorder))"
-where
-  lsp_nil [intro]: "lsp_p f [] ([], [], [])" |
-  lsp_single [intro]: "lsp_p f [x] ([], [x], [])" |
-  lsp_cons_1 [intro]: "\<lbrakk>xs \<noteq> []; lsp_p f xs (l, [m], r); f x \<ge> f m\<rbrakk> \<Longrightarrow> lsp_p f (x#xs) ([], [x], xs)" |
-  lsp_cons_2 [intro]: "\<lbrakk>xs \<noteq> []; lsp_p f xs (l, [m], r); f x < f m\<rbrakk> \<Longrightarrow> lsp_p f (x#xs) (x#l, [m], r)"
-
-lemma lsp_p_lsp_1: "lsp_p f x y \<Longrightarrow> y = lsp f x"
-proof (induct rule:lsp_p.induct)
-  case (lsp_cons_1 xs  l m r x)
-  assume lsp_xs [symmetric]: "(l, [m], r) = lsp f xs"
-    and le_mx: "f m \<le> f x"
-  show ?case (is "?L = ?R")
-  proof(cases xs, simp)
-    case (Cons v vs)
-    show ?thesis
-      apply (simp add:Cons)
-      apply (fold Cons)
-      by (simp add:lsp_xs le_mx)
-  qed
-next
-  case (lsp_cons_2 xs l m r x)
-  assume lsp_xs [symmetric]: "(l, [m], r) = lsp f xs"
-    and lt_xm: "f x < f m"
-  show ?case (is "?L = ?R")
-  proof(cases xs)
-    case (Cons v vs)
-    show ?thesis
-      apply (simp add:Cons)
-      apply (fold Cons)
-      apply (simp add:lsp_xs)
-      by (insert lt_xm, auto)
-  next
-    case Nil
-    from prems show ?thesis by simp
-  qed
-qed auto
-
-lemma lsp_mid_nil: "lsp f xs = (a, [], c) \<Longrightarrow> xs = []"
-  apply (induct xs arbitrary:a c, auto)
-  apply (case_tac xs, auto)
-  by (case_tac "(lsp f (ab # list))", auto split:if_splits list.splits)
-
-
-lemma lsp_mid_length: "lsp f x = (u, v, w) \<Longrightarrow> length v \<le> 1"
-proof(induct x arbitrary:u v w, simp)
-  case (Cons x xs)
-  assume ih: "\<And> u v w. lsp f xs = (u, v, w) \<Longrightarrow> length v \<le> 1"
-  and h: "lsp f (x # xs) = (u, v, w)"
-  show "length v \<le> 1" using h
-  proof(cases xs, simp add:h)
-    case (Cons z zs)
-    assume eq_xs: "xs = z # zs"
-    show ?thesis
-    proof(cases "lsp f xs")
-      fix l m r
-      assume eq_lsp: "lsp f xs = (l, m, r)"
-      show ?thesis
-      proof(cases m)
-        case Nil
-        from Nil and eq_lsp have "lsp f xs = (l, [], r)" by simp
-        from lsp_mid_nil [OF this] have "xs = []" .
-        with h show ?thesis by auto
-      next
-        case (Cons y ys)
-        assume eq_m: "m = y # ys"
-        from ih [OF eq_lsp] have eq_xs_1: "length m \<le> 1" .
-        show ?thesis
-        proof(cases "f x \<ge> f y")
-          case True
-          from eq_xs eq_xs_1 True h eq_lsp show ?thesis 
-            by (auto split:list.splits if_splits)
-        next
-          case False
-          from eq_xs eq_xs_1 False h eq_lsp show ?thesis 
-             by (auto split:list.splits if_splits)
-        qed
-      qed
-    qed
-  next
-    assume "[] = u \<and> [x] = v \<and> [] = w"
-    hence "v = [x]" by simp
-    thus "length v \<le> Suc 0" by simp
-  qed
-qed
-
-lemma lsp_p_lsp_2: "lsp_p f x (lsp f x)"
-proof(induct x, auto)
-  case (Cons x xs)
-  assume ih: "lsp_p f xs (lsp f xs)"
-  show ?case
-  proof(cases xs)
-    case Nil
-    thus ?thesis by auto
-  next
-    case (Cons v vs)
-    show ?thesis
-    proof(cases "xs")
-      case Nil
-      thus ?thesis by auto
-    next
-      case (Cons v vs)
-      assume eq_xs: "xs = v # vs"
-      show ?thesis
-      proof(cases "lsp f xs")
-        fix l m r
-        assume eq_lsp_xs: "lsp f xs = (l, m, r)"
-        show ?thesis
-        proof(cases m)
-          case Nil
-          from eq_lsp_xs and Nil have "lsp f xs = (l, [], r)" by simp
-          from lsp_mid_nil [OF this] have eq_xs: "xs = []" .
-          hence "lsp f (x#xs) = ([], [x], [])" by simp
-          with eq_xs show ?thesis by auto
-        next
-          case (Cons y ys)
-          assume eq_m: "m = y # ys"
-          show ?thesis
-          proof(cases "f x \<ge> f y")
-            case True
-            from eq_xs eq_lsp_xs Cons True
-            have eq_lsp: "lsp f (x#xs) = ([], [x], v # vs)" by simp
-            show ?thesis
-            proof (simp add:eq_lsp)
-              show "lsp_p f (x # xs) ([], [x], v # vs)"
-              proof(fold eq_xs, rule lsp_cons_1 [OF _])
-                from eq_xs show "xs \<noteq> []" by simp
-              next
-                from lsp_mid_length [OF eq_lsp_xs] and Cons
-                have "m = [y]" by simp
-                with eq_lsp_xs have "lsp f xs = (l, [y], r)" by simp
-                with ih show "lsp_p f xs (l, [y], r)" by simp
-              next
-                from True show "f y \<le> f x" by simp
-              qed
-            qed
-          next
-            case False
-            from eq_xs eq_lsp_xs Cons False
-            have eq_lsp: "lsp f (x#xs) = (x # l, y # ys, r) " by simp
-            show ?thesis
-            proof (simp add:eq_lsp)
-              from lsp_mid_length [OF eq_lsp_xs] and eq_m
-              have "ys = []" by simp
-              moreover have "lsp_p f (x # xs) (x # l, [y], r)"
-              proof(rule lsp_cons_2)
-                from eq_xs show "xs \<noteq> []" by simp
-              next
-                from lsp_mid_length [OF eq_lsp_xs] and Cons
-                have "m = [y]" by simp
-                with eq_lsp_xs have "lsp f xs = (l, [y], r)" by simp
-                with ih show "lsp_p f xs (l, [y], r)" by simp
-              next
-                from False show "f x < f y" by simp
-              qed
-              ultimately show "lsp_p f (x # xs) (x # l, y # ys, r)" by simp
-            qed
-          qed
-        qed
-      qed
-    qed
-  qed
-qed
-
-lemma lsp_induct:
-  fixes f x1 x2 P
-  assumes h: "lsp f x1 = x2"
-  and p1: "P [] ([], [], [])"
-  and p2: "\<And>x. P [x] ([], [x], [])"
-  and p3: "\<And>xs l m r x. \<lbrakk>xs \<noteq> []; lsp f xs = (l, [m], r); P xs (l, [m], r); f m \<le> f x\<rbrakk> \<Longrightarrow> P (x # xs) ([], [x], xs)"
-  and p4: "\<And>xs l m r x. \<lbrakk>xs \<noteq> []; lsp f xs = (l, [m], r); P xs (l, [m], r); f x < f m\<rbrakk> \<Longrightarrow> P (x # xs) (x # l, [m], r)"
-  shows "P x1 x2"
-proof(rule lsp_p.induct)
-  from lsp_p_lsp_2 and h
-  show "lsp_p f x1 x2" by metis
-next
-  from p1 show "P [] ([], [], [])" by metis
-next
-  from p2 show "\<And>x. P [x] ([], [x], [])" by metis
-next
-  fix xs l m r x 
-  assume h1: "xs \<noteq> []" and h2: "lsp_p f xs (l, [m], r)" and h3: "P xs (l, [m], r)" and h4: "f m \<le> f x"
-  show "P (x # xs) ([], [x], xs)" 
-  proof(rule p3 [OF h1 _ h3 h4])
-    from h2 and lsp_p_lsp_1 show "lsp f xs = (l, [m], r)" by metis
-  qed
-next
-  fix xs l m r x 
-  assume h1: "xs \<noteq> []" and h2: "lsp_p f xs (l, [m], r)" and h3: "P xs (l, [m], r)" and h4: "f x < f m"
-  show "P (x # xs) (x # l, [m], r)"
-  proof(rule p4 [OF h1 _ h3 h4])
-    from h2 and lsp_p_lsp_1 show "lsp f xs = (l, [m], r)" by metis
-  qed
-qed
-
-lemma lsp_set_eq: 
-  fixes f x u v w
-  assumes h: "lsp f x = (u, v, w)"
-  shows "x = u@v@w"
-proof -
-  have "\<And> f x r. lsp f x = r \<Longrightarrow> \<forall> u v w. (r = (u, v, w) \<longrightarrow> x = u@v@w)" 
-    by (erule lsp_induct, simp+)
-  from this [rule_format, OF h] show ?thesis by simp
-qed
-
-lemma lsp_set: 
-  assumes h: "(u, v, w) = lsp f x"
-  shows "set (u@v@w) = set x"
-proof -
-  from lsp_set_eq [OF h[symmetric]] 
-  show ?thesis by simp
-qed
-
-lemma max_insert_gt:
-  fixes S fx
-  assumes h: "fx < Max S"
-  and np: "S \<noteq> {}"
-  and fn: "finite S" 
-  shows "Max S = Max (insert fx S)"
-proof -
-  from Max_insert [OF fn np]
-  have "Max (insert fx S) = max fx (Max S)" .
-  moreover have "\<dots> = Max S"
-  proof(cases "fx \<le> Max S")
-    case False
-    with h
-    show ?thesis by (simp add:max_def)
-  next
-    case True
-    thus ?thesis by (simp add:max_def)
-  qed
-  ultimately show ?thesis by simp
-qed
-
-lemma max_insert_le: 
-  fixes S fx
-  assumes h: "Max S \<le> fx"
-  and fn: "finite S"
-  shows "fx = Max (insert fx S)"
-proof(cases "S = {}")
-  case True
-  thus ?thesis by simp
-next
-  case False
-  from Max_insert [OF fn False]
-  have "Max (insert fx S) = max fx (Max S)" .
-  moreover have "\<dots> = fx"
-  proof(cases "fx \<le> Max S")
-    case False
-    thus ?thesis by (simp add:max_def)
-  next
-    case True
-    have hh: "\<And> x y. \<lbrakk> x \<le> (y::('a::linorder)); y \<le> x\<rbrakk> \<Longrightarrow> x = y" by auto
-    from hh [OF True h]
-    have "fx = Max S" .
-    thus ?thesis by simp
-  qed
-  ultimately show ?thesis by simp
-qed
-  
-lemma lsp_max: 
-  fixes f x u m w
-  assumes h: "lsp f x = (u, [m], w)"
-  shows "f m = Max (f ` (set x))"
-proof -
-  { fix y
-    have "lsp f x = y \<Longrightarrow> \<forall> u m w. y = (u, [m], w) \<longrightarrow> f m = Max (f ` (set x))"
-    proof(erule lsp_induct, simp)
-      { fix x u m w
-        assume "(([]::'a list), ([x]::'a list), ([]::'a list)) = (u, [m], w)"
-        hence "f m = Max (f ` set [x])"  by simp
-      } thus "\<And>x. \<forall>u m w. ([], [x], []) = (u, [m], w) \<longrightarrow> f m = Max (f ` set [x])" by simp
-    next
-      fix xs l m r x
-      assume h1: "xs \<noteq> []"
-        and h2: " lsp f xs = (l, [m], r)"
-        and h3: "\<forall>u ma w. (l, [m], r) = (u, [ma], w) \<longrightarrow> f ma = Max (f ` set xs)"
-        and h4: "f m \<le> f x"
-      show " \<forall>u m w. ([], [x], xs) = (u, [m], w) \<longrightarrow> f m = Max (f ` set (x # xs))"
-      proof -
-        have "f x = Max (f ` set (x # xs))"
-        proof -
-          from h2 h3 have "f m = Max (f ` set xs)" by simp
-          with h4 show ?thesis
-            apply auto
-            by (rule_tac max_insert_le, auto)
-        qed
-        thus ?thesis by simp
-      qed
-    next
-      fix xs l m r x
-      assume h1: "xs \<noteq> []"
-        and h2: " lsp f xs = (l, [m], r)"
-        and h3: " \<forall>u ma w. (l, [m], r) = (u, [ma], w) \<longrightarrow> f ma = Max (f ` set xs)"
-        and h4: "f x < f m"
-      show "\<forall>u ma w. (x # l, [m], r) = (u, [ma], w) \<longrightarrow> f ma = Max (f ` set (x # xs))"
-      proof -
-        from h2 h3 have "f m = Max (f ` set xs)" by simp
-        with h4
-        have "f m =  Max (f ` set (x # xs))"
-          apply auto
-          apply (rule_tac max_insert_gt, simp+)
-          by (insert h1, simp+)
-        thus ?thesis by auto
-      qed
-    qed
-  } with h show ?thesis by metis
-qed
-
-end
--- a/prio/Attic/Prio.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2813 +0,0 @@
-theory Prio
-imports Precedence_ord Moment Lsp Happen_within
-begin
-
-type_synonym thread = nat
-type_synonym priority = nat
-type_synonym cs = nat
-
-datatype event = 
-  Create thread priority |
-  Exit thread |
-  P thread cs |
-  V thread cs |
-  Set thread priority
-
-datatype node = 
-   Th "thread" |
-   Cs "cs"
-
-type_synonym state = "event list"
-
-fun threads :: "state \<Rightarrow> thread set"
-where 
-  "threads [] = {}" |
-  "threads (Create thread prio#s) = {thread} \<union> threads s" |
-  "threads (Exit thread # s) = (threads s) - {thread}" |
-  "threads (e#s) = threads s"
-
-fun original_priority :: "thread \<Rightarrow> state \<Rightarrow> nat"
-where
-  "original_priority thread [] = 0" |
-  "original_priority thread (Create thread' prio#s) = 
-     (if thread' = thread then prio else original_priority thread s)" |
-  "original_priority thread (Set thread' prio#s) = 
-     (if thread' = thread then prio else original_priority thread s)" |
-  "original_priority thread (e#s) = original_priority thread s"
-
-fun birthtime :: "thread \<Rightarrow> state \<Rightarrow> nat"
-where
-  "birthtime thread [] = 0" |
-  "birthtime thread ((Create thread' prio)#s) = (if (thread = thread') then length s 
-                                                  else birthtime thread s)" |
-  "birthtime thread ((Set thread' prio)#s) = (if (thread = thread') then length s 
-                                                  else birthtime thread s)" |
-  "birthtime thread (e#s) = birthtime thread s"
-
-definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
-  where "preced thread s = Prc (original_priority thread s) (birthtime thread s)"
-
-consts holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
-       waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
-       depend :: "'b \<Rightarrow> (node \<times> node) set"
-       dependents :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
-
-defs (overloaded) cs_holding_def: "holding wq thread cs == (thread \<in> set (wq cs) \<and> thread = hd (wq cs))"
-                  cs_waiting_def: "waiting wq thread cs == (thread \<in> set (wq cs) \<and> thread \<noteq> hd (wq cs))"
-                  cs_depend_def: "depend (wq::cs \<Rightarrow> thread list) == {(Th t, Cs c) | t c. waiting wq t c} \<union> 
-                                               {(Cs c, Th t) | c t. holding wq t c}"
-                  cs_dependents_def: "dependents (wq::cs \<Rightarrow> thread list) th == {th' . (Th th', Th th) \<in> (depend wq)^+}"
-
-record schedule_state = 
-    waiting_queue :: "cs \<Rightarrow> thread list"
-    cur_preced :: "thread \<Rightarrow> precedence"
-
-
-definition cpreced :: "state \<Rightarrow> (cs \<Rightarrow> thread list) \<Rightarrow> thread \<Rightarrow> precedence"
-where "cpreced s wq = (\<lambda> th. Max ((\<lambda> th. preced th s) ` ({th} \<union> dependents wq th)))"
-
-fun schs :: "state \<Rightarrow> schedule_state"
-where
-   "schs [] = \<lparr>waiting_queue = \<lambda> cs. [], 
-               cur_preced = cpreced [] (\<lambda> cs. [])\<rparr>" |
-   "schs (e#s) = (let ps = schs s in
-                  let pwq = waiting_queue ps in
-                  let pcp = cur_preced ps in
-                  let nwq = case e of
-                             P thread cs \<Rightarrow>  pwq(cs:=(pwq cs @ [thread])) |
-                             V thread cs \<Rightarrow> let nq = case (pwq cs) of
-                                                      [] \<Rightarrow> [] | 
-                                                      (th#pq) \<Rightarrow> case (lsp pcp pq) of
-                                                                   (l, [], r) \<Rightarrow> []
-                                                                 | (l, m#ms, r) \<Rightarrow> m#(l@ms@r)
-                                            in pwq(cs:=nq)                 |
-                              _ \<Rightarrow> pwq
-                  in let ncp = cpreced (e#s) nwq in 
-                     \<lparr>waiting_queue = nwq, cur_preced = ncp\<rparr>
-                 )"
-
-definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list" 
-where "wq s == waiting_queue (schs s)"
-
-definition cp :: "state \<Rightarrow> thread \<Rightarrow> precedence"
-where "cp s = cur_preced (schs s)"
-
-defs (overloaded) s_holding_def: "holding (s::state) thread cs == (thread \<in> set (wq s cs) \<and> thread = hd (wq s cs))"
-                  s_waiting_def: "waiting (s::state) thread cs == (thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs))"
-                  s_depend_def: "depend (s::state) == {(Th t, Cs c) | t c. waiting (wq s) t c} \<union> 
-                                               {(Cs c, Th t) | c t. holding (wq s) t c}"
-                  s_dependents_def: "dependents (s::state) th == {th' . (Th th', Th th) \<in> (depend (wq s))^+}"
-
-definition readys :: "state \<Rightarrow> thread set"
-where
-  "readys s = 
-     {thread . thread \<in> threads s \<and> (\<forall> cs. \<not> waiting s thread cs)}"
-
-definition runing :: "state \<Rightarrow> thread set"
-where "runing s = {th . th \<in> readys s \<and> cp s th = Max ((cp s) ` (readys s))}"
-
-definition holdents :: "state \<Rightarrow> thread \<Rightarrow> cs set"
-  where "holdents s th = {cs . (Cs cs, Th th) \<in> depend s}"
-
-inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
-where
-  thread_create: "\<lbrakk>prio \<le> max_prio; thread \<notin> threads s\<rbrakk> \<Longrightarrow> step s (Create thread prio)" |
-  thread_exit: "\<lbrakk>thread \<in> runing s; holdents s thread = {}\<rbrakk> \<Longrightarrow> step s (Exit thread)" |
-  thread_P: "\<lbrakk>thread \<in> runing s;  (Cs cs, Th thread)  \<notin> (depend s)^+\<rbrakk> \<Longrightarrow> step s (P thread cs)" |
-  thread_V: "\<lbrakk>thread \<in> runing s;  holding s thread cs\<rbrakk> \<Longrightarrow> step s (V thread cs)" |
-  thread_set: "\<lbrakk>thread \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
-
-inductive vt :: "(state \<Rightarrow> event \<Rightarrow> bool) \<Rightarrow> state \<Rightarrow> bool"
- for cs
-where
-  vt_nil[intro]: "vt cs []" |
-  vt_cons[intro]: "\<lbrakk>vt cs s; cs s e\<rbrakk> \<Longrightarrow> vt cs (e#s)"
-
-lemma runing_ready: "runing s \<subseteq> readys s"
-  by (auto simp only:runing_def readys_def)
-
-lemma wq_v_eq_nil: 
-  fixes s cs thread rest
-  assumes eq_wq: "wq s cs = thread # rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [], r)"
-  shows "wq (V thread cs#s) cs = []"
-proof -
-  from prems show ?thesis
-    by (auto simp:wq_def Let_def cp_def split:list.splits)
-qed
-
-lemma wq_v_eq: 
-  fixes s cs thread rest
-  assumes eq_wq: "wq s cs = thread # rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
-  shows "wq (V thread cs#s) cs = th'#l@r"
-proof -
-  from prems show ?thesis
-    by (auto simp:wq_def Let_def cp_def split:list.splits)
-qed
-
-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 wq_distinct: "vt step s \<Longrightarrow> distinct (wq s cs)"
-proof(erule_tac vt.induct, simp add:wq_def)
-  fix s e
-  assume h1: "step s e"
-  and h2: "distinct (wq s cs)"
-  thus "distinct (wq (e # s) cs)"
-  proof(induct rule:step.induct, auto simp: wq_def Let_def split:list.splits)
-    fix thread s
-    assume h1: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
-      and h2: "thread \<in> set (waiting_queue (schs s) cs)"
-      and h3: "thread \<in> runing s"
-    show "False" 
-    proof -
-      from h3 have "\<And> cs. thread \<in>  set (waiting_queue (schs s) cs) \<Longrightarrow> 
-                             thread = hd ((waiting_queue (schs s) cs))" 
-        by (simp add:runing_def readys_def s_waiting_def wq_def)
-      from this [OF h2] have "thread = hd (waiting_queue (schs s) cs)" .
-      with h2
-      have "(Cs cs, Th thread) \<in> (depend s)"
-        by (simp add:s_depend_def s_holding_def wq_def cs_holding_def)
-      with h1 show False by auto
-    qed
-  next
-    fix thread s a list
-    assume h1: "thread \<in> runing s" 
-      and h2: "holding s thread cs"
-      and h3: "waiting_queue (schs s) cs = a # list"
-      and h4: "a \<notin> set list"
-      and h5: "distinct list"
-    thus "distinct
-           ((\<lambda>(l, a, r). case a of [] \<Rightarrow> [] | m # ms \<Rightarrow> m # l @ ms @ r)
-             (lsp (cur_preced (schs s)) list))"
-    apply (cases "(lsp (cur_preced (schs s)) list)", simp)
-    apply (case_tac b, simp)
-    by (drule_tac lsp_set_eq, simp)
-  qed
-qed
-
-lemma block_pre: 
-  fixes thread cs s
-  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 -
-  have ee: "\<And> x y. \<lbrakk>x = y\<rbrakk> \<Longrightarrow> set x = set y"
-    by auto
-  from s_ni s_i show ?thesis
-  proof (cases e, auto split:if_splits simp add:Let_def wq_def)
-    fix uu uub uuc uud uue
-    assume h: "(uuc, thread # uu, uub) = lsp (cur_preced (schs s)) uud"
-      and h1 [symmetric]: "uue # uud = waiting_queue (schs s) cs"
-      and h2: "thread \<notin> set (waiting_queue (schs s) cs)"
-    from lsp_set [OF h] have "set (uuc @ (thread # uu) @ uub) = set uud" .
-    hence "thread \<in> set uud" by auto
-    with h1 have "thread \<in> set (waiting_queue (schs s) cs)" by auto
-    with h2 show False by auto
-  next
-    fix uu uua uub uuc uud uue
-    assume h1: "thread \<notin> set (waiting_queue (schs s) cs)"
-      and h2: "uue # uud = waiting_queue (schs s) cs"
-      and h3: "(uuc, uua # uu, uub) = lsp (cur_preced (schs s)) uud"
-      and h4: "thread \<in> set uuc"
-    from lsp_set [OF h3] have "set (uuc @ (uua # uu) @ uub) = set uud" .
-    with h4 have "thread \<in> set uud" by auto
-    with h2 have "thread \<in> set (waiting_queue (schs s) cs)" 
-      apply (drule_tac ee) by auto
-    with h1 show "False" by fast
-  next
-    fix uu uua uub uuc uud uue
-    assume h1: "thread \<notin> set (waiting_queue (schs s) cs)"
-      and h2: "uue # uud = waiting_queue (schs s) cs"
-      and h3: "(uuc, uua # uu, uub) = lsp (cur_preced (schs s)) uud"
-      and h4: "thread \<in> set uu"
-    from lsp_set [OF h3] have "set (uuc @ (uua # uu) @ uub) = set uud" .
-    with h4 have "thread \<in> set uud" by auto
-    with h2 have "thread \<in> set (waiting_queue (schs s) cs)" 
-      apply (drule_tac ee) by auto
-    with h1 show "False" by fast
-  next
-    fix uu uua uub uuc uud uue
-    assume h1: "thread \<notin> set (waiting_queue (schs s) cs)"
-      and h2: "uue # uud = waiting_queue (schs s) cs"
-      and h3: "(uuc, uua # uu, uub) = lsp (cur_preced (schs s)) uud"
-      and h4: "thread \<in> set uub"
-    from lsp_set [OF h3] have "set (uuc @ (uua # uu) @ uub) = set uud" .
-    with h4 have "thread \<in> set uud" by auto
-    with h2 have "thread \<in> set (waiting_queue (schs s) cs)" 
-      apply (drule_tac ee) by auto
-    with h1 show "False" by fast
-  qed
-qed
-
-lemma p_pre: "\<lbrakk>vt step ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
-  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (depend s)^+"
-apply (ind_cases "vt step ((P thread cs)#s)")
-apply (ind_cases "step s (P thread cs)")
-by auto
-
-lemma abs1:
-  fixes e es
-  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 cs (a#s)"
-
-lemma abs2:
-  assumes vt: "vt step (e#s)"
-  and 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 -
-  have ee: "\<And> uuc thread uu uub s list. (uuc, thread # uu, uub) = lsp (cur_preced (schs s)) list \<Longrightarrow> 
-                 lsp (cur_preced (schs s)) list = (uuc, thread # uu, uub) 
-               " by simp
-  from prems show "False"
-    apply (cases e)
-    apply ((simp split:if_splits add:Let_def wq_def)[1])+
-    apply (insert abs1, fast)[1] 
-    apply ((simp split:if_splits add:Let_def)[1])+
-    apply (simp split:if_splits list.splits add:Let_def wq_def) 
-    apply (auto dest!:ee)
-    apply (drule_tac lsp_set_eq, simp)
-    apply (subgoal_tac "distinct (waiting_queue (schs s) cs)", simp, fold wq_def)
-    apply (rule_tac wq_distinct, auto)
-    apply (erule_tac evt_cons, auto)
-    apply (drule_tac lsp_set_eq, simp)
-    apply (subgoal_tac "distinct (wq s cs)", simp)
-    apply (rule_tac wq_distinct, auto)
-    apply (erule_tac evt_cons, auto)
-    apply (drule_tac lsp_set_eq, simp)
-    apply (subgoal_tac "distinct (wq s cs)", simp)
-    apply (rule_tac wq_distinct, auto)
-    apply (erule_tac evt_cons, auto)
-    apply (auto simp:wq_def Let_def split:if_splits prod.splits)
-    done
-qed
-
-lemma vt_moment: "\<And> t. \<lbrakk>vt cs s; t \<le> length s\<rbrakk> \<Longrightarrow> vt cs (moment t s)"
-proof(induct s, simp)
-  fix a s t
-  assume h: "\<And>t.\<lbrakk>vt cs s; t \<le> length s\<rbrakk> \<Longrightarrow> vt cs (moment t s)"
-    and vt_a: "vt cs (a # s)"
-    and le_t: "t \<le> length (a # s)"
-  show "vt cs (moment t (a # s))"
-  proof(cases "t = length (a#s)")
-    case True
-    from True have "moment t (a#s) = a#s" by simp
-    with vt_a show ?thesis by simp
-  next
-    case False
-    with le_t have le_t1: "t \<le> length s" by simp
-    from vt_a have "vt cs s"
-      by (erule_tac evt_cons, simp)
-    from h [OF this le_t1] have "vt cs (moment t s)" .
-    moreover have "moment t (a#s) = moment t s"
-    proof -
-      from moment_app [OF le_t1, of "[a]"] 
-      show ?thesis by simp
-    qed
-    ultimately show ?thesis by auto
-  qed
-qed
-
-(* Wrong:
-    lemma \<lbrakk>thread \<in> set (waiting_queue cs1 s); thread \<in> set (waiting_queue cs2 s)\<rbrakk> \<Longrightarrow> cs1 = cs2"
-*)
-
-lemma waiting_unique_pre:
-  fixes cs1 cs2 s thread
-  assumes vt: "vt step s"
-  and 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: "vt step (e#moment t2 s)"
-      proof -
-        from vt_moment [OF vt le_t3]
-        have "vt step (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      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 abs2 [OF vt_e True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from block_pre [OF False h1]
-        have "e = P thread cs2" .
-        with vt_e have "vt step ((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 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: "vt step (e#moment t1 s)"
-      proof -
-        from vt_moment [OF vt le_t3]
-        have "vt step (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      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 abs2 [OF vt_e True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from block_pre [OF False h1]
-        have "e = P thread cs1" .
-        with vt_e have "vt step ((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 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 step (e#moment t1 s)"
-      proof -
-        from vt_moment [OF vt le_t3]
-        have "vt step (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      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 abs2 [OF vt_e True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from 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 step (e#moment t2 s)" by simp 
-          from abs2 [OF this True eq_th h2 h1]
-          show ?thesis .
-        next
-          case False
-          from 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
-
-lemma waiting_unique:
-  assumes "vt step s"
-  and "waiting s th cs1"
-  and "waiting s th cs2"
-  shows "cs1 = cs2"
-proof -
-  from waiting_unique_pre and prems
-  show ?thesis
-    by (auto simp add:s_waiting_def)
-qed
-
-lemma holded_unique:
-  assumes "vt step s"
-  and "holding s th1 cs"
-  and "holding s th2 cs"
-  shows "th1 = th2"
-proof -
-  from prems show ?thesis
-    unfolding s_holding_def
-    by auto
-qed
-
-lemma birthtime_lt: "th \<in> threads s \<Longrightarrow> birthtime th s < length s"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits)
-
-lemma birthtime_unique: 
-  "\<lbrakk>birthtime th1 s = birthtime 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:birthtime_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 "birthtime th1 s = birthtime th2 s" by (simp add:preced_def)
-  from birthtime_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
-
-lemma unique_minus:
-  fixes x y z r
-  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:
-  fixes r x y z
-  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:
-  fixes r x y z
-  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
-
-lemma depend_set_unchanged: "(depend (Set th prio # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-lemma depend_create_unchanged: "(depend (Create th prio # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-lemma depend_exit_unchanged: "(depend (Exit th # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-definition head_of :: "('a \<Rightarrow> 'b::linorder) \<Rightarrow> 'a set \<Rightarrow> 'a set"
-  where "head_of f A = {a . a \<in> A \<and> f a = Max (f ` A)}"
-
-definition wq_head :: "state \<Rightarrow> cs \<Rightarrow> thread set"
-  where "wq_head s cs = head_of (cp s) (set (wq s cs))"
-
-lemma f_nil_simp: "\<lbrakk>f cs = []\<rbrakk> \<Longrightarrow> f(cs:=[]) = f"
-proof
-  fix x
-  assume h:"f cs = []"
-  show "(f(cs := [])) x = f x"
-  proof(cases "cs = x")
-    case True
-    with h show ?thesis by simp
-  next
-    case False
-    with h show ?thesis by simp
-  qed
-qed
-
-lemma step_back_vt: "vt ccs (e#s) \<Longrightarrow> vt ccs s"
-  by(ind_cases "vt ccs (e#s)", simp)
-
-lemma step_back_step: "vt ccs (e#s) \<Longrightarrow> ccs s e"
-  by(ind_cases "vt ccs (e#s)", simp)
-
-lemma holding_nil:
-    "\<lbrakk>wq s cs = []; holding (wq s) th cs\<rbrakk> \<Longrightarrow> False"
-  by (unfold cs_holding_def, auto)
-
-lemma waiting_kept_1: "
-       \<lbrakk>vt step (V th cs#s); wq s cs = a # list; waiting ((wq s)(cs := ab # aa @ lista @ ca)) t c;
-        lsp (cp s) list = (aa, ab # lista, ca)\<rbrakk>
-       \<Longrightarrow> waiting (wq s) t c"
-  apply (drule_tac step_back_vt, drule_tac wq_distinct[of _ cs])
-  apply(drule_tac lsp_set_eq)
-  by (unfold cs_waiting_def, auto split:if_splits)
- 
-lemma waiting_kept_2: 
-  "\<And>a list t c aa ca.
-       \<lbrakk>wq s cs = a # list; waiting ((wq s)(cs := [])) t c; lsp (cp s) list = (aa, [], ca)\<rbrakk>
-       \<Longrightarrow> waiting (wq s) t c"
-  apply(drule_tac lsp_set_eq)
-  by (unfold cs_waiting_def, auto split:if_splits)
-  
-
-lemma holding_nil_simp: "\<lbrakk>holding ((wq s)(cs := [])) t c\<rbrakk> \<Longrightarrow> holding (wq s) t c"
-  by(unfold cs_holding_def, auto)
-
-lemma step_wq_elim: "\<lbrakk>vt step (V th cs#s); wq s cs = a # list; a \<noteq> th\<rbrakk> \<Longrightarrow> False"
-  apply(drule_tac step_back_step)
-  apply(ind_cases "step s (V th cs)")
-  by(unfold s_holding_def, auto)
-
-lemma holding_cs_neq_simp: "c \<noteq> cs \<Longrightarrow> holding ((wq s)(cs := u)) t c = holding (wq s) t c"
-  by (unfold cs_holding_def, auto)
-
-lemma holding_th_neq_elim:
-  "\<And>a list c t aa ca ab lista.
-       \<lbrakk>\<not> holding (wq s) t c; holding ((wq s)(cs := ab # aa @ lista @ ca)) t c;
-         ab \<noteq> t\<rbrakk>
-       \<Longrightarrow> False"
-  by (unfold cs_holding_def, auto split:if_splits)
-
-lemma holding_nil_abs:
-  "\<not> holding ((wq s)(cs := [])) th cs"
-  by (unfold cs_holding_def, auto split:if_splits)
-
-lemma holding_abs: "\<lbrakk>holding ((wq s)(cs := ab # aa @ lista @ c)) th cs; ab \<noteq> th\<rbrakk>
-       \<Longrightarrow> False"
-    by (unfold cs_holding_def, auto split:if_splits)
-
-lemma waiting_abs: "\<not> waiting ((wq s)(cs := t # l @ r)) t cs"
-    by (unfold cs_waiting_def, auto split:if_splits)
-
-lemma waiting_abs_1: 
-  "\<lbrakk>\<not> waiting ((wq s)(cs := [])) t c; waiting (wq s) t c; c \<noteq> cs\<rbrakk>
-       \<Longrightarrow> False"
-    by (unfold cs_waiting_def, auto split:if_splits)
-
-lemma waiting_abs_2: "
-       \<lbrakk>\<not> waiting ((wq s)(cs := ab # aa @ lista @ ca)) t c; waiting (wq s) t c;
-        c \<noteq> cs\<rbrakk>
-       \<Longrightarrow> False"
-  by (unfold cs_waiting_def, auto split:if_splits)
-
-lemma waiting_abs_3:
-     "\<lbrakk>wq s cs = a # list; \<not> waiting ((wq s)(cs := [])) t c;
-        waiting (wq s) t c; lsp (cp s) list = (aa, [], ca)\<rbrakk>
-       \<Longrightarrow> False"
-  apply(drule_tac lsp_mid_nil, simp)
-  by(unfold cs_waiting_def, auto split:if_splits)
-
-lemma waiting_simp: "c \<noteq> cs \<Longrightarrow> waiting ((wq s)(cs:=z)) t c = waiting (wq s) t c"
-   by(unfold cs_waiting_def, auto split:if_splits)
-
-lemma holding_cs_eq:
-  "\<lbrakk>\<not> holding ((wq s)(cs := [])) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> c = cs"
-   by(unfold cs_holding_def, auto split:if_splits)
-
-lemma holding_cs_eq_1:
-  "\<lbrakk>\<not> holding ((wq s)(cs := ab # aa @ lista @ ca)) t c; holding (wq s) t c\<rbrakk>
-       \<Longrightarrow> c = cs"
-   by(unfold cs_holding_def, auto split:if_splits)
-
-lemma holding_th_eq: 
-       "\<lbrakk>vt step (V th cs#s); wq s cs = a # list; \<not> holding ((wq s)(cs := [])) t c; holding (wq s) t c;
-        lsp (cp s) list = (aa, [], ca)\<rbrakk>
-       \<Longrightarrow> t = th"
-  apply(drule_tac lsp_mid_nil, simp)
-  apply(unfold cs_holding_def, auto split:if_splits)
-  apply(drule_tac step_back_step)
-  apply(ind_cases "step s (V th cs)")
-  by (unfold s_holding_def, auto split:if_splits)
-
-lemma holding_th_eq_1:
-  "\<lbrakk>vt step (V th cs#s); 
-     wq s cs = a # list; \<not> holding ((wq s)(cs := ab # aa @ lista @ ca)) t c; holding (wq s) t c;
-        lsp (cp s) list = (aa, ab # lista, ca)\<rbrakk>
-       \<Longrightarrow> t = th"
-  apply(drule_tac step_back_step)
-  apply(ind_cases "step s (V th cs)")
-  apply(unfold s_holding_def cs_holding_def)
-  by (auto split:if_splits)
-
-lemma holding_th_eq_2: "\<lbrakk>holding ((wq s)(cs := ac # x)) th cs\<rbrakk>
-       \<Longrightarrow> ac = th"
-  by (unfold cs_holding_def, auto)
-
-lemma holding_th_eq_3: "
-       \<lbrakk>\<not> holding (wq s) t c;
-        holding ((wq s)(cs := ac # x)) t c\<rbrakk>
-       \<Longrightarrow> ac = t"
-  by (unfold cs_holding_def, auto)
-
-lemma holding_wq_eq: "holding ((wq s)(cs := th' # l @ r)) th' cs"
-   by (unfold cs_holding_def, auto)
-
-lemma waiting_th_eq: "
-       \<lbrakk>waiting (wq s) t c; wq s cs = a # list;
-        lsp (cp s) list = (aa, ac # lista, ba); \<not> waiting ((wq s)(cs := ac # aa @ lista @ ba)) t c\<rbrakk>
-       \<Longrightarrow> ac = t"
-  apply(drule_tac lsp_set_eq, simp)
-  by (unfold cs_waiting_def, auto split:if_splits)
-
-lemma step_depend_v:
-  "vt step (V th cs#s) \<Longrightarrow>
-  depend (V th cs # s) =
-  depend s - {(Cs cs, Th th)} -
-  {(Th th', Cs cs) |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))} \<union>
-  {(Cs cs, Th th') |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))}"
-  apply (unfold s_depend_def wq_def, 
-         auto split:list.splits simp:Let_def f_nil_simp holding_wq_eq, fold wq_def cp_def)
-  apply (auto split:list.splits prod.splits  
-               simp:Let_def f_nil_simp holding_nil_simp holding_cs_neq_simp holding_nil_abs
-                    waiting_abs waiting_simp holding_wq_eq
-               elim:holding_nil waiting_kept_1 waiting_kept_2 step_wq_elim holding_th_neq_elim 
-               holding_abs waiting_abs_1 waiting_abs_3 holding_cs_eq holding_cs_eq_1
-               holding_th_eq holding_th_eq_1 holding_th_eq_2 holding_th_eq_3 waiting_th_eq
-               dest:lsp_mid_length)
-  done
-
-lemma step_depend_p:
-  "vt step (P th cs#s) \<Longrightarrow>
-  depend (P th cs # s) =  (if (wq s cs = []) then depend s \<union> {(Cs cs, Th th)}
-                                             else depend s \<union> {(Th th, Cs cs)})"
-  apply(unfold s_depend_def wq_def)
-  apply (auto split:list.splits prod.splits simp:Let_def cs_waiting_def cs_holding_def)
-  apply(case_tac "c = cs", auto)
-  apply(fold wq_def)
-  apply(drule_tac step_back_step)
-  by (ind_cases " step s (P (hd (wq s cs)) cs)", 
-    auto simp:s_depend_def wq_def cs_holding_def)
-
-lemma simple_A:
-  fixes A
-  assumes h: "\<And> x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x = y"
-  shows "A = {} \<or> (\<exists> a. A = {a})"
-proof(cases "A = {}")
-  case True thus ?thesis by simp
-next
-  case False then obtain a where "a \<in> A" by auto
-  with h have "A = {a}" by auto
-  thus ?thesis by simp
-qed
-
-lemma depend_target_th: "(Th th, x) \<in> depend (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
-  by (unfold s_depend_def, auto)
-
-lemma acyclic_depend: 
-  fixes s
-  assumes vt: "vt step s"
-  shows "acyclic (depend s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    assume ih: "acyclic (depend s)"
-      and stp: "step s e"
-      and vt: "vt step s"
-    show ?case
-    proof(cases e)
-      case (Create th prio)
-      with ih
-      show ?thesis by (simp add:depend_create_unchanged)
-    next
-      case (Exit th)
-      with ih show ?thesis by (simp add:depend_exit_unchanged)
-    next
-      case (V th cs)
-      from V vt stp have vtt: "vt step (V th cs#s)" by auto
-      from step_depend_v [OF this]
-      have eq_de: "depend (e # s) = 
-        depend s - {(Cs cs, Th th)} -
-        {(Th th', Cs cs) |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))} \<union>
-        {(Cs cs, Th th') |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))}"
-        (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
-      from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
-      have "?D = {} \<or> (\<exists> a. ?D = {a})" by (rule simple_A, auto)
-      thus ?thesis
-      proof(cases "wq s cs")
-        case Nil
-        hence "?D = {}" by simp
-        with ac and eq_de show ?thesis by simp
-      next
-        case (Cons tth rest)
-        from stp and V have "step s (V th cs)" by simp
-        hence eq_wq: "wq s cs = th # rest"
-        proof -
-          show "step s (V th cs) \<Longrightarrow> wq s cs = th # rest"
-            apply(ind_cases "step s (V th cs)")
-            by(insert Cons, unfold s_holding_def, simp)
-        qed
-        show ?thesis
-        proof(cases "lsp (cp s) rest")
-          fix l b r
-          assume eq_lsp: "lsp (cp s) rest = (l, b, r) "
-          show ?thesis
-          proof(cases "b")
-            case Nil
-            with eq_lsp and eq_wq have "?D = {}" by simp
-            with ac and eq_de show ?thesis by simp
-          next
-            case (Cons th' m)
-            with eq_lsp 
-            have eq_lsp: "lsp (cp s) rest = (l, [th'], r)" 
-              apply simp
-              by (drule_tac lsp_mid_length, simp)
-            from eq_wq and eq_lsp
-            have eq_D: "?D = {(Cs cs, Th th')}" by auto
-            from eq_wq and eq_lsp
-            have eq_C: "?C = {(Th th', Cs cs)}" by 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 depend_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_depend_def s_waiting_def cs_waiting_def by simp
-              hence "cs' = cs"
-              proof(rule waiting_unique [OF vt])
-                from eq_wq eq_lsp wq_distinct[OF vt, of cs]
-                show "waiting s th' cs" by(unfold s_waiting_def, auto dest:lsp_set_eq)
-              qed
-              with th'_e eq_x have "(Th th', Cs cs) \<in> ?E" by simp
-              with eq_C show "False" by simp
-            qed
-            with acyclic_insert[symmetric] and ac and eq_D
-            and eq_de show ?thesis by simp
-          qed 
-        qed
-      qed
-    next
-      case (P th cs)
-      from P vt stp have vtt: "vt step (P th cs#s)" by auto
-      from step_depend_p [OF this] P
-      have "depend (e # s) = 
-              (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
-                                    depend 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 =  depend s \<union> {(Cs cs, Th th)}" by simp
-        have "(Th th, Cs cs) \<notin> (depend s)\<^sup>*"
-        proof
-          assume "(Th th, Cs cs) \<in> (depend s)\<^sup>*"
-          hence "(Th th, Cs cs) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-          from tranclD2 [OF this]
-          obtain x where "(x, Cs cs) \<in> depend s" by auto
-          with True show False by (auto simp:s_depend_def cs_waiting_def)
-        qed
-        with acyclic_insert ih eq_r show ?thesis by auto
-      next
-        case False
-        hence eq_r: "?R =  depend s \<union> {(Th th, Cs cs)}" by simp
-        have "(Cs cs, Th th) \<notin> (depend s)\<^sup>*"
-        proof
-          assume "(Cs cs, Th th) \<in> (depend s)\<^sup>*"
-          hence "(Cs cs, Th th) \<in> (depend 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> (depend 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 depend_set_unchanged
-      show ?thesis by (simp add:depend_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "acyclic (depend ([]::state))"
-      by (auto simp: s_depend_def cs_waiting_def 
-                      cs_holding_def wq_def acyclic_def)
-  qed
-qed
-
-lemma finite_depend: 
-  fixes s
-  assumes vt: "vt step s"
-  shows "finite (depend s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    assume ih: "finite (depend s)"
-      and stp: "step s e"
-      and vt: "vt step s"
-    show ?case
-    proof(cases e)
-      case (Create th prio)
-      with ih
-      show ?thesis by (simp add:depend_create_unchanged)
-    next
-      case (Exit th)
-      with ih show ?thesis by (simp add:depend_exit_unchanged)
-    next
-      case (V th cs)
-      from V vt stp have vtt: "vt step (V th cs#s)" by auto
-      from step_depend_v [OF this]
-      have eq_de: "depend (e # s) = 
-        depend s - {(Cs cs, Th th)} -
-        {(Th th', Cs cs) |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))} \<union>
-        {(Cs cs, Th th') |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))}"
-        (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 (rule simple_A, auto)
-        thus ?thesis
-        proof
-          assume h: "?D = {}"
-          show ?thesis by (unfold h, simp)
-        next
-          assume "\<exists> a. ?D = {a}"
-          thus ?thesis by auto
-        qed
-      qed
-      ultimately show ?thesis by simp
-    next
-      case (P th cs)
-      from P vt stp have vtt: "vt step (P th cs#s)" by auto
-      from step_depend_p [OF this] P
-      have "depend (e # s) = 
-              (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
-                                    depend 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 =  depend s \<union> {(Cs cs, Th th)}" by simp
-        with True and ih show ?thesis by auto
-      next
-        case False
-        hence "?R = depend 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:depend_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "finite (depend ([]::state))"
-      by (auto simp: s_depend_def cs_waiting_def 
-                   cs_holding_def wq_def acyclic_def)
-  qed
-qed
-
-text {* Several useful lemmas *}
-
-thm wf_trancl
-thm finite_acyclic_wf
-thm finite_acyclic_wf_converse
-thm wf_induct
-
-
-lemma wf_dep_converse: 
-  fixes s
-  assumes vt: "vt step s"
-  shows "wf ((depend s)^-1)"
-proof(rule finite_acyclic_wf_converse)
-  from finite_depend [OF vt]
-  show "finite (depend s)" .
-next
-  from acyclic_depend[OF vt]
-  show "acyclic (depend s)" .
-qed
-
-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> depend (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> depend s"
-  by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-
-lemma wq_threads: 
-  fixes s cs
-  assumes vt: "vt step s"
-  and 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)
-    assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
-      and stp: "step s e"
-      and vt: "vt step 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_depend_def s_holding_def cs_holding_def)
-        by (fold wq_def, auto)
-    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 (waiting_queue (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 = waiting_queue (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 "waiting_queue (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: "waiting_queue (schs s) cs' = a # rest"
-              with h V show ?thesis
-              proof(cases "(lsp (cur_preced (schs s)) rest)", unfold V)
-                fix l m r
-                assume eq_lsp: "lsp (cur_preced (schs s)) rest = (l, m, r)"
-                  and eq_wq: "waiting_queue (schs s) cs' = a # rest"
-                  and th_in_set: "th \<in> set (wq (V th' cs' # s) cs)"
-                show ?thesis
-                proof(cases "m")
-                  case Nil
-                  with eq_lsp have "rest = []" using lsp_mid_nil by auto
-                  with eq_wq have "waiting_queue (schs s) cs' = [a]" by simp
-                  with h[unfolded V wq_def] True 
-                  show ?thesis
-                    by (simp add:Let_def)
-                next
-                  case (Cons b rb)
-                  with lsp_mid_length[OF eq_lsp] have eq_m: "m = [b]" by auto
-                  with eq_lsp have "lsp (cur_preced (schs s)) rest = (l, [b], r)" by simp
-                  with h[unfolded V wq_def] True lsp_set_eq [OF this] eq_wq
-                  show ?thesis
-                    apply (auto simp:Let_def, fold wq_def)
-                    by (rule_tac ih [of _ cs'], auto)+
-                qed
-              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>vt step s; (Th th) \<in> Range (depend (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
-  apply(unfold s_depend_def cs_waiting_def cs_holding_def)
-  by (auto intro:wq_threads)
-
-lemma readys_v_eq:
-  fixes th thread cs rest
-  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 prems show ?thesis
-    apply (auto simp:readys_def)
-    apply (case_tac "cs = csa", simp add:s_waiting_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)
-    by (auto simp:s_waiting_def wq_def Let_def split:list.splits prod.splits 
-            dest:lsp_set_eq)
-qed
-
-lemma readys_v_eq_1:
-  fixes th thread cs rest
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
-  and neq_th': "th \<noteq> th'"
-  shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
-proof -
-  from prems show ?thesis
-    apply (auto simp:readys_def)
-    apply (case_tac "cs = csa", simp add:s_waiting_def)
-    apply (erule_tac x = cs in allE)
-    apply (simp add:s_waiting_def wq_def Let_def split:prod.splits list.splits)
-    apply (drule_tac lsp_mid_nil,simp, clarify, fold cp_def, clarsimp)
-    apply (frule_tac lsp_set_eq, simp)
-    apply (erule_tac x = csa in allE)
-    apply (subst (asm) (2) s_waiting_def, unfold wq_def)
-    apply (auto simp:Let_def split:list.splits prod.splits if_splits 
-            dest:lsp_set_eq)
-    apply (unfold s_waiting_def)
-    apply (fold wq_def, clarsimp)
-    apply (clarsimp)+
-    apply (case_tac "csa = cs", simp)
-    apply (erule_tac x = cs in allE, simp)
-    apply (unfold wq_def)
-    by (auto simp:Let_def split:list.splits prod.splits if_splits 
-            dest:lsp_set_eq)
-qed
-
-lemma readys_v_eq_2:
-  fixes th thread cs rest
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
-  and neq_th': "th = th'"
-  and vt: "vt step s"
-  shows "(th \<in> readys (V thread cs#s))"
-proof -
-  from prems show ?thesis
-    apply (auto simp:readys_def)
-    apply (rule_tac wq_threads [of s _ cs], auto dest:lsp_set_eq)
-    apply (unfold s_waiting_def wq_def)
-    apply (auto simp:Let_def split:list.splits prod.splits if_splits 
-            dest:lsp_set_eq lsp_mid_nil lsp_mid_length)
-    apply (fold cp_def, simp+, clarsimp)
-    apply (frule_tac lsp_set_eq, simp)
-    apply (fold wq_def)
-    apply (subgoal_tac "csa = cs", simp)
-    apply (rule_tac waiting_unique [of s th'], simp)
-    by (auto simp:s_waiting_def)
-qed
-
-lemma chain_building:
-  assumes vt: "vt step s"
-  shows "node \<in> Domain (depend s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (depend s)^+)"
-proof -
-  from wf_dep_converse [OF vt]
-  have h: "wf ((depend s)\<inverse>)" .
-  show ?thesis
-  proof(induct rule:wf_induct [OF h])
-    fix x
-    assume ih [rule_format]: 
-      "\<forall>y. (y, x) \<in> (depend s)\<inverse> \<longrightarrow> 
-           y \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (depend s)\<^sup>+)"
-    show "x \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+)"
-    proof
-      assume x_d: "x \<in> Domain (depend s)"
-      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+"
-      proof(cases x)
-        case (Th th)
-        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> depend s" by (auto simp:s_depend_def)
-        with Th have x_in_r: "(Cs cs, x) \<in> (depend s)^-1" by simp
-        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> depend s" by blast
-        hence "Cs cs \<in> Domain (depend 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> (depend s)\<^sup>+" by auto
-        have "(x, Th th') \<in> (depend 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> (depend s)^-1" by (auto simp:s_depend_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 [OF vt] have "th' \<in> threads s" by auto
-          with False have "Th th' \<in> Domain (depend s)" 
-            by (auto simp:readys_def s_waiting_def s_depend_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> (depend s)\<^sup>+" by auto
-          from th'_d and th''_in 
-          have "(x, Th th'') \<in> (depend s)\<^sup>+" by auto
-          with th''_r show ?thesis by auto
-        qed
-      qed
-    qed
-  qed
-qed
-
-lemma th_chain_to_ready:
-  fixes s th
-  assumes vt: "vt step s"
-  and th_in: "th \<in> threads s"
-  shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (depend 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 (depend s)" 
-    by (auto simp:readys_def s_waiting_def s_depend_def cs_waiting_def Domain_def)
-  from chain_building [rule_format, OF vt this]
-  show ?thesis by auto
-qed
-
-lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
-  by  (unfold s_waiting_def cs_waiting_def, auto)
-
-lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
-  by (unfold s_holding_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)
-
-lemma unique_depend: "\<lbrakk>vt step s; (n, n1) \<in> depend s; (n, n2) \<in> depend s\<rbrakk> \<Longrightarrow> n1 = n2"
-  apply(unfold s_depend_def, auto, fold waiting_eq holding_eq)
-  by(auto elim:waiting_unique holding_unique)
-
-lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
-by (induct rule:trancl_induct, auto)
-
-lemma dchain_unique:
-  assumes vt: "vt step s"
-  and th1_d: "(n, Th th1) \<in> (depend s)^+"
-  and th1_r: "th1 \<in> readys s"
-  and th2_d: "(n, Th th2) \<in> (depend 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_depend [OF vt]
-    have "(Th th1, Th th2) \<in> (depend s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (depend s)\<^sup>+" by auto
-    hence "False"
-    proof
-      assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th1, n) \<in> depend s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_depend_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_depend_def s_waiting_def cs_waiting_def)
-      with th1_r show ?thesis by auto
-    next
-      assume "(Th th2, Th th1) \<in> (depend s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th2, n) \<in> depend s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_depend_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 s_depend_def s_waiting_def cs_waiting_def)
-      with th2_r show ?thesis by auto
-    qed
-  } thus ?thesis by auto
-qed
-             
-definition count :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat"
-where "count Q l = length (filter Q l)"
-
-definition cntP :: "state \<Rightarrow> thread \<Rightarrow> nat"
-where "cntP s th = count (\<lambda> e. \<exists> cs. e = P th cs) s"
-
-definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
-where "cntV s th = count (\<lambda> e. \<exists> cs. e = V th cs) s"
-
-
-lemma step_holdents_p_add:
-  fixes th cs s
-  assumes vt: "vt step (P th cs#s)"
-  and "wq s cs = []"
-  shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
-proof -
-  from prems show ?thesis
-  unfolding  holdents_def step_depend_p[OF vt] by auto
-qed
-
-lemma step_holdents_p_eq:
-  fixes th cs s
-  assumes vt: "vt step (P th cs#s)"
-  and "wq s cs \<noteq> []"
-  shows "holdents (P th cs#s) th = holdents s th"
-proof -
-  from prems show ?thesis
-  unfolding  holdents_def step_depend_p[OF vt] by auto
-qed
-
-lemma step_holdents_v_minus:
-  fixes th cs s
-  assumes vt: "vt step (V th cs#s)"
-  shows "holdents (V th cs#s) th = holdents s th - {cs}"
-proof -
-  { fix rest l r
-    assume eq_wq: "wq s cs = th # rest" 
-      and eq_lsp: "lsp (cp s) rest = (l, [th], r)"
-    have "False" 
-    proof -
-      from lsp_set_eq [OF eq_lsp] have " rest = l @ [th] @ r" .
-      with eq_wq have "wq s cs = th#\<dots>" by simp
-      with wq_distinct [OF step_back_vt[OF vt], of cs]
-      show ?thesis by auto
-    qed
-  } thus ?thesis unfolding holdents_def step_depend_v[OF vt] by auto
-qed
-
-lemma step_holdents_v_add:
-  fixes th th' cs s rest l r
-  assumes vt: "vt step (V th' cs#s)"
-  and neq_th: "th \<noteq> th'" 
-  and eq_wq: "wq s cs = th' # rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [th], r)"
-  shows "holdents (V th' cs#s) th = holdents s th \<union> {cs}"
-proof -
-  from prems show ?thesis
-  unfolding  holdents_def step_depend_v[OF vt] by auto
-qed
-
-lemma step_holdents_v_eq:
-  fixes th th' cs s rest l r th''
-  assumes vt: "vt step (V th' cs#s)"
-  and neq_th: "th \<noteq> th'" 
-  and eq_wq: "wq s cs = th' # rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [th''], r)"
-  and neq_th': "th \<noteq> th''"
-  shows "holdents (V th' cs#s) th = holdents s th"
-proof -
-  from prems show ?thesis
-  unfolding  holdents_def step_depend_v[OF vt] by auto
-qed
-
-definition cntCS :: "state \<Rightarrow> thread \<Rightarrow> nat"
-where "cntCS s th = card (holdents s th)"
-
-lemma cntCS_v_eq:
-  fixes th thread cs rest
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and not_in: "th \<notin>  set rest"
-  and vtv: "vt step (V thread cs#s)"
-  shows "cntCS (V thread cs#s) th = cntCS s th"
-proof -
-  from prems show ?thesis
-    apply (unfold cntCS_def holdents_def step_depend_v)
-    apply auto
-    apply (subgoal_tac "\<not>  (\<exists>l r. lsp (cp s) rest = (l, [th], r))", auto)
-    by (drule_tac lsp_set_eq, auto)
-qed
-
-lemma cntCS_v_eq_1:
-  fixes th thread cs rest
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
-  and neq_th': "th \<noteq> th'"
-  and vtv: "vt step (V thread cs#s)"
-  shows "cntCS (V thread cs#s) th = cntCS s th"
-proof -
-  from prems show ?thesis
-    apply (unfold cntCS_def holdents_def step_depend_v)
-    by auto
-qed
-
-fun the_cs :: "node \<Rightarrow> cs"
-where "the_cs (Cs cs) = cs"
-
-lemma cntCS_v_eq_2:
-  fixes th thread cs rest
-  assumes neq_th: "th \<noteq> thread"
-  and eq_wq: "wq s cs = thread#rest"
-  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
-  and neq_th': "th = th'"
-  and vtv: "vt step (V thread cs#s)"
-  shows "cntCS (V thread cs#s) th = 1 + cntCS s th"
-proof -
-  have "card {csa. csa = cs \<or> (Cs csa, Th th') \<in> depend s} = 
-                     Suc (card {cs. (Cs cs, Th th') \<in> depend s})" 
-    (is "card ?A = Suc (card ?B)")
-  proof -
-    have h: "?A = insert cs ?B" by auto
-    moreover have h1: "?B = ?B - {cs}"
-    proof -
-      { assume "(Cs cs, Th th') \<in> depend s"
-        moreover have "(Th th', Cs cs) \<in> depend s"
-        proof -
-          from wq_distinct [OF step_back_vt[OF vtv], of cs]
-          eq_wq lsp_set_eq [OF eq_lsp] show ?thesis
-            apply (auto simp:s_depend_def)
-            by (unfold cs_waiting_def, auto)
-        qed
-        moreover note acyclic_depend [OF step_back_vt[OF vtv]]
-        ultimately have "False"
-          apply (auto simp:acyclic_def)
-          apply (erule_tac x="Cs cs" in allE)
-          apply (subgoal_tac "(Cs cs, Cs cs) \<in> (depend s)\<^sup>+", simp)
-          by (rule_tac trancl_into_trancl [where b = "Th th'"], auto)
-      } thus ?thesis by auto
-    qed
-    moreover have "card (insert cs ?B) = Suc (card (?B - {cs}))"
-    proof(rule card_insert)
-      from finite_depend [OF step_back_vt [OF vtv]]
-      have fnt: "finite (depend s)" .
-      show " finite {cs. (Cs cs, Th th') \<in> depend s}" (is "finite ?B")
-      proof -
-        have "?B \<subseteq>  (\<lambda> (a, b). the_cs a) ` (depend s)"
-          apply (auto simp:image_def)
-          by (rule_tac x = "(Cs x, Th th')" in bexI, auto)
-        with fnt show ?thesis by (auto intro:finite_subset)
-      qed
-    qed
-    ultimately show ?thesis by simp
-  qed
-  with prems show ?thesis
-    apply (unfold cntCS_def holdents_def step_depend_v[OF vtv])
-    by auto
-qed
-
-lemma finite_holding:
-  fixes s th cs
-  assumes vt: "vt step s"
-  shows "finite (holdents s th)"
-proof -
-  let ?F = "\<lambda> (x, y). the_cs x"
-  from finite_depend [OF vt]
-  have "finite (depend s)" .
-  hence "finite (?F `(depend s))" by simp
-  moreover have "{cs . (Cs cs, Th th) \<in> depend 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> depend s"
-      hence "?F (Cs x, Th th) \<in> ?F `(depend s)" by (rule h)
-      moreover have "?F (Cs x, Th th) = x" by simp
-      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` depend s" by simp 
-    } thus ?thesis by auto
-  qed
-  ultimately show ?thesis by (unfold holdents_def, auto intro:finite_subset)
-qed
-
-inductive_cases case_step_v: "step s (V thread cs)"
-
-lemma cntCS_v_dec: 
-  fixes s thread cs
-  assumes vtv: "vt step (V thread cs#s)"
-  shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
-proof -
-  have cs_in: "cs \<in> holdents s thread" using step_back_step[OF vtv]
-    apply (erule_tac case_step_v)
-    apply (unfold holdents_def s_depend_def, simp)
-    by (unfold cs_holding_def s_holding_def, auto)
-  moreover have cs_not_in: 
-    "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
-    apply (insert wq_distinct[OF step_back_vt[OF vtv], of cs])
-    by (unfold holdents_def, unfold step_depend_v[OF vtv], 
-            auto dest:lsp_set_eq)
-  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 finite_holding [OF vtv]
-    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 cnp_cnv_cncs:
-  fixes s th
-  assumes vt: "vt step s"
-  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)
-    assume vt: "vt step 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 prio max_prio thread)
-      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 wq_threads [OF vt 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_def
-          by (simp add:depend_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_def
-        by (simp add:depend_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 (subst (1 2) wq_def)
-          apply (simp add:Let_def)
-          apply (subst s_waiting_def, simp)
-          by (fold wq_def, simp)
-        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> (depend s)\<^sup>+"
-      from prems have vtp: "vt step (P thread cs#s)" by auto
-      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, clarify)
-            apply (intro iffI allI, clarify)
-            apply (erule_tac x = csa in allE, auto)
-            apply (subgoal_tac "waiting_queue (schs s) cs \<noteq> []", auto)
-            apply (erule_tac x = cs in allE, auto)
-            by (case_tac "(waiting_queue (schs s) cs)", auto)
-          moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
-            apply (simp add:cntCS_def holdents_def)
-            by (unfold  step_depend_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_depend_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> depend s} =
-                  Suc (card {cs. (Cs cs, Th thread) \<in> depend 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 finite_holding [OF vt, of thread]
-                    show " finite {cs. (Cs cs, Th thread) \<in> depend s}"
-                      by (unfold holdents_def, simp)
-                  qed
-                  moreover have "?R - {cs} = ?R"
-                  proof -
-                    have "cs \<notin> ?R"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> depend 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_def)
-                  by (unfold step_depend_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)
-                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 wq_distinct [OF vtp, 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_def eq_e step_depend_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 prems have vtv: "vt step (V thread cs # s)" by auto
-      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:s_holding_def)
-      have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
-      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 -
-                  have "thread \<notin> set (wq (e#s) cs1)"
-                  proof(cases "lsp (cp s) rest")
-                    fix l m r 
-                    assume h: "lsp (cp s) rest = (l, m, r)"
-                    show ?thesis
-                    proof(cases "m")
-                      case Nil
-                      from wq_v_eq_nil [OF eq_wq] h Nil eq_e
-                      have " wq (e # s) cs = []" by auto
-                      thus ?thesis using eq_cs by auto
-                    next
-                      case (Cons th' l')
-                      from lsp_mid_length [OF h] and Cons h
-                      have eqh: "lsp (cp s) rest = (l, [th'], r)" by auto
-                      from wq_v_eq [OF eq_wq this]
-                      have "wq (V thread cs # s) cs = th' # l @ r" .
-                      moreover from lsp_set_eq [OF eqh]
-                      have "set rest = set \<dots>" by auto
-                      moreover have "thread \<notin> set rest"
-                      proof -
-                        from wq_distinct [OF step_back_vt[OF vtv], of cs]
-                        and eq_wq show ?thesis by auto
-                      qed
-                      moreover note eq_e eq_cs
-                      ultimately show ?thesis by simp
-                    qed
-                  qed
-                  thus ?thesis by (simp add: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: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)"
-              by(unfold eq_e, rule readys_v_eq [OF neq_th eq_wq False])
-            moreover have "cntCS (e#s) th = cntCS s th"
-              by(unfold eq_e, rule cntCS_v_eq [OF neq_th eq_wq False vtv]) 
-            moreover note ih eq_cnp eq_cnv eq_threads
-            ultimately show ?thesis by auto
-          next
-            case True
-            obtain l m r where eq_lsp: "lsp (cp s) rest = (l, m, r)" 
-              by (cases "lsp (cp s) rest", auto)
-            with True have "m \<noteq> []" by  (auto dest:lsp_mid_nil)
-            with eq_lsp obtain th' where eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
-              by (case_tac m, auto dest:lsp_mid_length)
-            show ?thesis
-            proof(cases "th = th'")
-              case False
-              have "(th \<in> readys (e # s)) = (th \<in> readys s)"
-                by (unfold eq_e, rule readys_v_eq_1 [OF neq_th eq_wq eq_lsp False])
-              moreover have "cntCS (e#s) th = cntCS s th" 
-                by (unfold eq_e, rule cntCS_v_eq_1[OF neq_th eq_wq eq_lsp False vtv])
-              moreover note ih eq_cnp eq_cnv eq_threads
-              ultimately show ?thesis by auto
-            next
-              case True
-              have "th \<in> readys (e # s)"
-                by (unfold eq_e, rule readys_v_eq_2 [OF neq_th eq_wq eq_lsp True vt])
-              moreover have "cntP s th = cntV s th + cntCS s th + 1"
-              proof -
-                have "th \<notin> readys s" 
-                proof -
-                  from True eq_wq lsp_set_eq [OF eq_lsp] neq_th
-                  show ?thesis
-                    apply (unfold readys_def s_waiting_def, auto)
-                    by (rule_tac x = cs in exI, auto)
-                qed
-                moreover have "th \<in> threads s"
-                proof -
-                  from True eq_wq lsp_set_eq [OF eq_lsp] neq_th
-                  have "th \<in> set (wq s cs)" by simp
-                  from wq_threads [OF step_back_vt[OF vtv] this] 
-                  show ?thesis .
-                qed
-                ultimately show ?thesis using ih by auto
-              qed
-              moreover have "cntCS (e # s) th = 1 + cntCS s th"
-                by (unfold eq_e, rule cntCS_v_eq_2 [OF neq_th eq_wq eq_lsp True vtv])
-              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_def
-          by (simp add:depend_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_def s_depend_def wq_def cs_holding_def)
-  qed
-qed
-
-lemma not_thread_cncs:
-  fixes th s
-  assumes vt: "vt step s"
-  and 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)
-    assume vt: "vt step 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 prio max_prio thread)
-      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_def)
-        by (simp add:depend_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_def)
-        by (simp add:depend_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 prems have vtp: "vt step (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_def eq_e)
-        by (unfold step_depend_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 prems have vtv: "vt step (V thread cs#s)" by auto
-      from hold obtain rest 
-        where eq_wq: "wq s cs = thread # rest"
-        by (case_tac "wq s cs", auto simp:s_holding_def)
-      have "cntCS (e # s) th  = cntCS s th"
-      proof(unfold eq_e, rule cntCS_v_eq [OF neq_th eq_wq _ vtv])
-        show "th \<notin> set rest" 
-        proof
-          assume "th \<in> set rest"
-          with eq_wq have "th \<in> set (wq s cs)" by simp
-          from wq_threads [OF vt this] eq_e not_in 
-          show False by simp
-        qed
-      qed
-      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_def)
-        by (simp add:depend_set_unchanged)
-    qed
-    next
-      case vt_nil
-      show ?case
-      by (unfold cntCS_def, 
-        auto simp:count_def holdents_def s_depend_def wq_def cs_holding_def)
-  qed
-qed
-
-lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
-  by (auto simp:s_waiting_def cs_waiting_def)
-
-lemma dm_depend_threads:
-  fixes th s
-  assumes vt: "vt step s"
-  and in_dom: "(Th th) \<in> Domain (depend s)"
-  shows "th \<in> threads s"
-proof -
-  from in_dom obtain n where "(Th th, n) \<in> depend s" by auto
-  moreover from depend_target_th[OF this] obtain cs where "n = Cs cs" by auto
-  ultimately have "(Th th, Cs cs) \<in> depend s" by simp
-  hence "th \<in> set (wq s cs)"
-    by (unfold s_depend_def, auto simp:cs_waiting_def)
-  from wq_threads [OF vt this] show ?thesis .
-qed
-
-lemma cp_eq_cpreced: "cp s th = cpreced s (wq s) th"
-proof(unfold cp_def wq_def, induct s)
-  case (Cons e s')
-  show ?case
-    by (auto simp:Let_def)
-next
-  case Nil
-  show ?case by (auto simp:Let_def)
-qed
-
-fun the_th :: "node \<Rightarrow> thread"
-  where "the_th (Th th) = th"
-
-lemma runing_unique:
-  fixes th1 th2 s
-  assumes vt: "vt step s"
-  and 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"
-    by (unfold runing_def, simp)
-  hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1)) =
-                 Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependents (wq s) th2))"
-    (is "Max (?f ` ?A) = Max (?f ` ?B)")
-    by (unfold cp_eq_cpreced 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 (dependents (wq s) th1)"
-        proof-
-          have "finite {th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (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_depend[OF vt] have "finite (depend s)" .
-              hence "finite ((depend (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_depend_def cs_depend_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependents_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 by (auto intro:that)
-  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 (dependents (wq s) th2)"
-        proof-
-          have "finite {th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (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_depend[OF vt] have "finite (depend s)" .
-              hence "finite ((depend (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_depend_def cs_depend_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependents_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> dependents (wq s) th1)" by simp
-    thus "th1' \<in> threads s"
-    proof
-      assume "th1' \<in> dependents (wq s) th1"
-      hence "(Th th1') \<in> Domain ((depend s)^+)"
-        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
-        by (auto simp:Domain_def)
-      hence "(Th th1') \<in> Domain (depend s)" by (simp add:trancl_domain)
-      from dm_depend_threads[OF vt 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> dependents (wq s) th2)" by simp
-    thus "th2' \<in> threads s"
-    proof
-      assume "th2' \<in> dependents (wq s) th2"
-      hence "(Th th2') \<in> Domain ((depend s)^+)"
-        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
-        by (auto simp:Domain_def)
-      hence "(Th th2') \<in> Domain (depend s)" by (simp add:trancl_domain)
-      from dm_depend_threads[OF vt 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> dependents (wq s) th1" by simp
-  thus ?thesis
-  proof
-    assume eq_th': "th1' = th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
-    thus ?thesis
-    proof
-      assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
-    next
-      assume "th2' \<in> dependents (wq s) th2"
-      with eq_th12 eq_th' have "th1 \<in> dependents (wq s) th2" by simp
-      hence "(Th th1, Th th2) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      hence "Th th1 \<in> Domain ((depend s)^+)" using Domain_def [of "(depend s)^+"]
-        by auto
-      hence "Th th1 \<in> Domain (depend s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th1, n) \<in> depend s" by (auto simp:Domain_def)
-      from depend_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th1, Cs cs') \<in> depend s" by simp
-      with runing_1 have "False"
-        apply (unfold runing_def readys_def s_depend_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    qed
-  next
-    assume th1'_in: "th1' \<in> dependents (wq s) th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
-    thus ?thesis 
-    proof
-      assume "th2' = th2"
-      with th1'_in eq_th12 have "th2 \<in> dependents (wq s) th1" by simp
-      hence "(Th th2, Th th1) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      hence "Th th2 \<in> Domain ((depend s)^+)" using Domain_def [of "(depend s)^+"]
-        by auto
-      hence "Th th2 \<in> Domain (depend s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th2, n) \<in> depend s" by (auto simp:Domain_def)
-      from depend_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th2, Cs cs') \<in> depend s" by simp
-      with runing_2 have "False"
-        apply (unfold runing_def readys_def s_depend_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    next
-      assume "th2' \<in> dependents (wq s) th2"
-      with eq_th12 have "th1' \<in> dependents (wq s) th2" by simp
-      hence h1: "(Th th1', Th th2) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      from th1'_in have h2: "(Th th1', Th th1) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      show ?thesis
-      proof(rule dchain_unique[OF vt 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 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 prio max_prio thread)
-    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 length_down_to_in: 
-  assumes le_ij: "i \<le> j"
-    and le_js: "j \<le> length s"
-  shows "length (down_to j i s) = j - i"
-proof -
-  have "length (down_to j i s) = length (from_to i j (rev s))"
-    by (unfold down_to_def, auto)
-  also have "\<dots> = j - i"
-  proof(rule length_from_to_in[OF le_ij])
-    from le_js show "j \<le> length (rev s)" by simp
-  qed
-  finally show ?thesis .
-qed
-
-
-lemma moment_head: 
-  assumes le_it: "Suc i \<le> length t"
-  obtains e where "moment (Suc i) t = e#moment i t"
-proof -
-  have "i \<le> Suc i" by simp
-  from length_down_to_in [OF this le_it]
-  have "length (down_to (Suc i) i t) = 1" by auto
-  then obtain e where "down_to (Suc i) i t = [e]"
-    apply (cases "(down_to (Suc i) i t)") by auto
-  moreover have "down_to (Suc i) 0 t = down_to (Suc i) i t @ down_to i 0 t"
-    by (rule down_to_conc[symmetric], auto)
-  ultimately have eq_me: "moment (Suc i) t = e#(moment i t)"
-    by (auto simp:down_to_moment)
-  from that [OF this] show ?thesis .
-qed
-
-lemma cnp_cnv_eq:
-  fixes th s
-  assumes "vt step s"
-  and "th \<notin> threads s"
-  shows "cntP s th = cntV s th"
-proof -
-  from assms show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    have ih: "th \<notin> threads s \<Longrightarrow> cntP s th = cntV s th" by fact
-    have not_in: "th \<notin> threads (e # s)" by fact
-    have "step s e" by fact
-    thus ?case proof(cases)
-      case (thread_create prio max_prio thread)
-      assume eq_e: "e = Create thread prio"
-      hence "thread \<in> threads (e#s)" by simp
-      with not_in and eq_e have "th \<notin> threads s" by auto
-      from ih [OF this] show ?thesis using eq_e
-        by (auto simp:cntP_def cntV_def count_def)
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread"
-        and not_holding: "holdents s thread = {}"
-      have vt_s: "vt step s" by fact
-      from finite_holding[OF vt_s] have "finite (holdents s thread)" .
-      with not_holding have "cntCS s thread = 0" by (unfold cntCS_def, auto)
-      moreover have "thread \<in> readys s" using thread_exit by (auto simp:runing_def)
-      moreover note cnp_cnv_cncs[OF vt_s, of thread]
-      ultimately have eq_thread: "cntP s thread = cntV s thread" by auto
-      show ?thesis
-      proof(cases "th = thread")
-        case True
-        with eq_thread eq_e show ?thesis 
-          by (auto simp:cntP_def cntV_def count_def)
-      next
-        case False
-        with not_in and eq_e have "th \<notin> threads s" by simp
-        from ih[OF this] and eq_e show ?thesis 
-           by (auto simp:cntP_def cntV_def count_def)
-      qed
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-      have "thread \<in> runing s" by fact
-      with not_in eq_e have neq_th: "thread \<noteq> th" 
-        by (auto simp:runing_def readys_def)
-      from not_in eq_e have "th \<notin> threads s" by simp
-      from ih[OF this] and neq_th and eq_e show ?thesis
-        by (auto simp:cntP_def cntV_def count_def)
-    next
-      case (thread_V thread cs)
-      assume eq_e: "e = V thread cs"
-      have "thread \<in> runing s" by fact
-      with not_in eq_e have neq_th: "thread \<noteq> th" 
-        by (auto simp:runing_def readys_def)
-      from not_in eq_e have "th \<notin> threads s" by simp
-      from ih[OF this] and neq_th and eq_e show ?thesis
-        by (auto simp:cntP_def cntV_def count_def)
-    next
-      case (thread_set thread prio)
-      assume eq_e: "e = Set thread prio"
-        and "thread \<in> runing s"
-      hence "thread \<in> threads (e#s)" 
-        by (simp add:runing_def readys_def)
-      with not_in and eq_e have "th \<notin> threads s" by auto
-      from ih [OF this] show ?thesis using eq_e
-        by (auto simp:cntP_def cntV_def count_def)  
-    qed
-  next
-    case vt_nil
-    show ?case by (auto simp:cntP_def cntV_def count_def)
-  qed
-qed
-
-lemma eq_depend: 
-  "depend (wq s) = depend s"
-by (unfold cs_depend_def s_depend_def, auto)
-
-lemma count_eq_dependents:
-  assumes vt: "vt step s"
-  and eq_pv: "cntP s th = cntV s th"
-  shows "dependents (wq s) th = {}"
-proof -
-  from cnp_cnv_cncs[OF vt] and eq_pv
-  have "cntCS s th = 0" 
-    by (auto split:if_splits)
-  moreover have "finite {cs. (Cs cs, Th th) \<in> depend s}"
-  proof -
-    from finite_holding[OF vt, of th] show ?thesis
-      by (simp add:holdents_def)
-  qed
-  ultimately have h: "{cs. (Cs cs, Th th) \<in> depend s} = {}"
-    by (unfold cntCS_def holdents_def cs_dependents_def, auto)
-  show ?thesis
-  proof(unfold cs_dependents_def)
-    { assume "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}"
-      then obtain th' where "(Th th', Th th) \<in> (depend (wq s))\<^sup>+" by auto
-      hence "False"
-      proof(cases)
-        assume "(Th th', Th th) \<in> depend (wq s)"
-        thus "False" by (auto simp:cs_depend_def)
-      next
-        fix c
-        assume "(c, Th th) \<in> depend (wq s)"
-        with h and eq_depend show "False"
-          by (cases c, auto simp:cs_depend_def)
-      qed
-    } thus "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} = {}" by auto
-  qed
-qed
-
-lemma dependents_threads:
-  fixes s th
-  assumes vt: "vt step s"
-  shows "dependents (wq s) th \<subseteq> threads s"
-proof
-  { fix th th'
-    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (depend (wq s))\<^sup>+}"
-    have "Th th \<in> Domain (depend s)"
-    proof -
-      from h obtain th' where "(Th th, Th th') \<in> (depend (wq s))\<^sup>+" by auto
-      hence "(Th th) \<in> Domain ( (depend (wq s))\<^sup>+)" by (auto simp:Domain_def)
-      with trancl_domain have "(Th th) \<in> Domain (depend (wq s))" by simp
-      thus ?thesis using eq_depend by simp
-    qed
-    from dm_depend_threads[OF vt this]
-    have "th \<in> threads s" .
-  } note hh = this
-  fix th1 
-  assume "th1 \<in> dependents (wq s) th"
-  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (depend (wq s))\<^sup>+}"
-    by (unfold cs_dependents_def, simp)
-  from hh [OF this] show "th1 \<in> threads s" .
-qed
-
-lemma finite_threads:
-  assumes vt: "vt step s"
-  shows "finite (threads s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    assume vt: "vt step s"
-    and step: "step s e"
-    and ih: "finite (threads s)"
-    from step
-    show ?case
-    proof(cases)
-      case (thread_create prio max_prio thread)
-      assume eq_e: "e = Create thread prio"
-      with ih
-      show ?thesis by (unfold eq_e, auto)
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread"
-      with ih show ?thesis 
-        by (unfold eq_e, auto)
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-      with ih show ?thesis by (unfold eq_e, auto)
-    next
-      case (thread_V thread cs)
-      assume eq_e: "e = V thread cs"
-      with ih show ?thesis by (unfold eq_e, auto)
-    next 
-      case (thread_set thread prio)
-      from vt_cons thread_set show ?thesis by simp
-    qed
-  next
-    case vt_nil
-    show ?case by (auto)
-  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 cp_le:
-  assumes vt: "vt step s"
-  and 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_dependents_def)
-  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (depend (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> (depend (wq s))\<^sup>+} \<noteq> {}" by simp
-  next
-    from finite_threads [OF vt]
-    show "finite (threads s)" .
-  next
-    from th_in
-    show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> threads s"
-      apply (auto simp:Domain_def)
-      apply (rule_tac dm_depend_threads[OF vt])
-      apply (unfold trancl_domain [of "depend s", symmetric])
-      by (unfold cs_depend_def s_depend_def, auto simp:Domain_def)
-  qed
-qed
-
-lemma le_cp:
-  assumes vt: "vt step s"
-  shows "preced th s \<le> cp s th"
-proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
-  show "Prc (original_priority th s) (birthtime th s)
-    \<le> Max (insert (Prc (original_priority th s) (birthtime th s))
-            ((\<lambda>th. Prc (original_priority th s) (birthtime th s)) ` dependents (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> (depend (wq s))\<^sup>+}"
-        proof -
-          let ?F = "\<lambda> (x, y). the_th x"
-          have "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (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_depend[OF vt] have "finite (depend s)" .
-            hence "finite ((depend (wq s))\<^sup>+)"
-              apply (unfold finite_trancl)
-              by (auto simp: s_depend_def cs_depend_def wq_def)
-            thus ?thesis by auto
-          qed
-          ultimately show ?thesis by (auto intro:finite_subset)
-        qed
-        thus ?thesis by (simp add:cs_dependents_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: 
-  assumes vt: "vt step s"
-  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[OF vt] 
-    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 vt 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[OF vt]
-      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 vt, 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[OF vt]
-        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 vt: "vt step s"
-  and np: "threads s \<noteq> {}"
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(unfold max_cp_eq[OF vt])
-  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[OF vt] 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 vt tm_in]
-    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+)" .
-    thus ?thesis
-    proof
-      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+ "
-      then obtain th' where th'_in: "th' \<in> readys s" 
-        and tm_chain:"(Th tm, Th th') \<in> (depend s)\<^sup>+" by auto
-      have "cp s th' = ?f tm"
-      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
-        from dependents_threads[OF vt] finite_threads[OF vt]
-        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th'))" 
-          by (auto intro:finite_subset)
-      next
-        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependents (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[OF vt]
-          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        next
-          from p_in and th'_in and dependents_threads[OF vt, 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> dependents (wq s) th')"
-        proof -
-          from tm_chain
-          have "tm \<in> dependents (wq s) th'"
-            by (unfold cs_dependents_def s_depend_def cs_depend_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 dependents_threads [OF vt, of th1] th1_in
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (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> dependents (wq s) th1) \<noteq> {}" by simp
-        next
-          from finite_threads[OF vt] 
-          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        qed
-      next
-        from finite_threads[OF vt]
-        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> dependents (wq s) tm)"
-          show "y \<le> preced tm s"
-          proof -
-            { fix y'
-              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependents (wq s) tm)"
-              have "y' \<le> preced tm s"
-              proof(unfold tm_max, rule Max_ge)
-                from hy' dependents_threads[OF vt, of tm]
-                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
-              next
-                from finite_threads[OF vt] 
-                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-              qed
-            } with hy show ?thesis by auto
-          qed
-        next
-          from dependents_threads[OF vt, of tm] finite_threads[OF vt]
-          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm))"
-            by (auto intro:finite_subset)
-        next
-          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependents (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[OF vt]
-          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[OF vt]
-            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-          next
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<noteq> {}"
-              by simp
-          next
-            from dependents_threads[OF vt, of th1] th1_readys
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (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
-
-lemma max_cp_readys_threads:
-  assumes vt: "vt step s"
-  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 vt False])
-qed
-
-lemma readys_threads:
-  shows "readys s \<subseteq> threads s"
-proof
-  fix th
-  assume "th \<in> readys s"
-  thus "th \<in> threads s"
-    by (unfold readys_def, auto)
-qed
-
-lemma eq_holding: "holding (wq s) th cs = holding s th cs"
-  apply (unfold s_holding_def cs_holding_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
-
-end
\ No newline at end of file
--- a/prio/CpsG.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1997 +0,0 @@
-theory CpsG
-imports PrioG 
-begin
-
-lemma not_thread_holdents:
-  fixes th s
-  assumes vt: "vt s"
-  and not_in: "th \<notin> threads s" 
-  shows "holdents s th = {}"
-proof -
-  from vt not_in show ?thesis
-  proof(induct arbitrary:th)
-    case (vt_cons s e th)
-    assume vt: "vt s"
-      and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> holdents s th = {}"
-      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 "holdents (e # s) th = holdents s th"
-        apply (unfold eq_e holdents_test)
-        by (simp add:depend_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 = {}"
-      show ?thesis
-      proof(cases "th = thread")
-        case True
-        with nh eq_e
-        show ?thesis 
-          by (auto simp:holdents_test depend_exit_unchanged)
-      next
-        case False
-        with not_in and eq_e
-        have "th \<notin> threads s" by simp
-        from ih[OF this] False eq_e show ?thesis 
-          by (auto simp:holdents_test depend_exit_unchanged)
-      qed
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-      and is_runing: "thread \<in> runing s"
-      from assms thread_exit ih stp not_in vt eq_e 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 "holdents (e # s) th  = holdents s th "
-        apply (unfold cntCS_def holdents_test eq_e)
-        by (unfold step_depend_p[OF vtp], auto)
-      moreover have "holdents s th = {}"
-      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 eq_e ih stp not_in vt have vtv: "vt (V thread cs#s)" by auto
-      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 wq_distinct[OF step_back_vt[OF vtv], 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 "hd x \<in> set rest" by (cases x, auto)
-        qed
-        with eq_wq have "?t \<in> set (wq s cs)" by simp
-        from wq_threads[OF step_back_vt[OF vtv], OF this] and ni
-        show False by auto
-      qed
-      moreover note neq_th eq_wq
-      ultimately have "holdents (e # s) th  = holdents s th"
-        by (unfold eq_e cntCS_def holdents_test step_depend_v[OF vtv], auto)
-      moreover have "holdents s th = {}"
-      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:depend_set_unchanged)
-    qed
-    next
-      case vt_nil
-      show ?case
-      by (auto simp:count_def holdents_test s_depend_def wq_def cs_holding_def)
-  qed
-qed
-
-
-
-lemma next_th_neq: 
-  assumes vt: "vt s"
-  and nt: "next_th s th cs th'"
-  shows "th' \<noteq> th"
-proof -
-  from nt show ?thesis
-    apply (auto simp:next_th_def)
-  proof -
-    fix rest
-    assume eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
-      and ne: "rest \<noteq> []"
-    have "hd (SOME q. distinct q \<and> set q = set rest) \<in> set rest" 
-    proof(rule someI2)
-      from wq_distinct[OF vt, of cs] eq_wq
-      show "distinct rest \<and> set rest = set rest" by auto
-    next
-      fix x
-      assume "distinct x \<and> set x = set rest"
-      hence eq_set: "set x = set rest" by auto
-      with ne have "x \<noteq> []" by auto
-      hence "hd x \<in> set x" by auto
-      with eq_set show "hd x \<in> set rest" by auto
-    qed
-    with wq_distinct[OF vt, of cs] eq_wq show False by auto
-  qed
-qed
-
-lemma next_th_unique: 
-  assumes nt1: "next_th s th cs th1"
-  and nt2: "next_th s th cs th2"
-  shows "th1 = th2"
-proof -
-  from assms show ?thesis
-    by (unfold next_th_def, auto)
-qed
-
-lemma pp_sub: "(r^+)^+ \<subseteq> r^+"
-  by auto
-
-lemma wf_depend:
-  assumes vt: "vt s"
-  shows "wf (depend s)"
-proof(rule finite_acyclic_wf)
-  from finite_depend[OF vt] show "finite (depend s)" .
-next
-  from acyclic_depend[OF vt] show "acyclic (depend s)" .
-qed
-
-lemma Max_Union:
-  assumes fc: "finite C"
-  and ne: "C \<noteq> {}"
-  and fa: "\<And> A. A \<in> C \<Longrightarrow> finite A \<and> A \<noteq> {}"
-  shows "Max (\<Union> C) = Max (Max ` C)"
-proof -
-  from fc ne fa show ?thesis
-  proof(induct)
-    case (insert x F)
-    assume ih: "\<lbrakk>F \<noteq> {}; \<And>A. A \<in> F \<Longrightarrow> finite A \<and> A \<noteq> {}\<rbrakk> \<Longrightarrow> Max (\<Union>F) = Max (Max ` F)"
-    and h: "\<And> A. A \<in> insert x F \<Longrightarrow> finite A \<and> A \<noteq> {}"
-    show ?case (is "?L = ?R")
-    proof(cases "F = {}")
-      case False
-      from Union_insert have "?L = Max (x \<union> (\<Union> F))" by simp
-      also have "\<dots> = max (Max x) (Max(\<Union> F))"
-      proof(rule Max_Un)
-        from h[of x] show "finite x" by auto
-      next
-        from h[of x] show "x \<noteq> {}" by auto
-      next
-        show "finite (\<Union>F)"
-        proof(rule finite_Union)
-          show "finite F" by fact
-        next
-          from h show "\<And>M. M \<in> F \<Longrightarrow> finite M" by auto
-        qed
-      next
-        from False and h show "\<Union>F \<noteq> {}" by auto
-      qed
-      also have "\<dots> = ?R"
-      proof -
-        have "?R = Max (Max ` ({x} \<union> F))" by simp
-        also have "\<dots> = Max ({Max x} \<union> (Max ` F))" by simp
-        also have "\<dots> = max (Max x) (Max (\<Union>F))"
-        proof -
-          have "Max ({Max x} \<union> Max ` F) = max (Max {Max x}) (Max (Max ` F))"
-          proof(rule Max_Un)
-            show "finite {Max x}" by simp
-          next
-            show "{Max x} \<noteq> {}" by simp
-          next
-            from insert show "finite (Max ` F)" by auto
-          next
-            from False show "Max ` F \<noteq> {}" by auto
-          qed
-          moreover have "Max {Max x} = Max x" by simp
-          moreover have "Max (\<Union>F) = Max (Max ` F)"
-          proof(rule ih)
-            show "F \<noteq> {}" by fact
-          next
-            from h show "\<And>A. A \<in> F \<Longrightarrow> finite A \<and> A \<noteq> {}"
-              by auto
-          qed
-          ultimately show ?thesis by auto
-        qed
-        finally show ?thesis by simp
-      qed
-      finally show ?thesis by simp
-    next
-      case True
-      thus ?thesis by auto
-    qed
-  next
-    case empty
-    assume "{} \<noteq> {}" show ?case by auto
-  qed
-qed
-
-definition child :: "state \<Rightarrow> (node \<times> node) set"
-  where "child s \<equiv>
-            {(Th th', Th th) | th th'. \<exists> cs. (Th th', Cs cs) \<in> depend s \<and> (Cs cs, Th th) \<in> depend s}"
-
-definition children :: "state \<Rightarrow> thread \<Rightarrow> thread set"
-  where "children s th \<equiv> {th'. (Th th', Th th) \<in> child s}"
-
-lemma children_def2:
-  "children s th \<equiv> {th'. \<exists> cs. (Th th', Cs cs) \<in> depend s \<and> (Cs cs, Th th) \<in> depend s}"
-unfolding child_def children_def by simp
-
-lemma children_dependents: "children s th \<subseteq> dependents (wq s) th"
-  by (unfold children_def child_def cs_dependents_def, auto simp:eq_depend)
-
-lemma child_unique:
-  assumes vt: "vt s"
-  and ch1: "(Th th, Th th1) \<in> child s"
-  and ch2: "(Th th, Th th2) \<in> child s"
-  shows "th1 = th2"
-proof -
-  from ch1 ch2 show ?thesis
-  proof(unfold child_def, clarsimp)
-    fix cs csa
-    assume h1: "(Th th, Cs cs) \<in> depend s"
-      and h2: "(Cs cs, Th th1) \<in> depend s"
-      and h3: "(Th th, Cs csa) \<in> depend s"
-      and h4: "(Cs csa, Th th2) \<in> depend s"
-    from unique_depend[OF vt h1 h3] have "cs = csa" by simp
-    with h4 have "(Cs cs, Th th2) \<in> depend s" by simp
-    from unique_depend[OF vt h2 this]
-    show "th1 = th2" by simp
-  qed 
-qed
-
-
-lemma cp_eq_cpreced_f: "cp s = cpreced (wq s) s"
-proof -
-  from fun_eq_iff 
-  have h:"\<And>f g. (\<forall> x. f x = g x) \<Longrightarrow> f = g" by auto
-  show ?thesis
-  proof(rule h)
-    from cp_eq_cpreced show "\<forall>x. cp s x = cpreced (wq s) s x" by auto
-  qed
-qed
-
-lemma depend_children:
-  assumes h: "(Th th1, Th th2) \<in> (depend s)^+"
-  shows "th1 \<in> children s th2 \<or> (\<exists> th3. th3 \<in> children s th2 \<and> (Th th1, Th th3) \<in> (depend s)^+)"
-proof -
-  from h show ?thesis
-  proof(induct rule: tranclE)
-    fix c th2
-    assume h1: "(Th th1, c) \<in> (depend s)\<^sup>+"
-    and h2: "(c, Th th2) \<in> depend s"
-    from h2 obtain cs where eq_c: "c = Cs cs"
-      by (case_tac c, auto simp:s_depend_def)
-    show "th1 \<in> children s th2 \<or> (\<exists>th3. th3 \<in> children s th2 \<and> (Th th1, Th th3) \<in> (depend s)\<^sup>+)"
-    proof(rule tranclE[OF h1])
-      fix ca
-      assume h3: "(Th th1, ca) \<in> (depend s)\<^sup>+"
-        and h4: "(ca, c) \<in> depend s"
-      show "th1 \<in> children s th2 \<or> (\<exists>th3. th3 \<in> children s th2 \<and> (Th th1, Th th3) \<in> (depend s)\<^sup>+)"
-      proof -
-        from eq_c and h4 obtain th3 where eq_ca: "ca = Th th3"
-          by (case_tac ca, auto simp:s_depend_def)
-        from eq_ca h4 h2 eq_c
-        have "th3 \<in> children s th2" by (auto simp:children_def child_def)
-        moreover from h3 eq_ca have "(Th th1, Th th3) \<in> (depend s)\<^sup>+" by simp
-        ultimately show ?thesis by auto
-      qed
-    next
-      assume "(Th th1, c) \<in> depend s"
-      with h2 eq_c
-      have "th1 \<in> children s th2"
-        by (auto simp:children_def child_def)
-      thus ?thesis by auto
-    qed
-  next
-    assume "(Th th1, Th th2) \<in> depend s"
-    thus ?thesis
-      by (auto simp:s_depend_def)
-  qed
-qed
-
-lemma sub_child: "child s \<subseteq> (depend s)^+"
-  by (unfold child_def, auto)
-
-lemma wf_child: 
-  assumes vt: "vt s"
-  shows "wf (child s)"
-proof(rule wf_subset)
-  from wf_trancl[OF wf_depend[OF vt]]
-  show "wf ((depend s)\<^sup>+)" .
-next
-  from sub_child show "child s \<subseteq> (depend s)\<^sup>+" .
-qed
-
-lemma depend_child_pre:
-  assumes vt: "vt s"
-  shows
-  "(Th th, n) \<in> (depend s)^+ \<longrightarrow> (\<forall> th'. n = (Th th') \<longrightarrow> (Th th, Th th') \<in> (child s)^+)" (is "?P n")
-proof -
-  from wf_trancl[OF wf_depend[OF vt]]
-  have wf: "wf ((depend s)^+)" .
-  show ?thesis
-  proof(rule wf_induct[OF wf, of ?P], clarsimp)
-    fix th'
-    assume ih[rule_format]: "\<forall>y. (y, Th th') \<in> (depend s)\<^sup>+ \<longrightarrow>
-               (Th th, y) \<in> (depend s)\<^sup>+ \<longrightarrow> (\<forall>th'. y = Th th' \<longrightarrow> (Th th, Th th') \<in> (child s)\<^sup>+)"
-    and h: "(Th th, Th th') \<in> (depend s)\<^sup>+"
-    show "(Th th, Th th') \<in> (child s)\<^sup>+"
-    proof -
-      from depend_children[OF h]
-      have "th \<in> children s th' \<or> (\<exists>th3. th3 \<in> children s th' \<and> (Th th, Th th3) \<in> (depend s)\<^sup>+)" .
-      thus ?thesis
-      proof
-        assume "th \<in> children s th'"
-        thus "(Th th, Th th') \<in> (child s)\<^sup>+" by (auto simp:children_def)
-      next
-        assume "\<exists>th3. th3 \<in> children s th' \<and> (Th th, Th th3) \<in> (depend s)\<^sup>+"
-        then obtain th3 where th3_in: "th3 \<in> children s th'" 
-          and th_dp: "(Th th, Th th3) \<in> (depend s)\<^sup>+" by auto
-        from th3_in have "(Th th3, Th th') \<in> (depend s)^+" by (auto simp:children_def child_def)
-        from ih[OF this th_dp, of th3] have "(Th th, Th th3) \<in> (child s)\<^sup>+" by simp
-        with th3_in show "(Th th, Th th') \<in> (child s)\<^sup>+" by (auto simp:children_def)
-      qed
-    qed
-  qed
-qed
-
-lemma depend_child: "\<lbrakk>vt s; (Th th, Th th') \<in> (depend s)^+\<rbrakk> \<Longrightarrow> (Th th, Th th') \<in> (child s)^+"
-  by (insert depend_child_pre, auto)
-
-lemma child_depend_p:
-  assumes "(n1, n2) \<in> (child s)^+"
-  shows "(n1, n2) \<in> (depend s)^+"
-proof -
-  from assms show ?thesis
-  proof(induct)
-    case (base y)
-    with sub_child show ?case by auto
-  next
-    case (step y z)
-    assume "(y, z) \<in> child s"
-    with sub_child have "(y, z) \<in> (depend s)^+" by auto
-    moreover have "(n1, y) \<in> (depend s)^+" by fact
-    ultimately show ?case by auto
-  qed
-qed
-
-lemma child_depend_eq: 
-  assumes vt: "vt s"
-  shows 
-  "((Th th1, Th th2) \<in> (child s)^+) = 
-   ((Th th1, Th th2) \<in> (depend s)^+)"
-  by (auto intro: depend_child[OF vt] child_depend_p)
-
-lemma children_no_dep:
-  fixes s th th1 th2 th3
-  assumes vt: "vt s"
-  and ch1: "(Th th1, Th th) \<in> child s"
-  and ch2: "(Th th2, Th th) \<in> child s"
-  and ch3: "(Th th1, Th th2) \<in> (depend s)^+"
-  shows "False"
-proof -
-  from depend_child[OF vt ch3]
-  have "(Th th1, Th th2) \<in> (child s)\<^sup>+" .
-  thus ?thesis
-  proof(rule converse_tranclE)
-    thm tranclD
-    assume "(Th th1, Th th2) \<in> child s"
-    from child_unique[OF vt ch1 this] have "th = th2" by simp
-    with ch2 have "(Th th2, Th th2) \<in> child s" by simp
-    with wf_child[OF vt] show ?thesis by auto
-  next
-    fix c
-    assume h1: "(Th th1, c) \<in> child s"
-      and h2: "(c, Th th2) \<in> (child s)\<^sup>+"
-    from h1 obtain th3 where eq_c: "c = Th th3" by (unfold child_def, auto)
-    with h1 have "(Th th1, Th th3) \<in> child s" by simp
-    from child_unique[OF vt ch1 this] have eq_th3: "th3 = th" by simp
-    with eq_c and h2 have "(Th th, Th th2) \<in> (child s)\<^sup>+" by simp
-    with ch2 have "(Th th, Th th) \<in> (child s)\<^sup>+" by auto
-    moreover have "wf ((child s)\<^sup>+)"
-    proof(rule wf_trancl)
-      from wf_child[OF vt] show "wf (child s)" .
-    qed
-    ultimately show False by auto
-  qed
-qed
-
-lemma unique_depend_p:
-  assumes vt: "vt s"
-  and dp1: "(n, n1) \<in> (depend s)^+"
-  and dp2: "(n, n2) \<in> (depend s)^+"
-  and neq: "n1 \<noteq> n2"
-  shows "(n1, n2) \<in> (depend s)^+ \<or> (n2, n1) \<in> (depend s)^+"
-proof(rule unique_chain [OF _ dp1 dp2 neq])
-  from unique_depend[OF vt]
-  show "\<And>a b c. \<lbrakk>(a, b) \<in> depend s; (a, c) \<in> depend s\<rbrakk> \<Longrightarrow> b = c" by auto
-qed
-
-lemma dependents_child_unique:
-  fixes s th th1 th2 th3
-  assumes vt: "vt s"
-  and ch1: "(Th th1, Th th) \<in> child s"
-  and ch2: "(Th th2, Th th) \<in> child s"
-  and dp1: "th3 \<in> dependents s th1"
-  and dp2: "th3 \<in> dependents s th2"
-shows "th1 = th2"
-proof -
-  { assume neq: "th1 \<noteq> th2"
-    from dp1 have dp1: "(Th th3, Th th1) \<in> (depend s)^+" 
-      by (simp add:s_dependents_def eq_depend)
-    from dp2 have dp2: "(Th th3, Th th2) \<in> (depend s)^+" 
-      by (simp add:s_dependents_def eq_depend)
-    from unique_depend_p[OF vt dp1 dp2] and neq
-    have "(Th th1, Th th2) \<in> (depend s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (depend s)\<^sup>+" by auto
-    hence False
-    proof
-      assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+ "
-      from children_no_dep[OF vt ch1 ch2 this] show ?thesis .
-    next
-      assume " (Th th2, Th th1) \<in> (depend s)\<^sup>+"
-      from children_no_dep[OF vt ch2 ch1 this] show ?thesis .
-    qed
-  } thus ?thesis by auto
-qed
-
-lemma cp_rec:
-  fixes s th
-  assumes vt: "vt s"
-  shows "cp s th = Max ({preced th s} \<union> (cp s ` children s th))"
-proof(unfold cp_eq_cpreced_f cpreced_def)
-  let ?f = "(\<lambda>th. preced th s)"
-  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)) =
-        Max ({preced th s} \<union> (\<lambda>th. Max ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))) ` children s th)"
-  proof(cases " children s th = {}")
-    case False
-    have "(\<lambda>th. Max ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))) ` children s th = 
-          {Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) | th' . th' \<in> children s th}"
-      (is "?L = ?R")
-      by auto
-    also have "\<dots> = 
-      Max ` {((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) | th' . th' \<in> children s th}"
-      (is "_ = Max ` ?C")
-      by auto
-    finally have "Max ?L = Max (Max ` ?C)" by auto
-    also have "\<dots> = Max (\<Union> ?C)"
-    proof(rule Max_Union[symmetric])
-      from children_dependents[of s th] finite_threads[OF vt] and dependents_threads[OF vt, of th]
-      show "finite {(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th}"
-        by (auto simp:finite_subset)
-    next
-      from False
-      show "{(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th} \<noteq> {}"
-        by simp
-    next
-      show "\<And>A. A \<in> {(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th} \<Longrightarrow>
-        finite A \<and> A \<noteq> {}"
-        apply (auto simp:finite_subset)
-      proof -
-        fix th'
-        from finite_threads[OF vt] and dependents_threads[OF vt, of th']
-        show "finite ((\<lambda>th. preced th s) ` dependents (wq s) th')" by (auto simp:finite_subset)
-      qed
-    qed
-    also have "\<dots> = Max ((\<lambda>th. preced th s) ` dependents (wq s) th)"
-      (is "Max ?A = Max ?B")
-    proof -
-      have "?A = ?B"
-      proof
-        show "\<Union>{(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th}
-                    \<subseteq> (\<lambda>th. preced th s) ` dependents (wq s) th"
-        proof
-          fix x 
-          assume "x \<in> \<Union>{(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th}"
-          then obtain th' where 
-             th'_in: "th' \<in> children s th"
-            and x_in: "x \<in> ?f ` ({th'} \<union> dependents (wq s) th')" by auto
-          hence "x = ?f th' \<or> x \<in> (?f ` dependents (wq s) th')" by auto
-          thus "x \<in> ?f ` dependents (wq s) th"
-          proof
-            assume "x = preced th' s"
-            with th'_in and children_dependents
-            show "x \<in> (\<lambda>th. preced th s) ` dependents (wq s) th" by auto
-          next
-            assume "x \<in> (\<lambda>th. preced th s) ` dependents (wq s) th'"
-            moreover note th'_in
-            ultimately show " x \<in> (\<lambda>th. preced th s) ` dependents (wq s) th"
-              by (unfold cs_dependents_def children_def child_def, auto simp:eq_depend)
-          qed
-        qed
-      next
-        show "?f ` dependents (wq s) th
-           \<subseteq> \<Union>{?f ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th}"
-        proof
-          fix x 
-          assume x_in: "x \<in> (\<lambda>th. preced th s) ` dependents (wq s) th"
-          then obtain th' where
-            eq_x: "x = ?f th'" and dp: "(Th th', Th th) \<in> (depend s)^+" 
-            by (auto simp:cs_dependents_def eq_depend)
-          from depend_children[OF dp]
-          have "th' \<in> children s th \<or> (\<exists>th3. th3 \<in> children s th \<and> (Th th', Th th3) \<in> (depend s)\<^sup>+)" .
-          thus "x \<in> \<Union>{(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th}"
-          proof
-            assume "th' \<in> children s th"
-            with eq_x
-            show "x \<in> \<Union>{(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th}"
-              by auto
-          next
-            assume "\<exists>th3. th3 \<in> children s th \<and> (Th th', Th th3) \<in> (depend s)\<^sup>+"
-            then obtain th3 where th3_in: "th3 \<in> children s th"
-              and dp3: "(Th th', Th th3) \<in> (depend s)\<^sup>+" by auto
-            show "x \<in> \<Union>{(\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th') |th'. th' \<in> children s th}"
-            proof -
-              from dp3
-              have "th' \<in> dependents (wq s) th3"
-                by (auto simp:cs_dependents_def eq_depend)
-              with eq_x th3_in show ?thesis by auto
-            qed
-          qed          
-        qed
-      qed
-      thus ?thesis by simp
-    qed
-    finally have "Max ((\<lambda>th. preced th s) ` dependents (wq s) th) = Max (?L)" 
-      (is "?X = ?Y") by auto
-    moreover have "Max ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)) = 
-                   max (?f th) ?X" 
-    proof -
-      have "Max ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)) = 
-            Max ({?f th} \<union> ?f ` (dependents (wq s) th))" by simp
-      also have "\<dots> = max (Max {?f th}) (Max (?f ` (dependents (wq s) th)))"
-      proof(rule Max_Un, auto)
-        from finite_threads[OF vt] and dependents_threads[OF vt, of th]
-        show "finite ((\<lambda>th. preced th s) ` dependents (wq s) th)" by (auto simp:finite_subset)
-      next
-        assume "dependents (wq s) th = {}"
-        with False and children_dependents show False by auto
-      qed
-      also have "\<dots> = max (?f th) ?X" by simp
-      finally show ?thesis .
-    qed
-    moreover have "Max ({preced th s} \<union> 
-                     (\<lambda>th. Max ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))) ` children s th) = 
-                   max (?f th) ?Y"
-    proof -
-      have "Max ({preced th s} \<union> 
-                     (\<lambda>th. Max ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))) ` children s th) = 
-            max (Max {preced th s}) ?Y"
-      proof(rule Max_Un, auto)
-        from finite_threads[OF vt] dependents_threads[OF vt, of th] children_dependents [of s th]
-        show "finite ((\<lambda>th. Max (insert (preced th s) ((\<lambda>th. preced th s) ` dependents (wq s) th))) ` 
-                       children s th)" 
-          by (auto simp:finite_subset)
-      next
-        assume "children s th = {}"
-        with False show False by auto
-      qed
-      thus ?thesis by simp
-    qed
-    ultimately show ?thesis by auto
-  next
-    case True
-    moreover have "dependents (wq s) th = {}"
-    proof -
-      { fix th'
-        assume "th' \<in> dependents (wq s) th"
-        hence " (Th th', Th th) \<in> (depend s)\<^sup>+" by (simp add:cs_dependents_def eq_depend)
-        from depend_children[OF this] and True
-        have "False" by auto
-      } thus ?thesis by auto
-    qed
-    ultimately show ?thesis by auto
-  qed
-qed
-
-definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
-where "cps s = {(th, cp s th) | th . th \<in> threads s}"
-
-locale step_set_cps =
-  fixes s' th prio s 
-  defines s_def : "s \<equiv> (Set th prio#s')"
-  assumes vt_s: "vt s"
-
-context step_set_cps 
-begin
-
-lemma eq_preced:
-  fixes th'
-  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_dep: "depend s = depend s'"
-  by (unfold s_def depend_set_unchanged, auto)
-
-lemma eq_cp_pre:
-  fixes th' 
-  assumes neq_th: "th' \<noteq> th"
-  and nd: "th \<notin> dependents s th'"
-  shows "cp s th' = cp s' th'"
-  apply (unfold cp_eq_cpreced cpreced_def)
-proof -
-  have eq_dp: "\<And> th. dependents (wq s) th = dependents (wq s') th"
-    by (unfold cs_dependents_def, auto simp:eq_dep eq_depend)
-  moreover {
-    fix th1 
-    assume "th1 \<in> {th'} \<union> dependents (wq s') th'"
-    hence "th1 = th' \<or> th1 \<in> dependents (wq s') th'" by auto
-    hence "preced th1 s = preced th1 s'"
-    proof
-      assume "th1 = th'"
-      with eq_preced[OF neq_th]
-      show "preced th1 s = preced th1 s'" by simp
-    next
-      assume "th1 \<in> dependents (wq s') th'"
-      with nd and eq_dp have "th1 \<noteq> th"
-        by (auto simp:eq_depend cs_dependents_def s_dependents_def eq_dep)
-      from eq_preced[OF this] show "preced th1 s = preced th1 s'" by simp
-    qed
-  } ultimately have "((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) = 
-                     ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" 
-    by (auto simp:image_def)
-  thus "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) =
-        Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" by simp
-qed
-
-lemma no_dependents:
-  assumes "th' \<noteq> th"
-  shows "th \<notin> dependents s th'"
-proof
-  assume h: "th \<in> dependents s th'"
-  from step_back_step [OF vt_s[unfolded s_def]]
-  have "step s' (Set th prio)" .
-  hence "th \<in> runing s'" by (cases, simp)
-  hence rd_th: "th \<in> readys s'" 
-    by (simp add:readys_def runing_def)
-  from h have "(Th th, Th th') \<in> (depend s')\<^sup>+"
-    by (unfold s_dependents_def, unfold eq_depend, unfold eq_dep, auto)
-  from tranclD[OF this]
-  obtain z where "(Th th, z) \<in> depend s'" by auto
-  with rd_th show "False"
-    apply (case_tac z, auto simp:readys_def s_waiting_def s_depend_def s_waiting_def cs_waiting_def)
-    by (fold wq_def, blast)
-qed
-
-(* Result improved *)
-lemma eq_cp:
- fixes th' 
-  assumes neq_th: "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-proof(rule eq_cp_pre [OF neq_th])
-  from no_dependents[OF neq_th] 
-  show "th \<notin> dependents s th'" .
-qed
-
-lemma eq_up:
-  fixes th' th''
-  assumes dp1: "th \<in> dependents s th'"
-  and dp2: "th' \<in> dependents s th''"
-  and eq_cps: "cp s th' = cp s' th'"
-  shows "cp s th'' = cp s' th''"
-proof -
-  from dp2
-  have "(Th th', Th th'') \<in> (depend (wq s))\<^sup>+" by (simp add:s_dependents_def)
-  from depend_child[OF vt_s this[unfolded eq_depend]]
-  have ch_th': "(Th th', Th th'') \<in> (child s)\<^sup>+" .
-  moreover { fix n th''
-    have "\<lbrakk>(Th th', n) \<in> (child s)^+\<rbrakk> \<Longrightarrow>
-                   (\<forall> th'' . n = Th th'' \<longrightarrow> cp s th'' = cp s' th'')"
-    proof(erule trancl_induct, auto)
-      fix y th''
-      assume y_ch: "(y, Th th'') \<in> child s"
-        and ih: "\<forall>th''. y = Th th'' \<longrightarrow> cp s th'' = cp s' th''"
-        and ch': "(Th th', y) \<in> (child s)\<^sup>+"
-      from y_ch obtain thy where eq_y: "y = Th thy" by (auto simp:child_def)
-      with ih have eq_cpy:"cp s thy = cp s' thy" by blast
-      from dp1 have "(Th th, Th th') \<in> (depend s)^+" by (auto simp:s_dependents_def eq_depend)
-      moreover from child_depend_p[OF ch'] and eq_y
-      have "(Th th', Th thy) \<in> (depend s)^+" by simp
-      ultimately have dp_thy: "(Th th, Th thy) \<in> (depend s)^+" by auto
-      show "cp s th'' = cp s' th''"
-        apply (subst cp_rec[OF vt_s])
-      proof -
-        have "preced th'' s = preced th'' s'"
-        proof(rule eq_preced)
-          show "th'' \<noteq> th"
-          proof
-            assume "th'' = th"
-            with dp_thy y_ch[unfolded eq_y] 
-            have "(Th th, Th th) \<in> (depend s)^+"
-              by (auto simp:child_def)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
-            show False by auto
-          qed
-        qed
-        moreover { 
-          fix th1
-          assume th1_in: "th1 \<in> children s th''"
-          have "cp s th1 = cp s' th1"
-          proof(cases "th1 = thy")
-            case True
-            with eq_cpy show ?thesis by simp
-          next
-            case False
-            have neq_th1: "th1 \<noteq> th"
-            proof
-              assume eq_th1: "th1 = th"
-              with dp_thy have "(Th th1, Th thy) \<in> (depend s)^+" by simp
-              from children_no_dep[OF vt_s _ _ this] and 
-              th1_in y_ch eq_y show False by (auto simp:children_def)
-            qed
-            have "th \<notin> dependents s th1"
-            proof
-              assume h:"th \<in> dependents s th1"
-              from eq_y dp_thy have "th \<in> dependents s thy" by (auto simp:s_dependents_def eq_depend)
-              from dependents_child_unique[OF vt_s _ _ h this]
-              th1_in y_ch eq_y have "th1 = thy" by (auto simp:children_def child_def)
-              with False show False by auto
-            qed
-            from eq_cp_pre[OF neq_th1 this]
-            show ?thesis .
-          qed
-        }
-        ultimately have "{preced th'' s} \<union> (cp s ` children s th'') = 
-          {preced th'' s'} \<union> (cp s' ` children s th'')" by (auto simp:image_def)
-        moreover have "children s th'' = children s' th''"
-          by (unfold children_def child_def s_def depend_set_unchanged, simp)
-        ultimately show "Max ({preced th'' s} \<union> cp s ` children s th'') = cp s' th''"
-          by (simp add:cp_rec[OF step_back_vt[OF vt_s[unfolded s_def]]])
-      qed
-    next
-      fix th''
-      assume dp': "(Th th', Th th'') \<in> child s"
-      show "cp s th'' = cp s' th''"
-        apply (subst cp_rec[OF vt_s])
-      proof -
-        have "preced th'' s = preced th'' s'"
-        proof(rule eq_preced)
-          show "th'' \<noteq> th"
-          proof
-            assume "th'' = th"
-            with dp1 dp'
-            have "(Th th, Th th) \<in> (depend s)^+"
-              by (auto simp:child_def s_dependents_def eq_depend)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
-            show False by auto
-          qed
-        qed
-        moreover { 
-          fix th1
-          assume th1_in: "th1 \<in> children s th''"
-          have "cp s th1 = cp s' th1"
-          proof(cases "th1 = th'")
-            case True
-            with eq_cps show ?thesis by simp
-          next
-            case False
-            have neq_th1: "th1 \<noteq> th"
-            proof
-              assume eq_th1: "th1 = th"
-              with dp1 have "(Th th1, Th th') \<in> (depend s)^+" 
-                by (auto simp:s_dependents_def eq_depend)
-              from children_no_dep[OF vt_s _ _ this]
-              th1_in dp'
-              show False by (auto simp:children_def)
-            qed
-            thus ?thesis
-            proof(rule eq_cp_pre)
-              show "th \<notin> dependents s th1"
-              proof
-                assume "th \<in> dependents s th1"
-                from dependents_child_unique[OF vt_s _ _ this dp1]
-                th1_in dp' have "th1 = th'"
-                  by (auto simp:children_def)
-                with False show False by auto
-              qed
-            qed
-          qed
-        }
-        ultimately have "{preced th'' s} \<union> (cp s ` children s th'') = 
-          {preced th'' s'} \<union> (cp s' ` children s th'')" by (auto simp:image_def)
-        moreover have "children s th'' = children s' th''"
-          by (unfold children_def child_def s_def depend_set_unchanged, simp)
-        ultimately show "Max ({preced th'' s} \<union> cp s ` children s th'') = cp s' th''"
-          by (simp add:cp_rec[OF step_back_vt[OF vt_s[unfolded s_def]]])
-      qed     
-    qed
-  }
-  ultimately show ?thesis by auto
-qed
-
-lemma eq_up_self:
-  fixes th' th''
-  assumes dp: "th \<in> dependents s th''"
-  and eq_cps: "cp s th = cp s' th"
-  shows "cp s th'' = cp s' th''"
-proof -
-  from dp
-  have "(Th th, Th th'') \<in> (depend (wq s))\<^sup>+" by (simp add:s_dependents_def)
-  from depend_child[OF vt_s this[unfolded eq_depend]]
-  have ch_th': "(Th th, Th th'') \<in> (child s)\<^sup>+" .
-  moreover { fix n th''
-    have "\<lbrakk>(Th th, n) \<in> (child s)^+\<rbrakk> \<Longrightarrow>
-                   (\<forall> th'' . n = Th th'' \<longrightarrow> cp s th'' = cp s' th'')"
-    proof(erule trancl_induct, auto)
-      fix y th''
-      assume y_ch: "(y, Th th'') \<in> child s"
-        and ih: "\<forall>th''. y = Th th'' \<longrightarrow> cp s th'' = cp s' th''"
-        and ch': "(Th th, y) \<in> (child s)\<^sup>+"
-      from y_ch obtain thy where eq_y: "y = Th thy" by (auto simp:child_def)
-      with ih have eq_cpy:"cp s thy = cp s' thy" by blast
-      from child_depend_p[OF ch'] and eq_y
-      have dp_thy: "(Th th, Th thy) \<in> (depend s)^+" by simp
-      show "cp s th'' = cp s' th''"
-        apply (subst cp_rec[OF vt_s])
-      proof -
-        have "preced th'' s = preced th'' s'"
-        proof(rule eq_preced)
-          show "th'' \<noteq> th"
-          proof
-            assume "th'' = th"
-            with dp_thy y_ch[unfolded eq_y] 
-            have "(Th th, Th th) \<in> (depend s)^+"
-              by (auto simp:child_def)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
-            show False by auto
-          qed
-        qed
-        moreover { 
-          fix th1
-          assume th1_in: "th1 \<in> children s th''"
-          have "cp s th1 = cp s' th1"
-          proof(cases "th1 = thy")
-            case True
-            with eq_cpy show ?thesis by simp
-          next
-            case False
-            have neq_th1: "th1 \<noteq> th"
-            proof
-              assume eq_th1: "th1 = th"
-              with dp_thy have "(Th th1, Th thy) \<in> (depend s)^+" by simp
-              from children_no_dep[OF vt_s _ _ this] and 
-              th1_in y_ch eq_y show False by (auto simp:children_def)
-            qed
-            have "th \<notin> dependents s th1"
-            proof
-              assume h:"th \<in> dependents s th1"
-              from eq_y dp_thy have "th \<in> dependents s thy" by (auto simp:s_dependents_def eq_depend)
-              from dependents_child_unique[OF vt_s _ _ h this]
-              th1_in y_ch eq_y have "th1 = thy" by (auto simp:children_def child_def)
-              with False show False by auto
-            qed
-            from eq_cp_pre[OF neq_th1 this]
-            show ?thesis .
-          qed
-        }
-        ultimately have "{preced th'' s} \<union> (cp s ` children s th'') = 
-          {preced th'' s'} \<union> (cp s' ` children s th'')" by (auto simp:image_def)
-        moreover have "children s th'' = children s' th''"
-          by (unfold children_def child_def s_def depend_set_unchanged, simp)
-        ultimately show "Max ({preced th'' s} \<union> cp s ` children s th'') = cp s' th''"
-          by (simp add:cp_rec[OF step_back_vt[OF vt_s[unfolded s_def]]])
-      qed
-    next
-      fix th''
-      assume dp': "(Th th, Th th'') \<in> child s"
-      show "cp s th'' = cp s' th''"
-        apply (subst cp_rec[OF vt_s])
-      proof -
-        have "preced th'' s = preced th'' s'"
-        proof(rule eq_preced)
-          show "th'' \<noteq> th"
-          proof
-            assume "th'' = th"
-            with dp dp'
-            have "(Th th, Th th) \<in> (depend s)^+"
-              by (auto simp:child_def s_dependents_def eq_depend)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
-            show False by auto
-          qed
-        qed
-        moreover { 
-          fix th1
-          assume th1_in: "th1 \<in> children s th''"
-          have "cp s th1 = cp s' th1"
-          proof(cases "th1 = th")
-            case True
-            with eq_cps show ?thesis by simp
-          next
-            case False
-            assume neq_th1: "th1 \<noteq> th"
-            thus ?thesis
-            proof(rule eq_cp_pre)
-              show "th \<notin> dependents s th1"
-              proof
-                assume "th \<in> dependents s th1"
-                hence "(Th th, Th th1) \<in> (depend s)^+" by (auto simp:s_dependents_def eq_depend)
-                from children_no_dep[OF vt_s _ _ this]
-                and th1_in dp' show False
-                  by (auto simp:children_def)
-              qed
-            qed
-          qed
-        }
-        ultimately have "{preced th'' s} \<union> (cp s ` children s th'') = 
-          {preced th'' s'} \<union> (cp s' ` children s th'')" by (auto simp:image_def)
-        moreover have "children s th'' = children s' th''"
-          by (unfold children_def child_def s_def depend_set_unchanged, simp)
-        ultimately show "Max ({preced th'' s} \<union> cp s ` children s th'') = cp s' th''"
-          by (simp add:cp_rec[OF step_back_vt[OF vt_s[unfolded s_def]]])
-      qed     
-    qed
-  }
-  ultimately show ?thesis by auto
-qed
-end
-
-lemma next_waiting:
-  assumes vt: "vt s"
-  and nxt: "next_th s th cs th'"
-  shows "waiting s th' cs"
-proof -
-  from assms show ?thesis
-    apply (auto simp:next_th_def s_waiting_def[folded wq_def])
-  proof -
-    fix rest
-    assume ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
-      and eq_wq: "wq s cs = th # rest"
-      and ne: "rest \<noteq> []"
-    have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
-    proof(rule someI2)
-      from wq_distinct[OF vt, of cs] 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
-    with 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 wq_distinct[OF vt, of cs] eq_wq
-      show "distinct rest \<and> set rest = set rest" by auto
-    next
-      from ne show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> x \<noteq> []" by auto
-    qed
-    ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
-  next
-    fix rest
-    assume eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
-      and ne: "rest \<noteq> []"
-    have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-    proof(rule someI2)
-      from wq_distinct[OF vt, of cs] eq_wq
-      show "distinct rest \<and> set rest = set rest" by auto
-    next
-      from ne show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> 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 "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
-    proof(rule someI2)
-      from wq_distinct[OF vt, of cs] 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
-    ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<in> set rest" by simp
-    with eq_wq and wq_distinct[OF vt, of cs]
-    show False by auto
-  qed
-qed
-
-
-
-
-locale step_v_cps =
-  fixes s' th cs s 
-  defines s_def : "s \<equiv> (V th cs#s')"
-  assumes vt_s: "vt s"
-
-locale step_v_cps_nt = step_v_cps +
-  fixes th'
-  assumes nt: "next_th s' th cs th'"
-
-context step_v_cps_nt
-begin
-
-lemma depend_s:
-  "depend s = (depend s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union>
-                                         {(Cs cs, Th th')}"
-proof -
-  from step_depend_v[OF vt_s[unfolded s_def], folded s_def]
-    and nt show ?thesis  by (auto intro:next_th_unique)
-qed
-
-lemma dependents_kept:
-  fixes th''
-  assumes neq1: "th'' \<noteq> th"
-  and neq2: "th'' \<noteq> th'"
-  shows "dependents (wq s) th'' = dependents (wq s') th''"
-proof(auto)
-  fix x
-  assume "x \<in> dependents (wq s) th''"
-  hence dp: "(Th x, Th th'') \<in> (depend s)^+"
-    by (auto simp:cs_dependents_def eq_depend)
-  { fix n
-    have "(n, Th th'') \<in> (depend s)^+ \<Longrightarrow>  (n, Th th'') \<in> (depend s')^+"
-    proof(induct rule:converse_trancl_induct)
-      fix y 
-      assume "(y, Th th'') \<in> depend s"
-      with depend_s neq1 neq2
-      have "(y, Th th'') \<in> depend s'" by auto
-      thus "(y, Th th'') \<in> (depend s')\<^sup>+" by auto
-    next
-      fix y z 
-      assume yz: "(y, z) \<in> depend s"
-        and ztp: "(z, Th th'') \<in> (depend s)\<^sup>+"
-        and ztp': "(z, Th th'') \<in> (depend s')\<^sup>+"
-      have "y \<noteq> Cs cs \<and> y \<noteq> Th th'"
-      proof
-        show "y \<noteq> Cs cs"
-        proof
-          assume eq_y: "y = Cs cs"
-          with yz have dp_yz: "(Cs cs, z) \<in> depend s" by simp
-          from depend_s
-          have cst': "(Cs cs, Th th') \<in> depend s" by simp
-          from unique_depend[OF vt_s this dp_yz] 
-          have eq_z: "z = Th th'" by simp
-          with ztp have "(Th th', Th th'') \<in> (depend s)^+" by simp
-          from converse_tranclE[OF this]
-          obtain cs' where dp'': "(Th th', Cs cs') \<in> depend s"
-            by (auto simp:s_depend_def)
-          with depend_s have dp': "(Th th', Cs cs') \<in> depend s'" by auto
-          from dp'' eq_y yz eq_z have "(Cs cs, Cs cs') \<in> (depend s)^+" by auto
-          moreover have "cs' = cs"
-          proof -
-            from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt]
-            have "(Th th', Cs cs) \<in> depend s'"
-              by (auto simp:s_waiting_def wq_def s_depend_def cs_waiting_def)
-            from unique_depend[OF step_back_vt[OF vt_s[unfolded s_def]] this dp']
-            show ?thesis by simp
-          qed
-          ultimately have "(Cs cs, Cs cs) \<in> (depend s)^+" by simp
-          moreover note wf_trancl[OF wf_depend[OF vt_s]]
-          ultimately show False by auto
-        qed
-      next
-        show "y \<noteq> Th th'"
-        proof
-          assume eq_y: "y = Th th'"
-          with yz have dps: "(Th th', z) \<in> depend s" by simp
-          with depend_s have dps': "(Th th', z) \<in> depend s'" by auto
-          have "z = Cs cs"
-          proof -
-            from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt]
-            have "(Th th', Cs cs) \<in> depend s'"
-              by (auto simp:s_waiting_def wq_def s_depend_def cs_waiting_def)
-            from unique_depend[OF step_back_vt[OF vt_s[unfolded s_def]] dps' this]
-            show ?thesis .
-          qed
-          with dps depend_s show False by auto
-        qed
-      qed
-      with depend_s yz have "(y, z) \<in> depend s'" by auto
-      with ztp'
-      show "(y, Th th'') \<in> (depend s')\<^sup>+" by auto
-    qed    
-  }
-  from this[OF dp]
-  show "x \<in> dependents (wq s') th''" 
-    by (auto simp:cs_dependents_def eq_depend)
-next
-  fix x
-  assume "x \<in> dependents (wq s') th''"
-  hence dp: "(Th x, Th th'') \<in> (depend s')^+"
-    by (auto simp:cs_dependents_def eq_depend)
-  { fix n
-    have "(n, Th th'') \<in> (depend s')^+ \<Longrightarrow>  (n, Th th'') \<in> (depend s)^+"
-    proof(induct rule:converse_trancl_induct)
-      fix y 
-      assume "(y, Th th'') \<in> depend s'"
-      with depend_s neq1 neq2
-      have "(y, Th th'') \<in> depend s" by auto
-      thus "(y, Th th'') \<in> (depend s)\<^sup>+" by auto
-    next
-      fix y z 
-      assume yz: "(y, z) \<in> depend s'"
-        and ztp: "(z, Th th'') \<in> (depend s')\<^sup>+"
-        and ztp': "(z, Th th'') \<in> (depend s)\<^sup>+"
-      have "y \<noteq> Cs cs \<and> y \<noteq> Th th'"
-      proof
-        show "y \<noteq> Cs cs"
-        proof
-          assume eq_y: "y = Cs cs"
-          with yz have dp_yz: "(Cs cs, z) \<in> depend s'" by simp
-          from this have eq_z: "z = Th th"
-          proof -
-            from step_back_step[OF vt_s[unfolded s_def]]
-            have "(Cs cs, Th th) \<in> depend s'"
-              by(cases, auto simp: wq_def s_depend_def cs_holding_def s_holding_def)
-            from unique_depend[OF step_back_vt[OF vt_s[unfolded s_def]] this dp_yz]
-            show ?thesis by simp
-          qed
-          from converse_tranclE[OF ztp]
-          obtain u where "(z, u) \<in> depend s'" by auto
-          moreover 
-          from step_back_step[OF vt_s[unfolded s_def]]
-          have "th \<in> readys s'" by (cases, simp add:runing_def)
-          moreover note eq_z
-          ultimately show False 
-            by (auto simp:readys_def wq_def s_depend_def s_waiting_def cs_waiting_def)
-        qed
-      next
-        show "y \<noteq> Th th'"
-        proof
-          assume eq_y: "y = Th th'"
-          with yz have dps: "(Th th', z) \<in> depend s'" by simp
-          have "z = Cs cs"
-          proof -
-            from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt]
-            have "(Th th', Cs cs) \<in> depend s'"
-              by (auto simp:s_waiting_def wq_def s_depend_def cs_waiting_def)
-            from unique_depend[OF step_back_vt[OF vt_s[unfolded s_def]] dps this]
-            show ?thesis .
-          qed
-          with ztp have cs_i: "(Cs cs, Th th'') \<in>  (depend s')\<^sup>+" by simp
-          from step_back_step[OF vt_s[unfolded s_def]]
-          have cs_th: "(Cs cs, Th th) \<in> depend s'"
-            by(cases, auto simp: s_depend_def wq_def cs_holding_def s_holding_def)
-          have "(Cs cs, Th th'') \<notin>  depend s'"
-          proof
-            assume "(Cs cs, Th th'') \<in> depend s'"
-            from unique_depend[OF step_back_vt[OF vt_s[unfolded s_def]] this cs_th]
-            and neq1 show "False" by simp
-          qed
-          with converse_tranclE[OF cs_i]
-          obtain u where cu: "(Cs cs, u) \<in> depend s'"  
-            and u_t: "(u, Th th'') \<in> (depend s')\<^sup>+" by auto
-          have "u = Th th"
-          proof -
-            from unique_depend[OF step_back_vt[OF vt_s[unfolded s_def]] cu cs_th]
-            show ?thesis .
-          qed
-          with u_t have "(Th th, Th th'') \<in> (depend s')\<^sup>+" by simp
-          from converse_tranclE[OF this]
-          obtain v where "(Th th, v) \<in> (depend s')" by auto
-          moreover from step_back_step[OF vt_s[unfolded s_def]]
-          have "th \<in> readys s'" by (cases, simp add:runing_def)
-          ultimately show False 
-            by (auto simp:readys_def wq_def s_depend_def s_waiting_def cs_waiting_def)
-        qed
-      qed
-      with depend_s yz have "(y, z) \<in> depend s" by auto
-      with ztp'
-      show "(y, Th th'') \<in> (depend s)\<^sup>+" by auto
-    qed    
-  }
-  from this[OF dp]
-  show "x \<in> dependents (wq s) th''"
-    by (auto simp:cs_dependents_def eq_depend)
-qed
-
-lemma cp_kept:
-  fixes th''
-  assumes neq1: "th'' \<noteq> th"
-  and neq2: "th'' \<noteq> th'"
-  shows "cp s th'' = cp s' th''"
-proof -
-  from dependents_kept[OF neq1 neq2]
-  have "dependents (wq s) th'' = dependents (wq s') th''" .
-  moreover {
-    fix th1
-    assume "th1 \<in> dependents (wq s) th''"
-    have "preced th1 s = preced th1 s'" 
-      by (unfold s_def, auto simp:preced_def)
-  }
-  moreover have "preced th'' s = preced th'' s'" 
-    by (unfold s_def, auto simp:preced_def)
-  ultimately have "((\<lambda>th. preced th s) ` ({th''} \<union> dependents (wq s) th'')) = 
-    ((\<lambda>th. preced th s') ` ({th''} \<union> dependents (wq s') th''))"
-    by (auto simp:image_def)
-  thus ?thesis
-    by (unfold cp_eq_cpreced cpreced_def, simp)
-qed
-
-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 nw_cs: "(Th th1, Cs cs) \<notin> depend s'"
-proof
-  assume "(Th th1, Cs cs) \<in> depend s'"
-  thus "False"
-    apply (auto simp:s_depend_def cs_waiting_def)
-  proof -
-    assume h1: "th1 \<in> set (wq s' cs)"
-      and h2: "th1 \<noteq> hd (wq s' cs)"
-    from step_back_step[OF vt_s[unfolded s_def]]
-    show "False"
-    proof(cases)
-      assume "holding s' th cs" 
-      then obtain rest where
-        eq_wq: "wq s' cs = th#rest"
-        apply (unfold s_holding_def wq_def[symmetric])
-        by (case_tac "(wq s' cs)", auto)
-      with h1 h2 have ne: "rest \<noteq> []" by auto
-      with eq_wq
-      have "next_th s' th cs (hd (SOME q. distinct q \<and> set q = set rest))"
-        by(unfold next_th_def, rule_tac x = "rest" in exI, auto)
-      with nnt show ?thesis by auto
-    qed
-  qed
-qed
-
-lemma depend_s: "depend s = depend s' - {(Cs cs, Th th)}"
-proof -
-  from nnt and  step_depend_v[OF vt_s[unfolded s_def], folded s_def]
-  show ?thesis by auto
-qed
-
-lemma child_kept_left:
-  assumes 
-  "(n1, n2) \<in> (child s')^+"
-  shows "(n1, n2) \<in> (child s)^+"
-proof -
-  from assms show ?thesis 
-  proof(induct rule: converse_trancl_induct)
-    case (base y)
-    from base obtain th1 cs1 th2
-      where h1: "(Th th1, Cs cs1) \<in> depend s'"
-      and h2: "(Cs cs1, Th th2) \<in> depend s'"
-      and eq_y: "y = Th th1" and eq_n2: "n2 = Th th2"  by (auto simp:child_def)
-    have "cs1 \<noteq> cs"
-    proof
-      assume eq_cs: "cs1 = cs"
-      with h1 have "(Th th1, Cs cs1) \<in> depend s'" by simp
-      with nw_cs eq_cs show False by auto
-    qed
-    with h1 h2 depend_s have 
-      h1': "(Th th1, Cs cs1) \<in> depend s" and
-      h2': "(Cs cs1, Th th2) \<in> depend s" by auto
-    hence "(Th th1, Th th2) \<in> child s" by (auto simp:child_def)
-    with eq_y eq_n2 have "(y, n2) \<in> child s" by simp
-    thus ?case by auto
-  next
-    case (step y z)
-    have "(y, z) \<in> child s'" by fact
-    then obtain th1 cs1 th2
-      where h1: "(Th th1, Cs cs1) \<in> depend s'"
-      and h2: "(Cs cs1, Th th2) \<in> depend s'"
-      and eq_y: "y = Th th1" and eq_z: "z = Th th2"  by (auto simp:child_def)
-    have "cs1 \<noteq> cs"
-    proof
-      assume eq_cs: "cs1 = cs"
-      with h1 have "(Th th1, Cs cs1) \<in> depend s'" by simp
-      with nw_cs eq_cs show False by auto
-    qed
-    with h1 h2 depend_s have 
-      h1': "(Th th1, Cs cs1) \<in> depend s" and
-      h2': "(Cs cs1, Th th2) \<in> depend s" by auto
-    hence "(Th th1, Th th2) \<in> child s" by (auto simp:child_def)
-    with eq_y eq_z have "(y, z) \<in> child s" by simp
-    moreover have "(z, n2) \<in> (child s)^+" by fact
-    ultimately show ?case by auto
-  qed
-qed
-
-lemma  child_kept_right:
-  assumes
-  "(n1, n2) \<in> (child s)^+"
-  shows "(n1, n2) \<in> (child s')^+"
-proof -
-  from assms show ?thesis
-  proof(induct)
-    case (base y)
-    from base and depend_s 
-    have "(n1, y) \<in> child s'"
-      by (auto simp:child_def)
-    thus ?case by auto
-  next
-    case (step y z)
-    have "(y, z) \<in> child s" by fact
-    with depend_s have "(y, z) \<in> child s'"
-      by (auto simp:child_def)
-    moreover have "(n1, y) \<in> (child s')\<^sup>+" by fact
-    ultimately show ?case by auto
-  qed
-qed
-
-lemma eq_child: "(child s)^+ = (child s')^+"
-  by (insert child_kept_left child_kept_right, auto)
-
-lemma eq_cp:
-  fixes th' 
-  shows "cp s th' = cp s' th'"
-  apply (unfold cp_eq_cpreced cpreced_def)
-proof -
-  have eq_dp: "\<And> th. dependents (wq s) th = dependents (wq s') th"
-    apply (unfold cs_dependents_def, unfold eq_depend)
-  proof -
-    from eq_child
-    have "\<And>th. {th'. (Th th', Th th) \<in> (child s)\<^sup>+} = {th'. (Th th', Th th) \<in> (child s')\<^sup>+}"
-      by simp
-    with child_depend_eq[OF vt_s] child_depend_eq[OF step_back_vt[OF vt_s[unfolded s_def]]]
-    show "\<And>th. {th'. (Th th', Th th) \<in> (depend s)\<^sup>+} = {th'. (Th th', Th th) \<in> (depend s')\<^sup>+}"
-      by simp
-  qed
-  moreover {
-    fix th1 
-    assume "th1 \<in> {th'} \<union> dependents (wq s') th'"
-    hence "th1 = th' \<or> th1 \<in> dependents (wq s') th'" by auto
-    hence "preced th1 s = preced th1 s'"
-    proof
-      assume "th1 = th'"
-      show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def)
-    next
-      assume "th1 \<in> dependents (wq s') th'"
-      show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def)
-    qed
-  } ultimately have "((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) = 
-                     ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" 
-    by (auto simp:image_def)
-  thus "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) =
-        Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" by simp
-qed
-
-end
-
-locale step_P_cps =
-  fixes s' th cs s 
-  defines s_def : "s \<equiv> (P th cs#s')"
-  assumes vt_s: "vt s"
-
-locale step_P_cps_ne =step_P_cps +
-  assumes ne: "wq s' cs \<noteq> []"
-
-locale step_P_cps_e =step_P_cps +
-  assumes ee: "wq s' cs = []"
-
-context step_P_cps_e
-begin
-
-lemma depend_s: "depend s = depend s' \<union> {(Cs cs, Th th)}"
-proof -
-  from ee and  step_depend_p[OF vt_s[unfolded s_def], folded s_def]
-  show ?thesis by auto
-qed
-
-lemma child_kept_left:
-  assumes 
-  "(n1, n2) \<in> (child s')^+"
-  shows "(n1, n2) \<in> (child s)^+"
-proof -
-  from assms show ?thesis 
-  proof(induct rule: converse_trancl_induct)
-    case (base y)
-    from base obtain th1 cs1 th2
-      where h1: "(Th th1, Cs cs1) \<in> depend s'"
-      and h2: "(Cs cs1, Th th2) \<in> depend s'"
-      and eq_y: "y = Th th1" and eq_n2: "n2 = Th th2"  by (auto simp:child_def)
-    have "cs1 \<noteq> cs"
-    proof
-      assume eq_cs: "cs1 = cs"
-      with h1 have "(Th th1, Cs cs) \<in> depend s'" by simp
-      with ee show False
-        by (auto simp:s_depend_def cs_waiting_def)
-    qed
-    with h1 h2 depend_s have 
-      h1': "(Th th1, Cs cs1) \<in> depend s" and
-      h2': "(Cs cs1, Th th2) \<in> depend s" by auto
-    hence "(Th th1, Th th2) \<in> child s" by (auto simp:child_def)
-    with eq_y eq_n2 have "(y, n2) \<in> child s" by simp
-    thus ?case by auto
-  next
-    case (step y z)
-    have "(y, z) \<in> child s'" by fact
-    then obtain th1 cs1 th2
-      where h1: "(Th th1, Cs cs1) \<in> depend s'"
-      and h2: "(Cs cs1, Th th2) \<in> depend s'"
-      and eq_y: "y = Th th1" and eq_z: "z = Th th2"  by (auto simp:child_def)
-    have "cs1 \<noteq> cs"
-    proof
-      assume eq_cs: "cs1 = cs"
-      with h1 have "(Th th1, Cs cs) \<in> depend s'" by simp
-      with ee show False 
-        by (auto simp:s_depend_def cs_waiting_def)
-    qed
-    with h1 h2 depend_s have 
-      h1': "(Th th1, Cs cs1) \<in> depend s" and
-      h2': "(Cs cs1, Th th2) \<in> depend s" by auto
-    hence "(Th th1, Th th2) \<in> child s" by (auto simp:child_def)
-    with eq_y eq_z have "(y, z) \<in> child s" by simp
-    moreover have "(z, n2) \<in> (child s)^+" by fact
-    ultimately show ?case by auto
-  qed
-qed
-
-lemma  child_kept_right:
-  assumes
-  "(n1, n2) \<in> (child s)^+"
-  shows "(n1, n2) \<in> (child s')^+"
-proof -
-  from assms show ?thesis
-  proof(induct)
-    case (base y)
-    from base and depend_s
-    have "(n1, y) \<in> child s'"
-      apply (auto simp:child_def)
-      proof -
-        fix th'
-        assume "(Th th', Cs cs) \<in> depend s'"
-        with ee have "False"
-          by (auto simp:s_depend_def cs_waiting_def)
-        thus "\<exists>cs. (Th th', Cs cs) \<in> depend s' \<and> (Cs cs, Th th) \<in> depend s'" by auto 
-      qed
-    thus ?case by auto
-  next
-    case (step y z)
-    have "(y, z) \<in> child s" by fact
-    with depend_s have "(y, z) \<in> child s'"
-      apply (auto simp:child_def)
-      proof -
-        fix th'
-        assume "(Th th', Cs cs) \<in> depend s'"
-        with ee have "False"
-          by (auto simp:s_depend_def cs_waiting_def)
-        thus "\<exists>cs. (Th th', Cs cs) \<in> depend s' \<and> (Cs cs, Th th) \<in> depend s'" by auto 
-      qed
-    moreover have "(n1, y) \<in> (child s')\<^sup>+" by fact
-    ultimately show ?case by auto
-  qed
-qed
-
-lemma eq_child: "(child s)^+ = (child s')^+"
-  by (insert child_kept_left child_kept_right, auto)
-
-lemma eq_cp:
-  fixes th' 
-  shows "cp s th' = cp s' th'"
-  apply (unfold cp_eq_cpreced cpreced_def)
-proof -
-  have eq_dp: "\<And> th. dependents (wq s) th = dependents (wq s') th"
-    apply (unfold cs_dependents_def, unfold eq_depend)
-  proof -
-    from eq_child
-    have "\<And>th. {th'. (Th th', Th th) \<in> (child s)\<^sup>+} = {th'. (Th th', Th th) \<in> (child s')\<^sup>+}"
-      by auto
-    with child_depend_eq[OF vt_s] child_depend_eq[OF step_back_vt[OF vt_s[unfolded s_def]]]
-    show "\<And>th. {th'. (Th th', Th th) \<in> (depend s)\<^sup>+} = {th'. (Th th', Th th) \<in> (depend s')\<^sup>+}"
-      by simp
-  qed
-  moreover {
-    fix th1 
-    assume "th1 \<in> {th'} \<union> dependents (wq s') th'"
-    hence "th1 = th' \<or> th1 \<in> dependents (wq s') th'" by auto
-    hence "preced th1 s = preced th1 s'"
-    proof
-      assume "th1 = th'"
-      show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def)
-    next
-      assume "th1 \<in> dependents (wq s') th'"
-      show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def)
-    qed
-  } ultimately have "((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) = 
-                     ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" 
-    by (auto simp:image_def)
-  thus "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) =
-        Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" by simp
-qed
-
-end
-
-context step_P_cps_ne
-begin
-
-lemma depend_s: "depend s = depend s' \<union> {(Th th, Cs cs)}"
-proof -
-  from step_depend_p[OF vt_s[unfolded s_def]] and ne
-  show ?thesis by (simp add:s_def)
-qed
-
-lemma eq_child_left:
-  assumes nd: "(Th th, Th th') \<notin> (child s)^+"
-  shows "(n1, Th th') \<in> (child s)^+ \<Longrightarrow> (n1, Th th') \<in> (child s')^+"
-proof(induct rule:converse_trancl_induct)
-  case (base y)
-  from base obtain th1 cs1
-    where h1: "(Th th1, Cs cs1) \<in> depend s"
-    and h2: "(Cs cs1, Th th') \<in> depend s"
-    and eq_y: "y = Th th1"   by (auto simp:child_def)
-  have "th1 \<noteq> th"
-  proof
-    assume "th1 = th"
-    with base eq_y have "(Th th, Th th') \<in> child s" by simp
-    with nd show False by auto
-  qed
-  with h1 h2 depend_s 
-  have h1': "(Th th1, Cs cs1) \<in> depend s'" and 
-       h2': "(Cs cs1, Th th') \<in> depend s'" by auto
-  with eq_y show ?case by (auto simp:child_def)
-next
-  case (step y z)
-  have yz: "(y, z) \<in> child s" by fact
-  then obtain th1 cs1 th2
-    where h1: "(Th th1, Cs cs1) \<in> depend s"
-    and h2: "(Cs cs1, Th th2) \<in> depend s"
-    and eq_y: "y = Th th1" and eq_z: "z = Th th2"  by (auto simp:child_def)
-  have "th1 \<noteq> th"
-  proof
-    assume "th1 = th"
-    with yz eq_y have "(Th th, z) \<in> child s" by simp
-    moreover have "(z, Th th') \<in> (child s)^+" by fact
-    ultimately have "(Th th, Th th') \<in> (child s)^+" by auto
-    with nd show False by auto
-  qed
-  with h1 h2 depend_s have h1': "(Th th1, Cs cs1) \<in> depend s'"
-                       and h2': "(Cs cs1, Th th2) \<in> depend s'" by auto
-  with eq_y eq_z have "(y, z) \<in> child s'" by (auto simp:child_def)
-  moreover have "(z, Th th') \<in> (child s')^+" by fact
-  ultimately show ?case by auto
-qed
-
-lemma eq_child_right:
-  shows "(n1, Th th') \<in> (child s')^+ \<Longrightarrow> (n1, Th th') \<in> (child s)^+"
-proof(induct rule:converse_trancl_induct)
-  case (base y)
-  with depend_s show ?case by (auto simp:child_def)
-next
-  case (step y z)
-  have "(y, z) \<in> child s'" by fact
-  with depend_s have "(y, z) \<in> child s" by (auto simp:child_def)
-  moreover have "(z, Th th') \<in> (child s)^+" by fact
-  ultimately show ?case by auto
-qed
-
-lemma eq_child:
-  assumes nd: "(Th th, Th th') \<notin> (child s)^+"
-  shows "((n1, Th th') \<in> (child s)^+) = ((n1, Th th') \<in> (child s')^+)"
-  by (insert eq_child_left[OF nd] eq_child_right, auto)
-
-lemma eq_cp:
-  fixes th' 
-  assumes nd: "th \<notin> dependents s th'"
-  shows "cp s th' = cp s' th'"
-  apply (unfold cp_eq_cpreced cpreced_def)
-proof -
-  have nd': "(Th th, Th th') \<notin> (child s)^+"
-  proof
-    assume "(Th th, Th th') \<in> (child s)\<^sup>+"
-    with child_depend_eq[OF vt_s]
-    have "(Th th, Th th') \<in> (depend s)\<^sup>+" by simp
-    with nd show False 
-      by (simp add:s_dependents_def eq_depend)
-  qed
-  have eq_dp: "dependents (wq s) th' = dependents (wq s') th'"
-  proof(auto)
-    fix x assume " x \<in> dependents (wq s) th'"
-    thus "x \<in> dependents (wq s') th'"
-      apply (auto simp:cs_dependents_def eq_depend)
-    proof -
-      assume "(Th x, Th th') \<in> (depend s)\<^sup>+"
-      with  child_depend_eq[OF vt_s] have "(Th x, Th th') \<in> (child s)\<^sup>+" by simp
-      with eq_child[OF nd'] have "(Th x, Th th') \<in> (child s')^+" by simp
-      with child_depend_eq[OF step_back_vt[OF vt_s[unfolded s_def]]]
-      show "(Th x, Th th') \<in> (depend s')\<^sup>+" by simp
-    qed
-  next
-    fix x assume "x \<in> dependents (wq s') th'"
-    thus "x \<in> dependents (wq s) th'"
-      apply (auto simp:cs_dependents_def eq_depend)
-    proof -
-      assume "(Th x, Th th') \<in> (depend s')\<^sup>+"
-      with child_depend_eq[OF step_back_vt[OF vt_s[unfolded s_def]]] 
-      have "(Th x, Th th') \<in> (child s')\<^sup>+" by simp
-      with eq_child[OF nd'] have "(Th x, Th th') \<in> (child s)^+" by simp
-      with  child_depend_eq[OF vt_s]
-      show "(Th x, Th th') \<in> (depend s)\<^sup>+" by simp
-    qed
-  qed
-  moreover {
-    fix th1 have "preced th1 s = preced th1 s'" by (simp add:s_def preced_def)
-  } ultimately have "((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) = 
-                     ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" 
-    by (auto simp:image_def)
-  thus "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) =
-        Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" by simp
-qed
-
-lemma eq_up:
-  fixes th' th''
-  assumes dp1: "th \<in> dependents s th'"
-  and dp2: "th' \<in> dependents s th''"
-  and eq_cps: "cp s th' = cp s' th'"
-  shows "cp s th'' = cp s' th''"
-proof -
-  from dp2
-  have "(Th th', Th th'') \<in> (depend (wq s))\<^sup>+" by (simp add:s_dependents_def)
-  from depend_child[OF vt_s this[unfolded eq_depend]]
-  have ch_th': "(Th th', Th th'') \<in> (child s)\<^sup>+" .
-  moreover {
-    fix n th''
-    have "\<lbrakk>(Th th', n) \<in> (child s)^+\<rbrakk> \<Longrightarrow>
-                   (\<forall> th'' . n = Th th'' \<longrightarrow> cp s th'' = cp s' th'')"
-    proof(erule trancl_induct, auto)
-      fix y th''
-      assume y_ch: "(y, Th th'') \<in> child s"
-        and ih: "\<forall>th''. y = Th th'' \<longrightarrow> cp s th'' = cp s' th''"
-        and ch': "(Th th', y) \<in> (child s)\<^sup>+"
-      from y_ch obtain thy where eq_y: "y = Th thy" by (auto simp:child_def)
-      with ih have eq_cpy:"cp s thy = cp s' thy" by blast
-      from dp1 have "(Th th, Th th') \<in> (depend s)^+" by (auto simp:s_dependents_def eq_depend)
-      moreover from child_depend_p[OF ch'] and eq_y
-      have "(Th th', Th thy) \<in> (depend s)^+" by simp
-      ultimately have dp_thy: "(Th th, Th thy) \<in> (depend s)^+" by auto
-      show "cp s th'' = cp s' th''"
-        apply (subst cp_rec[OF vt_s])
-      proof -
-        have "preced th'' s = preced th'' s'"
-          by (simp add:s_def preced_def)
-        moreover { 
-          fix th1
-          assume th1_in: "th1 \<in> children s th''"
-          have "cp s th1 = cp s' th1"
-          proof(cases "th1 = thy")
-            case True
-            with eq_cpy show ?thesis by simp
-          next
-            case False
-            have neq_th1: "th1 \<noteq> th"
-            proof
-              assume eq_th1: "th1 = th"
-              with dp_thy have "(Th th1, Th thy) \<in> (depend s)^+" by simp
-              from children_no_dep[OF vt_s _ _ this] and 
-              th1_in y_ch eq_y show False by (auto simp:children_def)
-            qed
-            have "th \<notin> dependents s th1"
-            proof
-              assume h:"th \<in> dependents s th1"
-              from eq_y dp_thy have "th \<in> dependents s thy" by (auto simp:s_dependents_def eq_depend)
-              from dependents_child_unique[OF vt_s _ _ h this]
-              th1_in y_ch eq_y have "th1 = thy" by (auto simp:children_def child_def)
-              with False show False by auto
-            qed
-            from eq_cp[OF this]
-            show ?thesis .
-          qed
-        }
-        ultimately have "{preced th'' s} \<union> (cp s ` children s th'') = 
-          {preced th'' s'} \<union> (cp s' ` children s th'')" by (auto simp:image_def)
-        moreover have "children s th'' = children s' th''"
-          apply (unfold children_def child_def s_def depend_set_unchanged, simp)
-          apply (fold s_def, auto simp:depend_s)
-          proof -
-            assume "(Cs cs, Th th'') \<in> depend s'"
-            with depend_s have cs_th': "(Cs cs, Th th'') \<in> depend s" by auto
-            from dp1 have "(Th th, Th th') \<in> (depend s)^+"
-              by (auto simp:s_dependents_def eq_depend)
-            from converse_tranclE[OF this]
-            obtain cs1 where h1: "(Th th, Cs cs1) \<in> depend s"
-              and h2: "(Cs cs1 , Th th') \<in> (depend s)\<^sup>+"
-              by (auto simp:s_depend_def)
-            have eq_cs: "cs1 = cs" 
-            proof -
-              from depend_s have "(Th th, Cs cs) \<in> depend s" by simp
-              from unique_depend[OF vt_s this h1]
-              show ?thesis by simp
-            qed
-            have False
-            proof(rule converse_tranclE[OF h2])
-              assume "(Cs cs1, Th th') \<in> depend s"
-              with eq_cs have "(Cs cs, Th th') \<in> depend s" by simp
-              from unique_depend[OF vt_s this cs_th']
-              have "th' = th''" by simp
-              with ch' y_ch have "(Th th'', Th th'') \<in> (child s)^+" by auto
-              with wf_trancl[OF wf_child[OF vt_s]] 
-              show False by auto
-            next
-              fix y
-              assume "(Cs cs1, y) \<in> depend s"
-                and ytd: " (y, Th th') \<in> (depend s)\<^sup>+"
-              with eq_cs have csy: "(Cs cs, y) \<in> depend s" by simp
-              from unique_depend[OF vt_s this cs_th']
-              have "y = Th th''" .
-              with ytd have "(Th th'', Th th') \<in> (depend s)^+" by simp
-              from depend_child[OF vt_s this]
-              have "(Th th'', Th th') \<in> (child s)\<^sup>+" .
-              moreover from ch' y_ch have ch'': "(Th th', Th th'') \<in> (child s)^+" by auto
-              ultimately have "(Th th'', Th th'') \<in> (child s)^+" by auto 
-              with wf_trancl[OF wf_child[OF vt_s]] 
-              show False by auto
-            qed
-            thus "\<exists>cs. (Th th, Cs cs) \<in> depend s' \<and> (Cs cs, Th th'') \<in> depend s'" by auto
-          qed
-          ultimately show "Max ({preced th'' s} \<union> cp s ` children s th'') = cp s' th''"
-          by (simp add:cp_rec[OF step_back_vt[OF vt_s[unfolded s_def]]])
-      qed
-    next
-      fix th''
-      assume dp': "(Th th', Th th'') \<in> child s"
-      show "cp s th'' = cp s' th''"
-        apply (subst cp_rec[OF vt_s])
-      proof -
-        have "preced th'' s = preced th'' s'"
-          by (simp add:s_def preced_def)
-        moreover { 
-          fix th1
-          assume th1_in: "th1 \<in> children s th''"
-          have "cp s th1 = cp s' th1"
-          proof(cases "th1 = th'")
-            case True
-            with eq_cps show ?thesis by simp
-          next
-            case False
-            have neq_th1: "th1 \<noteq> th"
-            proof
-              assume eq_th1: "th1 = th"
-              with dp1 have "(Th th1, Th th') \<in> (depend s)^+" 
-                by (auto simp:s_dependents_def eq_depend)
-              from children_no_dep[OF vt_s _ _ this]
-              th1_in dp'
-              show False by (auto simp:children_def)
-            qed
-            show ?thesis
-            proof(rule eq_cp)
-              show "th \<notin> dependents s th1"
-              proof
-                assume "th \<in> dependents s th1"
-                from dependents_child_unique[OF vt_s _ _ this dp1]
-                th1_in dp' have "th1 = th'"
-                  by (auto simp:children_def)
-                with False show False by auto
-              qed
-            qed
-          qed
-        }
-        ultimately have "{preced th'' s} \<union> (cp s ` children s th'') = 
-          {preced th'' s'} \<union> (cp s' ` children s th'')" by (auto simp:image_def)
-        moreover have "children s th'' = children s' th''"
-          apply (unfold children_def child_def s_def depend_set_unchanged, simp)
-          apply (fold s_def, auto simp:depend_s)
-          proof -
-            assume "(Cs cs, Th th'') \<in> depend s'"
-            with depend_s have cs_th': "(Cs cs, Th th'') \<in> depend s" by auto
-            from dp1 have "(Th th, Th th') \<in> (depend s)^+"
-              by (auto simp:s_dependents_def eq_depend)
-            from converse_tranclE[OF this]
-            obtain cs1 where h1: "(Th th, Cs cs1) \<in> depend s"
-              and h2: "(Cs cs1 , Th th') \<in> (depend s)\<^sup>+"
-              by (auto simp:s_depend_def)
-            have eq_cs: "cs1 = cs" 
-            proof -
-              from depend_s have "(Th th, Cs cs) \<in> depend s" by simp
-              from unique_depend[OF vt_s this h1]
-              show ?thesis by simp
-            qed
-            have False
-            proof(rule converse_tranclE[OF h2])
-              assume "(Cs cs1, Th th') \<in> depend s"
-              with eq_cs have "(Cs cs, Th th') \<in> depend s" by simp
-              from unique_depend[OF vt_s this cs_th']
-              have "th' = th''" by simp
-              with dp' have "(Th th'', Th th'') \<in> (child s)^+" by auto
-              with wf_trancl[OF wf_child[OF vt_s]] 
-              show False by auto
-            next
-              fix y
-              assume "(Cs cs1, y) \<in> depend s"
-                and ytd: " (y, Th th') \<in> (depend s)\<^sup>+"
-              with eq_cs have csy: "(Cs cs, y) \<in> depend s" by simp
-              from unique_depend[OF vt_s this cs_th']
-              have "y = Th th''" .
-              with ytd have "(Th th'', Th th') \<in> (depend s)^+" by simp
-              from depend_child[OF vt_s this]
-              have "(Th th'', Th th') \<in> (child s)\<^sup>+" .
-              moreover from dp' have ch'': "(Th th', Th th'') \<in> (child s)^+" by auto
-              ultimately have "(Th th'', Th th'') \<in> (child s)^+" by auto 
-              with wf_trancl[OF wf_child[OF vt_s]] 
-              show False by auto
-            qed
-            thus "\<exists>cs. (Th th, Cs cs) \<in> depend s' \<and> (Cs cs, Th th'') \<in> depend s'" by auto
-          qed
-        ultimately show "Max ({preced th'' s} \<union> cp s ` children s th'') = cp s' th''"
-          by (simp add:cp_rec[OF step_back_vt[OF vt_s[unfolded s_def]]])
-      qed     
-    qed
-  }
-  ultimately show ?thesis by auto
-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"
-
-context step_create_cps
-begin
-
-lemma eq_dep: "depend s = depend s'"
-  by (unfold s_def depend_create_unchanged, auto)
-
-lemma eq_cp:
-  fixes th' 
-  assumes neq_th: "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-  apply (unfold cp_eq_cpreced cpreced_def)
-proof -
-  have nd: "th \<notin> dependents s th'"
-  proof
-    assume "th \<in> dependents s th'"
-    hence "(Th th, Th th') \<in> (depend s)^+" by (simp add:s_dependents_def eq_depend)
-    with eq_dep have "(Th th, Th th') \<in> (depend s')^+" by simp
-    from converse_tranclE[OF this]
-    obtain y where "(Th th, y) \<in> depend s'" by auto
-    with dm_depend_threads[OF step_back_vt[OF vt_s[unfolded s_def]]]
-    have in_th: "th \<in> threads s'" by auto
-    from step_back_step[OF vt_s[unfolded s_def]]
-    show False
-    proof(cases)
-      assume "th \<notin> threads s'" 
-      with in_th show ?thesis by simp
-    qed
-  qed
-  have eq_dp: "\<And> th. dependents (wq s) th = dependents (wq s') th"
-    by (unfold cs_dependents_def, auto simp:eq_dep eq_depend)
-  moreover {
-    fix th1 
-    assume "th1 \<in> {th'} \<union> dependents (wq s') th'"
-    hence "th1 = th' \<or> th1 \<in> dependents (wq s') th'" by auto
-    hence "preced th1 s = preced th1 s'"
-    proof
-      assume "th1 = th'"
-      with neq_th
-      show "preced th1 s = preced th1 s'" by (auto simp:s_def preced_def)
-    next
-      assume "th1 \<in> dependents (wq s') th'"
-      with nd and eq_dp have "th1 \<noteq> th"
-        by (auto simp:eq_depend cs_dependents_def s_dependents_def eq_dep)
-      thus "preced th1 s = preced th1 s'" by (auto simp:s_def preced_def)
-    qed
-  } ultimately have "((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) = 
-                     ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" 
-    by (auto simp:image_def)
-  thus "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) =
-        Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" by simp
-qed
-
-lemma nil_dependents: "dependents s th = {}"
-proof -
-  from step_back_step[OF vt_s[unfolded s_def]]
-  show ?thesis
-  proof(cases)
-    assume "th \<notin> threads s'"
-    from not_thread_holdents[OF step_back_vt[OF vt_s[unfolded s_def]] this]
-    have hdn: " holdents s' th = {}" .
-    have "dependents s' th = {}"
-    proof -
-      { assume "dependents s' th \<noteq> {}"
-        then obtain th' where dp: "(Th th', Th th) \<in> (depend s')^+"
-          by (auto simp:s_dependents_def eq_depend)
-        from tranclE[OF this] obtain cs' where 
-          "(Cs cs', Th th) \<in> depend s'" by (auto simp:s_depend_def)
-        with hdn
-        have False by (auto simp:holdents_test)
-      } thus ?thesis by auto
-    qed
-    thus ?thesis 
-      by (unfold s_def s_dependents_def eq_depend depend_create_unchanged, simp)
-  qed
-qed
-
-lemma eq_cp_th: "cp s th = preced th s"
-  apply (unfold cp_eq_cpreced cpreced_def)
-  by (insert nil_dependents, unfold s_dependents_def cs_dependents_def, auto)
-
-end
-
-
-locale step_exit_cps =
-  fixes s' th prio s 
-  defines s_def : "s \<equiv> (Exit th#s')"
-  assumes vt_s: "vt s"
-
-context step_exit_cps
-begin
-
-lemma eq_dep: "depend s = depend s'"
-  by (unfold s_def depend_exit_unchanged, auto)
-
-lemma eq_cp:
-  fixes th' 
-  assumes neq_th: "th' \<noteq> th"
-  shows "cp s th' = cp s' th'"
-  apply (unfold cp_eq_cpreced cpreced_def)
-proof -
-  have nd: "th \<notin> dependents s th'"
-  proof
-    assume "th \<in> dependents s th'"
-    hence "(Th th, Th th') \<in> (depend s)^+" by (simp add:s_dependents_def eq_depend)
-    with eq_dep have "(Th th, Th th') \<in> (depend s')^+" by simp
-    from converse_tranclE[OF this]
-    obtain cs' where bk: "(Th th, Cs cs') \<in> depend s'"
-      by (auto simp:s_depend_def)
-    from step_back_step[OF vt_s[unfolded s_def]]
-    show False
-    proof(cases)
-      assume "th \<in> runing s'"
-      with bk show ?thesis
-        apply (unfold runing_def readys_def s_waiting_def s_depend_def)
-        by (auto simp:cs_waiting_def wq_def)
-    qed
-  qed
-  have eq_dp: "\<And> th. dependents (wq s) th = dependents (wq s') th"
-    by (unfold cs_dependents_def, auto simp:eq_dep eq_depend)
-  moreover {
-    fix th1 
-    assume "th1 \<in> {th'} \<union> dependents (wq s') th'"
-    hence "th1 = th' \<or> th1 \<in> dependents (wq s') th'" by auto
-    hence "preced th1 s = preced th1 s'"
-    proof
-      assume "th1 = th'"
-      with neq_th
-      show "preced th1 s = preced th1 s'" by (auto simp:s_def preced_def)
-    next
-      assume "th1 \<in> dependents (wq s') th'"
-      with nd and eq_dp have "th1 \<noteq> th"
-        by (auto simp:eq_depend cs_dependents_def s_dependents_def eq_dep)
-      thus "preced th1 s = preced th1 s'" by (auto simp:s_def preced_def)
-    qed
-  } ultimately have "((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) = 
-                     ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" 
-    by (auto simp:image_def)
-  thus "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')) =
-        Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependents (wq s') th'))" by simp
-qed
-
-end
-end
-
--- a/prio/ExtGG.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1046 +0,0 @@
-theory ExtGG
-imports PrioG
-begin
-
-lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
-  apply (induct s, simp)
-proof -
-  fix a s
-  assume ih: "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
-    and eq_as: "a # s \<noteq> []"
-  show "birthtime th (a # s) < length (a # s)"
-  proof(cases "s \<noteq> []")
-    case False
-    from False show ?thesis
-      by (cases a, auto simp:birthtime.simps)
-  next
-    case True
-    from ih [OF True] show ?thesis
-      by (cases a, auto simp:birthtime.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)
-
-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)"
-  and preced_th: "preced th s = Prc prio tm"
-
-context highest_gen
-begin
-
-
-
-lemma lt_tm: "tm < length s"
-  by (insert preced_tm_lt[OF threads_s preced_th], simp)
-
-lemma eq_cp_s_th: "cp s th = preced th s"
-proof -
-  from highest and max_cp_eq[OF vt_s]
-  have is_max: "preced th s = Max ((\<lambda>th. preced th s) ` threads s)" by simp
-  have sbs: "({th} \<union> dependents (wq s) th) \<subseteq> threads s"
-  proof -
-    from threads_s and dependents_threads[OF vt_s, of th]
-    show ?thesis by auto
-  qed
-  show ?thesis
-  proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
-    show "preced th s \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)" by simp
-  next
-    fix y 
-    assume "y \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)"
-    then obtain th1 where th1_in: "th1 \<in> ({th} \<union> dependents (wq s) th)"
-      and eq_y: "y = preced th1 s" by auto
-    show "y \<le> preced th s"
-    proof(unfold is_max, rule Max_ge)
-      from finite_threads[OF vt_s] 
-      show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-    next
-      from sbs th1_in and eq_y 
-      show "y \<in> (\<lambda>th. preced th s) ` threads s" by auto
-    qed
-  next
-    from sbs and finite_threads[OF vt_s]
-    show "finite ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))"
-      by (auto intro:finite_subset)
-  qed
-qed
-
-lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp)
-
-lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
-  by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
-
-lemma highest': "cp s th = Max (cp s ` threads s)"
-proof -
-  from highest_cp_preced max_cp_eq[OF vt_s, symmetric]
-  show ?thesis by simp
-qed
-
-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"
-
-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
-
-context extend_highest_gen
-begin
-
-thm extend_highest_gen_axioms_def
-
-lemma 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)
-
-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 "th \<in> threads ([] @ s) \<and> preced th ([] @ s) = preced th s"
-      by auto
-  next
-    case (Cons e t)
-    show ?case
-    proof(cases e)
-      case (Create thread prio)
-      assume eq_e: " e = Create thread prio"
-      show ?thesis
-      proof -
-        from Cons and eq_e have "step (t@s) (Create thread prio)" by auto
-        hence "th \<noteq> thread"
-        proof(cases)
-          assume "thread \<notin> threads (t @ s)"
-          with Cons show ?thesis by auto
-        qed
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    next
-      case (Exit thread)
-      assume eq_e: "e = Exit thread"
-      from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
-      from extend_highest_gen.exit_diff [OF this] and eq_e
-      have neq_th: "thread \<noteq> th" by auto
-      with Cons
-      show ?thesis
-        by (unfold eq_e, auto simp:preced_def)
-    next
-      case (P thread cs)
-      assume eq_e: "e = P thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (V thread cs)
-      assume eq_e: "e = V thread cs"
-      with Cons
-      show ?thesis 
-        by (auto simp:eq_e preced_def)
-    next
-      case (Set thread prio')
-      assume eq_e: " e = Set thread prio'"
-      show ?thesis
-      proof -
-        from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
-        from extend_highest_gen.set_diff_low[OF this] and eq_e
-        have "th \<noteq> thread" by auto
-        hence "preced th ((e # t) @ s)  = preced th (t @ s)"
-          by (unfold eq_e, auto simp:preced_def)
-        moreover note Cons
-        ultimately show ?thesis
-          by (auto simp:eq_e)
-      qed
-    qed
-  qed
-qed
-
-lemma max_kept: "Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s"
-proof(induct rule:ind)
-  case Nil
-  from highest_preced_thread
-  show "Max ((\<lambda>th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s"
-    by simp
-next
-  case (Cons e t)
-  show ?case
-  proof(cases e)
-    case (Create thread prio')
-    assume eq_e: " e = Create thread prio'"
-    from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto
-    hence neq_thread: "thread \<noteq> th"
-    proof(cases)
-      assume "thread \<notin> threads (t @ s)"
-      moreover have "th \<in> threads (t@s)"
-      proof -
-        from Cons have "extend_highest_gen s th prio tm t" by auto
-        from extend_highest_gen.th_kept[OF this] show ?thesis by (simp)
-      qed
-      ultimately show ?thesis by auto
-    qed
-    from Cons have "extend_highest_gen s th prio tm t" by auto
-    from extend_highest_gen.th_kept[OF this]
-    have h': " th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" 
-      by (auto)
-    from stp
-    have thread_ts: "thread \<notin> threads (t @ s)"
-      by (cases, auto)
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))"
-        by (unfold eq_e, simp)
-      moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
-      proof(rule Max_insert)
-        from Cons have "vt (t @ s)" by auto
-        from finite_threads[OF this]
-        show "finite (?f ` (threads (t@s)))" by simp
-      next
-        from h' show "(?f ` (threads (t@s))) \<noteq> {}" by auto
-      qed
-      moreover have "(Max (?f ` (threads (t@s)))) = ?t"
-      proof -
-        have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
-          (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B")
-        proof -
-          { fix th' 
-            assume "th' \<in> ?B"
-            with thread_ts eq_e
-            have "?f1 th' = ?f2 th'" by (auto simp:preced_def)
-          } thus ?thesis 
-            apply (auto simp:Image_def)
-          proof -
-            fix th'
-            assume h: "\<And>th'. th' \<in> threads (t @ s) \<Longrightarrow> 
-              preced th' (e # t @ s) = preced th' (t @ s)"
-              and h1: "th' \<in> threads (t @ s)"
-            show "preced th' (t @ s) \<in> (\<lambda>th'. preced th' (e # t @ s)) ` threads (t @ s)"
-            proof -
-              from h1 have "?f1 th' \<in> ?f1 ` ?B" by auto
-              moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp
-              ultimately show ?thesis by simp
-            qed
-          qed
-        qed
-        with Cons show ?thesis by auto
-      qed
-      moreover have "?f thread < ?t"
-      proof -
-        from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
-        from extend_highest_gen.create_low[OF this] and eq_e
-        have "prio' \<le> prio" by auto
-        thus ?thesis
-        by (unfold preced_th, unfold eq_e, insert lt_tm, 
-          auto simp:preced_def precedence_less_def preced_th)
-    qed
-    ultimately show ?thesis by (auto simp:max_def)
-  qed
-next
-    case (Exit thread)
-    assume eq_e: "e = Exit thread"
-    from Cons have vt_e: "vt (e#(t @ s))" by auto
-    from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto
-    from stp have thread_ts: "thread \<in> threads (t @ s)"
-      by(cases, unfold runing_def readys_def, auto)
-    from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
-    from extend_highest_gen.exit_diff[OF this] and eq_e
-    have neq_thread: "thread \<noteq> th" by auto
-    from Cons have "extend_highest_gen s th prio tm t" by auto
-    from extend_highest_gen.th_kept[OF this]
-    have h': "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      have "threads (t@s) = insert thread ?A"
-        by (insert stp thread_ts, unfold eq_e, auto)
-      hence "Max (?f ` (threads (t@s))) = Max (?f ` \<dots>)" by simp
-      also from this have "\<dots> = Max (insert (?f thread) (?f ` ?A))" by simp
-      also have "\<dots> = max (?f thread) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        from finite_threads [OF vt_e]
-        show "finite (?f ` ?A)" by simp
-      next
-        from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
-        from extend_highest_gen.th_kept[OF this]
-        show "?f ` ?A \<noteq> {}" by  auto
-      qed
-      finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" .
-      moreover have "Max (?f ` (threads (t@s))) = ?t"
-      proof -
-        from Cons show ?thesis
-          by (unfold eq_e, auto simp:preced_def)
-      qed
-      ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp
-      moreover have "?f thread < ?t"
-      proof(unfold eq_e, simp add:preced_def, fold preced_def)
-        show "preced thread (t @ s) < ?t"
-        proof -
-          have "preced thread (t @ s) \<le> ?t" 
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" 
-              (is "?t = Max (?g ` ?B)") by simp
-            moreover have "?g thread \<le> \<dots>"
-            proof(rule Max_ge)
-              have "vt (t@s)" by fact
-              from finite_threads [OF this]
-              show "finite (?g ` ?B)" by simp
-            next
-              from thread_ts
-              show "?g thread \<in> (?g ` ?B)" by auto
-            qed
-            ultimately show ?thesis by auto
-          qed
-          moreover have "preced thread (t @ s) \<noteq> ?t"
-          proof
-            assume "preced thread (t @ s) = preced th s"
-            with h' have "preced thread (t @ s) = preced th (t@s)" by simp
-            from preced_unique [OF this] have "thread = th"
-            proof
-              from h' show "th \<in> threads (t @ s)" by simp
-            next
-              from thread_ts show "thread \<in> threads (t @ s)" .
-            qed(simp)
-            with neq_thread show "False" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      qed
-      ultimately show ?thesis 
-        by (auto simp:max_def split:if_splits)
-    qed
-  next
-    case (P thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (V thread cs)
-    with Cons
-    show ?thesis by (auto simp:preced_def)
-  next
-    case (Set thread prio')
-    show ?thesis (is "Max (?f ` ?A) = ?t")
-    proof -
-      let ?B = "threads (t@s)"
-      from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
-      from extend_highest_gen.set_diff_low[OF this] and Set
-      have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
-      from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
-      also have "\<dots> = ?t"
-      proof(rule Max_eqI)
-        fix y
-        assume y_in: "y \<in> ?f ` ?B"
-        then obtain th1 where 
-          th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
-        show "y \<le> ?t"
-        proof(cases "th1 = thread")
-          case True
-          with neq_thread le_p eq_y Set
-          show ?thesis
-            apply (subst preced_th, insert lt_tm)
-            by (auto simp:preced_def precedence_le_def)
-        next
-          case False
-          with Set eq_y
-          have "y  = preced th1 (t@s)"
-            by (simp add:preced_def)
-          moreover have "\<dots> \<le> ?t"
-          proof -
-            from Cons
-            have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
-              by auto
-            moreover have "preced th1 (t@s) \<le> \<dots>"
-            proof(rule Max_ge)
-              from th1_in 
-              show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
-                by simp
-            next
-              show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-              proof -
-                from Cons have "vt (t @ s)" by auto
-                from finite_threads[OF this] show ?thesis by auto
-              qed
-            qed
-            ultimately show ?thesis by auto
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from Cons and finite_threads
-        show "finite (?f ` ?B)" by auto
-      next
-        from Cons have "extend_highest_gen s th prio tm t" by auto
-        from extend_highest_gen.th_kept [OF this]
-        have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
-        show "?t \<in> (?f ` ?B)" 
-        proof -
-          from neq_thread Set h
-          have "?t = ?f th" by (auto simp:preced_def)
-          with h show ?thesis by auto
-        qed
-      qed
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
-  by (insert th_kept max_kept, auto)
-
-lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
-  (is "?L = ?R")
-proof -
-  have "?L = cpreced (wq (t@s)) (t@s) th" 
-    by (unfold cp_eq_cpreced, simp)
-  also have "\<dots> = ?R"
-  proof(unfold cpreced_def)
-    show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependents (wq (t @ s)) th)) =
-          Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
-      (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
-    proof(cases "?A = {}")
-      case False
-      have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
-      moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
-      proof(rule Max_insert)
-        show "finite (?f ` ?A)"
-        proof -
-          from dependents_threads[OF vt_t]
-          have "?A \<subseteq> threads (t@s)" .
-          moreover from finite_threads[OF vt_t] have "finite \<dots>" .
-          ultimately show ?thesis 
-            by (auto simp:finite_subset)
-        qed
-      next
-        from False show "(?f ` ?A) \<noteq> {}" by simp
-      qed
-      moreover have "\<dots> = Max (?f ` ?B)"
-      proof -
-        from max_preced have "?f th = Max (?f ` ?B)" .
-        moreover have "Max (?f ` ?A) \<le> \<dots>" 
-        proof(rule Max_mono)
-          from False show "(?f ` ?A) \<noteq> {}" by simp
-        next
-          show "?f ` ?A \<subseteq> ?f ` ?B" 
-          proof -
-            have "?A \<subseteq> ?B" by (rule dependents_threads[OF vt_t])
-            thus ?thesis by auto
-          qed
-        next
-          from finite_threads[OF vt_t] 
-          show "finite (?f ` ?B)" by simp
-        qed
-        ultimately show ?thesis
-          by (auto simp:max_def)
-      qed
-      ultimately show ?thesis by auto
-    next
-      case True
-      with max_preced show ?thesis by auto
-    qed
-  qed
-  finally show ?thesis .
-qed
-
-lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
-  by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)
-
-lemma th_cp_preced: "cp (t@s) th = preced th s"
-  by (fold max_kept, unfold th_cp_max_preced, simp)
-
-lemma preced_less:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  shows "preced th' s < preced th s"
-proof -
-  have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
-  proof(rule Max_ge)
-    from finite_threads [OF vt_s]
-    show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
-  next
-    from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
-      by simp
-  qed
-  moreover have "preced th' s \<noteq> preced th s"
-  proof
-    assume "preced th' s = preced th s"
-    from preced_unique[OF this th'_in] neq_th' threads_s
-    show "False" by  (auto simp:readys_def)
-  qed
-  ultimately show ?thesis using highest_preced_thread
-    by auto
-qed
-
-lemma pv_blocked_pre:
-  fixes th'
-  assumes th'_in: "th' \<in> threads (t@s)"
-  and neq_th': "th' \<noteq> th"
-  and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
-  shows "th' \<notin> runing (t@s)"
-proof
-  assume "th' \<in> runing (t@s)"
-  hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
-    by (auto simp:runing_def)
-  with max_cp_readys_threads [OF vt_t]
-  have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
-    by auto
-  moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
-  ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
-  moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
-    by simp
-  finally have h: "cp (t @ s) th' = preced th (t @ s)" .
-  show False
-  proof -
-    have "dependents (wq (t @ s)) th' = {}" 
-      by (rule count_eq_dependents [OF vt_t eq_pv])
-    moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
-    proof
-      assume "preced th' (t @ s) = preced th (t @ s)"
-      hence "th' = th"
-      proof(rule preced_unique)
-        from th_kept show "th \<in> threads (t @ s)" by simp
-      next
-        from th'_in show "th' \<in> threads (t @ s)" by simp
-      qed
-      with assms show False by simp
-    qed
-    ultimately show ?thesis
-      by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
-  qed
-qed
-
-lemmas pv_blocked = pv_blocked_pre[folded detached_eq [OF vt_t]]
-
-lemma runing_precond_pre:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<in> threads (t@s) \<and>
-         cntP (t@s) th' = cntV (t@s) th'"
-proof -
-  show ?thesis
-  proof(induct rule:ind)
-    case (Cons e t)
-    from Cons
-    have in_thread: "th' \<in> threads (t @ s)"
-      and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    from Cons have "extend_highest_gen s th prio tm t" by auto
-    then have not_runing: "th' \<notin> runing (t @ s)" 
-      apply(rule extend_highest_gen.pv_blocked) 
-      using Cons(1) in_thread neq_th' not_holding
-      apply(simp_all add: detached_eq)
-      done
-    show ?case
-    proof(cases e)
-      case (V thread cs)
-      from Cons and V have vt_v: "vt (V thread cs#(t@s))" by auto
-
-      show ?thesis
-      proof -
-        from Cons and V have "step (t@s) (V thread cs)" by auto
-        hence neq_th': "thread \<noteq> th'"
-        proof(cases)
-          assume "thread \<in> runing (t@s)"
-          moreover have "th' \<notin> runing (t@s)" by fact
-          ultimately show ?thesis by auto
-        qed
-        with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
-          by (unfold V, simp add:cntP_def cntV_def count_def)
-        moreover from in_thread
-        have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
-        ultimately show ?thesis by auto
-      qed
-    next
-      case (P thread cs)
-      from Cons and P have "step (t@s) (P thread cs)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t@s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Create thread prio')
-      from Cons and Create have "step (t@s) (Create thread prio')" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<notin> threads (t @ s)"
-        moreover have "th' \<in> threads (t@s)" by fact
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Create 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Create 
-      have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
-      ultimately show ?thesis by auto
-    next
-      case (Exit thread)
-      from Cons and Exit have "step (t@s) (Exit thread)" by auto
-      hence neq_th': "thread \<noteq> th'"
-      proof(cases)
-        assume "thread \<in> runing (t @ s)"
-        moreover note not_runing
-        ultimately show ?thesis by auto
-      qed
-      with Cons and Exit 
-      have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
-        by (auto simp:cntP_def cntV_def count_def)
-      moreover from Cons and Exit and neq_th' 
-      have in_thread': "th' \<in> threads ((e # t) @ s)"
-        by auto
-      ultimately show ?thesis by auto
-    next
-      case (Set thread prio')
-      with Cons
-      show ?thesis 
-        by (auto simp:cntP_def cntV_def count_def)
-    qed
-  next
-    case Nil
-    with assms
-    show ?case by auto
-  qed
-qed
-
-(*
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and eq_pv: "cntP s th' = cntV s th'"
-  and neq_th': "th' \<noteq> th"
-  shows "th' \<notin> runing (t@s)"
-proof -
-  from runing_precond_pre[OF th'_in eq_pv neq_th']
-  have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-  from pv_blocked[OF h1 neq_th' h2] 
-  show ?thesis .
-qed
-*)
-
-lemmas runing_precond_pre_dtc = runing_precond_pre[folded detached_eq[OF vt_t] detached_eq[OF vt_s]]
-
-lemma runing_precond:
-  fixes th'
-  assumes th'_in: "th' \<in> threads s"
-  and neq_th': "th' \<noteq> th"
-  and is_runing: "th' \<in> runing (t@s)"
-  shows "cntP s th' > cntV s th'"
-proof -
-  have "cntP s th' \<noteq> cntV s th'"
-  proof
-    assume eq_pv: "cntP s th' = cntV s th'"
-    from runing_precond_pre[OF th'_in eq_pv neq_th']
-    have h1: "th' \<in> threads (t @ s)"  
-      and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
-    from pv_blocked_pre[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
-    with is_runing show "False" by simp
-  qed
-  moreover from cnp_cnv_cncs[OF vt_s, of th'] 
-  have "cntV s th' \<le> cntP s th'" by auto
-  ultimately show ?thesis by auto
-qed
-
-lemma moment_blocked_pre:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
-         th' \<in> threads ((moment (i+j) t)@s)"
-proof(induct j)
-  case (Suc k)
-  show ?case
-  proof -
-    { assume True: "Suc (i+k) \<le> length t"
-      from moment_head [OF this] 
-      obtain e where
-        eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
-        by blast
-      from red_moment[of "Suc(i+k)"]
-      and eq_me have "extend_highest_gen s th prio tm (e # moment (i + k) t)" by simp
-      hence vt_e: "vt (e#(moment (i + k) t)@s)"
-        by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
-                          highest_gen_def, auto)
-      have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
-      proof -
-        show "th' \<notin> runing (moment (i + k) t @ s)"
-        proof(rule extend_highest_gen.pv_blocked)
-          from Suc show "th' \<in> threads (moment (i + k) t @ s)"
-            by simp
-        next
-          from neq_th' show "th' \<noteq> th" .
-        next
-          from red_moment show "extend_highest_gen s th prio tm (moment (i + k) t)" .
-        next
-          from Suc vt_e show "detached (moment (i + k) t @ s) th'"
-            apply(subst detached_eq)
-            apply(auto intro: vt_e evt_cons)
-            done
-        qed
-      qed
-      from step_back_step[OF vt_e]
-      have "step ((moment (i + k) t)@s) e" .
-      hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
-        th' \<in> threads (e#(moment (i + k) t)@s)"
-      proof(cases)
-        case (thread_create thread prio)
-        with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_exit thread)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_P thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_V thread cs)
-        moreover have "thread \<noteq> th'"
-        proof -
-          have "thread \<in> runing (moment (i + k) t @ s)" by fact
-          moreover note not_runing'
-          ultimately show ?thesis by auto
-        qed
-        moreover note Suc 
-        ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
-      next
-        case (thread_set thread prio')
-        with Suc show ?thesis
-          by (auto simp:cntP_def cntV_def count_def)
-      qed
-      with eq_me have ?thesis using eq_me by auto 
-    } note h = this
-    show ?thesis
-    proof(cases "Suc (i+k) \<le> length t")
-      case True
-      from h [OF this] show ?thesis .
-    next
-      case False
-      with moment_ge
-      have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
-      with Suc show ?thesis by auto
-    qed
-  qed
-next
-  case 0
-  from assms show ?case by auto
-qed
-
-lemma moment_blocked_eqpv:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
-  and le_ij: "i \<le> j"
-  shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
-         th' \<in> threads ((moment j t)@s) \<and>
-         th' \<notin> runing ((moment j t)@s)"
-proof -
-  from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
-  have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
-    and h2: "th' \<in> threads ((moment j t)@s)" by auto
-  with extend_highest_gen.pv_blocked 
-  show ?thesis 
-    using  red_moment [of j] h2 neq_th' h1
-    apply(auto)
-    by (metis extend_highest_gen.pv_blocked_pre)
-qed
-
-lemma moment_blocked:
-  assumes neq_th': "th' \<noteq> th"
-  and th'_in: "th' \<in> threads ((moment i t)@s)"
-  and dtc: "detached (moment i t @ s) th'"
-  and le_ij: "i \<le> j"
-  shows "detached (moment j t @ s) th' \<and>
-         th' \<in> threads ((moment j t)@s) \<and>
-         th' \<notin> runing ((moment j t)@s)"
-proof -
-  from vt_moment[OF vt_t, of "i+length s"] moment_prefix[of i t s]
-  have vt_i: "vt (moment i t @ s)" by auto
-  from vt_moment[OF vt_t, of "j+length s"] moment_prefix[of j t s]
-  have vt_j: "vt  (moment j t @ s)" by auto
-  from moment_blocked_eqpv [OF neq_th' th'_in detached_elim [OF vt_i dtc] le_ij, 
-  folded detached_eq[OF vt_j]]
-  show ?thesis .
-qed
-
-lemma runing_inversion_1:
-  assumes neq_th': "th' \<noteq> th"
-  and runing': "th' \<in> runing (t@s)"
-  shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
-proof(cases "th' \<in> threads s")
-  case True
-  with runing_precond [OF this neq_th' runing'] show ?thesis by simp
-next
-  case False
-  let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
-  let ?q = "moment 0 t"
-  from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
-  from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
-  from p_split_gen [of ?Q, OF this not_thread]
-  obtain i where lt_its: "i < length t"
-    and le_i: "0 \<le> i"
-    and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
-    and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
-  from lt_its have "Suc i \<le> length t" by auto
-  from moment_head[OF this] obtain e where 
-   eq_me: "moment (Suc i) t = e # moment i t" by blast
-  from red_moment[of "Suc i"] and eq_me
-  have "extend_highest_gen s th prio tm (e # moment i t)" by simp
-  hence vt_e: "vt (e#(moment i t)@s)"
-    by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
-      highest_gen_def, auto)
-  from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
-  from post[rule_format, of "Suc i"] and eq_me 
-  have not_in': "th' \<in> threads (e # moment i t@s)" by auto
-  from create_pre[OF stp_i pre this] 
-  obtain prio where eq_e: "e = Create th' prio" .
-  have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
-  proof(rule cnp_cnv_eq)
-    from step_back_vt [OF vt_e] 
-    show "vt (moment i t @ s)" .
-  next
-    from eq_e and stp_i 
-    have "step (moment i t @ s) (Create th' prio)" by simp
-    thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
-  qed
-  with eq_e
-  have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
-    by (simp add:cntP_def cntV_def count_def)
-  with eq_me[symmetric]
-  have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
-    by simp
-  from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
-  with eq_me [symmetric]
-  have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
-  from moment_blocked_eqpv [OF neq_th' h2 h1, of "length t"] and lt_its
-  and moment_ge
-  have "th' \<notin> runing (t @ s)" by auto
-  with runing'
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_2:
-  assumes runing': "th' \<in> runing (t@s)"
-  shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
-proof -
-  from runing_inversion_1[OF _ runing']
-  show ?thesis by auto
-qed
-
-lemma runing_preced_inversion:
-  assumes runing': "th' \<in> runing (t@s)"
-  shows "cp (t@s) th' = preced th s"
-proof -
-  from runing' have "cp (t@s) th' = Max (cp (t @ s) ` readys (t @ s))"
-    by (unfold runing_def, auto)
-  also have "\<dots> = preced th s"
-  proof -
-    from max_cp_readys_threads[OF vt_t]
-    have "\<dots> =  Max (cp (t @ s) ` threads (t @ s))" .
-    also have "\<dots> = preced th s"
-    proof -
-      from max_kept
-      and max_cp_eq [OF vt_t]
-      show ?thesis by auto
-    qed
-    finally show ?thesis .
-  qed
-  finally show ?thesis .
-qed
-
-lemma runing_inversion_3:
-  assumes runing': "th' \<in> runing (t@s)"
-  and neq_th: "th' \<noteq> th"
-  shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
-proof -
-  from runing_inversion_2 [OF runing'] 
-    and neq_th 
-    and runing_preced_inversion[OF runing']
-  show ?thesis by auto
-qed
-
-lemma runing_inversion_4:
-  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"
-using runing_inversion_3 [OF runing'] 
-  and neq_th 
-  and runing_preced_inversion[OF runing']
-apply(auto simp add: detached_eq[OF vt_s])
-done
-
-
-
-lemma live: "runing (t@s) \<noteq> {}"
-proof(cases "th \<in> runing (t@s)")
-  case True thus ?thesis by auto
-next
-  case False
-  then have not_ready: "th \<notin> readys (t@s)"
-    apply (unfold runing_def, 
-            insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
-    by auto
-  from th_kept have "th \<in> threads (t@s)" by auto
-  from th_chain_to_ready[OF vt_t this] and not_ready
-  obtain th' where th'_in: "th' \<in> readys (t@s)"
-    and dp: "(Th th, Th th') \<in> (depend (t @ s))\<^sup>+" by auto
-  have "th' \<in> runing (t@s)"
-  proof -
-    have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
-    proof -
-      have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')) = 
-               preced th (t@s)"
-      proof(rule Max_eqI)
-        fix y
-        assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-        then obtain th1 where
-          h1: "th1 = th' \<or> th1 \<in>  dependents (wq (t @ s)) th'"
-          and eq_y: "y = preced th1 (t@s)" by auto
-        show "y \<le> preced th (t @ s)"
-        proof -
-          from max_preced
-          have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
-          moreover have "y \<le> \<dots>"
-          proof(rule Max_ge)
-            from h1
-            have "th1 \<in> threads (t@s)"
-            proof
-              assume "th1 = th'"
-              with th'_in show ?thesis by (simp add:readys_def)
-            next
-              assume "th1 \<in> dependents (wq (t @ s)) th'"
-              with dependents_threads [OF vt_t]
-              show "th1 \<in> threads (t @ s)" by auto
-            qed
-            with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
-          next
-            from finite_threads[OF vt_t]
-            show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
-          qed
-          ultimately show ?thesis by auto
-        qed
-      next
-        from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th']
-        show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th'))"
-          by (auto intro:finite_subset)
-      next
-        from dp
-        have "th \<in> dependents (wq (t @ s)) th'" 
-          by (unfold cs_dependents_def, auto simp:eq_depend)
-        thus "preced th (t @ s) \<in> 
-                (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
-          by auto
-      qed
-      moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
-      proof -
-        from max_preced and max_cp_eq[OF vt_t, symmetric]
-        have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
-        with max_cp_readys_threads[OF vt_t] show ?thesis by simp
-      qed
-      ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
-    qed
-    with th'_in show ?thesis by (auto simp:runing_def)
-  qed
-  thus ?thesis by auto
-qed
-
-end
-end
-
-
-
--- a/prio/IsaMakefile	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,51 +0,0 @@
-
-## targets
-
-default: itp
-all: session itp slides1
-
-## global settings
-
-SRC = $(ISABELLE_HOME)/src
-OUT = $(ISABELLE_OUTPUT)
-LOG = $(OUT)/log
-
-
-USEDIR = $(ISABELLE_TOOL) usedir -v true -t true 
-
-
-## Slides
-
-session1: Slides/ROOT1.ML \
-	Slides/document/root* \
-	Slides/Slides1.thy
-	@$(USEDIR) -D generated -f ROOT1.ML HOL Slides
-
-slides1: session1
-	rm -f Slides/generated/*.aux # otherwise latex will fall over
-	cd Slides/generated ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
-	cp Slides/generated/root.beamer.pdf Slides/slides.pdf
-
-# main files                        
-
-session: ./ROOT.ML ./*.thy
-	@$(USEDIR) -b -D generated -f ROOT.ML HOL Prio
-
-
-# itp paper
-
-itp: Paper/*.thy Paper/*.ML 
-	@$(USEDIR) -D generated -f ROOT.ML Prio Paper
-	rm -f Paper/generated/*.aux # otherwise latex will fall over  
-	cd Paper/generated ; $(ISABELLE_TOOL) latex -o pdf root.tex 
-	cd Paper/generated ; bibtex root
-	cd Paper/generated ; $(ISABELLE_TOOL) latex -o pdf root.tex
-	cp Paper/generated/root.pdf paper.pdf 
-
-
-slides: Slides/ROOT1.ML Slides/*.thy 
-	@$(USEDIR) -D generated -f ROOT1.ML Prio Slides
-	rm -f Slides/generated/*.aux # otherwise latex will fall over
-	cd Slides/generated ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
-	cp Slides/generated/root.beamer.pdf Slides/slides.pdf
-
--- a/prio/Moment.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,783 +0,0 @@
-theory Moment
-imports Main
-begin
-
-fun firstn :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "firstn 0 s = []" |
-  "firstn (Suc n) [] = []" |
-  "firstn (Suc n) (e#s) = e#(firstn n s)"
-
-fun restn :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where "restn n s = rev (firstn (length s - n) (rev s))"
-
-definition moment :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where "moment n s = rev (firstn n (rev s))"
-
-definition restm :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where "restm n s = rev (restn n (rev s))"
-
-definition from_to :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-  where "from_to i j s = firstn (j - i) (restn i s)"
-
-definition down_to :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where "down_to j i s = rev (from_to i j (rev s))"
-
-(*
-value "down_to 6 2 [10, 9, 8, 7, 6, 5, 4, 3, 2, 1, 0]"
-value "from_to 2 6 [0, 1, 2, 3, 4, 5, 6, 7]"
-*)
-
-lemma length_eq_elim_l: "\<lbrakk>length xs = length ys; xs@us = ys@vs\<rbrakk> \<Longrightarrow> xs = ys \<and> us = vs"
-  by auto
-
-lemma length_eq_elim_r: "\<lbrakk>length us = length vs; xs@us = ys@vs\<rbrakk> \<Longrightarrow> xs = ys \<and> us = vs"
-  by simp
-
-lemma firstn_nil [simp]: "firstn n [] = []"
-  by (cases n, simp+)
-
-(*
-value "from_to 0 2 [0, 1, 2, 3, 4, 5, 6, 7, 8, 9] @ 
-       from_to 2 5 [0, 1, 2, 3, 4, 5, 6, 7, 8, 9]"
-*)
-
-lemma firstn_le: "\<And> n s'. n \<le> length s \<Longrightarrow> firstn n (s@s') = firstn n s"
-proof (induct s, simp)
-  fix a s n s'
-  assume ih: "\<And>n s'. n \<le> length s \<Longrightarrow> firstn n (s @ s') = firstn n s"
-  and le_n: " n \<le> length (a # s)"
-  show "firstn n ((a # s) @ s') = firstn n (a # s)"
-  proof(cases n, simp)
-    fix k
-    assume eq_n: "n = Suc k"
-    with le_n have "k \<le> length s" by auto
-    from ih [OF this] and eq_n
-    show "firstn n ((a # s) @ s') = firstn n (a # s)" by auto
-  qed
-qed
-
-lemma firstn_ge [simp]: "\<And>n. length s \<le> n \<Longrightarrow> firstn n s = s"
-proof(induct s, simp)
-  fix a s n
-  assume ih: "\<And>n. length s \<le> n \<Longrightarrow> firstn n s = s"
-    and le: "length (a # s) \<le> n"
-  show "firstn n (a # s) = a # s"
-  proof(cases n)
-    assume eq_n: "n = 0" with le show ?thesis by simp
-  next
-    fix k
-    assume eq_n: "n = Suc k"
-    with le have le_k: "length s \<le> k" by simp
-    from ih [OF this] have "firstn k s = s" .
-    from eq_n and this
-    show ?thesis by simp
-  qed
-qed
-
-lemma firstn_eq [simp]: "firstn (length s) s = s"
-  by simp
-
-lemma firstn_restn_s: "(firstn n (s::'a list)) @ (restn n s) = s"
-proof(induct n arbitrary:s, simp)
-  fix n s
-  assume ih: "\<And>t. firstn n (t::'a list) @ restn n t = t"
-  show "firstn (Suc n) (s::'a list) @ restn (Suc n) s = s"
-  proof(cases s, simp)
-    fix x xs
-    assume eq_s: "s = x#xs"
-    show "firstn (Suc n) s @ restn (Suc n) s = s"
-    proof -
-      have "firstn (Suc n) s @ restn (Suc n) s =  x # (firstn n xs @ restn n xs)"
-      proof -
-        from eq_s have "firstn (Suc n) s =  x # firstn n xs" by simp
-        moreover have "restn (Suc n) s = restn n xs"
-        proof -
-          from eq_s have "restn (Suc n) s = rev (firstn (length xs - n) (rev xs @ [x]))" by simp
-          also have "\<dots> = restn n xs"
-          proof -
-            have "(firstn (length xs - n) (rev xs @ [x])) = (firstn (length xs - n) (rev xs))"
-              by(rule firstn_le, simp)
-            hence "rev (firstn (length xs - n) (rev xs @ [x])) = 
-              rev (firstn (length xs - n) (rev xs))" by simp
-            also have "\<dots> = rev (firstn (length (rev xs) - n) (rev xs))" by simp
-            finally show ?thesis by simp
-          qed
-          finally show ?thesis by simp
-        qed
-        ultimately show ?thesis by simp
-      qed with ih eq_s show ?thesis by simp
-    qed
-  qed
-qed
-
-lemma moment_restm_s: "(restm n s)@(moment n s) = s"
-proof -
-  have " rev  ((firstn n (rev s)) @ (restn n (rev s))) = s" (is "rev ?x = s")
-  proof -
-    have "?x = rev s" by (simp only:firstn_restn_s)
-    thus ?thesis by auto
-  qed
-  thus ?thesis 
-    by (auto simp:restm_def moment_def)
-qed
-
-declare restn.simps [simp del] firstn.simps[simp del]
-
-lemma length_firstn_ge: "length s \<le> n \<Longrightarrow> length (firstn n s) = length s"
-proof(induct n arbitrary:s, simp add:firstn.simps)
-  case (Suc k)
-  assume ih: "\<And> s. length (s::'a list) \<le> k \<Longrightarrow> length (firstn k s) = length s"
-  and le: "length s \<le> Suc k"
-  show ?case
-  proof(cases s)
-    case Nil
-    from Nil show ?thesis by simp
-  next
-    case (Cons x xs)
-    from le and Cons have "length xs \<le> k" by simp
-    from ih [OF this] have "length (firstn k xs) = length xs" .
-    moreover from Cons have "length (firstn (Suc k) s) = Suc (length (firstn k xs))" 
-      by (simp add:firstn.simps)
-    moreover note Cons
-    ultimately show ?thesis by simp
-  qed
-qed
-
-lemma length_firstn_le: "n \<le> length s \<Longrightarrow> length (firstn n s) = n"
-proof(induct n arbitrary:s, simp add:firstn.simps)
-  case (Suc k)
-  assume ih: "\<And>s. k \<le> length (s::'a list) \<Longrightarrow> length (firstn k s) = k"
-    and le: "Suc k \<le> length s"
-  show ?case
-  proof(cases s)
-    case Nil
-    from Nil and le show ?thesis by auto
-  next
-    case (Cons x xs)
-    from le and Cons have "k \<le> length xs" by simp
-    from ih [OF this] have "length (firstn k xs) = k" .
-    moreover from Cons have "length (firstn (Suc k) s) = Suc (length (firstn k xs))" 
-      by (simp add:firstn.simps)
-    ultimately show ?thesis by simp
-  qed
-qed
-
-lemma app_firstn_restn: 
-  fixes s1 s2
-  shows "s1 = firstn (length s1) (s1 @ s2) \<and> s2 = restn (length s1) (s1 @ s2)"
-proof(rule length_eq_elim_l)
-  have "length s1 \<le> length (s1 @ s2)" by simp
-  from length_firstn_le [OF this]
-  show "length s1 = length (firstn (length s1) (s1 @ s2))" by simp
-next
-  from firstn_restn_s 
-  show "s1 @ s2 = firstn (length s1) (s1 @ s2) @ restn (length s1) (s1 @ s2)"
-    by metis
-qed
-
-
-lemma length_moment_le:
-  fixes k s
-  assumes le_k: "k \<le> length s"
-  shows "length (moment k s) = k"
-proof -
-  have "length (rev (firstn k (rev s))) = k"
-  proof -
-    have "length (rev (firstn k (rev s))) = length (firstn k (rev s))" by simp
-    also have "\<dots> = k" 
-    proof(rule length_firstn_le)
-      from le_k show "k \<le> length (rev s)" by simp
-    qed
-    finally show ?thesis .
-  qed
-  thus ?thesis by (simp add:moment_def)
-qed
-
-lemma app_moment_restm: 
-  fixes s1 s2
-  shows "s1 = restm (length s2) (s1 @ s2) \<and> s2 = moment (length s2) (s1 @ s2)"
-proof(rule length_eq_elim_r)
-  have "length s2 \<le> length (s1 @ s2)" by simp
-  from length_moment_le [OF this]
-  show "length s2 = length (moment (length s2) (s1 @ s2))" by simp
-next
-  from moment_restm_s 
-  show "s1 @ s2 = restm (length s2) (s1 @ s2) @ moment (length s2) (s1 @ s2)"
-    by metis
-qed
-
-lemma length_moment_ge:
-  fixes k s
-  assumes le_k: "length s \<le> k"
-  shows "length (moment k s) = (length s)"
-proof -
-  have "length (rev (firstn k (rev s))) = length s"
-  proof -
-    have "length (rev (firstn k (rev s))) = length (firstn k (rev s))" by simp
-    also have "\<dots> = length s" 
-    proof -
-      have "\<dots> = length (rev s)"
-      proof(rule length_firstn_ge)
-        from le_k show "length (rev s) \<le> k" by simp
-      qed
-      also have "\<dots> = length s" by simp
-      finally show ?thesis .
-    qed
-    finally show ?thesis .
-  qed
-  thus ?thesis by (simp add:moment_def)
-qed
-
-lemma length_firstn: "(length (firstn n s) = length s) \<or> (length (firstn n s) = n)"
-proof(cases "n \<le> length s")
-  case True
-  from length_firstn_le [OF True] show ?thesis by auto
-next
-  case False
-  from False have "length s \<le> n" by simp
-  from firstn_ge [OF this] show ?thesis by auto
-qed
-
-lemma firstn_conc: 
-  fixes m n
-  assumes le_mn: "m \<le> n"
-  shows "firstn m s = firstn m (firstn n  s)"
-proof(cases "m \<le> length s")
-  case True
-  have "s = (firstn n s) @ (restn n s)" by (simp add:firstn_restn_s)
-  hence "firstn m s = firstn m \<dots>" by simp
-  also have "\<dots> = firstn m (firstn n s)" 
-  proof -
-    from length_firstn [of n s]
-    have "m \<le> length (firstn n s)"
-    proof
-      assume "length (firstn n s) = length s" with True show ?thesis by simp
-    next
-      assume "length (firstn n s) = n " with le_mn show ?thesis by simp
-    qed
-    from firstn_le [OF this, of "restn n s"]
-    show ?thesis .
-  qed
-  finally show ?thesis by simp
-next
-  case False
-  from False and le_mn have "length s \<le> n"  by simp
-  from firstn_ge [OF this] show ?thesis by simp
-qed
-
-lemma restn_conc: 
-  fixes i j k s
-  assumes eq_k: "j + i = k"
-  shows "restn k s = restn j (restn i s)"
-proof -
-  have "(firstn (length s - k) (rev s)) =
-        (firstn (length (rev (firstn (length s - i) (rev s))) - j) 
-                            (rev (rev (firstn (length s - i) (rev s)))))"
-  proof  -
-    have "(firstn (length s - k) (rev s)) =
-            (firstn (length (rev (firstn (length s - i) (rev s))) - j) 
-                                           (firstn (length s - i) (rev s)))"
-    proof -
-      have " (length (rev (firstn (length s - i) (rev s))) - j) = length s - k"
-      proof -
-        have "(length (rev (firstn (length s - i) (rev s))) - j) = (length s - i) - j"
-        proof -
-          have "(length (rev (firstn (length s - i) (rev s))) - j) = 
-                                         length ((firstn (length s - i) (rev s))) - j"
-            by simp
-          also have "\<dots> = length ((firstn (length (rev s) - i) (rev s))) - j" by simp
-          also have "\<dots> = (length (rev s) - i) - j" 
-          proof -
-            have "length ((firstn (length (rev s) - i) (rev s))) = (length (rev s) - i)"
-              by (rule length_firstn_le, simp)
-            thus ?thesis by simp
-          qed
-          also have "\<dots> = (length s - i) - j" by simp
-          finally show ?thesis .
-        qed
-        with eq_k show ?thesis by auto
-      qed
-      moreover have "(firstn (length s - k) (rev s)) =
-                             (firstn (length s - k) (firstn (length s - i) (rev s)))"
-      proof(rule firstn_conc)
-        from eq_k show "length s - k \<le> length s - i" by simp
-      qed
-      ultimately show ?thesis by simp
-    qed
-    thus ?thesis by simp
-  qed
-  thus ?thesis by (simp only:restn.simps)
-qed
-
-(*
-value "down_to 2 0 [5, 4, 3, 2, 1, 0]"
-value "moment 2 [5, 4, 3, 2, 1, 0]"
-*)
-
-lemma from_to_firstn: "from_to 0 k s = firstn k s"
-by (simp add:from_to_def restn.simps)
-
-lemma moment_app [simp]:
-  assumes 
-  ile: "i \<le> length s"
-  shows "moment i (s'@s) = moment i s"
-proof -
-  have "moment i (s'@s) = rev (firstn i (rev (s'@s)))" by (simp add:moment_def)
-  moreover have "firstn i (rev (s'@s)) = firstn i (rev s @ rev s')" by simp
-  moreover have "\<dots> = firstn i (rev s)"
-  proof(rule firstn_le)
-    have "length (rev s) = length s" by simp
-    with ile show "i \<le> length (rev s)" by simp
-  qed
-  ultimately show ?thesis by (simp add:moment_def)
-qed
-
-lemma moment_eq [simp]: "moment (length s) (s'@s) = s"
-proof -
-  have "length s \<le> length s" by simp
-  from moment_app [OF this, of s'] 
-  have " moment (length s) (s' @ s) = moment (length s) s" .
-  moreover have "\<dots> = s" by (simp add:moment_def)
-  ultimately show ?thesis by simp
-qed
-
-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 firstn.simps)
-
-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: 
-  "\<And> s Q. \<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_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
-qed
-
-lemma from_to_conc:
-  fixes i j k s
-  assumes le_ij: "i \<le> j"
-  and le_jk: "j \<le> k"
-  shows "from_to i j s @ from_to j k s = from_to i k s"
-proof -
-  let ?ris = "restn i s"
-  have "firstn (j - i) (restn i s) @ firstn (k - j) (restn j s) =
-         firstn (k - i) (restn i s)" (is "?x @ ?y = ?z")
-  proof -
-    let "firstn (k-j) ?u" = "?y"
-    let ?rst = " restn (k - j) (restn (j - i) ?ris)"
-    let ?rst' = "restn (k - i) ?ris"
-    have "?u = restn (j-i) ?ris"
-    proof(rule restn_conc)
-      from le_ij show "j - i + i = j" by simp
-    qed
-    hence "?x @ ?y = ?x @ firstn (k-j) (restn (j-i) ?ris)" by simp
-    moreover have "firstn (k - j) (restn (j - i) (restn i s)) @ ?rst = 
-                      restn (j-i) ?ris" by (simp add:firstn_restn_s)
-    ultimately have "?x @ ?y @ ?rst = ?x @ (restn (j-i) ?ris)" by simp
-    also have "\<dots> = ?ris" by (simp add:firstn_restn_s)
-    finally have "?x @ ?y @ ?rst = ?ris" .
-    moreover have "?z @ ?rst = ?ris"
-    proof -
-      have "?z @ ?rst' = ?ris" by (simp add:firstn_restn_s)
-      moreover have "?rst' = ?rst"
-      proof(rule restn_conc)
-        from le_ij le_jk show "k - j + (j - i) = k - i" by auto
-      qed
-      ultimately show ?thesis by simp
-    qed
-    ultimately have "?x @ ?y @ ?rst = ?z @ ?rst" by simp
-    thus ?thesis by auto    
-  qed
-  thus ?thesis by (simp only:from_to_def)
-qed
-
-lemma down_to_conc:
-  fixes i j k s
-  assumes le_ij: "i \<le> j"
-  and le_jk: "j \<le> k"
-  shows "down_to k j s @ down_to j i s = down_to k i s"
-proof -
-  have "rev (from_to j k (rev s)) @ rev (from_to i j (rev s)) = rev (from_to i k (rev s))"
-    (is "?L = ?R")
-  proof -
-    have "?L = rev (from_to i j (rev s) @ from_to j k (rev s))" by simp
-    also have "\<dots> = ?R" (is "rev ?x = rev ?y")
-    proof -
-      have "?x = ?y" by (rule from_to_conc[OF le_ij le_jk])
-      thus ?thesis by simp
-    qed
-    finally show ?thesis .
-  qed
-  thus ?thesis by (simp add:down_to_def)
-qed
-
-lemma restn_ge:
-  fixes s k
-  assumes le_k: "length s \<le> k"
-  shows "restn k s = []"
-proof -
-  from firstn_restn_s [of k s, symmetric] have "s = (firstn k s) @ (restn k s)" .
-  hence "length s = length \<dots>" by simp
-  also have "\<dots> = length (firstn k s) + length (restn k s)" by simp
-  finally have "length s = ..." by simp
-  moreover from length_firstn_ge and le_k 
-  have "length (firstn k s) = length s" by simp
-  ultimately have "length (restn k s) = 0" by auto
-  thus ?thesis by auto
-qed
-
-lemma from_to_ge: "length s \<le> k \<Longrightarrow> from_to k j s = []"
-proof(simp only:from_to_def)
-  assume "length s \<le> k"
-  from restn_ge [OF this] 
-  show "firstn (j - k) (restn k s) = []" by simp
-qed
-
-(*
-value "from_to 2 5 [0, 1, 2, 3, 4]"
-value "restn 2  [0, 1, 2, 3, 4]"
-*)
-
-lemma from_to_restn: 
-  fixes k j s
-  assumes le_j: "length s \<le> j"
-  shows "from_to k j s = restn k s"
-proof -
-  have "from_to 0 k s @ from_to k j s = from_to 0 j s"
-  proof(cases "k \<le> j")
-    case True
-    from from_to_conc True show ?thesis by auto
-  next
-    case False
-    from False le_j have lek: "length s \<le>  k" by auto
-    from from_to_ge [OF this] have "from_to k j s = []" .
-    hence "from_to 0 k s @ from_to k j s = from_to 0 k s" by simp
-    also have "\<dots> = s"
-    proof -
-      from from_to_firstn [of k s]
-      have "\<dots> = firstn k s" .
-      also have "\<dots> = s" by (rule firstn_ge [OF lek])
-      finally show ?thesis .
-    qed
-    finally have "from_to 0 k s @ from_to k j s = s" .
-    moreover have "from_to 0 j s = s"
-    proof -
-      have "from_to 0 j s = firstn j s" by (simp add:from_to_firstn)
-      also have "\<dots> = s"
-      proof(rule firstn_ge)
-        from le_j show "length s \<le> j " by simp
-      qed
-      finally show ?thesis .
-    qed
-    ultimately show ?thesis by auto
-  qed
-  also have "\<dots> = s" 
-  proof -
-    from from_to_firstn have "\<dots> = firstn j s" .
-    also have "\<dots> = s"
-    proof(rule firstn_ge)
-      from le_j show "length s \<le> j" by simp
-    qed
-    finally show ?thesis .
-  qed
-  finally have "from_to 0 k s @ from_to k j s = s" .
-  moreover have "from_to 0 k s @ restn k s = s"
-  proof -
-    from from_to_firstn [of k s]
-    have "from_to 0 k s = firstn k s" .
-    thus ?thesis by (simp add:firstn_restn_s)
-  qed
-  ultimately have "from_to 0 k s @ from_to k j s  = 
-                    from_to 0 k s @ restn k s" by simp
-  thus ?thesis by auto
-qed
-
-lemma down_to_moment: "down_to k 0 s = moment k s"
-proof -
-  have "rev (from_to 0 k (rev s)) = rev (firstn k (rev s))" 
-    using from_to_firstn by metis
-  thus ?thesis by (simp add:down_to_def moment_def)
-qed
-
-lemma down_to_restm:
-  assumes le_s: "length s \<le> j"
-  shows "down_to j k s = restm k s"
-proof -
-  have "rev (from_to k j (rev s)) = rev (restn k (rev s))" (is "?L = ?R")
-  proof -
-    from le_s have "length (rev s) \<le> j" by simp
-    from from_to_restn [OF this, of k] show ?thesis by simp
-  qed
-  thus ?thesis by (simp add:down_to_def restm_def)
-qed
-
-lemma moment_split: "moment (m+i) s = down_to (m+i) i s @down_to i 0 s"
-proof -
-  have "moment (m + i) s = down_to (m+i) 0 s" using down_to_moment by metis
-  also have "\<dots> = (down_to (m+i) i s) @ (down_to i 0 s)" 
-    by(rule down_to_conc[symmetric], auto)
-  finally show ?thesis .
-qed
-
-lemma length_restn: "length (restn i s) = length s - i"
-proof(cases "i \<le> length s")
-  case True
-  from length_firstn_le [OF this] have "length (firstn i s) = i" .
-  moreover have "length s = length (firstn i s) + length (restn i s)"
-  proof -
-    have "s = firstn i s @ restn i s" using firstn_restn_s by metis
-    hence "length s = length \<dots>" by simp
-    thus ?thesis by simp
-  qed
-  ultimately show ?thesis by simp
-next 
-  case False
-  hence "length s \<le> i" by simp
-  from restn_ge [OF this] have "restn i s = []" .
-  with False show ?thesis by simp
-qed
-
-lemma length_from_to_in:
-  fixes i j s
-  assumes le_ij: "i \<le> j"
-  and le_j: "j \<le> length s"
-  shows "length (from_to i j s) = j - i"
-proof -
-  have "from_to 0 j s = from_to 0 i s @ from_to i j s"
-    by (rule from_to_conc[symmetric, OF _ le_ij], simp)
-  moreover have "length (from_to 0 j s) = j"
-  proof -
-    have "from_to 0 j s = firstn j s" using from_to_firstn by metis
-    moreover have "length \<dots> = j" by (rule length_firstn_le [OF le_j])
-    ultimately show ?thesis by simp
-  qed
-  moreover have "length (from_to 0 i s) = i"
-  proof -
-    have "from_to 0 i s = firstn i s" using from_to_firstn by metis
-    moreover have "length \<dots> = i" 
-    proof (rule length_firstn_le)
-      from le_ij le_j show "i \<le> length s" by simp
-    qed
-    ultimately show ?thesis by simp
-  qed
-  ultimately show ?thesis by auto
-qed
-
-lemma firstn_restn_from_to: "from_to i (m + i) s = firstn m (restn i s)"
-proof(cases "m+i \<le> length s")
-  case True
-  have "restn i s = from_to i (m+i) s @ from_to (m+i) (length s) s"
-  proof -
-    have "restn i s = from_to i (length s) s"
-      by(rule from_to_restn[symmetric], simp)
-    also have "\<dots> = from_to i (m+i) s @ from_to (m+i) (length s) s"
-      by(rule from_to_conc[symmetric, OF _ True], simp)
-    finally show ?thesis .
-  qed
-  hence "firstn m (restn i s) = firstn m \<dots>" by simp
-  moreover have "\<dots> = firstn (length (from_to i (m+i) s)) 
-                    (from_to i (m+i) s @ from_to (m+i) (length s) s)"
-  proof -
-    have "length (from_to i (m+i) s) = m"
-    proof -
-      have "length (from_to i (m+i) s) = (m+i) - i"
-        by(rule length_from_to_in [OF _ True], simp)
-      thus ?thesis by simp
-    qed
-    thus ?thesis by simp
-  qed
-  ultimately show ?thesis using app_firstn_restn by metis
-next
-  case False
-  hence "length s \<le> m + i" by simp
-  from from_to_restn [OF this]
-  have "from_to i (m + i) s = restn i s" .
-  moreover have "firstn m (restn i s) = restn i s" 
-  proof(rule firstn_ge)
-    show "length (restn i s) \<le> m"
-    proof -
-      have "length (restn i s) = length s - i" using length_restn by metis
-      with False show ?thesis by simp
-    qed
-  qed
-  ultimately show ?thesis by simp
-qed
-
-lemma down_to_moment_restm:
-  fixes m i s
-  shows "down_to (m + i) i s = moment m (restm i s)"
-  by (simp add:firstn_restn_from_to down_to_def moment_def restm_def)
-
-lemma moment_plus_split:
-  fixes m i s
-  shows "moment (m + i) s = moment m (restm i s) @ moment i s"
-proof -
-  from moment_split [of m i s]
-  have "moment (m + i) s = down_to (m + i) i s @ down_to i 0 s" .
-  also have "\<dots> = down_to (m+i) i s @ moment i s" using down_to_moment by simp
-  also from down_to_moment_restm have "\<dots> = moment m (restm i s) @ moment i s"
-    by simp
-  finally show ?thesis .
-qed
-
-lemma length_restm: "length (restm i s) = length s - i"
-proof -
-  have "length (rev (restn i (rev s))) = length s - i" (is "?L = ?R")
-  proof -
-    have "?L = length (restn i (rev s))" by simp
-    also have "\<dots>  = length (rev s) - i" using length_restn by metis
-    also have "\<dots> = ?R" by simp
-    finally show ?thesis .
-  qed
-  thus ?thesis by (simp add:restm_def)
-qed
-
-lemma moment_prefix: "(moment i t @ s) = moment (i + length s) (t @ s)"
-proof -
-  from moment_plus_split [of i "length s" "t@s"]
-  have " moment (i + length s) (t @ s) = moment i (restm (length s) (t @ s)) @ moment (length s) (t @ s)"
-    by auto
-  with app_moment_restm[of t s]
-  have "moment (i + length s) (t @ s) = moment i t @ moment (length s) (t @ s)" by simp
-  with moment_app show ?thesis by auto
-qed
-
-end
\ No newline at end of file
--- a/prio/Paper/Paper.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1345 +0,0 @@
-(*<*)
-theory Paper
-imports "../CpsG" "../ExtGG" "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-(*
-find_unused_assms CpsG 
-find_unused_assms ExtGG 
-find_unused_assms Moment 
-find_unused_assms Precedence_ord 
-find_unused_assms PrioG 
-find_unused_assms PrioGDef
-*)
-
-ML {*
-  open Printer;
-  show_question_marks_default := false;
-  *}
-
-notation (latex output)
-  Cons ("_::_" [78,77] 73) and
-  vt ("valid'_state") and
-  runing ("running") and
-  birthtime ("last'_set") and
-  If  ("(\<^raw:\textrm{>if\<^raw:}> (_)/ \<^raw:\textrm{>then\<^raw:}> (_)/ \<^raw:\textrm{>else\<^raw:}> (_))" 10) and
-  Prc ("'(_, _')") and
-  holding ("holds") and
-  waiting ("waits") and
-  Th ("T") and
-  Cs ("C") and
-  readys ("ready") and
-  depend ("RAG") and 
-  preced ("prec") and
-  cpreced ("cprec") and
-  dependents ("dependants") and
-  cp ("cprec") and
-  holdents ("resources") and
-  original_priority ("priority") and
-  DUMMY  ("\<^raw:\mbox{$\_\!\_$}>")
-
-(*abbreviation
- "detached s th \<equiv> cntP s th = cntV s th"
-*)
-(*>*)
-
-section {* Introduction *}
-
-text {*
-  Many real-time systems need to support threads involving priorities and
-  locking of resources. Locking of resources ensures mutual exclusion
-  when accessing shared data or devices that cannot be
-  preempted. Priorities allow scheduling of threads that need to
-  finish their work within deadlines.  Unfortunately, both features
-  can interact in subtle ways leading to a problem, called
-  \emph{Priority Inversion}. Suppose three threads having priorities
-  $H$(igh), $M$(edium) and $L$(ow). We would expect that the thread
-  $H$ blocks any other thread with lower priority and the thread itself cannot
-  be blocked indefinitely by threads with lower priority. Alas, in a naive
-  implementation of resource locking and priorities this property can
-  be violated. For this let $L$ be in the
-  possession of a lock for a resource that $H$ also needs. $H$ must
-  therefore wait for $L$ to exit the critical section and release this
-  lock. The problem is that $L$ might in turn be blocked by any
-  thread with priority $M$, and so $H$ sits there potentially waiting
-  indefinitely. Since $H$ is blocked by threads with lower
-  priorities, the problem is called Priority Inversion. It was first
-  described in \cite{Lampson80} in the context of the
-  Mesa programming language designed for concurrent programming.
-
-  If the problem of Priority Inversion is ignored, real-time systems
-  can become unpredictable and resulting bugs can be hard to diagnose.
-  The classic example where this happened is the software that
-  controlled the Mars Pathfinder mission in 1997 \cite{Reeves98}.
-  Once the spacecraft landed, the software shut down at irregular
-  intervals leading to loss of project time as normal operation of the
-  craft could only resume the next day (the mission and data already
-  collected were fortunately not lost, because of a clever system
-  design).  The reason for the shutdowns was that the scheduling
-  software fell victim to Priority Inversion: a low priority thread
-  locking a resource prevented a high priority thread from running in
-  time, leading to a system reset. Once the problem was found, it was
-  rectified by enabling the \emph{Priority Inheritance Protocol} (PIP)
-  \cite{Sha90}\footnote{Sha et al.~call it the \emph{Basic Priority
-  Inheritance Protocol} \cite{Sha90} and others sometimes also call it
-  \emph{Priority Boosting} or \emph{Priority Donation}.} in the scheduling software.
-
-  The idea behind PIP is to let the thread $L$ temporarily inherit
-  the high priority from $H$ until $L$ leaves the critical section
-  unlocking the resource. This solves the problem of $H$ having to
-  wait indefinitely, because $L$ cannot be blocked by threads having
-  priority $M$. While a few other solutions exist for the Priority
-  Inversion problem, PIP is one that is widely deployed and
-  implemented. This includes VxWorks (a proprietary real-time OS used
-  in the Mars Pathfinder mission, in Boeing's 787 Dreamliner, Honda's
-  ASIMO robot, etc.), but also the POSIX 1003.1c Standard realised for
-  example in libraries for FreeBSD, Solaris and Linux. 
-
-  One advantage of PIP is that increasing the priority of a thread
-  can be dynamically calculated by the scheduler. This is in contrast
-  to, for example, \emph{Priority Ceiling} \cite{Sha90}, another
-  solution to the Priority Inversion problem, which requires static
-  analysis of the program in order to prevent Priority
-  Inversion. However, there has also been strong criticism against
-  PIP. For instance, PIP cannot prevent deadlocks when lock
-  dependencies are circular, and also blocking times can be
-  substantial (more than just the duration of a critical section).
-  Though, most criticism against PIP centres around unreliable
-  implementations and PIP being too complicated and too inefficient.
-  For example, Yodaiken writes in \cite{Yodaiken02}:
-
-  \begin{quote}
-  \it{}``Priority inheritance is neither efficient nor reliable. Implementations
-  are either incomplete (and unreliable) or surprisingly complex and intrusive.''
-  \end{quote}
-
-  \noindent
-  He suggests avoiding PIP altogether by designing the system so that no 
-  priority inversion may happen in the first place. However, such ideal designs may 
-  not always be achievable in practice.
-
-  In our opinion, there is clearly a need for investigating correct
-  algorithms for PIP. A few specifications for PIP exist (in English)
-  and also a few high-level descriptions of implementations (e.g.~in
-  the textbook \cite[Section 5.6.5]{Vahalia96}), but they help little
-  with actual implementations. That this is a problem in practice is
-  proved by an email by Baker, who wrote on 13 July 2009 on the Linux
-  Kernel mailing list:
-
-  \begin{quote}
-  \it{}``I observed in the kernel code (to my disgust), the Linux PIP
-  implementation is a nightmare: extremely heavy weight, involving
-  maintenance of a full wait-for graph, and requiring updates for a
-  range of events, including priority changes and interruptions of
-  wait operations.''
-  \end{quote}
-
-  \noindent
-  The criticism by Yodaiken, Baker and others suggests another look
-  at PIP from a more abstract level (but still concrete enough
-  to inform an implementation), and makes PIP a good candidate for a
-  formal verification. An additional reason is that the original
-  presentation of PIP~\cite{Sha90}, despite being informally
-  ``proved'' correct, is actually \emph{flawed}. 
-
-  Yodaiken \cite{Yodaiken02} points to a subtlety that had been
-  overlooked in the informal proof by Sha et al. They specify in
-  \cite{Sha90} that after the thread (whose priority has been raised)
-  completes its critical section and releases the lock, it ``returns
-  to its original priority level.'' This leads them to believe that an
-  implementation of PIP is ``rather straightforward''~\cite{Sha90}.
-  Unfortunately, as Yodaiken points out, this behaviour is too
-  simplistic.  Consider the case where the low priority thread $L$
-  locks \emph{two} resources, and two high-priority threads $H$ and
-  $H'$ each wait for one of them.  If $L$ releases one resource
-  so that $H$, say, can proceed, then we still have Priority Inversion
-  with $H'$ (which waits for the other resource). The correct
-  behaviour for $L$ is to switch to the highest remaining priority of
-  the threads that it blocks. The advantage of formalising the
-  correctness of a high-level specification of PIP in a theorem prover
-  is that such issues clearly show up and cannot be overlooked as in
-  informal reasoning (since we have to analyse all possible behaviours
-  of threads, i.e.~\emph{traces}, that could possibly happen).\medskip
-
-  \noindent
-  {\bf Contributions:} There have been earlier formal investigations
-  into PIP \cite{Faria08,Jahier09,Wellings07}, but they employ model
-  checking techniques. This paper presents a formalised and
-  mechanically checked proof for the correctness of PIP (to our
-  knowledge the first one).  In contrast to model checking, our
-  formalisation provides insight into why PIP is correct and allows us
-  to prove stronger properties that, as we will show, can help with an
-  efficient implementation of PIP in the educational PINTOS operating
-  system \cite{PINTOS}.  For example, we found by ``playing'' with the
-  formalisation that the choice of the next thread to take over a lock
-  when a resource is released is irrelevant for PIP being correct---a
-  fact that has not been mentioned in the literature and not been used
-  in the reference implementation of PIP in PINTOS.  This fact, however, is important
-  for an efficient implementation of PIP, because we can give the lock
-  to the thread with the highest priority so that it terminates more
-  quickly.
-*}
-
-section {* Formal Model of the Priority Inheritance Protocol *}
-
-text {*
-  The Priority Inheritance Protocol, short PIP, is a scheduling
-  algorithm for a single-processor system.\footnote{We shall come back
-  later to the case of PIP on multi-processor systems.} 
-  Following good experience in earlier work \cite{Wang09},  
-  our model of PIP is based on Paulson's inductive approach to protocol
-  verification \cite{Paulson98}. In this approach a \emph{state} of a system is
-  given by a list of events that happened so far (with new events prepended to the list). 
-  \emph{Events} of PIP fall
-  into five categories defined as the datatype:
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \mbox{\begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{7mm}}l}
-  \isacommand{datatype} event 
-  & @{text "="} & @{term "Create thread priority"}\\
-  & @{text "|"} & @{term "Exit thread"} \\
-  & @{text "|"} & @{term "Set thread priority"} & {\rm reset of the priority for} @{text thread}\\
-  & @{text "|"} & @{term "P thread cs"} & {\rm request of resource} @{text "cs"} {\rm by} @{text "thread"}\\
-  & @{text "|"} & @{term "V thread cs"} & {\rm release of resource} @{text "cs"} {\rm by} @{text "thread"}
-  \end{tabular}}
-  \end{isabelle}
-
-  \noindent
-  whereby threads, priorities and (critical) resources are represented
-  as natural numbers. The event @{term Set} models the situation that
-  a thread obtains a new priority given by the programmer or
-  user (for example via the {\tt nice} utility under UNIX).  As in Paulson's work, we
-  need to define functions that allow us to make some observations
-  about states.  One, called @{term threads}, calculates the set of
-  ``live'' threads that we have seen so far:
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \mbox{\begin{tabular}{lcl}
-  @{thm (lhs) threads.simps(1)} & @{text "\<equiv>"} & 
-    @{thm (rhs) threads.simps(1)}\\
-  @{thm (lhs) threads.simps(2)[where thread="th"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) threads.simps(2)[where thread="th"]}\\
-  @{thm (lhs) threads.simps(3)[where thread="th"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) threads.simps(3)[where thread="th"]}\\
-  @{term "threads (DUMMY#s)"} & @{text "\<equiv>"} & @{term "threads s"}\\
-  \end{tabular}}
-  \end{isabelle}
-
-  \noindent
-  In this definition @{term "DUMMY # DUMMY"} stands for list-cons.
-  Another function calculates the priority for a thread @{text "th"}, which is 
-  defined as
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \mbox{\begin{tabular}{lcl}
-  @{thm (lhs) original_priority.simps(1)[where thread="th"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) original_priority.simps(1)[where thread="th"]}\\
-  @{thm (lhs) original_priority.simps(2)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) original_priority.simps(2)[where thread="th" and thread'="th'"]}\\
-  @{thm (lhs) original_priority.simps(3)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) original_priority.simps(3)[where thread="th" and thread'="th'"]}\\
-  @{term "original_priority th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "original_priority th s"}\\
-  \end{tabular}}
-  \end{isabelle}
-
-  \noindent
-  In this definition we set @{text 0} as the default priority for
-  threads that have not (yet) been created. The last function we need 
-  calculates the ``time'', or index, at which time a process had its 
-  priority last set.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \mbox{\begin{tabular}{lcl}
-  @{thm (lhs) birthtime.simps(1)[where thread="th"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) birthtime.simps(1)[where thread="th"]}\\
-  @{thm (lhs) birthtime.simps(2)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) birthtime.simps(2)[where thread="th" and thread'="th'"]}\\
-  @{thm (lhs) birthtime.simps(3)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & 
-    @{thm (rhs) birthtime.simps(3)[where thread="th" and thread'="th'"]}\\
-  @{term "birthtime th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "birthtime th s"}\\
-  \end{tabular}}
-  \end{isabelle}
-
-  \noindent
-  In this definition @{term "length s"} stands for the length of the list
-  of events @{text s}. Again the default value in this function is @{text 0}
-  for threads that have not been created yet. A \emph{precedence} of a thread @{text th} in a 
-  state @{text s} is the pair of natural numbers defined as
-  
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{thm preced_def[where thread="th"]}
-  \end{isabelle}
-
-  \noindent
-  The point of precedences is to schedule threads not according to priorities (because what should
-  we do in case two threads have the same priority), but according to precedences. 
-  Precedences allow us to always discriminate between two threads with equal priority by 
-  taking into account the time when the priority was last set. We order precedences so 
-  that threads with the same priority get a higher precedence if their priority has been 
-  set earlier, since for such threads it is more urgent to finish their work. In an implementation
-  this choice would translate to a quite natural FIFO-scheduling of processes with 
-  the same priority.
-
-  Next, we introduce the concept of \emph{waiting queues}. They are
-  lists of threads associated with every resource. The first thread in
-  this list (i.e.~the head, or short @{term hd}) is chosen to be the one 
-  that is in possession of the
-  ``lock'' of the corresponding resource. We model waiting queues as
-  functions, below abbreviated as @{text wq}. They take a resource as
-  argument and return a list of threads.  This allows us to define
-  when a thread \emph{holds}, respectively \emph{waits} for, a
-  resource @{text cs} given a waiting queue function @{text wq}.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm cs_holding_def[where thread="th"]}\\
-  @{thm cs_waiting_def[where thread="th"]}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  In this definition we assume @{text "set"} converts a list into a set.
-  At the beginning, that is in the state where no thread is created yet, 
-  the waiting queue function will be the function that returns the
-  empty list for every resource.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{abbrev all_unlocked}\hfill\numbered{allunlocked}
-  \end{isabelle}
-
-  \noindent
-  Using @{term "holding"} and @{term waiting}, we can introduce \emph{Resource Allocation Graphs} 
-  (RAG), which represent the dependencies between threads and resources.
-  We represent RAGs as relations using pairs of the form
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{term "(Th th, Cs cs)"} \hspace{5mm}{\rm and}\hspace{5mm}
-  @{term "(Cs cs, Th th)"}
-  \end{isabelle}
-
-  \noindent
-  where the first stands for a \emph{waiting edge} and the second for a 
-  \emph{holding edge} (@{term Cs} and @{term Th} are constructors of a 
-  datatype for vertices). Given a waiting queue function, a RAG is defined 
-  as the union of the sets of waiting and holding edges, namely
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{thm cs_depend_def}
-  \end{isabelle}
-
-  \noindent
-  Given four threads and three resources, an instance of a RAG can be pictured 
-  as follows:
-
-  \begin{center}
-  \newcommand{\fnt}{\fontsize{7}{8}\selectfont}
-  \begin{tikzpicture}[scale=1]
-  %%\draw[step=2mm] (-3,2) grid (1,-1);
-
-  \node (A) at (0,0) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>0"}};
-  \node (B) at (2,0) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>1"}};
-  \node (C) at (4,0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>1"}};
-  \node (D) at (4,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>2"}};
-  \node (E) at (6,-0.7) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>2"}};
-  \node (E1) at (6, 0.2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>3"}};
-  \node (F) at (8,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>3"}};
-
-  \draw [<-,line width=0.6mm] (A) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding}  (B);
-  \draw [->,line width=0.6mm] (C) to node [pos=0.4,sloped,above=-0.5mm] {\fnt{}waiting}  (B);
-  \draw [->,line width=0.6mm] (D) to node [pos=0.4,sloped,below=-0.5mm] {\fnt{}waiting}  (B);
-  \draw [<-,line width=0.6mm] (D) to node [pos=0.54,sloped,below=-0.5mm] {\fnt{}holding}  (E);
-  \draw [<-,line width=0.6mm] (D) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding}  (E1);
-  \draw [->,line width=0.6mm] (F) to node [pos=0.45,sloped,below=-0.5mm] {\fnt{}waiting}  (E);
-  \end{tikzpicture}
-  \end{center}
-
-  \noindent
-  The use of relations for representing RAGs allows us to conveniently define
-  the notion of the \emph{dependants} of a thread using the transitive closure
-  operation for relations. This gives
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{thm cs_dependents_def}
-  \end{isabelle}
-
-  \noindent
-  This definition needs to account for all threads that wait for a thread to
-  release a resource. This means we need to include threads that transitively
-  wait for a resource being released (in the picture above this means the dependants
-  of @{text "th\<^isub>0"} are @{text "th\<^isub>1"} and @{text "th\<^isub>2"}, which wait for resource @{text "cs\<^isub>1"}, 
-  but also @{text "th\<^isub>3"}, 
-  which cannot make any progress unless @{text "th\<^isub>2"} makes progress, which
-  in turn needs to wait for @{text "th\<^isub>0"} to finish). If there is a circle of dependencies 
-  in a RAG, then clearly
-  we have a deadlock. Therefore when a thread requests a resource,
-  we must ensure that the resulting RAG is not circular. In practice, the 
-  programmer has to ensure this.
-
-
-  Next we introduce the notion of the \emph{current precedence} of a thread @{text th} in a 
-  state @{text s}. It is defined as
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{thm cpreced_def2}\hfill\numbered{cpreced}
-  \end{isabelle}
-
-  \noindent
-  where the dependants of @{text th} are given by the waiting queue function.
-  While the precedence @{term prec} of a thread is determined statically 
-  (for example when the thread is
-  created), the point of the current precedence is to let the scheduler increase this
-  precedence, if needed according to PIP. Therefore the current precedence of @{text th} is
-  given as the maximum of the precedence @{text th} has in state @{text s} \emph{and} all 
-  threads that are dependants of @{text th}. Since the notion @{term "dependants"} is
-  defined as the transitive closure of all dependent threads, we deal correctly with the 
-  problem in the informal algorithm by Sha et al.~\cite{Sha90} where a priority of a thread is
-  lowered prematurely.
-  
-  The next function, called @{term schs}, defines the behaviour of the scheduler. It will be defined
-  by recursion on the state (a list of events); this function returns a \emph{schedule state}, which 
-  we represent as a record consisting of two
-  functions:
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{text "\<lparr>wq_fun, cprec_fun\<rparr>"}
-  \end{isabelle}
-
-  \noindent
-  The first function is a waiting queue function (that is, it takes a
-  resource @{text "cs"} and returns the corresponding list of threads
-  that lock, respectively wait for, it); the second is a function that
-  takes a thread and returns its current precedence (see
-  the definition in \eqref{cpreced}). We assume the usual getter and setter methods for
-  such records.
-
-  In the initial state, the scheduler starts with all resources unlocked (the corresponding 
-  function is defined in \eqref{allunlocked}) and the
-  current precedence of every thread is initialised with @{term "Prc 0 0"}; that means 
-  \mbox{@{abbrev initial_cprec}}. Therefore
-  we have for the initial shedule state
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm (lhs) schs.simps(1)} @{text "\<equiv>"}\\ 
-  \hspace{5mm}@{term "(|wq_fun = all_unlocked, cprec_fun = (\<lambda>_::thread. Prc 0 0)|)"}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  The cases for @{term Create}, @{term Exit} and @{term Set} are also straightforward:
-  we calculate the waiting queue function of the (previous) state @{text s}; 
-  this waiting queue function @{text wq} is unchanged in the next schedule state---because
-  none of these events lock or release any resource; 
-  for calculating the next @{term "cprec_fun"}, we use @{text wq} and 
-  @{term cpreced}. This gives the following three clauses for @{term schs}:
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm (lhs) schs.simps(2)} @{text "\<equiv>"}\\ 
-  \hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
-  \hspace{8mm}@{term "(|wq_fun = wq\<iota>, cprec_fun = cpreced wq\<iota> (Create th prio # s)|)"}\smallskip\\
-  @{thm (lhs) schs.simps(3)} @{text "\<equiv>"}\\
-  \hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
-  \hspace{8mm}@{term "(|wq_fun = wq\<iota>, cprec_fun = cpreced wq\<iota> (Exit th # s)|)"}\smallskip\\
-  @{thm (lhs) schs.simps(4)} @{text "\<equiv>"}\\ 
-  \hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
-  \hspace{8mm}@{term "(|wq_fun = wq\<iota>, cprec_fun = cpreced wq\<iota> (Set th prio # s)|)"}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent 
-  More interesting are the cases where a resource, say @{text cs}, is locked or released. In these cases
-  we need to calculate a new waiting queue function. For the event @{term "P th cs"}, we have to update
-  the function so that the new thread list for @{text cs} is the old thread list plus the thread @{text th} 
-  appended to the end of that list (remember the head of this list is assigned to be in the possession of this
-  resource). This gives the clause
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm (lhs) schs.simps(5)} @{text "\<equiv>"}\\ 
-  \hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
-  \hspace{5mm}@{text "let"} @{text "new_wq = wq(cs := (wq cs @ [th]))"} @{text "in"}\\
-  \hspace{8mm}@{term "(|wq_fun = new_wq, cprec_fun = cpreced new_wq (P th cs # s)|)"}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  The clause for event @{term "V th cs"} is similar, except that we need to update the waiting queue function
-  so that the thread that possessed the lock is deleted from the corresponding thread list. For this 
-  list transformation, we use
-  the auxiliary function @{term release}. A simple version of @{term release} would
-  just delete this thread and return the remaining threads, namely
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}lcl}
-  @{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\
-  @{term "release (DUMMY # qs)"} & @{text "\<equiv>"} & @{term "qs"}\\
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  In practice, however, often the thread with the highest precedence in the list will get the
-  lock next. We have implemented this choice, but later found out that the choice 
-  of which thread is chosen next is actually irrelevant for the correctness of PIP.
-  Therefore we prove the stronger result where @{term release} is defined as
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}lcl}
-  @{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\
-  @{term "release (DUMMY # qs)"} & @{text "\<equiv>"} & @{term "SOME qs'. distinct qs' \<and> set qs' = set qs"}\\
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  where @{text "SOME"} stands for Hilbert's epsilon and implements an arbitrary
-  choice for the next waiting list. It just has to be a list of distinctive threads and
-  contain the same elements as @{text "qs"}. This gives for @{term V} the clause:
- 
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm (lhs) schs.simps(6)} @{text "\<equiv>"}\\
-  \hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
-  \hspace{5mm}@{text "let"} @{text "new_wq = release (wq cs)"} @{text "in"}\\
-  \hspace{8mm}@{term "(|wq_fun = new_wq, cprec_fun = cpreced new_wq (V th cs # s)|)"}
-  \end{tabular}
-  \end{isabelle}
-
-  Having the scheduler function @{term schs} at our disposal, we can ``lift'', or
-  overload, the notions
-  @{term waiting}, @{term holding}, @{term depend} and @{term cp} to operate on states only.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}rcl}
-  @{thm (lhs) s_holding_abv} & @{text "\<equiv>"} & @{thm (rhs) s_holding_abv}\\
-  @{thm (lhs) s_waiting_abv} & @{text "\<equiv>"} & @{thm (rhs) s_waiting_abv}\\
-  @{thm (lhs) s_depend_abv}  & @{text "\<equiv>"} & @{thm (rhs) s_depend_abv}\\
-  @{thm (lhs) cp_def}        & @{text "\<equiv>"} & @{thm (rhs) cp_def}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  With these abbreviations in place we can introduce 
-  the notion of a thread being @{term ready} in a state (i.e.~threads
-  that do not wait for any resource) and the running thread.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm readys_def}\\
-  @{thm runing_def}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  In the second definition @{term "DUMMY ` DUMMY"} stands for the image of a set under a function.
-  Note that in the initial state, that is where the list of events is empty, the set 
-  @{term threads} is empty and therefore there is neither a thread ready nor running.
-  If there is one or more threads ready, then there can only be \emph{one} thread
-  running, namely the one whose current precedence is equal to the maximum of all ready 
-  threads. We use sets to capture both possibilities.
-  We can now also conveniently define the set of resources that are locked by a thread in a
-  given state and also when a thread is detached that state (meaning the thread neither 
-  holds nor waits for a resource):
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm holdents_def}\\
-  @{thm detached_def}
-  \end{tabular}
-  \end{isabelle}
-
-  %\noindent
-  %The second definition states that @{text th}  in @{text s}.
-  
-  Finally we can define what a \emph{valid state} is in our model of PIP. For
-  example we cannot expect to be able to exit a thread, if it was not
-  created yet. 
-  These validity constraints on states are characterised by the
-  inductive predicate @{term "step"} and @{term vt}. We first give five inference rules
-  for @{term step} relating a state and an event that can happen next.
-
-  \begin{center}
-  \begin{tabular}{c}
-  @{thm[mode=Rule] thread_create[where thread=th]}\hspace{1cm}
-  @{thm[mode=Rule] thread_exit[where thread=th]}
-  \end{tabular}
-  \end{center}
-
-  \noindent
-  The first rule states that a thread can only be created, if it is not alive yet.
-  Similarly, the second rule states that a thread can only be terminated if it was
-  running and does not lock any resources anymore (this simplifies slightly our model;
-  in practice we would expect the operating system releases all locks held by a
-  thread that is about to exit). The event @{text Set} can happen
-  if the corresponding thread is running. 
-
-  \begin{center}
-  @{thm[mode=Rule] thread_set[where thread=th]}
-  \end{center}
-
-  \noindent
-  If a thread wants to lock a resource, then the thread needs to be
-  running and also we have to make sure that the resource lock does
-  not lead to a cycle in the RAG. In practice, ensuring the latter
-  is the responsibility of the programmer.  In our formal
-  model we brush aside these problematic cases in order to be able to make
-  some meaningful statements about PIP.\footnote{This situation is
-  similar to the infamous \emph{occurs check} in Prolog: In order to say
-  anything meaningful about unification, one needs to perform an occurs
-  check. But in practice the occurs check is omitted and the
-  responsibility for avoiding problems rests with the programmer.}
-
- 
-  \begin{center}
-  @{thm[mode=Rule] thread_P[where thread=th]}
-  \end{center}
- 
-  \noindent
-  Similarly, if a thread wants to release a lock on a resource, then
-  it must be running and in the possession of that lock. This is
-  formally given by the last inference rule of @{term step}.
- 
-  \begin{center}
-  @{thm[mode=Rule] thread_V[where thread=th]}
-  \end{center}
-
-  \noindent
-  A valid state of PIP can then be conveniently be defined as follows:
-
-  \begin{center}
-  \begin{tabular}{c}
-  @{thm[mode=Axiom] vt_nil}\hspace{1cm}
-  @{thm[mode=Rule] vt_cons}
-  \end{tabular}
-  \end{center}
-
-  \noindent
-  This completes our formal model of PIP. In the next section we present
-  properties that show our model of PIP is correct.
-*}
-
-section {* The Correctness Proof *}
-
-(*<*)
-context extend_highest_gen
-begin
-(*>*)
-text {* 
-  Sha et al.~state their first correctness criterion for PIP in terms
-  of the number of low-priority threads \cite[Theorem 3]{Sha90}: if
-  there are @{text n} low-priority threads, then a blocked job with
-  high priority can only be blocked a maximum of @{text n} times.
-  Their second correctness criterion is given
-  in terms of the number of critical resources \cite[Theorem 6]{Sha90}: if there are
-  @{text m} critical resources, then a blocked job with high priority
-  can only be blocked a maximum of @{text m} times. Both results on their own, strictly speaking, do
-  \emph{not} prevent indefinite, or unbounded, Priority Inversion,
-  because if a low-priority thread does not give up its critical
-  resource (the one the high-priority thread is waiting for), then the
-  high-priority thread can never run.  The argument of Sha et al.~is
-  that \emph{if} threads release locked resources in a finite amount
-  of time, then indefinite Priority Inversion cannot occur---the high-priority
-  thread is guaranteed to run eventually. The assumption is that
-  programmers must ensure that threads are programmed in this way.  However, even
-  taking this assumption into account, the correctness properties of
-  Sha et al.~are
-  \emph{not} true for their version of PIP---despite being ``proved''. As Yodaiken
-  \cite{Yodaiken02} pointed out: If a low-priority thread possesses
-  locks to two resources for which two high-priority threads are
-  waiting for, then lowering the priority prematurely after giving up
-  only one lock, can cause indefinite Priority Inversion for one of the
-  high-priority threads, invalidating their two bounds.
-
-  Even when fixed, their proof idea does not seem to go through for
-  us, because of the way we have set up our formal model of PIP.  One
-  reason is that we allow critical sections, which start with a @{text P}-event
-  and finish with a corresponding @{text V}-event, to arbitrarily overlap
-  (something Sha et al.~explicitly exclude).  Therefore we have
-  designed a different correctness criterion for PIP. The idea behind
-  our criterion is as follows: for all states @{text s}, we know the
-  corresponding thread @{text th} with the highest precedence; we show
-  that in every future state (denoted by @{text "s' @ s"}) in which
-  @{text th} is still alive, either @{text th} is running or it is
-  blocked by a thread that was alive in the state @{text s} and was waiting 
-  for or in the possession of a lock in @{text s}. Since in @{text s}, as in
-  every state, the set of alive threads is finite, @{text th} can only
-  be blocked a finite number of times. This is independent of how many
-  threads of lower priority are created in @{text "s'"}. We will actually prove a
-  stronger statement where we also provide the current precedence of
-  the blocking thread. However, this correctness criterion hinges upon
-  a number of assumptions about the states @{text s} and @{text "s' @
-  s"}, the thread @{text th} and the events happening in @{text
-  s'}. We list them next:
-
-  \begin{quote}
-  {\bf Assumptions on the states {\boldmath@{text s}} and 
-  {\boldmath@{text "s' @ s"}:}} We need to require that @{text "s"} and 
-  @{text "s' @ s"} are valid states:
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{l}
-  @{term "vt s"}, @{term "vt (s' @ s)"} 
-  \end{tabular}
-  \end{isabelle}
-  \end{quote}
-
-  \begin{quote}
-  {\bf Assumptions on the thread {\boldmath@{text "th"}:}} 
-  The thread @{text th} must be alive in @{text s} and 
-  has the highest precedence of all alive threads in @{text s}. Furthermore the
-  priority of @{text th} is @{text prio} (we need this in the next assumptions).
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{l}
-  @{term "th \<in> threads s"}\\
-  @{term "prec th s = Max (cprec s ` threads s)"}\\
-  @{term "prec th s = (prio, DUMMY)"}
-  \end{tabular}
-  \end{isabelle}
-  \end{quote}
-  
-  \begin{quote}
-  {\bf Assumptions on the events in {\boldmath@{text "s'"}:}} We want to prove that @{text th} cannot
-  be blocked indefinitely. Of course this can happen if threads with higher priority
-  than @{text th} are continuously created in @{text s'}. Therefore we have to assume that  
-  events in @{text s'} can only create (respectively set) threads with equal or lower 
-  priority than @{text prio} of @{text th}. We also need to assume that the
-  priority of @{text "th"} does not get reset and also that @{text th} does
-  not get ``exited'' in @{text "s'"}. This can be ensured by assuming the following three implications. 
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{l}
-  {If}~~@{text "Create th' prio' \<in> set s'"}~~{then}~~@{text "prio' \<le> prio"}\\
-  {If}~~@{text "Set th' prio' \<in> set s'"}~~{then}~~@{text "th' \<noteq> th"}~~{and}~~@{text "prio' \<le> prio"}\\
-  {If}~~@{text "Exit th' \<in> set s'"}~~{then}~~@{text "th' \<noteq> th"}\\
-  \end{tabular}
-  \end{isabelle}
-  \end{quote}
-
-  \noindent
-  The locale mechanism of Isabelle helps us to manage conveniently such assumptions~\cite{Haftmann08}.
-  Under these assumptions we shall prove the following correctness property:
-
-  \begin{theorem}\label{mainthm}
-  Given the assumptions about states @{text "s"} and @{text "s' @ s"},
-  the thread @{text th} and the events in @{text "s'"},
-  if @{term "th' \<in> running (s' @ s)"} and @{text "th' \<noteq> th"} then
-  @{text "th' \<in> threads s"}, @{text "\<not> detached s th'"} and 
-  @{term "cp (s' @ s) th' = prec th s"}.
-  \end{theorem}
-
-  \noindent
-  This theorem ensures that the thread @{text th}, which has the
-  highest precedence in the state @{text s}, can only be blocked in
-  the state @{text "s' @ s"} by a thread @{text th'} that already
-  existed in @{text s} and requested or had a lock on at least 
-  one resource---that means the thread was not \emph{detached} in @{text s}. 
-  As we shall see shortly, that means there are only finitely 
-  many threads that can block @{text th} in this way and then they 
-  need to run with the same current precedence as @{text th}.
-
-  Like in the argument by Sha et al.~our
-  finite bound does not guarantee absence of indefinite Priority
-  Inversion. For this we further have to assume that every thread
-  gives up its resources after a finite amount of time. We found that
-  this assumption is awkward to formalise in our model. Therefore we
-  leave it out and let the programmer assume the responsibility to
-  program threads in such a benign manner (in addition to causing no 
-  circularity in the @{text RAG}). In this detail, we do not
-  make any progress in comparison with the work by Sha et al.
-  However, we are able to combine their two separate bounds into a
-  single theorem improving their bound.
-
-  In what follows we will describe properties of PIP that allow us to prove 
-  Theorem~\ref{mainthm} and, when instructive, briefly describe our argument. 
-  It is relatively easy to see that
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{text "running s \<subseteq> ready s \<subseteq> threads s"}\\
-  @{thm[mode=IfThen]  finite_threads}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  The second property is by induction of @{term vt}. The next three
-  properties are 
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm[mode=IfThen] waiting_unique[of _ _ "cs\<^isub>1" "cs\<^isub>2"]}\\
-  @{thm[mode=IfThen] held_unique[of _ "th\<^isub>1" _ "th\<^isub>2"]}\\
-  @{thm[mode=IfThen] runing_unique[of _ "th\<^isub>1" "th\<^isub>2"]}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  The first property states that every waiting thread can only wait for a single
-  resource (because it gets suspended after requesting that resource); the second 
-  that every resource can only be held by a single thread; 
-  the third property establishes that in every given valid state, there is
-  at most one running thread. We can also show the following properties 
-  about the @{term RAG} in @{text "s"}.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{text If}~@{thm (prem 1) acyclic_depend}~@{text "then"}:\\
-  \hspace{5mm}@{thm (concl) acyclic_depend},
-  @{thm (concl) finite_depend} and
-  @{thm (concl) wf_dep_converse},\\
-  \hspace{5mm}@{text "if"}~@{thm (prem 2) dm_depend_threads}~@{text "then"}~@{thm (concl) dm_depend_threads}
-  and\\
-  \hspace{5mm}@{text "if"}~@{thm (prem 2) range_in}~@{text "then"}~@{thm (concl) range_in}.
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  The acyclicity property follows from how we restricted the events in
-  @{text step}; similarly the finiteness and well-foundedness property.
-  The last two properties establish that every thread in a @{text "RAG"}
-  (either holding or waiting for a resource) is a live thread.
-
-  The key lemma in our proof of Theorem~\ref{mainthm} is as follows:
-
-  \begin{lemma}\label{mainlem}
-  Given the assumptions about states @{text "s"} and @{text "s' @ s"},
-  the thread @{text th} and the events in @{text "s'"},
-  if @{term "th' \<in> threads (s' @ s)"}, @{text "th' \<noteq> th"} and @{text "detached (s' @ s) th'"}\\
-  then @{text "th' \<notin> running (s' @ s)"}.
-  \end{lemma}
-
-  \noindent
-  The point of this lemma is that a thread different from @{text th} (which has the highest
-  precedence in @{text s}) and not holding any resource, cannot be running 
-  in the state @{text "s' @ s"}.
-
-  \begin{proof}
-  Since thread @{text "th'"} does not hold any resource, no thread can depend on it. 
-  Therefore its current precedence @{term "cp (s' @ s) th'"} equals its own precedence
-  @{term "prec th' (s' @ s)"}. Since @{text "th"} has the highest precedence in the 
-  state @{text "(s' @ s)"} and precedences are distinct among threads, we have
-  @{term "prec th' (s' @s ) < prec th (s' @ s)"}. From this 
-  we have @{term "cp (s' @ s) th' < prec th (s' @ s)"}.
-  Since @{text "prec th (s' @ s)"} is already the highest 
-  @{term "cp (s' @ s) th"} can not be higher than this and can not be lower either (by 
-  definition of @{term "cp"}). Consequently, we have @{term "prec th (s' @ s) = cp (s' @ s) th"}.
-  Finally we have @{term "cp (s' @ s) th' < cp (s' @ s) th"}.
-  By defintion of @{text "running"}, @{text "th'"} can not be running in state
-  @{text "s' @ s"}, as we had to show.\qed
-  \end{proof}
-
-  \noindent
-  Since @{text "th'"} is not able to run in state @{text "s' @ s"}, it is not able to 
-  issue a @{text "P"} or @{text "V"} event. Therefore if @{text "s' @ s"} is extended
-  one step further, @{text "th'"} still cannot hold any resource. The situation will 
-  not change in further extensions as long as @{text "th"} holds the highest precedence.
-
-  From this lemma we can deduce Theorem~\ref{mainthm}: that @{text th} can only be 
-  blocked by a thread @{text th'} that
-  held some resource in state @{text s} (that is not @{text "detached"}). And furthermore
-  that the current precedence of @{text th'} in state @{text "(s' @ s)"} must be equal to the 
-  precedence of @{text th} in @{text "s"}.
-  We show this theorem by induction on @{text "s'"} using Lemma~\ref{mainlem}.
-  This theorem gives a stricter bound on the threads that can block @{text th} than the
-  one obtained by Sha et al.~\cite{Sha90}:
-  only threads that were alive in state @{text s} and moreover held a resource.
-  This means our bound is in terms of both---alive threads in state @{text s}
-  and number of critical resources. Finally, the theorem establishes that the blocking threads have the
-  current precedence raised to the precedence of @{text th}.
-
-  We can furthermore prove that under our assumptions no deadlock exists in the state @{text "s' @ s"}
-  by showing that @{text "running (s' @ s)"} is not empty.
-
-  \begin{lemma}
-  Given the assumptions about states @{text "s"} and @{text "s' @ s"},
-  the thread @{text th} and the events in @{text "s'"},
-  @{term "running (s' @ s) \<noteq> {}"}.
-  \end{lemma}
-
-  \begin{proof}
-  If @{text th} is blocked, then by following its dependants graph, we can always 
-  reach a ready thread @{text th'}, and that thread must have inherited the 
-  precedence of @{text th}.\qed
-  \end{proof}
-
-
-  %The following lemmas show how every node in RAG can be chased to ready threads:
-  %\begin{enumerate}
-  %\item Every node in RAG can be chased to a ready thread (@{text "chain_building"}):
-  %  @   {thm [display] chain_building[rule_format]}
-  %\item The ready thread chased to is unique (@{text "dchain_unique"}):
-  %  @   {thm [display] dchain_unique[of _ _ "th\<^isub>1" "th\<^isub>2"]}
-  %\end{enumerate}
-
-  %Some deeper results about the system:
-  %\begin{enumerate}
-  %\item The maximum of @{term "cp"} and @{term "preced"} are equal (@{text "max_cp_eq"}):
-  %@  {thm [display] max_cp_eq}
-  %\item There must be one ready thread having the max @{term "cp"}-value 
-  %(@{text "max_cp_readys_threads"}):
-  %@  {thm [display] max_cp_readys_threads}
-  %\end{enumerate}
-
-  %The relationship between the count of @{text "P"} and @{text "V"} and the number of 
-  %critical resources held by a thread is given as follows:
-  %\begin{enumerate}
-  %\item The @{term "V"}-operation decreases the number of critical resources 
-  %  one thread holds (@{text "cntCS_v_dec"})
-  %   @  {thm [display]  cntCS_v_dec}
-  %\item The number of @{text "V"} never exceeds the number of @{text "P"} 
-  %  (@  {text "cnp_cnv_cncs"}):
-  %  @  {thm [display]  cnp_cnv_cncs}
-  %\item The number of @{text "V"} equals the number of @{text "P"} when 
-  %  the relevant thread is not living:
-  %  (@{text "cnp_cnv_eq"}):
-  %  @  {thm [display]  cnp_cnv_eq}
-  %\item When a thread is not living, it does not hold any critical resource 
-  %  (@{text "not_thread_holdents"}):
-  %  @  {thm [display] not_thread_holdents}
-  %\item When the number of @{text "P"} equals the number of @{text "V"}, the relevant 
-  %  thread does not hold any critical resource, therefore no thread can depend on it
-  %  (@{text "count_eq_dependents"}):
-  %  @  {thm [display] count_eq_dependents}
-  %\end{enumerate}
-
-  %The reason that only threads which already held some resoures
-  %can be runing and block @{text "th"} is that if , otherwise, one thread 
-  %does not hold any resource, it may never have its prioirty raised
-  %and will not get a chance to run. This fact is supported by 
-  %lemma @{text "moment_blocked"}:
-  %@   {thm [display] moment_blocked}
-  %When instantiating  @{text "i"} to @{text "0"}, the lemma means threads which did not hold any
-  %resource in state @{text "s"} will not have a change to run latter. Rephrased, it means 
-  %any thread which is running after @{text "th"} became the highest must have already held
-  %some resource at state @{text "s"}.
-
-
-  %When instantiating @{text "i"} to a number larger than @{text "0"}, the lemma means 
-  %if a thread releases all its resources at some moment in @{text "t"}, after that, 
-  %it may never get a change to run. If every thread releases its resource in finite duration,
-  %then after a while, only thread @{text "th"} is left running. This shows how indefinite 
-  %priority inversion can be avoided. 
-
-  %All these assumptions are put into a predicate @{term "extend_highest_gen"}. 
-  %It can be proved that @{term "extend_highest_gen"} holds 
-  %for any moment @{text "i"} in it @{term "t"} (@{text "red_moment"}):
-  %@   {thm [display] red_moment}
-  
-  %From this, an induction principle can be derived for @{text "t"}, so that 
-  %properties already derived for @{term "t"} can be applied to any prefix 
-  %of @{text "t"} in the proof of new properties 
-  %about @{term "t"} (@{text "ind"}):
-  %\begin{center}
-  %@   {thm[display] ind}
-  %\end{center}
-
-  %The following properties can be proved about @{term "th"} in @{term "t"}:
-  %\begin{enumerate}
-  %\item In @{term "t"}, thread @{term "th"} is kept live and its 
-  %  precedence is preserved as well
-  %  (@{text "th_kept"}): 
-  %  @   {thm [display] th_kept}
-  %\item In @{term "t"}, thread @{term "th"}'s precedence is always the maximum among 
-  %  all living threads
-  %  (@{text "max_preced"}): 
-  %  @   {thm [display] max_preced}
-  %\item In @{term "t"}, thread @{term "th"}'s current precedence is always the maximum precedence
-  %  among all living threads
-  %  (@{text "th_cp_max_preced"}): 
-  %  @   {thm [display] th_cp_max_preced}
-  %\item In @{term "t"}, thread @{term "th"}'s current precedence is always the maximum current 
-  %  precedence among all living threads
-  %  (@{text "th_cp_max"}): 
-  %  @   {thm [display] th_cp_max}
-  %\item In @{term "t"}, thread @{term "th"}'s current precedence equals its precedence at moment 
-  %  @{term "s"}
-  %  (@{text "th_cp_preced"}): 
-  %  @   {thm [display] th_cp_preced}
-  %\end{enumerate}
-
-  %The main theorem of this part is to characterizing the running thread during @{term "t"} 
-  %(@{text "runing_inversion_2"}):
-  %@   {thm [display] runing_inversion_2}
-  %According to this, if a thread is running, it is either @{term "th"} or was
-  %already live and held some resource 
-  %at moment @{text "s"} (expressed by: @{text "cntV s th' < cntP s th'"}).
-
-  %Since there are only finite many threads live and holding some resource at any moment,
-  %if every such thread can release all its resources in finite duration, then after finite
-  %duration, none of them may block @{term "th"} anymore. So, no priority inversion may happen
-  %then.
-  *}
-(*<*)
-end
-(*>*)
-
-section {* Properties for an Implementation\label{implement} *}
-
-text {*
-  While our formalised proof gives us confidence about the correctness of our model of PIP, 
-  we found that the formalisation can even help us with efficiently implementing it.
-
-  For example Baker complained that calculating the current precedence
-  in PIP is quite ``heavy weight'' in Linux (see the Introduction).
-  In our model of PIP the current precedence of a thread in a state @{text s}
-  depends on all its dependants---a ``global'' transitive notion,
-  which is indeed heavy weight (see Def.~shown in \eqref{cpreced}).
-  We can however improve upon this. For this let us define the notion
-  of @{term children} of a thread @{text th} in a state @{text s} as
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm children_def2}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  where a child is a thread that is only one ``hop'' away from the thread
-  @{text th} in the @{term RAG} (and waiting for @{text th} to release
-  a resource). We can prove the following lemma.
-
-  \begin{lemma}\label{childrenlem}
-  @{text "If"} @{thm (prem 1) cp_rec} @{text "then"}
-  \begin{center}
-  @{thm (concl) cp_rec}.
-  \end{center}
-  \end{lemma}
-  
-  \noindent
-  That means the current precedence of a thread @{text th} can be
-  computed locally by considering only the children of @{text th}. In
-  effect, it only needs to be recomputed for @{text th} when one of
-  its children changes its current precedence.  Once the current 
-  precedence is computed in this more efficient manner, the selection
-  of the thread with highest precedence from a set of ready threads is
-  a standard scheduling operation implemented in most operating
-  systems.
-
-  Of course the main work for implementing PIP involves the
-  scheduler and coding how it should react to events.  Below we
-  outline how our formalisation guides this implementation for each
-  kind of events.\smallskip
-*}
-
-(*<*)
-context step_create_cps
-begin
-(*>*)
-text {*
-  \noindent
-  \colorbox{mygrey}{@{term "Create th prio"}:} We assume that the current state @{text s'} and 
-  the next state @{term "s \<equiv> Create th prio#s'"} are both valid (meaning the event
-  is allowed to occur). In this situation we can show that
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm eq_dep},\\
-  @{thm eq_cp_th}, and\\
-  @{thm[mode=IfThen] eq_cp}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  This means in an implementation we do not have recalculate the @{text RAG} and also none of the
-  current precedences of the other threads. The current precedence of the created
-  thread @{text th} is just its precedence, namely the pair @{term "(prio, length (s::event list))"}.
-  \smallskip
-  *}
-(*<*)
-end
-context step_exit_cps
-begin
-(*>*)
-text {*
-  \noindent
-  \colorbox{mygrey}{@{term "Exit th"}:} We again assume that the current state @{text s'} and 
-  the next state @{term "s \<equiv> Exit th#s'"} are both valid. We can show that
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm eq_dep}, and\\
-  @{thm[mode=IfThen] eq_cp}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  This means again we do not have to recalculate the @{text RAG} and
-  also not the current precedences for the other threads. Since @{term th} is not
-  alive anymore in state @{term "s"}, there is no need to calculate its
-  current precedence.
-  \smallskip
-*}
-(*<*)
-end
-context step_set_cps
-begin
-(*>*)
-text {*
-  \noindent
-  \colorbox{mygrey}{@{term "Set th prio"}:} We assume that @{text s'} and 
-  @{term "s \<equiv> Set th prio#s'"} are both valid. We can show that
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm[mode=IfThen] eq_dep}, and\\
-  @{thm[mode=IfThen] eq_cp_pre}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  The first property is again telling us we do not need to change the @{text RAG}. 
-  The second shows that the @{term cp}-values of all threads other than @{text th} 
-  are unchanged. The reason is that @{text th} is running; therefore it is not in 
-  the @{term dependants} relation of any other thread. This in turn means that the 
-  change of its priority cannot affect other threads.
-
-  %The second
-  %however states that only threads that are \emph{not} dependants of @{text th} have their
-  %current precedence unchanged. For the others we have to recalculate the current
-  %precedence. To do this we can start from @{term "th"} 
-  %and follow the @{term "depend"}-edges to recompute  using Lemma~\ref{childrenlem} 
-  %the @{term "cp"} of every 
-  %thread encountered on the way. Since the @{term "depend"}
-  %is assumed to be loop free, this procedure will always stop. The following two lemmas show, however, 
-  %that this procedure can actually stop often earlier without having to consider all
-  %dependants.
-  %
-  %\begin{isabelle}\ \ \ \ \ %%%
-  %\begin{tabular}{@ {}l}
-  %@{thm[mode=IfThen] eq_up_self}\\
-  %@{text "If"} @{thm (prem 1) eq_up}, @{thm (prem 2) eq_up} and @{thm (prem 3) eq_up}\\
-  %@{text "then"} @{thm (concl) eq_up}.
-  %\end{tabular}
-  %\end{isabelle}
-  %
-  %\noindent
-  %The first lemma states that if the current precedence of @{text th} is unchanged,
-  %then the procedure can stop immediately (all dependent threads have their @{term cp}-value unchanged).
-  %The second states that if an intermediate @{term cp}-value does not change, then
-  %the procedure can also stop, because none of its dependent threads will
-  %have their current precedence changed.
-  \smallskip
-  *}
-(*<*)
-end
-context step_v_cps_nt
-begin
-(*>*)
-text {*
-  \noindent
-  \colorbox{mygrey}{@{term "V th cs"}:} We assume that @{text s'} and 
-  @{term "s \<equiv> V th cs#s'"} are both valid. We have to consider two
-  subcases: one where there is a thread to ``take over'' the released
-  resource @{text cs}, and one where there is not. Let us consider them
-  in turn. Suppose in state @{text s}, the thread @{text th'} takes over
-  resource @{text cs} from thread @{text th}. We can prove
-
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{thm depend_s}
-  \end{isabelle}
-  
-  \noindent
-  which shows how the @{text RAG} needs to be changed. The next lemma suggests
-  how the current precedences need to be recalculated. For threads that are
-  not @{text "th"} and @{text "th'"} nothing needs to be changed, since we
-  can show
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  @{thm[mode=IfThen] cp_kept}
-  \end{isabelle}
-  
-  \noindent
-  For @{text th} and @{text th'} we need to use Lemma~\ref{childrenlem} to
-  recalculate their current precedence since their children have changed. *}(*<*)end context step_v_cps_nnt begin (*>*)text {*
-  \noindent
-  In the other case where there is no thread that takes over @{text cs}, we can show how
-  to recalculate the @{text RAG} and also show that no current precedence needs
-  to be recalculated.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm depend_s}\\
-  @{thm eq_cp}
-  \end{tabular}
-  \end{isabelle}
-  *}
-(*<*)
-end
-context step_P_cps_e
-begin
-(*>*)
-text {*
-  \noindent
-  \colorbox{mygrey}{@{term "P th cs"}:} We assume that @{text s'} and 
-  @{term "s \<equiv> P th cs#s'"} are both valid. We again have to analyse two subcases, namely
-  the one where @{text cs} is not locked, and one where it is. We treat the former case
-  first by showing that
-  
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm depend_s}\\
-  @{thm eq_cp}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  This means we need to add a holding edge to the @{text RAG} and no
-  current precedence needs to be recalculated.*}(*<*)end context step_P_cps_ne begin(*>*) text {*
-  \noindent
-  In the second case we know that resource @{text cs} is locked. We can show that
-  
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm depend_s}\\
-  @{thm[mode=IfThen] eq_cp}
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  That means we have to add a waiting edge to the @{text RAG}. Furthermore
-  the current precedence for all threads that are not dependants of @{text th}
-  are unchanged. For the others we need to follow the edges 
-  in the @{text RAG} and recompute the @{term "cp"}. To do this we can start from @{term "th"} 
-  and follow the @{term "depend"}-edges to recompute  using Lemma~\ref{childrenlem} 
-  the @{term "cp"} of every 
-  thread encountered on the way. Since the @{term "depend"}
-  is loop free, this procedure will always stop. The following lemma shows, however, 
-  that this procedure can actually stop often earlier without having to consider all
-  dependants.
-
-  \begin{isabelle}\ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  %%@ {t hm[mode=IfThen] eq_up_self}\\
-  @{text "If"} @{thm (prem 1) eq_up}, @{thm (prem 2) eq_up} and @{thm (prem 3) eq_up}\\
-  @{text "then"} @{thm (concl) eq_up}.
-  \end{tabular}
-  \end{isabelle}
-
-  \noindent
-  This lemma states that if an intermediate @{term cp}-value does not change, then
-  the procedure can also stop, because none of its dependent threads will
-  have their current precedence changed.
-  *}
-(*<*)
-end
-(*>*)
-text {*
-  \noindent
-  As can be seen, a pleasing byproduct of our formalisation is that the properties in
-  this section closely inform an implementation of PIP, namely whether the
-  @{text RAG} needs to be reconfigured or current precedences need to
-  be recalculated for an event. This information is provided by the lemmas we proved.
-  We confirmed that our observations translate into practice by implementing
-  our version of PIP on top of PINTOS, a small operating system written in C and used for teaching at 
-  Stanford University \cite{PINTOS}. To implement PIP, we only need to modify the kernel 
-  functions corresponding to the events in our formal model. The events translate to the following 
-  function interface in PINTOS:
-
-  \begin{center}
-  \begin{tabular}{|l@ {\hspace{2mm}}|l@ {\hspace{2mm}}|}
-  \hline
-  {\bf Event} & {\bf PINTOS function} \\
-  \hline
-  @{text Create} & @{text "thread_create"}\\
-  @{text Exit}   & @{text "thread_exit"}\\
-  @{text Set}    & @{text "thread_set_priority"}\\
-  @{text P}      & @{text "lock_acquire"}\\
-  @{text V}      & @{text "lock_release"}\\ 
-  \hline
-  \end{tabular}
-  \end{center}
-
-  \noindent
-  Our implicit assumption that every event is an atomic operation is ensured by the architecture of 
-  PINTOS. The case where an unlocked resource is given next to the waiting thread with the
-  highest precedence is realised in our implementation by priority queues. We implemented
-  them as \emph{Braun trees} \cite{Paulson96}, which provide efficient @{text "O(log n)"}-operations
-  for accessing and updating. Apart from having to implement relatively complex data\-structures in C
-  using pointers, our experience with the implementation has been very positive: our specification 
-  and formalisation of PIP translates smoothly to an efficent implementation in PINTOS. 
-*}
-
-section {* Conclusion *}
-
-text {* 
-  The Priority Inheritance Protocol (PIP) is a classic textbook
-  algorithm used in many real-time operating systems in order to avoid the problem of
-  Priority Inversion.  Although classic and widely used, PIP does have
-  its faults: for example it does not prevent deadlocks in cases where threads
-  have circular lock dependencies.
-
-  We had two goals in mind with our formalisation of PIP: One is to
-  make the notions in the correctness proof by Sha et al.~\cite{Sha90}
-  precise so that they can be processed by a theorem prover. The reason is
-  that a mechanically checked proof avoids the flaws that crept into their
-  informal reasoning. We achieved this goal: The correctness of PIP now
-  only hinges on the assumptions behind our formal model. The reasoning, which is
-  sometimes quite intricate and tedious, has been checked by Isabelle/HOL. 
-  We can also confirm that Paulson's
-  inductive method for protocol verification~\cite{Paulson98} is quite
-  suitable for our formal model and proof. The traditional application
-  area of this method is security protocols. 
-
-  The second goal of our formalisation is to provide a specification for actually
-  implementing PIP. Textbooks, for example \cite[Section 5.6.5]{Vahalia96},
-  explain how to use various implementations of PIP and abstractly
-  discuss their properties, but surprisingly lack most details important for a
-  programmer who wants to implement PIP (similarly Sha et al.~\cite{Sha90}).  
-  That this is an issue in practice is illustrated by the
-  email from Baker we cited in the Introduction. We achieved also this
-  goal: The formalisation allowed us to efficently implement our version
-  of PIP on top of PINTOS \cite{PINTOS}, a simple instructional operating system for the x86 
-  architecture. It also gives the first author enough data to enable
-  his undergraduate students to implement PIP (as part of their OS course).
-  A byproduct of our formalisation effort is that nearly all
-  design choices for the PIP scheduler are backed up with a proved
-  lemma. We were also able to establish the property that the choice of
-  the next thread which takes over a lock is irrelevant for the correctness
-  of PIP. 
-
-  PIP is a scheduling algorithm for single-processor systems. We are
-  now living in a multi-processor world. Priority Inversion certainly
-  occurs also there.  However, there is very little ``foundational''
-  work about PIP-algorithms on multi-processor systems.  We are not
-  aware of any correctness proofs, not even informal ones. There is an
-  implementation of a PIP-algorithm for multi-processors as part of the
-  ``real-time'' effort in Linux, including an informal description of the implemented scheduling
-  algorithm given in \cite{LINUX}.  We estimate that the formal
-  verification of this algorithm, involving more fine-grained events,
-  is a magnitude harder than the one we presented here, but still
-  within reach of current theorem proving technology. We leave this
-  for future work.
-
-  The most closely related work to ours is the formal verification in
-  PVS of the Priority Ceiling Protocol done by Dutertre
-  \cite{dutertre99b}---another solution to the Priority Inversion
-  problem, which however needs static analysis of programs in order to
-  avoid it. There have been earlier formal investigations
-  into PIP \cite{Faria08,Jahier09,Wellings07}, but they employ model
-  checking techniques. The results obtained by them apply,
-  however, only to systems with a fixed size, such as a fixed number of 
-  events and threads. In contrast, our result applies to systems of arbitrary
-  size. Moreover, our result is a good 
-  witness for one of the major reasons to be interested in machine checked 
-  reasoning: gaining deeper understanding of the subject matter.
-
-  Our formalisation
-  consists of around 210 lemmas and overall 6950 lines of readable Isabelle/Isar
-  code with a few apply-scripts interspersed. The formal model of PIP
-  is 385 lines long; the formal correctness proof 3800 lines. Some auxiliary
-  definitions and proofs span over 770 lines of code. The properties relevant
-  for an implementation require 2000 lines. 
-  %The code of our formalisation 
-  %can be downloaded from
-  %\url{http://www.inf.kcl.ac.uk/staff/urbanc/pip.html}.\medskip
-
-  \noindent
-  {\bf Acknowledgements:}
-  We are grateful for the comments we received from anonymous
-  referees.
-
-  \bibliographystyle{plain}
-  \bibliography{root}
-*}
-
-
-(*<*)
-end
-(*>*)
\ No newline at end of file
--- a/prio/Paper/PrioGDef.tex	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,488 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{PrioGDef}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-In this section, the formal model of Priority Inheritance is presented. First, the identifiers of {\em threads},
-  {\em priority} and {\em critical resources } (abbreviated as \isa{cs}) are all represented as natural numbers,
-  i.e. standard Isabelle/HOL type \isa{nat}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}\isamarkupfalse%
-\ thread\ {\isaliteral{3D}{\isacharequal}}\ nat\ %
-\isamarkupcmt{Type for thread identifiers.%
-}
-\isanewline
-\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}\isamarkupfalse%
-\ priority\ {\isaliteral{3D}{\isacharequal}}\ nat\ \ %
-\isamarkupcmt{Type for priorities.%
-}
-\isanewline
-\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}\isamarkupfalse%
-\ cs\ {\isaliteral{3D}{\isacharequal}}\ nat\ %
-\isamarkupcmt{Type for critical sections (or critical resources).%
-}
-%
-\begin{isamarkuptext}%
-Priority Inheritance protocol is modeled as an event driven system, where every event represents an 
-  system call. Event format is given by the following type definition:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{datatype}\isamarkupfalse%
-\ event\ {\isaliteral{3D}{\isacharequal}}\ \isanewline
-\ \ Create\ thread\ priority\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{Thread \isa{thread} is created with priority \isa{priority}.%
-}
-\isanewline
-\ \ Exit\ thread\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{Thread \isa{thread} finishing its execution.%
-}
-\isanewline
-\ \ P\ thread\ cs\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{Thread \isa{thread} requesting critical resource \isa{cs}.%
-}
-\isanewline
-\ \ V\ thread\ cs\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{Thread \isa{thread}  releasing critical resource \isa{cs}.%
-}
-\isanewline
-\ \ Set\ thread\ priority\ %
-\isamarkupcmt{Thread \isa{thread} resets its priority to \isa{priority}.%
-}
-%
-\begin{isamarkuptext}%
-Resource Allocation Graph (RAG for short) is used extensively in the analysis of Priority Inheritance. 
-  The following type \isa{node} is used to model nodes in RAG.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{datatype}\isamarkupfalse%
-\ node\ {\isaliteral{3D}{\isacharequal}}\ \isanewline
-\ \ \ Th\ {\isaliteral{22}{\isachardoublequoteopen}}thread{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{Node for thread.%
-}
-\isanewline
-\ \ \ Cs\ {\isaliteral{22}{\isachardoublequoteopen}}cs{\isaliteral{22}{\isachardoublequoteclose}}\ %
-\isamarkupcmt{Node for critical resource.%
-}
-%
-\begin{isamarkuptext}%
-The protocol is analyzed using Paulson's inductive protocol verification method, where 
-  the state of the system is modelled as the list of events happened so far with the latest 
-  event at the head. Therefore, the state of the system is represented by the following
-  type \isa{state}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}\isamarkupfalse%
-\ state\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequoteopen}}event\ list{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The following \isa{threads} is used to calculate the set of live threads (\isa{threads\ s})
-  in state \isa{s}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{fun}\isamarkupfalse%
-\ threads\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ \isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}threads\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{At the start of the system, the set of threads is empty.%
-}
-\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}threads\ {\isaliteral{28}{\isacharparenleft}}Create\ thread\ prio{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}thread{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ threads\ s{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{New thread is added to the \isa{threads}.%
-}
-\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}threads\ {\isaliteral{28}{\isacharparenleft}}Exit\ thread\ {\isaliteral{23}{\isacharhash}}\ s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}threads\ s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}\ {\isaliteral{7B}{\isacharbraceleft}}thread{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{Finished thread is removed.%
-}
-\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}threads\ {\isaliteral{28}{\isacharparenleft}}e{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ threads\ s{\isaliteral{22}{\isachardoublequoteclose}}\ %
-\isamarkupcmt{other kind of events does not affect the value of \isa{threads}.%
-}
-%
-\begin{isamarkuptext}%
-Functions such as \isa{threads}, which extract information out of system states, are called
-  {\em observing functions}. A series of observing functions will be defined in the sequel in order to 
-  model the protocol. 
-  Observing function \isa{original{\isaliteral{5F}{\isacharunderscore}}priority} calculates 
-  the {\em original priority} of thread \isa{th} in state \isa{s}, expressed as
-  : \isa{original{\isaliteral{5F}{\isacharunderscore}}priority\ th\ s}. The {\em original priority} is the priority 
-  assigned to a thread when it is created or when it is reset by system call \isa{Set\ thread\ priority}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{fun}\isamarkupfalse%
-\ original{\isaliteral{5F}{\isacharunderscore}}priority\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ priority{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ %
-\isamarkupcmt{\isa{{\isadigit{0}}} is assigned to threads which have never been created.%
-}
-\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ {\isaliteral{28}{\isacharparenleft}}Create\ thread{\isaliteral{27}{\isacharprime}}\ prio{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ \isanewline
-\ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}if\ thread{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ thread\ then\ prio\ else\ original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ {\isaliteral{28}{\isacharparenleft}}Set\ thread{\isaliteral{27}{\isacharprime}}\ prio{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ \isanewline
-\ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}if\ thread{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ thread\ then\ prio\ else\ original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ {\isaliteral{28}{\isacharparenleft}}e{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ s{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\isa{birthtime\ th\ s} is the time when thread \isa{th} is created, observed from state \isa{s}.
-  The time in the system is measured by the number of events happened so far since the very beginning.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{fun}\isamarkupfalse%
-\ birthtime\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}birthtime\ thread\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}birthtime\ thread\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}Create\ thread{\isaliteral{27}{\isacharprime}}\ prio{\isaliteral{29}{\isacharparenright}}{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ {\isaliteral{28}{\isacharparenleft}}thread\ {\isaliteral{3D}{\isacharequal}}\ thread{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}\ then\ length\ s\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ birthtime\ thread\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}birthtime\ thread\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}Set\ thread{\isaliteral{27}{\isacharprime}}\ prio{\isaliteral{29}{\isacharparenright}}{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ {\isaliteral{28}{\isacharparenleft}}thread\ {\isaliteral{3D}{\isacharequal}}\ thread{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}\ then\ length\ s\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ birthtime\ thread\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}birthtime\ thread\ {\isaliteral{28}{\isacharparenleft}}e{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ birthtime\ thread\ s{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-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 birth time}. The intention is
-  to discriminate threads with the same priority by giving threads with the earlier assigned priority
-  higher precedence in scheduling. This explains the following definition:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ preced\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ precedence{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}preced\ thread\ s\ {\isaliteral{3D}{\isacharequal}}\ Prc\ {\isaliteral{28}{\isacharparenleft}}original{\isaliteral{5F}{\isacharunderscore}}priority\ thread\ s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}birthtime\ thread\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-A number of important notions are defined here:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ \isanewline
-\ \ holding\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}b\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \isanewline
-\ \ \ \ \ \ \ waiting\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}b\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \ \ \ depend\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}b\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}node\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ node{\isaliteral{29}{\isacharparenright}}\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \ \ \ dependents\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}b\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ set{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The definition of the following several functions, it is supposed that
-  the waiting queue of every critical resource is given by a waiting queue 
-  function \isa{wq}, which servers as arguments of these functions.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{defs}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{overloaded}{\isaliteral{29}{\isacharparenright}}\ \isanewline
-\ \ %
-\isamarkupcmt{\begin{minipage}{0.8\textwidth}
-  We define that the thread which is at the head of waiting queue of resource \isa{cs}
-  is holding the resource. This definition is slightly different from tradition where
-  all threads in the waiting queue are considered as waiting for the resource.
-  This notion is reflected in the definition of \isa{holding\ wq\ th\ cs} as follows:
-  \end{minipage}%
-}
-\isanewline
-\ \ cs{\isaliteral{5F}{\isacharunderscore}}holding{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}holding\ wq\ thread\ cs\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ {\isaliteral{28}{\isacharparenleft}}wq\ cs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ thread\ {\isaliteral{3D}{\isacharequal}}\ hd\ {\isaliteral{28}{\isacharparenleft}}wq\ cs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ %
-\isamarkupcmt{\begin{minipage}{0.8\textwidth}
-  In accordance with the definition of \isa{holding\ wq\ th\ cs}, 
-  a thread \isa{th} is considered waiting for \isa{cs} if 
-  it is in the {\em waiting queue} of critical resource \isa{cs}, but not at the head.
-  This is reflected in the definition of \isa{waiting\ wq\ th\ cs} as follows:
-  \end{minipage}%
-}
-\isanewline
-\ \ cs{\isaliteral{5F}{\isacharunderscore}}waiting{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}waiting\ wq\ thread\ cs\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ {\isaliteral{28}{\isacharparenleft}}wq\ cs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ thread\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ hd\ {\isaliteral{28}{\isacharparenleft}}wq\ cs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ %
-\isamarkupcmt{\begin{minipage}{0.8\textwidth}
-  \isa{depend\ wq} represents the Resource Allocation Graph of the system under the waiting 
-  queue function \isa{wq}.
-  \end{minipage}%
-}
-\isanewline
-\ \ cs{\isaliteral{5F}{\isacharunderscore}}depend{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}depend\ {\isaliteral{28}{\isacharparenleft}}wq{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ list{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}Th\ t{\isaliteral{2C}{\isacharcomma}}\ Cs\ c{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ t\ c{\isaliteral{2E}{\isachardot}}\ waiting\ wq\ t\ c{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}Cs\ c{\isaliteral{2C}{\isacharcomma}}\ Th\ t{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ c\ t{\isaliteral{2E}{\isachardot}}\ holding\ wq\ t\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ %
-\isamarkupcmt{\begin{minipage}{0.8\textwidth}
-  \isa{dependents\ wq\ th} represents the set of threads which are depending on
-  thread \isa{th} in Resource Allocation Graph \isa{depend\ wq}:
-  \end{minipage}%
-}
-\isanewline
-\ \ cs{\isaliteral{5F}{\isacharunderscore}}dependents{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}dependents\ {\isaliteral{28}{\isacharparenleft}}wq{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ list{\isaliteral{29}{\isacharparenright}}\ th\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}th{\isaliteral{27}{\isacharprime}}\ {\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Th\ th{\isaliteral{27}{\isacharprime}}{\isaliteral{2C}{\isacharcomma}}\ Th\ th{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{28}{\isacharparenleft}}depend\ wq{\isaliteral{29}{\isacharparenright}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-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 and a function assigning precedence to 
-  threads:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{record}\isamarkupfalse%
-\ schedule{\isaliteral{5F}{\isacharunderscore}}state\ {\isaliteral{3D}{\isacharequal}}\ \isanewline
-\ \ \ \ waiting{\isaliteral{5F}{\isacharunderscore}}queue\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ list{\isaliteral{22}{\isachardoublequoteclose}}\ %
-\isamarkupcmt{The function assigning waiting queue.%
-}
-\isanewline
-\ \ \ \ cur{\isaliteral{5F}{\isacharunderscore}}preced\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ precedence{\isaliteral{22}{\isachardoublequoteclose}}\ %
-\isamarkupcmt{The function assigning precedence.%
-}
-%
-\begin{isamarkuptext}%
-\isa{cpreced\ s\ th} gives the {\em current precedence} of thread \isa{th} under
-  state \isa{s}. The definition of \isa{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 dependents, i.e. the threads which are waiting 
-  directly or indirectly waiting for some resources from it. If no such thread exits, 
-  \isa{th}'s {\em current precedence} equals its original precedence, i.e. 
-  \isa{preced\ th\ s}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ cpreced\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ list{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ precedence{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}cpreced\ s\ wq\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\ th{\isaliteral{2E}{\isachardot}}\ Max\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\ th{\isaliteral{2E}{\isachardot}}\ preced\ th\ s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{60}{\isacharbackquote}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{7B}{\isacharbraceleft}}th{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ dependents\ wq\ th{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The following function \isa{schs} is used to calculate the schedule state \isa{schs\ s}.
-  It is the key function to model Priority Inheritance:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{fun}\isamarkupfalse%
-\ schs\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ schedule{\isaliteral{5F}{\isacharunderscore}}state{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-\ \ \ {\isaliteral{22}{\isachardoublequoteopen}}schs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}waiting{\isaliteral{5F}{\isacharunderscore}}queue\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\ cs{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2C}{\isacharcomma}}\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ cur{\isaliteral{5F}{\isacharunderscore}}preced\ {\isaliteral{3D}{\isacharequal}}\ cpreced\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\ cs{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ %
-\isamarkupcmt{\begin{minipage}{0.8\textwidth}
-  \begin{enumerate}
-  \item \isa{ps} is the schedule state of last moment.
-  \item \isa{pwq} is the waiting queue function of last moment.
-  \item \isa{pcp} is the precedence function of last moment. 
-  \item \isa{nwq} is the new waiting queue function. It is calculated using a \isa{case} statement:
-  \begin{enumerate}
-      \item If the happening event is \isa{P\ thread\ cs}, \isa{thread} is added to 
-            the end of \isa{cs}'s waiting queue.
-      \item If the happening event is \isa{V\ thread\ cs} and \isa{s} is a legal state,
-            \isa{th{\isaliteral{27}{\isacharprime}}} must equal to \isa{thread}, 
-            because \isa{thread} is the one currently holding \isa{cs}. 
-            The case \isa{{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}} may never be executed in a legal state.
-            the \isa{{\isaliteral{28}{\isacharparenleft}}SOME\ q{\isaliteral{2E}{\isachardot}}\ distinct\ q\ {\isaliteral{5C3C616E643E}{\isasymand}}\ set\ q\ {\isaliteral{3D}{\isacharequal}}\ set\ qs{\isaliteral{29}{\isacharparenright}}} is used to choose arbitrarily one 
-            thread in waiting to take over the released resource \isa{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 \isa{ncp} is new precedence function, it is calculated from the newly updated waiting queue 
-        function. The dependency 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}
-  \end{minipage}%
-}
-\isanewline
-\ \ \ {\isaliteral{22}{\isachardoublequoteopen}}schs\ {\isaliteral{28}{\isacharparenleft}}e{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}let\ ps\ {\isaliteral{3D}{\isacharequal}}\ schs\ s\ in\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ let\ pwq\ {\isaliteral{3D}{\isacharequal}}\ waiting{\isaliteral{5F}{\isacharunderscore}}queue\ ps\ in\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ let\ pcp\ {\isaliteral{3D}{\isacharequal}}\ cur{\isaliteral{5F}{\isacharunderscore}}preced\ ps\ in\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ let\ nwq\ {\isaliteral{3D}{\isacharequal}}\ case\ e\ of\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ P\ thread\ cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ \ pwq{\isaliteral{28}{\isacharparenleft}}cs{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{28}{\isacharparenleft}}pwq\ cs\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{5B}{\isacharbrackleft}}thread{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ V\ thread\ cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ let\ nq\ {\isaliteral{3D}{\isacharequal}}\ case\ {\isaliteral{28}{\isacharparenleft}}pwq\ cs{\isaliteral{29}{\isacharparenright}}\ of\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{7C}{\isacharbar}}\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}th{\isaliteral{27}{\isacharprime}}{\isaliteral{23}{\isacharhash}}qs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}SOME\ q{\isaliteral{2E}{\isachardot}}\ distinct\ q\ {\isaliteral{5C3C616E643E}{\isasymand}}\ set\ q\ {\isaliteral{3D}{\isacharequal}}\ set\ qs{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ in\ pwq{\isaliteral{28}{\isacharparenleft}}cs{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}nq{\isaliteral{29}{\isacharparenright}}\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ pwq\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ in\ let\ ncp\ {\isaliteral{3D}{\isacharequal}}\ cpreced\ {\isaliteral{28}{\isacharparenleft}}e{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}\ nwq\ in\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}waiting{\isaliteral{5F}{\isacharunderscore}}queue\ {\isaliteral{3D}{\isacharequal}}\ nwq{\isaliteral{2C}{\isacharcomma}}\ cur{\isaliteral{5F}{\isacharunderscore}}preced\ {\isaliteral{3D}{\isacharequal}}\ ncp{\isaliteral{5C3C72706172723E}{\isasymrparr}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\isa{wq} is a shorthand for \isa{waiting{\isaliteral{5F}{\isacharunderscore}}queue}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ wq\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ list{\isaliteral{22}{\isachardoublequoteclose}}\ \isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}wq\ s\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ waiting{\isaliteral{5F}{\isacharunderscore}}queue\ {\isaliteral{28}{\isacharparenleft}}schs\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\isa{cp} is a shorthand for \isa{cur{\isaliteral{5F}{\isacharunderscore}}preced}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ cp\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ precedence{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}cp\ s\ {\isaliteral{3D}{\isacharequal}}\ cur{\isaliteral{5F}{\isacharunderscore}}preced\ {\isaliteral{28}{\isacharparenleft}}schs\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-Functions \isa{holding}, \isa{waiting}, \isa{depend} and \isa{dependents} still have the 
-  same meaning, but redefined so that they no longer depend on the fictitious {\em waiting queue function}
-  \isa{wq}, but on system state \isa{s}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{defs}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{overloaded}{\isaliteral{29}{\isacharparenright}}\ \isanewline
-\ \ s{\isaliteral{5F}{\isacharunderscore}}holding{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}holding\ {\isaliteral{28}{\isacharparenleft}}s{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}state{\isaliteral{29}{\isacharparenright}}\ thread\ cs\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ {\isaliteral{28}{\isacharparenleft}}wq\ s\ cs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ thread\ {\isaliteral{3D}{\isacharequal}}\ hd\ {\isaliteral{28}{\isacharparenleft}}wq\ s\ cs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ s{\isaliteral{5F}{\isacharunderscore}}waiting{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}waiting\ {\isaliteral{28}{\isacharparenleft}}s{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}state{\isaliteral{29}{\isacharparenright}}\ thread\ cs\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ {\isaliteral{28}{\isacharparenleft}}wq\ s\ cs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ thread\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ hd\ {\isaliteral{28}{\isacharparenleft}}wq\ s\ cs{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ s{\isaliteral{5F}{\isacharunderscore}}depend{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}depend\ {\isaliteral{28}{\isacharparenleft}}s{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}Th\ t{\isaliteral{2C}{\isacharcomma}}\ Cs\ c{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ t\ c{\isaliteral{2E}{\isachardot}}\ waiting\ {\isaliteral{28}{\isacharparenleft}}wq\ s{\isaliteral{29}{\isacharparenright}}\ t\ c{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}Cs\ c{\isaliteral{2C}{\isacharcomma}}\ Th\ t{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ c\ t{\isaliteral{2E}{\isachardot}}\ holding\ {\isaliteral{28}{\isacharparenleft}}wq\ s{\isaliteral{29}{\isacharparenright}}\ t\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ s{\isaliteral{5F}{\isacharunderscore}}dependents{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}dependents\ {\isaliteral{28}{\isacharparenleft}}s{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}state{\isaliteral{29}{\isacharparenright}}\ th\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}th{\isaliteral{27}{\isacharprime}}\ {\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Th\ th{\isaliteral{27}{\isacharprime}}{\isaliteral{2C}{\isacharcomma}}\ Th\ th{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{28}{\isacharparenleft}}depend\ {\isaliteral{28}{\isacharparenleft}}wq\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The following function \isa{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.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ readys\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}readys\ s\ {\isaliteral{3D}{\isacharequal}}\ \isanewline
-\ \ \ \ \ {\isaliteral{7B}{\isacharbraceleft}}thread\ {\isaliteral{2E}{\isachardot}}\ thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ threads\ s\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}\ cs{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ waiting\ s\ thread\ cs{\isaliteral{29}{\isacharparenright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The following function \isa{runing} calculates the set of running thread, which is the ready 
-  thread with the highest precedence.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ runing\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}runing\ s\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}th\ {\isaliteral{2E}{\isachardot}}\ th\ {\isaliteral{5C3C696E3E}{\isasymin}}\ readys\ s\ {\isaliteral{5C3C616E643E}{\isasymand}}\ cp\ s\ th\ {\isaliteral{3D}{\isacharequal}}\ Max\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}cp\ s{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{60}{\isacharbackquote}}\ {\isaliteral{28}{\isacharparenleft}}readys\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The following function \isa{holdents\ s\ th} returns the set of resources held by thread 
-  \isa{th} in state \isa{s}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ holdents\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ cs\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}holdents\ s\ th\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}cs\ {\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}Cs\ cs{\isaliteral{2C}{\isacharcomma}}\ Th\ th{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ depend\ s{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\isa{cntCS\ s\ th} returns the number of resources held by thread \isa{th} in
-  state \isa{s}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ cntCS\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}cntCS\ s\ th\ {\isaliteral{3D}{\isacharequal}}\ card\ {\isaliteral{28}{\isacharparenleft}}holdents\ s\ th{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The fact that event \isa{e} is eligible to happen next in state \isa{s} 
-  is expressed as \isa{step\ s\ e}. The predicate \isa{step} is inductively defined as 
-  follows:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{inductive}\isamarkupfalse%
-\ step\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ event\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-\ \ %
-\isamarkupcmt{A thread can be created if it is not a live thread:%
-}
-\isanewline
-\ \ thread{\isaliteral{5F}{\isacharunderscore}}create{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}thread\ {\isaliteral{5C3C6E6F74696E3E}{\isasymnotin}}\ threads\ s{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ step\ s\ {\isaliteral{28}{\isacharparenleft}}Create\ thread\ prio{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ %
-\isamarkupcmt{A thread can exit if it no longer hold any resource:%
-}
-\isanewline
-\ \ thread{\isaliteral{5F}{\isacharunderscore}}exit{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ runing\ s{\isaliteral{3B}{\isacharsemicolon}}\ holdents\ s\ thread\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ step\ s\ {\isaliteral{28}{\isacharparenleft}}Exit\ thread{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ %
-\isamarkupcmt{A thread can request for an critical resource \isa{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.%
-}
-\isanewline
-\ \ thread{\isaliteral{5F}{\isacharunderscore}}P{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ runing\ s{\isaliteral{3B}{\isacharsemicolon}}\ \ {\isaliteral{28}{\isacharparenleft}}Cs\ cs{\isaliteral{2C}{\isacharcomma}}\ Th\ thread{\isaliteral{29}{\isacharparenright}}\ \ {\isaliteral{5C3C6E6F74696E3E}{\isasymnotin}}\ {\isaliteral{28}{\isacharparenleft}}depend\ s{\isaliteral{29}{\isacharparenright}}{\isaliteral{5E}{\isacharcircum}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ step\ s\ {\isaliteral{28}{\isacharparenleft}}P\ thread\ cs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ %
-\isamarkupcmt{A thread can release a critical resource \isa{cs} if it is running and holding that resource.%
-}
-\isanewline
-\ \ thread{\isaliteral{5F}{\isacharunderscore}}V{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ runing\ s{\isaliteral{3B}{\isacharsemicolon}}\ \ holding\ s\ thread\ cs{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ step\ s\ {\isaliteral{28}{\isacharparenleft}}V\ thread\ cs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ %
-\isamarkupcmt{A thread can adjust its own priority as long as it is current running.%
-}
-\ \ \isanewline
-\ \ thread{\isaliteral{5F}{\isacharunderscore}}set{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}thread\ {\isaliteral{5C3C696E3E}{\isasymin}}\ runing\ s{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ step\ s\ {\isaliteral{28}{\isacharparenleft}}Set\ thread\ prio{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-With predicate \isa{step}, the fact that \isa{s} is a legal state in 
-  Priority Inheritance protocol can be expressed as: \isa{vt\ step\ s}, where
-  the predicate \isa{vt} can be defined as the following:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{inductive}\isamarkupfalse%
-\ vt\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ event\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \isakeyword{for}\ cs\ %
-\isamarkupcmt{\isa{cs} is an argument representing any step predicate.%
-}
-\isanewline
-\isakeyword{where}\isanewline
-\ \ %
-\isamarkupcmt{Empty list \isa{{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}} is a legal state in any protocol:%
-}
-\isanewline
-\ \ vt{\isaliteral{5F}{\isacharunderscore}}nil{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}vt\ cs\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\ \ %
-\isamarkupcmt{If \isa{s} a legal state, and event \isa{e} is eligible to happen
-        in state \isa{s}, then \isa{e{\isaliteral{23}{\isacharhash}}{\isaliteral{23}{\isacharhash}}s} is a legal state as well:%
-}
-\isanewline
-\ \ vt{\isaliteral{5F}{\isacharunderscore}}cons{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}vt\ cs\ s{\isaliteral{3B}{\isacharsemicolon}}\ cs\ s\ e{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ vt\ cs\ {\isaliteral{28}{\isacharparenleft}}e{\isaliteral{23}{\isacharhash}}s{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-It is easy to see that the definition of \isa{vt} is generic. It can be applied to 
-  any step predicate to get the set of legal states.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The following two functions \isa{the{\isaliteral{5F}{\isacharunderscore}}cs} and \isa{the{\isaliteral{5F}{\isacharunderscore}}th} are used to extract
-  critical resource and thread respectively out of RAG nodes.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{fun}\isamarkupfalse%
-\ the{\isaliteral{5F}{\isacharunderscore}}cs\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}node\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ cs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}the{\isaliteral{5F}{\isacharunderscore}}cs\ {\isaliteral{28}{\isacharparenleft}}Cs\ cs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ cs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-\isacommand{fun}\isamarkupfalse%
-\ the{\isaliteral{5F}{\isacharunderscore}}th\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}node\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}the{\isaliteral{5F}{\isacharunderscore}}th\ {\isaliteral{28}{\isacharparenleft}}Th\ th{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ th{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The following predicate \isa{next{\isaliteral{5F}{\isacharunderscore}}th} describe the next thread to 
-  take over when a critical resource is released. In \isa{next{\isaliteral{5F}{\isacharunderscore}}th\ s\ th\ cs\ t}, 
-  \isa{th} is the thread to release, \isa{t} is the one to take over.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ next{\isaliteral{5F}{\isacharunderscore}}th{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ cs\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}next{\isaliteral{5F}{\isacharunderscore}}th\ s\ th\ cs\ t\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}\ rest{\isaliteral{2E}{\isachardot}}\ wq\ s\ cs\ {\isaliteral{3D}{\isacharequal}}\ th{\isaliteral{23}{\isacharhash}}rest\ {\isaliteral{5C3C616E643E}{\isasymand}}\ rest\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ t\ {\isaliteral{3D}{\isacharequal}}\ hd\ {\isaliteral{28}{\isacharparenleft}}SOME\ q{\isaliteral{2E}{\isachardot}}\ distinct\ q\ {\isaliteral{5C3C616E643E}{\isasymand}}\ set\ q\ {\isaliteral{3D}{\isacharequal}}\ set\ rest{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The function \isa{count\ Q\ l} is used to count the occurrence of situation \isa{Q}
-  in list \isa{l}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ count\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}count\ Q\ l\ {\isaliteral{3D}{\isacharequal}}\ length\ {\isaliteral{28}{\isacharparenleft}}filter\ Q\ l{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\isa{cntP\ s} returns the number of operation \isa{P} happened 
-  before reaching state \isa{s}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ cntP\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}cntP\ s\ th\ {\isaliteral{3D}{\isacharequal}}\ count\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\ e{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}\ cs{\isaliteral{2E}{\isachardot}}\ e\ {\isaliteral{3D}{\isacharequal}}\ P\ th\ cs{\isaliteral{29}{\isacharparenright}}\ s{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\isa{cntV\ s} returns the number of operation \isa{V} happened 
-  before reaching state \isa{s}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{definition}\isamarkupfalse%
-\ cntV\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}state\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ thread\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}cntV\ s\ th\ {\isaliteral{3D}{\isacharequal}}\ count\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\ e{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}\ cs{\isaliteral{2E}{\isachardot}}\ e\ {\isaliteral{3D}{\isacharequal}}\ V\ th\ cs{\isaliteral{29}{\isacharparenright}}\ s{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-\isacommand{end}\isamarkupfalse%
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-\isanewline
-%
-\endisadelimtheory
-\isanewline
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/prio/Paper/ROOT.ML	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-use_thy "Paper";
\ No newline at end of file
--- a/prio/Paper/document/llncs.cls	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1189 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%%  {Upper-case    \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%%   Lower-case    \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%%   Digits        \0\1\2\3\4\5\6\7\8\9
-%%   Exclamation   \!     Double quote  \"     Hash (number) \#
-%%   Dollar        \$     Percent       \%     Ampersand     \&
-%%   Acute accent  \'     Left paren    \(     Right paren   \)
-%%   Asterisk      \*     Plus          \+     Comma         \,
-%%   Minus         \-     Point         \.     Solidus       \/
-%%   Colon         \:     Semicolon     \;     Less than     \<
-%%   Equals        \=     Greater than  \>     Question mark \?
-%%   Commercial at \@     Left bracket  \[     Backslash     \\
-%%   Right bracket \]     Circumflex    \^     Underscore    \_
-%%   Grave accent  \`     Left brace    \{     Vertical bar  \|
-%%   Right brace   \}     Tilde         \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2002/01/28 v2.13
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
-  \ifnum #1>\c@tocdepth \else
-    \vskip \z@ \@plus.2\p@
-    {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
-               \parfillskip -\rightskip \pretolerance=10000
-     \parindent #2\relax\@afterindenttrue
-     \interlinepenalty\@M
-     \leavevmode
-     \@tempdima #3\relax
-     \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
-     {#4}\nobreak
-     \leaders\hbox{$\m@th
-        \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
-        mu$}\hfill
-     \nobreak
-     \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
-     \par}%
-  \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Key words:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
-   \@setfontsize\small\@ixpt{11}%
-   \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
-   \abovedisplayshortskip \z@ \@plus2\p@
-   \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
-   \def\@listi{\leftmargin\leftmargini
-               \parsep 0\p@ \@plus1\p@ \@minus\p@
-               \topsep 8\p@ \@plus2\p@ \@minus4\p@
-               \itemsep0\p@}%
-   \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin   {63\p@}
-\setlength\evensidemargin  {63\p@}
-\setlength\marginparwidth  {90\p@}
-
-\setlength\headsep   {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep   {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter      {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
-            \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
-       \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
-      \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
-                 \thispagestyle{empty}%
-                 \if@twocolumn
-                     \onecolumn
-                     \@tempswatrue
-                   \else
-                     \@tempswafalse
-                 \fi
-                 \null\vfil
-                 \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
-    \ifnum \c@secnumdepth >-2\relax
-      \refstepcounter{part}%
-      \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
-    \else
-      \addcontentsline{toc}{part}{#1}%
-    \fi
-    \markboth{}{}%
-    {\centering
-     \interlinepenalty \@M
-     \normalfont
-     \ifnum \c@secnumdepth >-2\relax
-       \huge\bfseries \partname~\thepart
-       \par
-       \vskip 20\p@
-     \fi
-     \Huge \bfseries #2\par}%
-    \@endpart}
-\def\@spart#1{%
-    {\centering
-     \interlinepenalty \@M
-     \normalfont
-     \Huge \bfseries #1\par}%
-    \@endpart}
-\def\@endpart{\vfil\newpage
-              \if@twoside
-                \null
-                \thispagestyle{empty}%
-                \newpage
-              \fi
-              \if@tempswa
-                \twocolumn
-              \fi}
-
-\newcommand\chapter{\clearpage
-                    \thispagestyle{empty}%
-                    \global\@topnum\z@
-                    \@afterindentfalse
-                    \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
-                       \if@mainmatter
-                         \refstepcounter{chapter}%
-                         \typeout{\@chapapp\space\thechapter.}%
-                         \addcontentsline{toc}{chapter}%
-                                  {\protect\numberline{\thechapter}#1}%
-                       \else
-                         \addcontentsline{toc}{chapter}{#1}%
-                       \fi
-                    \else
-                      \addcontentsline{toc}{chapter}{#1}%
-                    \fi
-                    \chaptermark{#1}%
-                    \addtocontents{lof}{\protect\addvspace{10\p@}}%
-                    \addtocontents{lot}{\protect\addvspace{10\p@}}%
-                    \if@twocolumn
-                      \@topnewpage[\@makechapterhead{#2}]%
-                    \else
-                      \@makechapterhead{#2}%
-                      \@afterheading
-                    \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
-  {\centering
-    \ifnum \c@secnumdepth >\m@ne
-      \if@mainmatter
-        \large\bfseries \@chapapp{} \thechapter
-        \par\nobreak
-        \vskip 20\p@
-      \fi
-    \fi
-    \interlinepenalty\@M
-    \Large \bfseries #1\par\nobreak
-    \vskip 40\p@
-  }}
-\def\@schapter#1{\if@twocolumn
-                   \@topnewpage[\@makeschapterhead{#1}]%
-                 \else
-                   \@makeschapterhead{#1}%
-                   \@afterheading
-                 \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
-  {\centering
-    \normalfont
-    \interlinepenalty\@M
-    \Large \bfseries  #1\par\nobreak
-    \vskip 40\p@
-  }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
-                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {12\p@ \@plus 4\p@ \@minus 4\p@}%
-                       {\normalfont\large\bfseries\boldmath
-                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
-                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {8\p@ \@plus 4\p@ \@minus 4\p@}%
-                       {\normalfont\normalsize\bfseries\boldmath
-                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
-                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {-0.5em \@plus -0.22em \@minus -0.1em}%
-                       {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
-                       {-12\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {-0.5em \@plus -0.22em \@minus -0.1em}%
-                       {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
-                  \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle     \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle        \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle      \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini  {17\p@}
-\setlength\leftmargin    {\leftmargini}
-\setlength\leftmarginii  {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv  {\leftmargini}
-\setlength  \labelsep  {.5em}
-\setlength  \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
-            \parsep 0\p@ \@plus1\p@ \@minus\p@
-            \topsep 8\p@ \@plus2\p@ \@minus4\p@
-            \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
-              \labelwidth\leftmarginii
-              \advance\labelwidth-\labelsep
-              \topsep    0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
-              \labelwidth\leftmarginiii
-              \advance\labelwidth-\labelsep
-              \topsep    0\p@ \@plus\p@\@minus\p@
-              \parsep    \z@
-              \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
-                                                    {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
-                 \unskip{} \andname\
-              \else
-                 \unskip \lastandname\
-              \fi}%
- \def\and{\stepcounter{@auth}\relax
-          \ifnum\value{@auth}=\value{auco}%
-             \lastand
-          \else
-             \unskip,
-          \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
-   \addvspace{2em plus\p@}%  % space above part line
-   \begingroup
-     \parindent \z@
-     \rightskip \z@ plus 5em
-     \hrule\vskip5pt
-     \large               % same size as for a contribution heading
-     \bfseries\boldmath   % set line in boldface
-     \leavevmode          % TeX command to enter horizontal mode.
-     #1\par
-     \vskip5pt
-     \hrule
-     \vskip1pt
-     \nobreak             % Never break after part entry
-   \endgroup}
-
-\def\@dotsep{2}
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{chapter.\thechapter}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
-                     {\thechapter}#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
-  \or    \def\@gtempa{\addcontentsmark}%
-  \or    \def\@gtempa{\addcontentsmarkwop}%
-  \fi\@gtempa{toc}{chapter}}
-\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
-      \nobreak
-      \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
-      \@dotsep mu$}\hfill
-      \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=\z@ %15\p@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-%\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@            % no chapter numbers
-\tocsecnum=15\p@          % section 88. plus 2.222pt
-\tocsubsecnum=23\p@       % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@    % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@         % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@      % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
-    \section*{\listfigurename
-      \@mkboth{\listfigurename}{\listfigurename}}%
-    \@starttoc{lof}%
-    }
-
-\renewcommand\listoftables{%
-    \section*{\listtablename
-      \@mkboth{\listtablename}{\listtablename}}%
-    \@starttoc{lot}%
-    }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
-     {\section*{\refname}
-      \def\@biblabel##1{##1.}
-      \small
-      \list{\@biblabel{\@arabic\c@enumiv}}%
-           {\settowidth\labelwidth{\@biblabel{#1}}%
-            \leftmargin\labelwidth
-            \advance\leftmargin\labelsep
-            \if@openbib
-              \advance\leftmargin\bibindent
-              \itemindent -\bibindent
-              \listparindent \itemindent
-              \parsep \z@
-            \fi
-            \usecounter{enumiv}%
-            \let\p@enumiv\@empty
-            \renewcommand\theenumiv{\@arabic\c@enumiv}}%
-      \if@openbib
-        \renewcommand\newblock{\par}%
-      \else
-        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
-      \fi
-      \sloppy\clubpenalty4000\widowpenalty4000%
-      \sfcode`\.=\@m}
-     {\def\@noitemerr
-       {\@latex@warning{Empty `thebibliography' environment}}%
-      \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
-     {\let\protect\noexpand\immediate
-     \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
-  \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
-    {\@ifundefined
-       {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
-        ?}\@warning
-       {Citation `\@citeb' on page \thepage \space undefined}}%
-    {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
-     \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
-       \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
-     \else
-      \advance\@tempcntb\@ne
-      \ifnum\@tempcntb=\@tempcntc
-      \else\advance\@tempcntb\m@ne\@citeo
-      \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
-               \@citea\def\@citea{,\,\hskip\z@skip}%
-               \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
-               {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
-                \def\@citea{--}\fi
-      \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
-     {\section*{\refname}
-      \small
-      \list{}%
-           {\settowidth\labelwidth{}%
-            \leftmargin\parindent
-            \itemindent=-\parindent
-            \labelsep=\z@
-            \if@openbib
-              \advance\leftmargin\bibindent
-              \itemindent -\bibindent
-              \listparindent \itemindent
-              \parsep \z@
-            \fi
-            \usecounter{enumiv}%
-            \let\p@enumiv\@empty
-            \renewcommand\theenumiv{}}%
-      \if@openbib
-        \renewcommand\newblock{\par}%
-      \else
-        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
-      \fi
-      \sloppy\clubpenalty4000\widowpenalty4000%
-      \sfcode`\.=\@m}
-     {\def\@noitemerr
-       {\@latex@warning{Empty `thebibliography' environment}}%
-      \endlist}
-      \def\@cite#1{#1}%
-      \def\@lbibitem[#1]#2{\item[]\if@filesw
-        {\def\protect##1{\string ##1\space}\immediate
-      \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-   \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
-                \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
-                \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
-               {\@mkboth{\indexname}{\indexname}%
-                \thispagestyle{empty}\parindent\z@
-                \parskip\z@ \@plus .3\p@\relax
-                \let\item\par
-                \def\,{\relax\ifmmode\mskip\thinmuskip
-                             \else\hskip0.2em\ignorespaces\fi}%
-                \normalfont\small
-                \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
-                }
-                {\end{multicols}}
-
-\renewcommand\footnoterule{%
-  \kern-3\p@
-  \hrule\@width 2truecm
-  \kern2.6\p@}
-  \newdimen\fnindent
-  \fnindent1em
-\long\def\@makefntext#1{%
-    \parindent \fnindent%
-    \leftskip \fnindent%
-    \noindent
-    \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
-  \vskip\abovecaptionskip
-  \sbox\@tempboxa{{\bfseries #1.} #2}%
-  \ifdim \wd\@tempboxa >\hsize
-    {\bfseries #1.} #2\par
-  \else
-    \global \@minipagefalse
-    \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
-  \fi
-  \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
-        \reset@font
-        \small
-        \@setnobreak
-        \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
-               {\setlength\abovecaptionskip{0\p@}%
-                \setlength\belowcaptionskip{10\p@}%
-                \@float{table}}
-               {\end@float}
-\renewenvironment{table*}
-               {\setlength\abovecaptionskip{0\p@}%
-                \setlength\belowcaptionskip{10\p@}%
-                \@dblfloat{table}}
-               {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
-  ext@#1\endcsname}{#1}{\protect\numberline{\csname
-  the#1\endcsname}{\ignorespaces #2}}\begingroup
-    \@parboxrestore
-    \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
-  \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
-                   \gdef\@title{No Title Given}%
-                   \gdef\@subtitle{}%
-                   \gdef\@institute{No Institute Given}%
-                   \gdef\@thanks{}%
-                   \global\titlerunning={}\global\authorrunning={}%
-                   \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
-   \gdef\fnnstart{0}%
- \else
-   \xdef\fnnstart{\c@@inst}%
-   \setcounter{@inst}{1}%
-   \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
-   {\star\star\star}\or \dagger\or \ddagger\or
-   \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
-   \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-
-\renewcommand\maketitle{\newpage
-  \refstepcounter{chapter}%
-  \stepcounter{section}%
-  \setcounter{section}{0}%
-  \setcounter{subsection}{0}%
-  \setcounter{figure}{0}
-  \setcounter{table}{0}
-  \setcounter{equation}{0}
-  \setcounter{footnote}{0}%
-  \begingroup
-    \parindent=\z@
-    \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
-    \if@twocolumn
-      \ifnum \col@number=\@ne
-        \@maketitle
-      \else
-        \twocolumn[\@maketitle]%
-      \fi
-    \else
-      \newpage
-      \global\@topnum\z@   % Prevents figures from going at top of page.
-      \@maketitle
-    \fi
-    \thispagestyle{empty}\@thanks
-%
-    \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
-    \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
-    \instindent=\hsize
-    \advance\instindent by-\headlineindent
-%    \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
-%       \addcontentsline{toc}{title}{\the\toctitle}\fi
-    \if@runhead
-       \if!\the\titlerunning!\else
-         \edef\@title{\the\titlerunning}%
-       \fi
-       \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
-       \ifdim\wd\titrun>\instindent
-          \typeout{Title too long for running head. Please supply}%
-          \typeout{a shorter form with \string\titlerunning\space prior to
-                   \string\maketitle}%
-          \global\setbox\titrun=\hbox{\small\rm
-          Title Suppressed Due to Excessive Length}%
-       \fi
-       \xdef\@title{\copy\titrun}%
-    \fi
-%
-    \if!\the\tocauthor!\relax
-      {\def\and{\noexpand\protect\noexpand\and}%
-      \protected@xdef\toc@uthor{\@author}}%
-    \else
-      \def\\{\noexpand\protect\noexpand\newline}%
-      \protected@xdef\scratch{\the\tocauthor}%
-      \protected@xdef\toc@uthor{\scratch}%
-    \fi
-%    \addcontentsline{toc}{author}{\toc@uthor}%
-    \if@runhead
-       \if!\the\authorrunning!
-         \value{@inst}=\value{@auth}%
-         \setcounter{@auth}{1}%
-       \else
-         \edef\@author{\the\authorrunning}%
-       \fi
-       \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
-       \ifdim\wd\authrun>\instindent
-          \typeout{Names of authors too long for running head. Please supply}%
-          \typeout{a shorter form with \string\authorrunning\space prior to
-                   \string\maketitle}%
-          \global\setbox\authrun=\hbox{\small\rm
-          Authors Suppressed Due to Excessive Length}%
-       \fi
-       \xdef\@author{\copy\authrun}%
-       \markboth{\@author}{\@title}%
-     \fi
-  \endgroup
-  \setcounter{footnote}{\fnnstart}%
-  \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
-                 \unskip{} \andname\
-              \else
-                 \unskip \lastandname\
-              \fi}%
- \def\and{\stepcounter{@auth}\relax
-          \ifnum\value{@auth}=\value{@inst}%
-             \lastand
-          \else
-             \unskip,
-          \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
-  \pretolerance=10000
-  \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
-  \vskip -.65cm
-  \pretolerance=10000
-  \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-%     \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or  \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or  \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
-  \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
-   {\@definecounter{#1}\@addtoreset{#1}{#3}%
-   \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
-     \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
-   \expandafter\xdef\csname #1name\endcsname{#2}%
-   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
-                              \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
-   {\@definecounter{#1}%
-   \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
-   \expandafter\xdef\csname #1name\endcsname{#2}%
-   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
-                               \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
-  \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
-  {\expandafter\@ifdefinable\csname #1\endcsname
-  {\global\@namedef{the#1}{\@nameuse{the#2}}%
-  \expandafter\xdef\csname #1name\endcsname{#3}%
-  \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
-  \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
-                    \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
-       the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
-                 \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
-      \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
-   {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
-    \expandafter\xdef\csname #1name\endcsname{#2}%
-    \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
-       {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
-                           \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
-      \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
-   \def\@thmcountersep{.}
-   \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
-   \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
-   \if@envcntreset
-      \@addtoreset{theorem}{section}
-   \else
-      \@addtoreset{theorem}{chapter}
-   \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
-   \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
-   \if@envcntsect % mit section numeriert
-      \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
-   \else % nicht mit section numeriert
-      \if@envcntreset
-         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
-                                   \@addtoreset{#1}{section}}
-      \else
-         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
-                                   \@addtoreset{#1}{chapter}}%
-      \fi
-   \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
-    \def\@tempa{#1}%
-    \let\@tempd\@elt
-    \def\@elt##1{%
-        \def\@tempb{##1}%
-        \ifx\@tempa\@tempb\else
-            \@addtoreset{##1}{#2}%
-        \fi}%
-    \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
-    \expandafter\def\csname cl@#2\endcsname{}%
-    \@tempc
-    \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
-      \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
-                  \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
-      \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
-      }
-
-\renewenvironment{abstract}{%
-      \list{}{\advance\topsep by0.35cm\relax\small
-      \leftmargin=1cm
-      \labelwidth=\z@
-      \listparindent=\z@
-      \itemindent\listparindent
-      \rightmargin\leftmargin}\item[\hskip\labelsep
-                                    \bfseries\abstractname]}
-    {\endlist}
-
-\newdimen\headlineindent             % dimension for space between
-\headlineindent=1.166cm              % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
-   \let\@oddfoot\@empty\let\@evenfoot\@empty
-   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
-                  \leftmark\hfil}
-   \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
-                 \llap{\thepage}}
-   \def\chaptermark##1{}%
-   \def\sectionmark##1{}%
-   \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
-   \let\@oddfoot\@empty\let\@evenfoot\@empty
-   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
-                  \hfil}
-   \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
-                 \llap{\thepage}}
-   \def\chaptermark##1{}%
-   \def\sectionmark##1{}%
-   \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/prio/Paper/document/mathpartir.sty	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,446 +0,0 @@
-%  Mathpartir --- Math Paragraph for Typesetting Inference Rules
-%
-%  Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy
-%
-%  Author         : Didier Remy 
-%  Version        : 1.2.0
-%  Bug Reports    : to author
-%  Web Site       : http://pauillac.inria.fr/~remy/latex/
-% 
-%  Mathpartir is free software; you can redistribute it and/or modify
-%  it under the terms of the GNU General Public License as published by
-%  the Free Software Foundation; either version 2, or (at your option)
-%  any later version.
-%  
-%  Mathpartir is distributed in the hope that it will be useful,
-%  but WITHOUT ANY WARRANTY; without even the implied warranty of
-%  MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
-%  GNU General Public License for more details 
-%  (http://pauillac.inria.fr/~remy/license/GPL).
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%  File mathpartir.sty (LaTeX macros)
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\NeedsTeXFormat{LaTeX2e}
-\ProvidesPackage{mathpartir}
-    [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules]
-
-%%
-
-%% Identification
-%% Preliminary declarations
-
-\RequirePackage {keyval}
-
-%% Options
-%% More declarations
-
-%% PART I: Typesetting maths in paragraphe mode
-
-%% \newdimen \mpr@tmpdim
-%% Dimens are a precious ressource. Uses seems to be local.
-\let \mpr@tmpdim \@tempdima
-
-% To ensure hevea \hva compatibility, \hva should expands to nothing 
-% in mathpar or in inferrule
-\let \mpr@hva \empty
-
-%% normal paragraph parametters, should rather be taken dynamically
-\def \mpr@savepar {%
-  \edef \MathparNormalpar
-     {\noexpand \lineskiplimit \the\lineskiplimit
-      \noexpand \lineskip \the\lineskip}%
-  }
-
-\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em}
-\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em}
-\def \mpr@lineskip  {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em}
-\let \MathparLineskip \mpr@lineskip
-\def \mpr@paroptions {\MathparLineskip}
-\let \mpr@prebindings \relax
-
-\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em
-
-\def \mpr@goodbreakand
-   {\hskip -\mpr@andskip  \penalty -1000\hskip \mpr@andskip}
-\def \mpr@and {\hskip \mpr@andskip}
-\def \mpr@andcr {\penalty 50\mpr@and}
-\def \mpr@cr {\penalty -10000\mpr@and}
-\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10}
-
-\def \mpr@bindings {%
-  \let \and \mpr@andcr
-  \let \par \mpr@andcr
-  \let \\\mpr@cr
-  \let \eqno \mpr@eqno
-  \let \hva \mpr@hva
-  } 
-\let \MathparBindings \mpr@bindings
-
-% \@ifundefined {ignorespacesafterend}
-%    {\def \ignorespacesafterend {\aftergroup \ignorespaces}
-
-\newenvironment{mathpar}[1][]
-  {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering
-     \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else
-     \noindent $\displaystyle\fi
-     \MathparBindings}
-  {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend}
-
-\newenvironment{mathparpagebreakable}[1][]
-  {\begingroup 
-   \par
-   \mpr@savepar \parskip 0em \hsize \linewidth \centering
-      \mpr@prebindings \mpr@paroptions #1%
-      \vskip \abovedisplayskip \vskip -\lineskip%
-     \ifmmode  \else  $\displaystyle\fi
-     \MathparBindings
-  }
-  {\unskip
-   \ifmmode $\fi \par\endgroup
-   \vskip \belowdisplayskip
-   \noindent
-  \ignorespacesafterend}
-
-% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum
-%     \wd0 < \hsize  $$\box0$$\else \bmathpar #1\emathpar \fi}
-
-%%% HOV BOXES
-
-\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip 
-  \vbox \bgroup \tabskip 0em \let \\ \cr
-  \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup
-  \egroup}
-
-\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize
-      \box0\else \mathvbox {#1}\fi}
-
-
-%% Part II -- operations on lists
-
-\newtoks \mpr@lista
-\newtoks \mpr@listb
-
-\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
-{#2}\edef #2{\the \mpr@lista \the \mpr@listb}}
-
-\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
-{#2}\edef #2{\the \mpr@listb\the\mpr@lista}}
-
-\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb
-\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}}
-
-\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2}
-\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}}
-
-\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2}
-\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}}
-
-\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}%
-   \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the
-   \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty 
-   \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop
-     \mpr@flatten \mpr@all \mpr@to \mpr@one
-     \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof
-     \mpr@all \mpr@stripend  
-     \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi
-     \ifx 1\mpr@isempty
-   \repeat
-}
-
-\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty
-   \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp}
-
-%% Part III -- Type inference rules
-
-\newif \if@premisse
-\newbox \mpr@hlist
-\newbox \mpr@vlist
-\newif \ifmpr@center \mpr@centertrue
-\def \mpr@htovlist {%
-   \setbox \mpr@hlist
-      \hbox {\strut
-             \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi
-             \unhbox \mpr@hlist}%
-   \setbox \mpr@vlist
-      \vbox {\if@premisse  \box \mpr@hlist \unvbox \mpr@vlist
-             \else \unvbox \mpr@vlist \box \mpr@hlist
-             \fi}%
-}
-% OLD version
-% \def \mpr@htovlist {%
-%    \setbox \mpr@hlist
-%       \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}%
-%    \setbox \mpr@vlist
-%       \vbox {\if@premisse  \box \mpr@hlist \unvbox \mpr@vlist
-%              \else \unvbox \mpr@vlist \box \mpr@hlist
-%              \fi}%
-% }
-
-\def \mpr@item #1{$\displaystyle #1$}
-\def \mpr@sep{2em}
-\def \mpr@blank { }
-\def \mpr@hovbox #1#2{\hbox
-  \bgroup
-  \ifx #1T\@premissetrue
-  \else \ifx #1B\@premissefalse
-  \else
-     \PackageError{mathpartir}
-       {Premisse orientation should either be T or B}
-       {Fatal error in Package}%
-  \fi \fi
-  \def \@test {#2}\ifx \@test \mpr@blank\else
-  \setbox \mpr@hlist \hbox {}%
-  \setbox \mpr@vlist \vbox {}%
-  \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi
-  \let \@hvlist \empty \let \@rev \empty
-  \mpr@tmpdim 0em
-  \expandafter \mpr@makelist #2\mpr@to \mpr@flat
-  \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi
-  \def \\##1{%
-     \def \@test {##1}\ifx \@test \empty
-        \mpr@htovlist
-        \mpr@tmpdim 0em %%% last bug fix not extensively checked
-     \else
-      \setbox0 \hbox{\mpr@item {##1}}\relax
-      \advance \mpr@tmpdim by \wd0
-      %\mpr@tmpdim 1.02\mpr@tmpdim
-      \ifnum \mpr@tmpdim < \hsize
-         \ifnum \wd\mpr@hlist > 0
-           \if@premisse
-             \setbox \mpr@hlist 
-                \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}%
-           \else
-             \setbox \mpr@hlist
-                \hbox {\unhbox \mpr@hlist  \hskip \mpr@sep \unhbox0}%
-           \fi
-         \else 
-         \setbox \mpr@hlist \hbox {\unhbox0}%
-         \fi
-      \else
-         \ifnum \wd \mpr@hlist > 0
-            \mpr@htovlist 
-            \mpr@tmpdim \wd0
-         \fi
-         \setbox \mpr@hlist \hbox {\unhbox0}%
-      \fi
-      \advance \mpr@tmpdim by \mpr@sep
-   \fi
-   }%
-   \@rev
-   \mpr@htovlist
-   \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist
-   \fi
-   \egroup
-}
-
-%%% INFERENCE RULES
-
-\@ifundefined{@@over}{%
-    \let\@@over\over % fallback if amsmath is not loaded
-    \let\@@overwithdelims\overwithdelims
-    \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims
-    \let\@@above\above \let\@@abovewithdelims\abovewithdelims
-  }{}
-
-%% The default
-
-\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em
-    $\displaystyle {#1\mpr@over #2}$}}
-\def \mpr@@nofraction #1#2{\hbox {\advance \hsize by -0.5em
-    $\displaystyle {#1\@@atop #2}$}}
-
-\let \mpr@fraction \mpr@@fraction
-
-%% A generic solution to arrow
-
-\def \mpr@make@fraction #1#2#3#4#5{\hbox {%
-     \def \mpr@tail{#1}%
-     \def \mpr@body{#2}%
-     \def \mpr@head{#3}%
-     \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}%
-     \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}%
-     \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}%
-     \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax
-     \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax
-     \setbox0=\hbox {$\box1 \@@atop \box2$}%
-     \dimen0=\wd0\box0
-     \box0 \hskip -\dimen0\relax
-     \hbox to \dimen0 {$%
-       \mathrel{\mpr@tail}\joinrel
-       \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}%
-     $}}}
-
-%% Old stuff should be removed in next version
-\def \mpr@@nothing #1#2
-    {$\lower 0.01pt \mpr@@nofraction {#1}{#2}$}
-\def \mpr@@reduce #1#2{\hbox
-    {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}}
-\def \mpr@@rewrite #1#2#3{\hbox
-    {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}}
-\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}}
-
-\def \mpr@empty {}
-\def \mpr@inferrule
-  {\bgroup
-     \ifnum \linewidth<\hsize \hsize \linewidth\fi
-     \mpr@rulelineskip
-     \let \and \qquad
-     \let \hva \mpr@hva
-     \let \@rulename \mpr@empty
-     \let \@rule@options \mpr@empty
-     \let \mpr@over \@@over
-     \mpr@inferrule@}
-\newcommand {\mpr@inferrule@}[3][]
-  {\everymath={\displaystyle}%       
-   \def \@test {#2}\ifx \empty \@test
-      \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}%
-   \else 
-   \def \@test {#3}\ifx \empty \@test
-      \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}%
-   \else
-   \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}%
-   \fi \fi
-   \def \@test {#1}\ifx \@test\empty \box0
-   \else \vbox 
-%%% Suggestion de Francois pour les etiquettes longues
-%%%   {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi
-      {\hbox {\RefTirName {#1}}\box0}\fi
-   \egroup}
-
-\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}}
-
-% They are two forms
-% \inferrule [label]{[premisses}{conclusions}
-% or
-% \inferrule* [options]{[premisses}{conclusions}
-%
-% Premisses and conclusions are lists of elements separated by \\
-% Each \\ produces a break, attempting horizontal breaks if possible, 
-% and  vertical breaks if needed. 
-% 
-% An empty element obtained by \\\\ produces a vertical break in all cases. 
-%
-% The former rule is aligned on the fraction bar. 
-% The optional label appears on top of the rule
-% The second form to be used in a derivation tree is aligned on the last
-% line of its conclusion
-% 
-% The second form can be parameterized, using the key=val interface. The
-% folloiwng keys are recognized:
-%       
-%  width                set the width of the rule to val
-%  narrower             set the width of the rule to val\hsize
-%  before               execute val at the beginning/left
-%  lab                  put a label [Val] on top of the rule
-%  lskip                add negative skip on the right
-%  left                 put a left label [Val]
-%  Left                 put a left label [Val],  ignoring its width 
-%  right                put a right label [Val]
-%  Right                put a right label [Val], ignoring its width
-%  leftskip             skip negative space on the left-hand side
-%  rightskip            skip negative space on the right-hand side
-%  vdots                lift the rule by val and fill vertical space with dots
-%  after                execute val at the end/right
-%  
-%  Note that most options must come in this order to avoid strange
-%  typesetting (in particular  leftskip must preceed left and Left and
-%  rightskip must follow Right or right; vdots must come last 
-%  or be only followed by rightskip. 
-%  
-
-%% Keys that make sence in all kinds of rules
-\def \mprset #1{\setkeys{mprset}{#1}}
-\define@key {mprset}{andskip}[]{\mpr@andskip=#1}
-\define@key {mprset}{lineskip}[]{\lineskip=#1}
-\define@key {mprset}{flushleft}[]{\mpr@centerfalse}
-\define@key {mprset}{center}[]{\mpr@centertrue}
-\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
-\define@key {mprset}{atop}[]{\let \mpr@fraction \mpr@@nofraction}
-\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1}
-\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
-\define@key {mprset}{sep}{\def\mpr@sep{#1}}
-
-\newbox \mpr@right
-\define@key {mpr}{flushleft}[]{\mpr@centerfalse}
-\define@key {mpr}{center}[]{\mpr@centertrue}
-\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
-\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1}
-\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
-\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
-     \advance \hsize by -\wd0\box0}
-\define@key {mpr}{width}{\hsize #1}
-\define@key {mpr}{sep}{\def\mpr@sep{#1}}
-\define@key {mpr}{before}{#1}
-\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
-\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
-\define@key {mpr}{narrower}{\hsize #1\hsize}
-\define@key {mpr}{leftskip}{\hskip -#1}
-\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce}
-\define@key {mpr}{rightskip}
-  {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}}
-\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax
-     \advance \hsize by -\wd0\box0}
-\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
-     \advance \hsize by -\wd0\box0}
-\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}}
-\define@key {mpr}{right}
-  {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0
-   \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
-\define@key {mpr}{RIGHT}
-  {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0
-   \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
-\define@key {mpr}{Right}
-  {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}}
-\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}}
-\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}}
-
-\newcommand \mpr@inferstar@ [3][]{\setbox0
-  \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax
-         \setbox \mpr@right \hbox{}%
-         $\setkeys{mpr}{#1}%
-          \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else
-          \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi
-          \box \mpr@right \mpr@vdots$}
-  \setbox1 \hbox {\strut}
-  \@tempdima \dp0 \advance \@tempdima by -\dp1
-  \raise \@tempdima \box0}
-
-\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}}
-\newcommand \mpr@err@skipargs[3][]{}
-\def \mpr@inferstar*{\ifmmode 
-    \let \@do \mpr@inferstar@
-  \else 
-    \let \@do \mpr@err@skipargs
-    \PackageError {mathpartir}
-      {\string\inferrule* can only be used in math mode}{}%
-  \fi \@do}
-
-
-%%% Exports
-
-% Envirnonment mathpar
-
-\let \inferrule \mpr@infer
-
-% make a short name \infer is not already defined
-\@ifundefined {infer}{\let \infer \mpr@infer}{}
-
-\def \TirNameStyle #1{\small \textsc{#1}}
-\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}}
-\let \TirName \tir@name
-\let \DefTirName \TirName
-\let \RefTirName \TirName
-
-%%% Other Exports
-
-% \let \listcons \mpr@cons
-% \let \listsnoc \mpr@snoc
-% \let \listhead \mpr@head
-% \let \listmake \mpr@makelist
-
-
-
-
-\endinput
--- a/prio/Paper/document/root.bib	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,205 +0,0 @@
-
-@Book{Paulson96,
-  author = 	 {L.~C.~Paulson},
-  title = 	 {{ML} for the {W}orking {P}rogrammer},
-  publisher = 	 {Cambridge University Press},
-  year = 	 {1996}
-}
-
-
-@Manual{LINUX,
-  author =       {S.~Rostedt},
-  title =        {{RT}-{M}utex {I}mplementation {D}esign},
-  note =         {Linux Kernel Distribution at, 
-                  www.kernel.org/doc/Documentation/rt-mutex-design.txt}
-}
-
-@Misc{PINTOS,
-  author = {B.~Pfaff},
-  title = {{PINTOS}}, 
-  note = {\url{http://www.stanford.edu/class/cs140/projects/}},
-}
-
-
-@inproceedings{Haftmann08,
-  author    = {F.~Haftmann and M.~Wenzel},
-  title     = {{L}ocal {T}heory {S}pecifications in {I}sabelle/{I}sar},
-  booktitle = {Proc.~of the International Conference on Types, Proofs and Programs (TYPES)},
-  year      = {2008},
-  pages     = {153-168},
-  series    = {LNCS},
-  volume    = {5497}
-}
-
-
-@TechReport{Yodaiken02,
-  author =       {V.~Yodaiken},
-  title =        {{A}gainst {P}riority {I}nheritance},
-  institution =  {Finite State Machine Labs (FSMLabs)},
-  year =         {2004}
-}
-
-
-@Book{Vahalia96,
-  author =       {U.~Vahalia},
-  title =        {{UNIX} {I}nternals: {T}he {N}ew {F}rontiers},
-  publisher =    {Prentice-Hall},
-  year =         {1996}
-}
-
-@Article{Reeves98,
-  title =	"{R}e: {W}hat {R}eally {H}appened on {M}ars?",
-  author =	"G.~E.~Reeves",
-  journal =	"Risks Forum",
-  year = 	"1998",
-  number =	"54",
-  volume =	"19"
-}
-
-@Article{Sha90,
-  title =	"{P}riority {I}nheritance {P}rotocols: {A}n {A}pproach to
-		 {R}eal-{T}ime {S}ynchronization",
-  author =	"L.~Sha and R.~Rajkumar and J.~P.~Lehoczky",
-  journal =	"IEEE Transactions on Computers",
-  year = 	"1990",
-  number =	"9",
-  volume =	"39",
-  pages =	"1175--1185"
-}
-
-
-@Article{Lampson80,
-  author =	"B.~W.~Lampson and D.~D.~Redell",
-  title =	"{{E}xperiences with {P}rocesses and {M}onitors in {M}esa}",
-  journal =	"Communications of the ACM",
-  volume =	"23",
-  number =	"2",
-  pages =	"105--117",
-  year = 	"1980"
-}
-
-@inproceedings{Wang09,
-  author    = {J.~Wang and H.~Yang and X.~Zhang},
-  title     = {{L}iveness {R}easoning with {I}sabelle/{HOL}},
-  booktitle = {Proc.~of the 22nd International Conference on Theorem Proving in 
-               Higher Order Logics (TPHOLs)},
-  year      = {2009},
-  pages     = {485--499},
-  volume    = {5674},
-  series    = {LNCS}
-}
-
-@PhdThesis{Faria08,
-  author =       {J.~M.~S.~Faria},
-  title =        {{F}ormal {D}evelopment of {S}olutions for {R}eal-{T}ime {O}perating {S}ystems 
-                 with {TLA+/TLC}},
-  school =       {University of Porto},
-  year =         {2008}
-}
-
-
-@Article{Paulson98,
-  author =       {L.~C.~Paulson},
-  title =        {{T}he {I}nductive {A}pproach to {V}erifying {C}ryptographic {P}rotocols},
-  journal =      {Journal of Computer Security},
-  year =         {1998},
-  volume =       {6},
-  number =       {1--2},
-  pages =        {85--128}
-}
-
-@MISC{locke-july02,
-author = {D. Locke},
-title = {Priority Inheritance: The Real Story},
-month = July,
-year = {2002},
-howpublished={\url{http://www.math.unipd.it/~tullio/SCD/2007/Materiale/Locke.pdf}},
-}
-
-
-@InProceedings{Steinberg10,
-  author =       {U.~Steinberg and A.~B\"otcher and B.~Kauer},
-  title =        {{T}imeslice {D}onation in {C}omponent-{B}ased {S}ystems},
-  booktitle = {Proc.~of the 6th International Workshop on Operating Systems
-               Platforms for Embedded Real-Time Applications (OSPERT)},
-  pages =     {16--23},
-  year =      {2010}
-}
-
-@inproceedings{dutertre99b,
-  title =	"{T}he {P}riority {C}eiling {P}rotocol: {F}ormalization and
-		 {A}nalysis {U}sing {PVS}",
-  author =	"B.~Dutertre",
-  booktitle = {Proc.~of the 21st IEEE Conference on Real-Time Systems Symposium (RTSS)},
-  year = {2000},
-  pages = {151--160},
-  publisher = {IEEE Computer Society}
-}
-
-@InProceedings{Jahier09,
-  title =	"{S}ynchronous {M}odeling and {V}alidation of {P}riority
-		 {I}nheritance {S}chedulers",
-  author =	"E.~Jahier and B.~Halbwachs and P.~Raymond",
-  booktitle =	"Proc.~of the 12th International Conference on Fundamental 
-                 Approaches to Software Engineering (FASE)",
-  year = 	"2009",
-  volume =	"5503",
-  series =      "LNCS",
-  pages =	"140--154"
-}
-
-@InProceedings{Wellings07,
-  title =	"{I}ntegrating {P}riority {I}nheritance {A}lgorithms in the {R}eal-{T}ime 
-                 {S}pecification for {J}ava",
-  author =	"A.~Wellings and A.~Burns and O.~M.~Santos and B.~M.~Brosgol",
-  publisher =	"IEEE Computer Society",
-  year = 	"2007",
-  booktitle =	"Proc.~of the 10th IEEE International Symposium on Object 
-                and Component-Oriented Real-Time Distributed Computing (ISORC)",
-  pages =	"115--123"
-}
-
-@Article{Wang:2002:SGP,
-  author =	"Y. Wang and E. Anceaume and F. Brasileiro and F.
-		 Greve and M. Hurfin",
-  title =	"Solving the group priority inversion problem in a
-		 timed asynchronous system",
-  journal =	"IEEE Transactions on Computers",
-  volume =	"51",
-  number =	"8",
-  pages =	"900--915",
-  month =	aug,
-  year = 	"2002",
-  CODEN =	"ITCOB4",
-  doi =  	"http://dx.doi.org/10.1109/TC.2002.1024738",
-  ISSN = 	"0018-9340 (print), 1557-9956 (electronic)",
-  issn-l =	"0018-9340",
-  bibdate =	"Tue Jul 5 09:41:56 MDT 2011",
-  bibsource =	"http://www.computer.org/tc/;
-		 http://www.math.utah.edu/pub/tex/bib/ieeetranscomput2000.bib",
-  URL =  	"http://ieeexplore.ieee.org/stamp/stamp.jsp?tp=&arnumber=1024738",
-  acknowledgement = "Nelson H. F. Beebe, University of Utah, Department
-		 of Mathematics, 110 LCB, 155 S 1400 E RM 233, Salt Lake
-		 City, UT 84112-0090, USA, Tel: +1 801 581 5254, FAX: +1
-		 801 581 4148, e-mail: \path|beebe@math.utah.edu|,
-		 \path|beebe@acm.org|, \path|beebe@computer.org|
-		 (Internet), URL:
-		 \path|http://www.math.utah.edu/~beebe/|",
-  fjournal =	"IEEE Transactions on Computers",
-  doi-url =	"http://dx.doi.org/10.1109/TC.2002.1024738",
-}
-
-@Article{journals/rts/BabaogluMS93,
-  title =	"A Formalization of Priority Inversion",
-  author =	"{\"O} Babaoglu and K. Marzullo and F. B. Schneider",
-  journal =	"Real-Time Systems",
-  year = 	"1993",
-  number =	"4",
-  volume =	"5",
-  bibdate =	"2011-06-03",
-  bibsource =	"DBLP,
-		 http://dblp.uni-trier.de/db/journals/rts/rts5.html#BabaogluMS93",
-  pages =	"285--303",
-  URL =  	"http://dx.doi.org/10.1007/BF01088832",
-}
-
--- a/prio/Paper/document/root.tex	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,79 +0,0 @@
-\documentclass[runningheads]{llncs}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-\usepackage{mathpartir}
-\usepackage{tikz}
-\usepackage{pgf}
-%\usetikzlibrary{arrows,automata,decorations,fit,calc}
-%\usetikzlibrary{shapes,shapes.arrows,snakes,positioning}
-%\usepgflibrary{shapes.misc} % LATEX and plain TEX and pure pgf
-%\usetikzlibrary{matrix}
-\usepackage{pdfsetup}
-\usepackage{ot1patch}
-\usepackage{times}
-%%\usepackage{proof}
-%%\usepackage{mathabx}
-\usepackage{stmaryrd}
-\usepackage{url}
-\usepackage{color}
-\titlerunning{Proving the Priority Inheritance Protocol Correct}
-
-
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyleminor}{\it}%
-\renewcommand{\isastyle}{\normalsize\it}%
-
-
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-\renewcommand{\isasymiota}{}
-
-\newcommand{\numbered}[1]{\refstepcounter{equation}{\rm(\arabic{equation})}\label{#1}}
-\definecolor{mygrey}{rgb}{.80,.80,.80}
-
-\begin{document}
-
-\title{Priority Inheritance Protocol Proved Correct}
-\author{Xingyuan Zhang\inst{1} \and Christian Urban\inst{2} \and Chunhan Wu\inst{1}}
-\institute{PLA University of Science and Technology, China \and 
-           King's College London, United Kingdom}
-\maketitle
-
-\begin{abstract}
-In real-time systems with threads, resource locking and 
-priority sched\-uling, one faces the problem of Priority
-Inversion. This problem can make the behaviour of threads
-unpredictable and the resulting bugs can be hard to find.  The
-Priority Inheritance Protocol is one solution implemented in many
-systems for solving this problem, but the correctness of this solution
-has never been formally verified in a theorem prover. As already
-pointed out in the literature, the original informal investigation of
-the Property Inheritance Protocol presents a correctness ``proof'' for
-an \emph{incorrect} algorithm. In this paper we fix the problem of
-this proof by making all notions precise and implementing a variant of
-a solution proposed earlier. Our formalisation in Isabelle/HOL
-uncovers facts not mentioned in the literature, but also shows how to
-efficiently implement this protocol. Earlier correct implementations
-were criticised as too inefficient. Our formalisation is based on
-Paulson's inductive approach to verifying protocols.\medskip
-
-{\bf Keywords:} Priority Inheritance Protocol, formal correctness proof, 
-real-time systems, Isabelle/HOL
-\end{abstract}
-
-\input{session}
-
-%\bibliographystyle{plain}
-%\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/prio/Paper/tt.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,94 +0,0 @@
-
-There are several works on inversion avoidance:
-\begin{enumerate}
-\item {\em Solving the group priority inversion problem in a timed asynchronous system}. 
-The notion of \<exclamdown>\<degree>Group Priority Inversion\<exclamdown>\<plusminus> is introduced. The main strategy is still inversion avoidance. 
-The method is by reordering requests in the setting of Client-Server.
-\item {\em Examples of inaccurate specification of the protocol}.
-\end{enumerate}
-
-
-
-
-
-
-section{* Related works *}
-
-text {*
-1.	<<Integrating Priority Inheritance Algorithms in the Real-Time Specification for Java>> models and 
-verifies the combination of Priority Inheritance (PI) and Priority Ceiling Emulation (PCE) protocols in 
-the setting of Java virtual machine using extended Timed Automata(TA) formalism of the UPPAAL tool. 
-Although a detailed formal model of combined PI and PCE is given, the number of properties is quite 
-small and the focus is put on the harmonious working of PI and PCE. Most key features of PI 
-(as well as PCE) are not shown. Because of the limitation of the model checking technique
- used there, properties are shown only for a small number of scenarios. Therefore, the verification 
-does not show the correctness of the formal model itself in a convincing way.  
-2.	<< Formal Development of Solutions for Real-Time Operating Systems with TLA+/TLC>>. A formal model 
-of PI is given in TLA+. Only 3 properties are shown for PI using model checking. The limitation of 
-model checking is intrinsic to the work.
-3.	<<Synchronous modeling and validation of priority inheritance schedulers>>. Gives a formal model 
-of PI and PCE in AADL (Architecture Analysis & Design Language) and checked several properties 
-using model checking. The number of properties shown there is less than here and the scale 
-is also limited by the model checking technique. 
-
-
-There are several works on inversion avoidance:
-1.	<<Solving the group priority inversion problem in a timed asynchronous system>>. 
-The notion of \<exclamdown>\<degree>Group Priority Inversion\<exclamdown>\<plusminus> is introduced. The main strategy is still inversion avoidance. 
-The method is by reordering requests in the setting of Client-Server.
-
-
-<<Examples of inaccurate specification of the protocol>>.
-
-*}
-
-text {*
-
-\section{An overview of priority inversion and priority inheritance}
-
-Priority inversion refers to the phenomenon when a thread with high priority is blocked 
-by a thread with low priority. Priority happens when the high priority thread requests 
-for some critical resource already taken by the low priority thread. Since the high 
-priority thread has to wait for the low priority thread to complete, it is said to be 
-blocked by the low priority thread. Priority inversion might prevent high priority 
-thread from fulfill its task in time if the duration of priority inversion is indefinite 
-and unpredictable. Indefinite priority inversion happens when indefinite number 
-of threads with medium priorities is activated during the period when the high 
-priority thread is blocked by the low priority thread. Although these medium 
-priority threads can not preempt the high priority thread directly, they are able 
-to preempt the low priority threads and cause it to stay in critical section for 
-an indefinite long duration. In this way, the high priority thread may be blocked indefinitely. 
-
-Priority inheritance is one protocol proposed to avoid indefinite priority inversion. 
-The basic idea is to let the high priority thread donate its priority to the low priority 
-thread holding the critical resource, so that it will not be preempted by medium priority 
-threads. The thread with highest priority will not be blocked unless it is requesting 
-some critical resource already taken by other threads. Viewed from a different angle, 
-any thread which is able to block the highest priority threads must already hold some 
-critical resource. Further more, it must have hold some critical resource at the 
-moment the highest priority is created, otherwise, it may never get change to run and 
-get hold. Since the number of such resource holding lower priority threads is finite, 
-if every one of them finishes with its own critical section in a definite duration, 
-the duration the highest priority thread is blocked is definite as well. The key to 
-guarantee lower priority threads to finish in definite is to donate them the highest 
-priority. In such cases, the lower priority threads is said to have inherited the 
-highest priority. And this explains the name of the protocol: 
-{\em Priority Inheritance} and how Priority Inheritance prevents indefinite delay.
-
-The objectives of this paper are:
-\begin{enumerate}
-\item Build the above mentioned idea into formal model and prove a series of properties 
-until we are convinced that the formal model does fulfill the original idea. 
-\item Show how formally derived properties can be used as guidelines for correct 
-and efficient implementation.
-\end{enumerate}.
-The proof is totally formal in the sense that every detail is reduced to the 
-very first principles of Higher Order Logic. The nature of interactive theorem 
-proving is for the human user to persuade computer program to accept its arguments. 
-A clear and simple understanding of the problem at hand is both a prerequisite and a 
-byproduct of such an effort, because everything has finally be reduced to the very 
-first principle to be checked mechanically. The former intuitive explanation of 
-Priority Inheritance is just such a byproduct. 
-*}
-
-*)
--- a/prio/Precedence_ord.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,34 +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)
-
-end
--- a/prio/PrioG.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2864 +0,0 @@
-theory PrioG
-imports PrioGDef 
-begin
-
-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 wq_distinct: "vt s \<Longrightarrow> distinct (wq s cs)"
-proof(erule_tac vt.induct, simp add:wq_def)
-  fix s e
-  assume h1: "step s e"
-  and h2: "distinct (wq s cs)"
-  thus "distinct (wq (e # s) cs)"
-  proof(induct rule:step.induct, auto simp: wq_def Let_def split:list.splits)
-    fix thread s
-    assume h1: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
-      and h2: "thread \<in> set (wq_fun (schs s) cs)"
-      and h3: "thread \<in> runing s"
-    show "False" 
-    proof -
-      from h3 have "\<And> cs. thread \<in>  set (wq_fun (schs s) cs) \<Longrightarrow> 
-                             thread = hd ((wq_fun (schs s) cs))" 
-        by (simp add:runing_def readys_def s_waiting_def wq_def)
-      from this [OF h2] have "thread = hd (wq_fun (schs s) cs)" .
-      with h2
-      have "(Cs cs, Th thread) \<in> (depend s)"
-        by (simp add:s_depend_def s_holding_def wq_def cs_holding_def)
-      with h1 show False by auto
-    qed
-  next
-    fix thread s a list
-    assume dst: "distinct list"
-    show "distinct (SOME q. distinct q \<and> set q = set list)"
-    proof(rule someI2)
-      from dst show  "distinct list \<and> set list = set list" by auto
-    next
-      fix q assume "distinct q \<and> set q = set list"
-      thus "distinct q" by auto
-    qed
-  qed
-qed
-
-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)
-
-lemma block_pre: 
-  fixes thread cs s
-  assumes vt_e: "vt (e#s)"
-  and s_ni: "thread \<notin>  set (wq s cs)"
-  and s_i: "thread \<in> set (wq (e#s) cs)"
-  shows "e = P thread cs"
-proof -
-  show ?thesis
-  proof(cases e)
-    case (P th cs)
-    with assms
-    show ?thesis
-      by (auto simp:wq_def Let_def split:if_splits)
-  next
-    case (Create th prio)
-    with assms show ?thesis
-      by (auto simp:wq_def Let_def split:if_splits)
-  next
-    case (Exit th)
-    with assms show ?thesis
-      by (auto simp:wq_def Let_def split:if_splits)
-  next
-    case (Set th prio)
-    with assms show ?thesis
-      by (auto simp:wq_def Let_def split:if_splits)
-  next
-    case (V th cs)
-    with assms show ?thesis
-      apply (auto simp:wq_def Let_def split:if_splits)
-    proof -
-      fix q qs
-      assume h1: "thread \<notin> set (wq_fun (schs s) cs)"
-        and h2: "q # qs = wq_fun (schs s) cs"
-        and h3: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
-        and vt: "vt (V th cs # s)"
-      from h1 and h2[symmetric] have "thread \<notin> set (q # 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 step_back_vt[OF vt], of cs]
-          and h2[symmetric, 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 h3 show ?thesis by simp
-      qed
-      ultimately show "False" by auto
-      qed
-  qed
-qed
-
-lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
-  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (depend s)^+"
-apply (ind_cases "vt ((P thread cs)#s)")
-apply (ind_cases "step s (P thread cs)")
-by auto
-
-lemma abs1:
-  fixes e es
-  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)"
-
-lemma abs2:
-  assumes vt: "vt (e#s)"
-  and 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 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 step_back_vt[OF vt], 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 step_back_vt[OF vt], 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
-
-lemma vt_moment: "\<And> t. \<lbrakk>vt s\<rbrakk> \<Longrightarrow> vt (moment t s)"
-proof(induct s, simp)
-  fix a s t
-  assume h: "\<And>t.\<lbrakk>vt s\<rbrakk> \<Longrightarrow> vt (moment t s)"
-    and vt_a: "vt (a # s)"
-  show "vt (moment t (a # s))"
-  proof(cases "t \<ge> length (a#s)")
-    case True
-    from True have "moment t (a#s) = a#s" by simp
-    with vt_a show ?thesis by simp
-  next
-    case False
-    hence le_t1: "t \<le> length s" by simp
-    from vt_a have "vt s"
-      by (erule_tac evt_cons, simp)
-    from h [OF this] have "vt (moment t s)" .
-    moreover have "moment t (a#s) = moment t s"
-    proof -
-      from moment_app [OF le_t1, of "[a]"] 
-      show ?thesis by simp
-    qed
-    ultimately show ?thesis by auto
-  qed
-qed
-
-(* Wrong:
-    lemma \<lbrakk>thread \<in> set (wq_fun cs1 s); thread \<in> set (wq_fun cs2 s)\<rbrakk> \<Longrightarrow> cs1 = cs2"
-*)
-
-lemma waiting_unique_pre:
-  fixes cs1 cs2 s thread
-  assumes vt: "vt s"
-  and 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: "vt (e#moment t2 s)"
-      proof -
-        from vt_moment [OF vt]
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      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 abs2 [OF vt_e True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from block_pre [OF vt_e False h1]
-        have "e = P thread cs2" .
-        with 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: "vt  (e#moment t1 s)"
-      proof -
-        from vt_moment [OF vt]
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      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 abs2 [OF vt_e True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from block_pre [OF vt_e False h1]
-        have "e = P thread cs1" .
-        with 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 [OF vt]
-        have "vt (moment ?t3 s)" .
-        with eq_m show ?thesis by simp
-      qed
-      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 abs2 [OF vt_e True eq_th h2 h1]
-        show ?thesis by auto
-      next
-        case False
-        from block_pre [OF vt_e 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 
-          from abs2 [OF this True eq_th h2 h1]
-          show ?thesis .
-        next
-          case False
-          have vt_e: "vt (e#moment t2 s)"
-          proof -
-            from vt_moment [OF vt] eqt12
-            have "vt (moment (Suc t2) s)" by auto
-            with eq_m eqt12 show ?thesis by simp
-          qed
-          from block_pre [OF vt_e False h1]
-          have "e = P thread cs2" .
-          with eq_e1 neq12 show ?thesis by auto
-        qed
-      qed
-    } ultimately show ?thesis by arith
-  qed
-qed
-
-lemma waiting_unique:
-  fixes s cs1 cs2
-  assumes "vt s"
-  and "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
-
-(* not used *)
-lemma held_unique:
-  fixes s::"state"
-  assumes "holding s th1 cs"
-  and "holding s th2 cs"
-  shows "th1 = th2"
-using assms
-unfolding s_holding_def
-by auto
-
-
-lemma birthtime_lt: "th \<in> threads s \<Longrightarrow> birthtime th s < length s"
-  apply (induct s, auto)
-  by (case_tac a, auto split:if_splits)
-
-lemma birthtime_unique: 
-  "\<lbrakk>birthtime th1 s = birthtime 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:birthtime_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 "birthtime th1 s = birthtime th2 s" by (simp add:preced_def)
-  from birthtime_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
-
-lemma unique_minus:
-  fixes x y z r
-  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:
-  fixes r x y z
-  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:
-  fixes r x y z
-  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
-
-lemma depend_set_unchanged: "(depend (Set th prio # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-lemma depend_create_unchanged: "(depend (Create th prio # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-lemma depend_exit_unchanged: "(depend (Exit th # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
-by (simp add:Let_def)
-
-
-
-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
-
-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"
-  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 wq_distinct[OF step_back_vt[OF vt], 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 wq_distinct[OF step_back_vt[OF vt], 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 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 wq_distinct[OF step_back_vt[OF vt], 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 wq_distinct[OF step_back_vt[OF vt], 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)"
-  have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
-  proof(rule someI2)
-    from wq_distinct[OF step_back_vt[OF vt], 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"
-  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 wq_distinct [OF step_back_vt[OF vt], 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 wq_distinct [OF step_back_vt[OF vt], 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 wq_distinct [OF step_back_vt[OF vt], of cs, unfolded wq_def]
-      show False by auto
-    qed
-  qed
-qed
-
-lemma step_depend_v:
-fixes th::thread
-assumes vt:
-  "vt (V th cs#s)"
-shows "
-  depend (V th cs # s) =
-  depend 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_depend_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
-
-lemma step_depend_p:
-  "vt (P th cs#s) \<Longrightarrow>
-  depend (P th cs # s) =  (if (wq s cs = []) then depend s \<union> {(Cs cs, Th th)}
-                                             else depend s \<union> {(Th th, Cs cs)})"
-  apply(simp only: s_depend_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(auto simp:s_depend_def wq_def cs_holding_def)
-  done
-
-lemma simple_A:
-  fixes A
-  assumes h: "\<And> x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x = y"
-  shows "A = {} \<or> (\<exists> a. A = {a})"
-proof(cases "A = {}")
-  case True thus ?thesis by simp
-next
-  case False then obtain a where "a \<in> A" by auto
-  with h have "A = {a}" by auto
-  thus ?thesis by simp
-qed
-
-lemma depend_target_th: "(Th th, x) \<in> depend (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
-  by (unfold s_depend_def, auto)
-
-lemma acyclic_depend: 
-  fixes s
-  assumes vt: "vt s"
-  shows "acyclic (depend s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    assume ih: "acyclic (depend 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:depend_create_unchanged)
-    next
-      case (Exit th)
-      with ih show ?thesis by (simp add:depend_exit_unchanged)
-    next
-      case (V th cs)
-      from V vt stp have vtt: "vt (V th cs#s)" by auto
-      from step_depend_v [OF this]
-      have eq_de: 
-        "depend (e # s) = 
-            depend 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 depend_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_depend_def s_waiting_def cs_waiting_def wq_def by simp
-            hence "cs' = cs"
-            proof(rule waiting_unique [OF vt])
-              from eq_wq wq_distinct[OF vt, 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 wq_distinct[OF vt, 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 wq_distinct[OF vt, 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 wq_distinct[OF vt, 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 wq_distinct[OF vt, 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_depend_p [OF this] P
-    have "depend (e # s) = 
-      (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
-      depend 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 =  depend s \<union> {(Cs cs, Th th)}" by simp
-      have "(Th th, Cs cs) \<notin> (depend s)\<^sup>*"
-      proof
-        assume "(Th th, Cs cs) \<in> (depend s)\<^sup>*"
-        hence "(Th th, Cs cs) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
-        from tranclD2 [OF this]
-        obtain x where "(x, Cs cs) \<in> depend s" by auto
-        with True show False by (auto simp:s_depend_def cs_waiting_def)
-      qed
-      with acyclic_insert ih eq_r show ?thesis by auto
-    next
-      case False
-      hence eq_r: "?R =  depend s \<union> {(Th th, Cs cs)}" by simp
-      have "(Cs cs, Th th) \<notin> (depend s)\<^sup>*"
-      proof
-        assume "(Cs cs, Th th) \<in> (depend s)\<^sup>*"
-        hence "(Cs cs, Th th) \<in> (depend 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> (depend 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 depend_set_unchanged
-      show ?thesis by (simp add:depend_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "acyclic (depend ([]::state))"
-      by (auto simp: s_depend_def cs_waiting_def 
-                      cs_holding_def wq_def acyclic_def)
-  qed
-qed
-
-lemma finite_depend: 
-  fixes s
-  assumes vt: "vt s"
-  shows "finite (depend s)"
-proof -
-  from vt show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    assume ih: "finite (depend 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:depend_create_unchanged)
-    next
-      case (Exit th)
-      with ih show ?thesis by (simp add:depend_exit_unchanged)
-    next
-      case (V th cs)
-      from V vt stp have vtt: "vt (V th cs#s)" by auto
-      from step_depend_v [OF this]
-      have eq_de: "depend (e # s) = 
-                   depend 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 auto
-        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_depend_p [OF this] P
-      have "depend (e # s) = 
-              (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
-                                    depend 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 =  depend s \<union> {(Cs cs, Th th)}" by simp
-        with True and ih show ?thesis by auto
-      next
-        case False
-        hence "?R = depend 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:depend_set_unchanged)
-    qed
-  next
-    case vt_nil
-    show "finite (depend ([]::state))"
-      by (auto simp: s_depend_def cs_waiting_def 
-                   cs_holding_def wq_def acyclic_def)
-  qed
-qed
-
-text {* Several useful lemmas *}
-
-lemma wf_dep_converse: 
-  fixes s
-  assumes vt: "vt s"
-  shows "wf ((depend s)^-1)"
-proof(rule finite_acyclic_wf_converse)
-  from finite_depend [OF vt]
-  show "finite (depend s)" .
-next
-  from acyclic_depend[OF vt]
-  show "acyclic (depend s)" .
-qed
-
-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> depend (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> depend s"
-  by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
-
-lemma wq_threads: 
-  fixes s cs
-  assumes vt: "vt s"
-  and 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)
-    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_depend_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 wq_distinct[OF vt, 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>vt s; (Th th) \<in> Range (depend (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
-  apply(unfold s_depend_def cs_waiting_def cs_holding_def)
-  by (auto intro:wq_threads)
-
-lemma readys_v_eq:
-  fixes th thread cs rest
-  assumes vt: "vt s"
-  and 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 vt, 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
-
-lemma chain_building:
-  assumes vt: "vt s"
-  shows "node \<in> Domain (depend s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (depend s)^+)"
-proof -
-  from wf_dep_converse [OF vt]
-  have h: "wf ((depend s)\<inverse>)" .
-  show ?thesis
-  proof(induct rule:wf_induct [OF h])
-    fix x
-    assume ih [rule_format]: 
-      "\<forall>y. (y, x) \<in> (depend s)\<inverse> \<longrightarrow> 
-           y \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (depend s)\<^sup>+)"
-    show "x \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+)"
-    proof
-      assume x_d: "x \<in> Domain (depend s)"
-      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+"
-      proof(cases x)
-        case (Th th)
-        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> depend s" by (auto simp:s_depend_def)
-        with Th have x_in_r: "(Cs cs, x) \<in> (depend s)^-1" by simp
-        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> depend s" by blast
-        hence "Cs cs \<in> Domain (depend 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> (depend s)\<^sup>+" by auto
-        have "(x, Th th') \<in> (depend 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> (depend s)^-1" by (auto simp:s_depend_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 [OF vt] have "th' \<in> threads s" by auto
-          with False have "Th th' \<in> Domain (depend s)" 
-            by (auto simp:readys_def wq_def s_waiting_def s_depend_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> (depend s)\<^sup>+" by auto
-          from th'_d and th''_in 
-          have "(x, Th th'') \<in> (depend s)\<^sup>+" by auto
-          with th''_r show ?thesis by auto
-        qed
-      qed
-    qed
-  qed
-qed
-
-lemma th_chain_to_ready:
-  fixes s th
-  assumes vt: "vt s"
-  and th_in: "th \<in> threads s"
-  shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (depend 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 (depend s)" 
-    by (auto simp:readys_def s_waiting_def s_depend_def wq_def cs_waiting_def Domain_def)
-  from chain_building [rule_format, OF vt this]
-  show ?thesis by auto
-qed
-
-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)
-
-lemma unique_depend: "\<lbrakk>vt s; (n, n1) \<in> depend s; (n, n2) \<in> depend s\<rbrakk> \<Longrightarrow> n1 = n2"
-  apply(unfold s_depend_def, auto, fold waiting_eq holding_eq)
-  by(auto elim:waiting_unique holding_unique)
-
-lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
-by (induct rule:trancl_induct, auto)
-
-lemma dchain_unique:
-  assumes vt: "vt s"
-  and th1_d: "(n, Th th1) \<in> (depend s)^+"
-  and th1_r: "th1 \<in> readys s"
-  and th2_d: "(n, Th th2) \<in> (depend 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_depend [OF vt]
-    have "(Th th1, Th th2) \<in> (depend s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (depend s)\<^sup>+" by auto
-    hence "False"
-    proof
-      assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th1, n) \<in> depend s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_depend_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_depend_def wq_def s_waiting_def cs_waiting_def)
-      with th1_r show ?thesis by auto
-    next
-      assume "(Th th2, Th th1) \<in> (depend s)\<^sup>+"
-      from trancl_split [OF this]
-      obtain n where dd: "(Th th2, n) \<in> depend s" by auto
-      then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_depend_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_depend_def s_waiting_def cs_waiting_def)
-      with th2_r show ?thesis by auto
-    qed
-  } thus ?thesis by auto
-qed
-             
-
-lemma step_holdents_p_add:
-  fixes th cs s
-  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_depend_p[OF vt] by (auto)
-qed
-
-lemma step_holdents_p_eq:
-  fixes th cs s
-  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_depend_p[OF vt] by auto
-qed
-
-
-lemma finite_holding:
-  fixes s th cs
-  assumes vt: "vt s"
-  shows "finite (holdents s th)"
-proof -
-  let ?F = "\<lambda> (x, y). the_cs x"
-  from finite_depend [OF vt]
-  have "finite (depend s)" .
-  hence "finite (?F `(depend s))" by simp
-  moreover have "{cs . (Cs cs, Th th) \<in> depend 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> depend s"
-      hence "?F (Cs x, Th th) \<in> ?F `(depend s)" by (rule h)
-      moreover have "?F (Cs x, Th th) = x" by simp
-      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` depend s" by simp 
-    } thus ?thesis by auto
-  qed
-  ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
-qed
-
-lemma cntCS_v_dec: 
-  fixes s thread cs
-  assumes vtv: "vt (V thread cs#s)"
-  shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
-proof -
-  from step_back_step[OF vtv]
-  have cs_in: "cs \<in> holdents s thread" 
-    apply (cases, unfold holdents_test s_depend_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 wq_distinct[OF step_back_vt[OF vtv], of cs])
-    apply (unfold holdents_test, unfold step_depend_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> depend 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 finite_holding [OF vtv]
-    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 cnp_cnv_cncs:
-  fixes s th
-  assumes vt: "vt s"
-  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)
-    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 wq_threads [OF vt 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:depend_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:depend_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> (depend s)\<^sup>+"
-      from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
-      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, clarify)
-            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_depend_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_depend_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> depend s} =
-                  Suc (card {cs. (Cs cs, Th thread) \<in> depend 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 finite_holding [OF vt, of thread]
-                    show " finite {cs. (Cs cs, Th thread) \<in> depend 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> depend 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_depend_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 wq_distinct [OF vtp, 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_depend_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
-      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 wq_distinct[OF step_back_vt[OF vtv], 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
-      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 wq_distinct[OF step_back_vt[OF vtv], of cs]
-                    and eq_wq show False by auto
-                  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 (unfold eq_e, rule readys_v_eq [OF _ neq_th eq_wq False], auto)
-            moreover have "cntCS (e#s) th = cntCS s th"
-              apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_depend_v[OF vtv], auto)
-              proof -
-                have "{csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
-                      {cs. (Cs cs, Th th) \<in> depend s}"
-                proof -
-                  from False eq_wq
-                  have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> depend 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 wq_distinct[OF step_back_vt[OF vtv], 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> depend s"
-                      by auto
-                  qed    
-                  thus ?thesis by auto
-                qed
-                thus "card {csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
-                             card {cs. (Cs cs, Th th) \<in> depend 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 wq_distinct[OF step_back_vt[OF vtv], 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_depend_v[OF vtv], 
-                         auto simp:next_th_def eq_set s_depend_def holdents_test wq_def
-                                   Let_def cs_holding_def)
-                  by (insert wq_distinct[OF step_back_vt[OF vtv], 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 wq_threads[OF step_back_vt[OF vtv]])
-                  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 wq_distinct[OF step_back_vt[OF vtv], 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 waiting_unique[OF step_back_vt[OF vtv] 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 wq_threads [OF step_back_vt[OF vtv] 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_depend_v[OF vtv], auto)
-              proof -
-                show "card {csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs} =
-                               Suc (card {cs. (Cs cs, Th th) \<in> depend 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) ` depend s)" 
-                      apply (auto simp:image_def)
-                      by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
-                    with finite_depend[OF step_back_vt[OF vtv]]
-                    show "finite {cs. (Cs cs, Th th) \<in> depend s}" by (auto intro:finite_subset)
-                  next
-                    show "cs \<notin> {cs. (Cs cs, Th th) \<in> depend s}"
-                    proof
-                      assume "cs \<in> {cs. (Cs cs, Th th) \<in> depend s}"
-                      hence "(Cs cs, Th th) \<in> depend s" by simp
-                      with True neq_th eq_wq show False
-                        by (auto simp:next_th_def s_depend_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:depend_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_depend_def wq_def cs_holding_def)
-  qed
-qed
-
-lemma not_thread_cncs:
-  fixes th s
-  assumes vt: "vt s"
-  and 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)
-    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:depend_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:depend_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_depend_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
-      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 wq_distinct[OF step_back_vt[OF vtv], 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 "hd x \<in> set rest" by (cases x, auto)
-        qed
-        with eq_wq have "?t \<in> set (wq s cs)" by simp
-        from wq_threads[OF step_back_vt[OF vtv], OF this] and ni
-        show False by auto
-      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_depend_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:depend_set_unchanged)
-    qed
-    next
-      case vt_nil
-      show ?case
-      by (unfold cntCS_def, 
-        auto simp:count_def holdents_test s_depend_def wq_def cs_holding_def)
-  qed
-qed
-
-lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
-  by (auto simp:s_waiting_def cs_waiting_def wq_def)
-
-lemma dm_depend_threads:
-  fixes th s
-  assumes vt: "vt s"
-  and in_dom: "(Th th) \<in> Domain (depend s)"
-  shows "th \<in> threads s"
-proof -
-  from in_dom obtain n where "(Th th, n) \<in> depend s" by auto
-  moreover from depend_target_th[OF this] obtain cs where "n = Cs cs" by auto
-  ultimately have "(Th th, Cs cs) \<in> depend s" by simp
-  hence "th \<in> set (wq s cs)"
-    by (unfold s_depend_def, auto simp:cs_waiting_def)
-  from wq_threads [OF vt 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 runing_unique:
-  fixes th1 th2 s
-  assumes vt: "vt s"
-  and 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"
-    by (unfold runing_def, simp)
-  hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1)) =
-                 Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependents (wq s) th2))"
-    (is "Max (?f ` ?A) = Max (?f ` ?B)")
-    by (unfold cp_eq_cpreced 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 (dependents (wq s) th1)"
-        proof-
-          have "finite {th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (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_depend[OF vt] have "finite (depend s)" .
-              hence "finite ((depend (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_depend_def cs_depend_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependents_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 by (auto intro:that)
-  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 (dependents (wq s) th2)"
-        proof-
-          have "finite {th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+}"
-          proof -
-            let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (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_depend[OF vt] have "finite (depend s)" .
-              hence "finite ((depend (wq s))\<^sup>+)"
-                apply (unfold finite_trancl)
-                by (auto simp: s_depend_def cs_depend_def wq_def)
-              thus ?thesis by auto
-            qed
-            ultimately show ?thesis by (auto intro:finite_subset)
-          qed
-          thus ?thesis by (simp add:cs_dependents_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> dependents (wq s) th1)" by simp
-    thus "th1' \<in> threads s"
-    proof
-      assume "th1' \<in> dependents (wq s) th1"
-      hence "(Th th1') \<in> Domain ((depend s)^+)"
-        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
-        by (auto simp:Domain_def)
-      hence "(Th th1') \<in> Domain (depend s)" by (simp add:trancl_domain)
-      from dm_depend_threads[OF vt 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> dependents (wq s) th2)" by simp
-    thus "th2' \<in> threads s"
-    proof
-      assume "th2' \<in> dependents (wq s) th2"
-      hence "(Th th2') \<in> Domain ((depend s)^+)"
-        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
-        by (auto simp:Domain_def)
-      hence "(Th th2') \<in> Domain (depend s)" by (simp add:trancl_domain)
-      from dm_depend_threads[OF vt 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> dependents (wq s) th1" by simp
-  thus ?thesis
-  proof
-    assume eq_th': "th1' = th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
-    thus ?thesis
-    proof
-      assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
-    next
-      assume "th2' \<in> dependents (wq s) th2"
-      with eq_th12 eq_th' have "th1 \<in> dependents (wq s) th2" by simp
-      hence "(Th th1, Th th2) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      hence "Th th1 \<in> Domain ((depend s)^+)" 
-        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
-        by (auto simp:Domain_def)
-      hence "Th th1 \<in> Domain (depend s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th1, n) \<in> depend s" by (auto simp:Domain_def)
-      from depend_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th1, Cs cs') \<in> depend s" by simp
-      with runing_1 have "False"
-        apply (unfold runing_def readys_def s_depend_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    qed
-  next
-    assume th1'_in: "th1' \<in> dependents (wq s) th1"
-    from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
-    thus ?thesis 
-    proof
-      assume "th2' = th2"
-      with th1'_in eq_th12 have "th2 \<in> dependents (wq s) th1" by simp
-      hence "(Th th2, Th th1) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      hence "Th th2 \<in> Domain ((depend s)^+)" 
-        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
-        by (auto simp:Domain_def)
-      hence "Th th2 \<in> Domain (depend s)" by (simp add:trancl_domain)
-      then obtain n where d: "(Th th2, n) \<in> depend s" by (auto simp:Domain_def)
-      from depend_target_th [OF this]
-      obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th2, Cs cs') \<in> depend s" by simp
-      with runing_2 have "False"
-        apply (unfold runing_def readys_def s_depend_def)
-        by (auto simp:eq_waiting)
-      thus ?thesis by simp
-    next
-      assume "th2' \<in> dependents (wq s) th2"
-      with eq_th12 have "th1' \<in> dependents (wq s) th2" by simp
-      hence h1: "(Th th1', Th th2) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      from th1'_in have h2: "(Th th1', Th th1) \<in> (depend s)^+"
-        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
-      show ?thesis
-      proof(rule dchain_unique[OF vt 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 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 length_down_to_in: 
-  assumes le_ij: "i \<le> j"
-    and le_js: "j \<le> length s"
-  shows "length (down_to j i s) = j - i"
-proof -
-  have "length (down_to j i s) = length (from_to i j (rev s))"
-    by (unfold down_to_def, auto)
-  also have "\<dots> = j - i"
-  proof(rule length_from_to_in[OF le_ij])
-    from le_js show "j \<le> length (rev s)" by simp
-  qed
-  finally show ?thesis .
-qed
-
-
-lemma moment_head: 
-  assumes le_it: "Suc i \<le> length t"
-  obtains e where "moment (Suc i) t = e#moment i t"
-proof -
-  have "i \<le> Suc i" by simp
-  from length_down_to_in [OF this le_it]
-  have "length (down_to (Suc i) i t) = 1" by auto
-  then obtain e where "down_to (Suc i) i t = [e]"
-    apply (cases "(down_to (Suc i) i t)") by auto
-  moreover have "down_to (Suc i) 0 t = down_to (Suc i) i t @ down_to i 0 t"
-    by (rule down_to_conc[symmetric], auto)
-  ultimately have eq_me: "moment (Suc i) t = e#(moment i t)"
-    by (auto simp:down_to_moment)
-  from that [OF this] show ?thesis .
-qed
-
-lemma cnp_cnv_eq:
-  fixes th s
-  assumes "vt s"
-  and "th \<notin> threads s"
-  shows "cntP s th = cntV s th"
-proof -
-  from assms show ?thesis
-  proof(induct)
-    case (vt_cons s e)
-    have ih: "th \<notin> threads s \<Longrightarrow> cntP s th = cntV s th" by fact
-    have not_in: "th \<notin> threads (e # s)" by fact
-    have "step s e" by fact
-    thus ?case proof(cases)
-      case (thread_create thread prio)
-      assume eq_e: "e = Create thread prio"
-      hence "thread \<in> threads (e#s)" by simp
-      with not_in and eq_e have "th \<notin> threads s" by auto
-      from ih [OF this] show ?thesis using eq_e
-        by (auto simp:cntP_def cntV_def count_def)
-    next
-      case (thread_exit thread)
-      assume eq_e: "e = Exit thread"
-        and not_holding: "holdents s thread = {}"
-      have vt_s: "vt s" by fact
-      from finite_holding[OF vt_s] have "finite (holdents s thread)" .
-      with not_holding have "cntCS s thread = 0" by (unfold cntCS_def, auto)
-      moreover have "thread \<in> readys s" using thread_exit by (auto simp:runing_def)
-      moreover note cnp_cnv_cncs[OF vt_s, of thread]
-      ultimately have eq_thread: "cntP s thread = cntV s thread" by auto
-      show ?thesis
-      proof(cases "th = thread")
-        case True
-        with eq_thread eq_e show ?thesis 
-          by (auto simp:cntP_def cntV_def count_def)
-      next
-        case False
-        with not_in and eq_e have "th \<notin> threads s" by simp
-        from ih[OF this] and eq_e show ?thesis 
-           by (auto simp:cntP_def cntV_def count_def)
-      qed
-    next
-      case (thread_P thread cs)
-      assume eq_e: "e = P thread cs"
-      have "thread \<in> runing s" by fact
-      with not_in eq_e have neq_th: "thread \<noteq> th" 
-        by (auto simp:runing_def readys_def)
-      from not_in eq_e have "th \<notin> threads s" by simp
-      from ih[OF this] and neq_th and eq_e show ?thesis
-        by (auto simp:cntP_def cntV_def count_def)
-    next
-      case (thread_V thread cs)
-      assume eq_e: "e = V thread cs"
-      have "thread \<in> runing s" by fact
-      with not_in eq_e have neq_th: "thread \<noteq> th" 
-        by (auto simp:runing_def readys_def)
-      from not_in eq_e have "th \<notin> threads s" by simp
-      from ih[OF this] and neq_th and eq_e show ?thesis
-        by (auto simp:cntP_def cntV_def count_def)
-    next
-      case (thread_set thread prio)
-      assume eq_e: "e = Set thread prio"
-        and "thread \<in> runing s"
-      hence "thread \<in> threads (e#s)" 
-        by (simp add:runing_def readys_def)
-      with not_in and eq_e have "th \<notin> threads s" by auto
-      from ih [OF this] show ?thesis using eq_e
-        by (auto simp:cntP_def cntV_def count_def)  
-    qed
-  next
-    case vt_nil
-    show ?case by (auto simp:cntP_def cntV_def count_def)
-  qed
-qed
-
-lemma eq_depend: 
-  "depend (wq s) = depend s"
-by (unfold cs_depend_def s_depend_def, auto)
-
-lemma count_eq_dependents:
-  assumes vt: "vt s"
-  and eq_pv: "cntP s th = cntV s th"
-  shows "dependents (wq s) th = {}"
-proof -
-  from cnp_cnv_cncs[OF vt] and eq_pv
-  have "cntCS s th = 0" 
-    by (auto split:if_splits)
-  moreover have "finite {cs. (Cs cs, Th th) \<in> depend s}"
-  proof -
-    from finite_holding[OF vt, of th] show ?thesis
-      by (simp add:holdents_test)
-  qed
-  ultimately have h: "{cs. (Cs cs, Th th) \<in> depend s} = {}"
-    by (unfold cntCS_def holdents_test cs_dependents_def, auto)
-  show ?thesis
-  proof(unfold cs_dependents_def)
-    { assume "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}"
-      then obtain th' where "(Th th', Th th) \<in> (depend (wq s))\<^sup>+" by auto
-      hence "False"
-      proof(cases)
-        assume "(Th th', Th th) \<in> depend (wq s)"
-        thus "False" by (auto simp:cs_depend_def)
-      next
-        fix c
-        assume "(c, Th th) \<in> depend (wq s)"
-        with h and eq_depend show "False"
-          by (cases c, auto simp:cs_depend_def)
-      qed
-    } thus "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} = {}" by auto
-  qed
-qed
-
-lemma dependents_threads:
-  fixes s th
-  assumes vt: "vt s"
-  shows "dependents (wq s) th \<subseteq> threads s"
-proof
-  { fix th th'
-    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (depend (wq s))\<^sup>+}"
-    have "Th th \<in> Domain (depend s)"
-    proof -
-      from h obtain th' where "(Th th, Th th') \<in> (depend (wq s))\<^sup>+" by auto
-      hence "(Th th) \<in> Domain ( (depend (wq s))\<^sup>+)" by (auto simp:Domain_def)
-      with trancl_domain have "(Th th) \<in> Domain (depend (wq s))" by simp
-      thus ?thesis using eq_depend by simp
-    qed
-    from dm_depend_threads[OF vt this]
-    have "th \<in> threads s" .
-  } note hh = this
-  fix th1 
-  assume "th1 \<in> dependents (wq s) th"
-  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (depend (wq s))\<^sup>+}"
-    by (unfold cs_dependents_def, simp)
-  from hh [OF this] show "th1 \<in> threads s" .
-qed
-
-lemma finite_threads:
-  assumes vt: "vt s"
-  shows "finite (threads s)"
-using vt
-by (induct) (auto elim: step.cases)
-
-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 cp_le:
-  assumes vt: "vt s"
-  and 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_dependents_def)
-  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (depend (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> (depend (wq s))\<^sup>+} \<noteq> {}" by simp
-  next
-    from finite_threads [OF vt]
-    show "finite (threads s)" .
-  next
-    from th_in
-    show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> threads s"
-      apply (auto simp:Domain_def)
-      apply (rule_tac dm_depend_threads[OF vt])
-      apply (unfold trancl_domain [of "depend s", symmetric])
-      by (unfold cs_depend_def s_depend_def, auto simp:Domain_def)
-  qed
-qed
-
-lemma le_cp:
-  assumes vt: "vt s"
-  shows "preced th s \<le> cp s th"
-proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
-  show "Prc (original_priority th s) (birthtime th s)
-    \<le> Max (insert (Prc (original_priority th s) (birthtime th s))
-            ((\<lambda>th. Prc (original_priority th s) (birthtime th s)) ` dependents (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> (depend (wq s))\<^sup>+}"
-        proof -
-          let ?F = "\<lambda> (x, y). the_th x"
-          have "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (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_depend[OF vt] have "finite (depend s)" .
-            hence "finite ((depend (wq s))\<^sup>+)"
-              apply (unfold finite_trancl)
-              by (auto simp: s_depend_def cs_depend_def wq_def)
-            thus ?thesis by auto
-          qed
-          ultimately show ?thesis by (auto intro:finite_subset)
-        qed
-        thus ?thesis by (simp add:cs_dependents_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: 
-  assumes vt: "vt s"
-  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[OF vt] 
-    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 vt 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[OF vt]
-      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 vt, 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[OF vt]
-        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 vt: "vt s"
-  and np: "threads s \<noteq> {}"
-  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
-proof(unfold max_cp_eq[OF vt])
-  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[OF vt] 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 vt tm_in]
-    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+)" .
-    thus ?thesis
-    proof
-      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+ "
-      then obtain th' where th'_in: "th' \<in> readys s" 
-        and tm_chain:"(Th tm, Th th') \<in> (depend s)\<^sup>+" by auto
-      have "cp s th' = ?f tm"
-      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
-        from dependents_threads[OF vt] finite_threads[OF vt]
-        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th'))" 
-          by (auto intro:finite_subset)
-      next
-        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependents (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[OF vt]
-          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        next
-          from p_in and th'_in and dependents_threads[OF vt, 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> dependents (wq s) th')"
-        proof -
-          from tm_chain
-          have "tm \<in> dependents (wq s) th'"
-            by (unfold cs_dependents_def s_depend_def cs_depend_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 dependents_threads [OF vt, of th1] th1_in
-          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (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> dependents (wq s) th1) \<noteq> {}" by simp
-        next
-          from finite_threads[OF vt] 
-          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
-        qed
-      next
-        from finite_threads[OF vt]
-        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> dependents (wq s) tm)"
-          show "y \<le> preced tm s"
-          proof -
-            { fix y'
-              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependents (wq s) tm)"
-              have "y' \<le> preced tm s"
-              proof(unfold tm_max, rule Max_ge)
-                from hy' dependents_threads[OF vt, of tm]
-                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
-              next
-                from finite_threads[OF vt] 
-                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-              qed
-            } with hy show ?thesis by auto
-          qed
-        next
-          from dependents_threads[OF vt, of tm] finite_threads[OF vt]
-          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm))"
-            by (auto intro:finite_subset)
-        next
-          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependents (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[OF vt]
-          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[OF vt]
-            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
-          next
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<noteq> {}"
-              by simp
-          next
-            from dependents_threads[OF vt, of th1] th1_readys
-            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (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
-
-lemma max_cp_readys_threads:
-  assumes vt: "vt s"
-  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 vt False])
-qed
-
-
-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 (depend s))"
-apply(simp add: detached_def Field_def)
-apply(simp add: s_depend_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
-
-lemma detached_intro:
-  fixes s th
-  assumes vt: "vt s"
-  and eq_pv: "cntP s th = cntV s th"
-  shows "detached s th"
-proof -
- from cnp_cnv_cncs[OF vt]
-  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[OF vt] dm_depend_threads[OF vt]
-    show ?thesis
-      by (auto simp add: detached_def s_depend_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 (depend s)"
-    proof -
-      from card_0_eq [OF finite_holding [OF vt]] 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_depend_def)
-        done
-    qed
-    ultimately show ?thesis
-      by (auto simp add: detached_def s_depend_def s_waiting_abv s_holding_abv wq_def readys_def)
-  qed
-qed
-
-lemma detached_elim:
-  fixes s th
-  assumes vt: "vt s"
-  and dtc: "detached s th"
-  shows "cntP s th = cntV s th"
-proof -
-  from cnp_cnv_cncs[OF vt]
-  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_depend_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_depend_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:
-  fixes s th
-  assumes vt: "vt s"
-  shows "(detached s th) = (cntP s th = cntV s th)"
-  by (insert vt, auto intro:detached_intro detached_elim)
-
-end
\ No newline at end of file
--- a/prio/PrioGDef.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,483 +0,0 @@
-(*<*)
-theory PrioGDef
-imports Precedence_ord Moment
-begin
-(*>*)
-
-text {*
-  In this section, the formal model of Priority Inheritance 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. 
-
-  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
-  Every event in the system corresponds to a system call, the formats of which are
-  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 {* 
-\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 {* 
-  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
-  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
-  Functions such as @{text "threads"}, which extract information out of system states, are called
-  {\em observing functions}. A series of observing functions will be defined in the sequel in order to 
-  model the protocol. 
-  Observing function @{text "original_priority"} calculates 
-  the {\em original priority} of thread @{text "th"} in state @{text "s"}, expressed as
-  : @{text "original_priority th s" }. The {\em original priority} is the priority 
-  assigned to a thread when it is created or when it is reset by system call 
-  @{text "Set thread priority"}.
-*}
-
-fun original_priority :: "thread \<Rightarrow> state \<Rightarrow> priority"
-  where
-  -- {* @{text "0"} is assigned to threads which have never been created: *}
-  "original_priority thread [] = 0" |
-  "original_priority thread (Create thread' prio#s) = 
-     (if thread' = thread then prio else original_priority thread s)" |
-  "original_priority thread (Set thread' prio#s) = 
-     (if thread' = thread then prio else original_priority thread s)" |
-  "original_priority thread (e#s) = original_priority thread s"
-
-text {*
-  \noindent
-  In the following,
-  @{text "birthtime th s"} is the time when thread @{text "th"} is created, 
-  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 birthtime :: "thread \<Rightarrow> state \<Rightarrow> nat"
-  where
-  "birthtime thread [] = 0" |
-  "birthtime thread ((Create thread' prio)#s) = 
-       (if (thread = thread') then length s else birthtime thread s)" |
-  "birthtime thread ((Set thread' prio)#s) = 
-       (if (thread = thread') then length s else birthtime thread s)" |
-  "birthtime thread (e#s) = birthtime 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 birth time}. 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 (original_priority thread s) (birthtime thread s)"
-
-
-text {*
-  \noindent
-  A number of important notions are defined here:
-  *}
-
-consts 
-  holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool" 
-  waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
-  depend :: "'b \<Rightarrow> (node \<times> node) set"
-  dependents :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
-
-text {*
-  \noindent
-  In the definition of the following several functions, it is supposed that
-  the waiting queue of every critical resource is given by a waiting queue 
-  function @{text "wq"}, which servers as arguments of these functions.
-  *}
-defs (overloaded) 
-  -- {* 
-  \begin{minipage}{0.9\textwidth}
-  We define that the thread which is at the head of waiting queue of resource @{text "cs"}
-  is holding the resource. This definition is slightly different from tradition where
-  all threads in the waiting queue are considered as waiting for the resource.
-  This notion is reflected in the definition of @{text "holding wq th cs"} as follows:
-  \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 "depend wq"} represents the Resource Allocation Graph of the system under the waiting 
-  queue function @{text "wq"}.
-  \end{minipage}
-  *}
-  cs_depend_def: 
-  "depend (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 "dependents wq th"} represents the set of threads which are depending on
-  thread @{text "th"} in Resource Allocation Graph @{text "depend wq"}:
-  \end{minipage}
-  *}
-  cs_dependents_def: 
-  "dependents (wq::cs \<Rightarrow> thread list) th \<equiv> {th' . (Th th', Th th) \<in> (depend wq)^+}"
-
-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 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
-  @{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 dependents, 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> dependents wq th)))"
-
-(*<*)
-lemma 
-  cpreced_def2:
-  "cpreced wq s th \<equiv> Max ({preced th s} \<union> {preced th' s | th'. th' \<in> dependents wq th})"
-  unfolding cpreced_def image_def
-  apply(rule eq_reflection)
-  apply(rule_tac f="Max" in arg_cong)
-  by (auto)
-(*>*)
-
-abbreviation
-  "all_unlocked \<equiv> \<lambda>_::cs. ([]::thread list)"
-
-abbreviation 
-  "initial_cprec \<equiv> \<lambda>_::thread. Prc 0 0"
- 
-abbreviation
-  "release qs \<equiv> case qs of
-             [] => [] 
-          |  (_#qs) => (SOME q. distinct q \<and> set q = set qs)"
-
-text {* \noindent
-  The following function @{text "schs"} is used to calculate the schedule state @{text "schs s"}.
-  It is the key function to model Priority Inheritance:
-  *}
-fun schs :: "state \<Rightarrow> schedule_state"
-  where 
-  "schs [] = (| wq_fun = \<lambda> cs. [],  cprec_fun = (\<lambda>_. Prc 0 0) |)" |
-
-  -- {*
-  \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 dependency 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}
-  \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)|))"
-|  "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_dependents_def cs_depend_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 "depend"} and 
-  @{text "dependents"} still have the 
-  same meaning, but redefined so that they no longer depend 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_depend_abv: 
-  "depend (s::state) \<equiv> depend (wq_fun (schs s))"
-  s_dependents_abv: 
-  "dependents (s::state) \<equiv> dependents (wq_fun (schs s))"
-
-
-text {* 
-  The following lemma can be proved easily:
-  *}
-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_depend_def: 
-  "depend (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_depend_abv wq_def cs_depend_def)
-
-lemma
-  s_dependents_def: 
-  "dependents (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (depend (wq s))^+}"
-  by (auto simp:s_dependents_abv wq_def cs_dependents_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
-  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> depend s}"
-unfolding holdents_def
-unfolding s_depend_def
-unfolding s_holding_abv
-unfolding wq_def
-by (simp)
-
-text {* \noindent
-  @{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
-  The fact that event @{text "e"} is eligible to happen next in state @{text "s"} 
-  is expressed as @{text "step s e"}. The predicate @{text "step"} is inductively defined as 
-  follows:
-  *}
-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> (depend 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)" |
-  -- {*
-  A thread can adjust its own priority as long as it is current running:
-  *}  
-  thread_set: "\<lbrakk>thread \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
-
-text {* \noindent
-  With predicate @{text "step"}, the fact that @{text "s"} is a legal state in 
-  Priority Inheritance protocol can be 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, and event @{text "e"} is eligible to happen
-  in state @{text "s"}, then @{text "e#s"} is a legal state as well:
-  \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 step predicate to get the set of legal states.
-  *}
-
-text {* \noindent
-  The following two 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.
-  *}
-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 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 @{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 @{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/prio/README	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,14 +0,0 @@
-Theories:
-=========
-
- Precedence_ord.thy      A theory of precedences.
- Moment.thy              The notion of moment.
- PrioGDef.thy            The formal definition of the PIP-model.
- PrioG.thy               Basic properties of the PIP-model.
- ExtGG.thy               The correctness proof of the PIP-model.
- CpsG.thy                Properties interesting for an implementation.
-
-The repository can be checked using Isabelle 2011-1.
-
-  isabelle make session
-
--- a/prio/ROOT.ML	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2 +0,0 @@
-use_thy "CpsG"; 
-use_thy "ExtGG";
--- a/prio/Slides/ROOT1.ML	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,7 +0,0 @@
-(*show_question_marks := false;*)
-
-no_document use_thy "../CpsG"; 
-no_document use_thy "../ExtGG"; 
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
-quick_and_dirty := true;
-use_thy "Slides1"
\ No newline at end of file
--- a/prio/Slides/Slides1.thy	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,669 +0,0 @@
-(*<*)
-theory Slides1
-imports "../CpsG" "../ExtGG" "~~/src/HOL/Library/LaTeXsugar"
-begin
-
-notation (latex output)
-  set ("_") and
-  Cons  ("_::/_" [66,65] 65) 
-
-ML {*
-  open Printer;
-  show_question_marks_default := false;
-  *}
-
-notation (latex output)
-  Cons ("_::_" [78,77] 73) and
-  vt ("valid'_state") and
-  runing ("running") and
-  birthtime ("last'_set") and
-  If  ("(\<^raw:\textrm{>if\<^raw:}> (_)/ \<^raw:\textrm{>then\<^raw:}> (_)/ \<^raw:\textrm{>else\<^raw:}> (_))" 10) and
-  Prc ("'(_, _')") and
-  holding ("holds") and
-  waiting ("waits") and
-  Th ("T") and
-  Cs ("C") and
-  readys ("ready") and
-  depend ("RAG") and 
-  preced ("prec") and
-  cpreced ("cprec") and
-  dependents ("dependants") and
-  cp ("cprec") and
-  holdents ("resources") and
-  original_priority ("priority") and
-  DUMMY  ("\<^raw:\mbox{$\_\!\_$}>")
-
-(*>*)
-
-
-
-text_raw {*
-  \renewcommand{\slidecaption}{Nanjing, P.R. China, 1 August 2012}
-  \newcommand{\bl}[1]{#1}                        
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}
-  \frametitle{%
-  \begin{tabular}{@ {}c@ {}}
-  \\[-3mm]
-  \Large Priority Inheritance Protocol \\[-3mm] 
-  \Large Proved Correct \\[0mm]
-  \end{tabular}}
-  
-  \begin{center}
-  \small Xingyuan Zhang \\
-  \small \mbox{PLA University of Science and Technology} \\
-  \small \mbox{Nanjing, China}
-  \end{center}
-
-  \begin{center}
-  \small joint work with \\
-  Christian Urban \\
-  Kings College, University of London, U.K.\\
-  Chunhan Wu \\
-  My Ph.D. student now working for Christian\\
-  \end{center}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{\large Prioirty Inheritance Protocol (PIP)}
-  \large
-
-  \begin{itemize} 
-  \item Widely used in Real-Time OSs \pause
-  \item One solution of \textcolor{red}{`Priority Inversion'} \pause
-  \item A flawed manual correctness proof (1990)\pause
-        \begin{itemize} \large
-          \item {Notations with no precise definition}
-          \item {Resorts to intuitions}
-        \end{itemize} \pause
-  \item Formal treatments using model-checking \pause
-        \begin{itemize} \large
-          \item {Applicable to small size system models}
-          \item { Unhelpful for human understanding } 
-        \end{itemize} \pause
-  \item Verification of PCP in PVS (2000)\pause
-        \begin{itemize} \large
-           \item {A related protocol}
-           \item {Priority Ceiling Protocol}
-        \end{itemize}
-  \end{itemize}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Our Motivation}
-  \large
-
-  \begin{itemize}
-  \item Undergraduate OS course in our university \pause
-        \begin{itemize}
-           \item {\large Experiments using instrutional OSs }
-           \item {\large PINTOS (Stanford) is chosen }
-           \item {\large Core project: Implementing PIP in it}
-        \end{itemize} \pause
-  \item Understanding is crucial for the implemention \pause
-  \item Existing literature of little help \pause
-  \item Some mention the complication
-  \end{itemize}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{\mbox{Some excerpts}}
-  
-  \begin{quote}
-  ``Priority inheritance is neither ef$\!$ficient nor reliable. 
-  Implementations are either incomplete (and unreliable) 
-  or surprisingly complex and intrusive.''
-  \end{quote}\medskip
-
-  \pause
-
-  \begin{quote}
-  ``I observed in the kernel code (to my disgust), the Linux 
-  PIP implementation is a nightmare: extremely heavy weight, 
-  involving maintenance of a full wait-for graph, and requiring 
-  updates for a range of events, including priority changes and 
-  interruptions of wait operations.''
-  \end{quote}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Our Aims}
-  \large
-
-  \begin{itemize}
-  \item Formal specification at appropriate abstract level,
-        convenient for:
-  \begin{itemize} \large
-    \item Constructing interactive proofs
-    \item Clarifying the underlying ideas
-  \end{itemize} \pause
-  \item Theorems usable to guide implementation, critical point:
-    \begin{itemize} \large
-      \item Understanding the relationship with real OS code \pause
-      \item Not yet formalized
-    \end{itemize}
-  \end{itemize}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Real-Time OSes}
-  \large
-
-  \begin{itemize}
-  \item Purpose: gurantee the most urgent task to be processed in time
-  \item Processes have priorities\\
-  \item Resources can be locked and unlocked
-  \end{itemize}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Problem}
-  \Large
-
-  \begin{center}
-  \begin{tabular}{l}
-  \alert{H}igh-priority process\\[4mm]
-  \onslide<2->{\alert{M}edium-priority process}\\[4mm]
-  \alert{L}ow-priority process\\[4mm]
-  \end{tabular}
-  \end{center}
-
-  \onslide<3->{
-  \begin{itemize}
-  \item \alert{Priority Inversion} @{text "\<equiv>"} \alert{H $<$ L}
-  \item<4> avoid indefinite priority inversion
-  \end{itemize}}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% 
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Priority Inversion}
-
-  \begin{center}
-  \includegraphics[scale=0.4]{PriorityInversion.png}
-  \end{center}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Mars Pathfinder Mission 1997}
-  \Large
-
-  \begin{center}
-  \includegraphics[scale=0.2]{marspath1.png}
-  \includegraphics[scale=0.22]{marspath3.png}
-  \includegraphics[scale=0.4]{marsrover.png}
-  \end{center}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Solution}
-  \Large
-
-  \alert{Priority Inheritance Protocol (PIP):}
-
-  \begin{center}
-  \begin{tabular}{l}
-  \alert{H}igh-priority process\\[4mm]
-  \textcolor{gray}{Medium-priority process}\\[4mm]
-  \alert{L}ow-priority process\\[21mm]
-  {\normalsize (temporarily raise its priority)}
-  \end{tabular}
-  \end{center}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{A Correctness ``Proof'' in 1990}
-  \Large
-
-  \begin{itemize}
-  \item a paper$^\star$ 
-  in 1990 ``proved'' the correctness of an algorithm for PIP\\[5mm]
-  \end{itemize}
-
-  \normalsize
-  \begin{quote}
-  \ldots{}after the thread (whose priority has been raised) completes its 
-  critical section and releases the lock, it ``returns to its original 
-  priority level''.
-  \end{quote}\bigskip
-
-  \small
-  $^\star$ in IEEE Transactions on Computers
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{}
-  \Large
-
-  \begin{center}
-  \begin{tabular}{l}
-  \alert{H}igh-priority process 1\\[2mm]
-  \alert{H}igh-priority process 2\\[8mm]
-  \alert{L}ow-priority process
-  \end{tabular}
-  \end{center}
-
-  \onslide<2->{
-  \begin{itemize}
-  \item Solution: \\Return to highest \alert{remaining} priority
-  \end{itemize}}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% 
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Event Abstraction}
-
-  \begin{itemize}\large
-     \item Use Inductive Approach of L. Paulson \pause
-     \item System is event-driven \pause
-     \item A \alert{state} is a list of events 
-  \end{itemize}
-
-  \pause
-
-  \begin{center}
-  \includegraphics[scale=0.4]{EventAbstract.png}
-  \end{center}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Events}
-  \Large
-
-  \begin{center}
-  \begin{tabular}{l}
-  Create \textcolor{gray}{thread priority}\\
-  Exit \textcolor{gray}{thread}\\
-  Set \textcolor{gray}{thread priority}\\
-  Lock \textcolor{gray}{thread cs}\\
-  Unlock \textcolor{gray}{thread cs}\\
-  \end{tabular}
-  \end{center}\medskip
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Precedences}
-  \large
-
-  \begin{center}
-  \begin{tabular}{l}
-  @{thm preced_def[where thread="th"]} 
-  \end{tabular}
-  \end{center}
-
-  
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{RAGs}
-
-\begin{center}
-  \newcommand{\fnt}{\fontsize{7}{8}\selectfont}
-  \begin{tikzpicture}[scale=1]
-  %%\draw[step=2mm] (-3,2) grid (1,-1);
-
-  \node (A) at (0,0) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>0"}};
-  \node (B) at (2,0) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>1"}};
-  \node (C) at (4,0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>1"}};
-  \node (D) at (4,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>2"}};
-  \node (E) at (6,-0.7) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>2"}};
-  \node (E1) at (6, 0.2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>3"}};
-  \node (F) at (8,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>3"}};
-
-  \draw [<-,line width=0.6mm] (A) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding}  (B);
-  \draw [->,line width=0.6mm] (C) to node [pos=0.4,sloped,above=-0.5mm] {\fnt{}waiting}  (B);
-  \draw [->,line width=0.6mm] (D) to node [pos=0.4,sloped,below=-0.5mm] {\fnt{}waiting}  (B);
-  \draw [<-,line width=0.6mm] (D) to node [pos=0.54,sloped,below=-0.5mm] {\fnt{}holding}  (E);
-  \draw [<-,line width=0.6mm] (D) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding}  (E1);
-  \draw [->,line width=0.6mm] (F) to node [pos=0.45,sloped,below=-0.5mm] {\fnt{}waiting}  (E);
-  \end{tikzpicture}
-  \end{center}\bigskip
-
-  \begin{center}
-  \begin{minipage}{0.8\linewidth}
-  \raggedleft
-  @{thm cs_depend_def}
-  \end{minipage}
-  \end{center}\pause
-
-  
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Good Next Events}
-  %%\large
-
-  \begin{center}
-  @{thm[mode=Rule] thread_create[where thread=th]}\bigskip
-
-  @{thm[mode=Rule] thread_exit[where thread=th]}\bigskip
-
-  @{thm[mode=Rule] thread_set[where thread=th]}
-  \end{center}
-
-  
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Good Next Events}
-  %%\large
-
-  \begin{center}
-  @{thm[mode=Rule] thread_P[where thread=th]}\bigskip
-  
-  @{thm[mode=Rule] thread_V[where thread=th]}\bigskip
-  \end{center}
-
-  
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-(*<*)
-context extend_highest_gen
-begin
-(*>*)
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{\mbox{\large Theorem: ``No indefinite priority inversion''}}
-
-  \pause
-
-  Theorem $^\star$: If th is the thread with the highest precedence in state 
-  @{text "s"}: \pause
-  \begin{center}
-  \textcolor{red}{@{thm highest})}
-  \end{center}
-  \pause
-  and @{text "th"} is blocked by a thread @{text "th'"} in 
-  a future state @{text "s'"} (with @{text "s' = t@s"}): \pause
-  \begin{center}
-  \textcolor{red}{@{text "th' \<in> running (t@s)"} and @{text "th' \<noteq> th"}} \pause
-  \end{center}
-  \fbox{ \hspace{1em} \pause
-  \begin{minipage}{0.95\textwidth}
-  \begin{itemize}
-  \item @{text "th'"} did not hold or wait for a resource in s:
-  \begin{center}
-  \textcolor{red}{@{text "\<not>detached s th'"}}  
-  \end{center} \pause
-  \item @{text "th'"} is running with the precedence of @{text "th"}:
-  \begin{center}
-  \textcolor{red}{@{text "cp (t@s) th' = preced th s"}} 
-  \end{center} 
-  \end{itemize}
-  \end{minipage}}
-  \pause
-  \small
-  $^\star$ modulo some further assumptions\bigskip\pause
-  It does not matter which process gets a released lock.
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[t]
-  \frametitle{Implementation}
-  
-  s $=$ current state; @{text "s'"} $=$ next state $=$ @{text "e#s"}\bigskip\bigskip
-
-  When @{text "e"} = \alert{Create th prio}, \alert{Exit th} 
-  
-  \begin{itemize}
-  \item @{text "RAG s' = RAG s"}
-  \item No precedence needs to be recomputed
-  \end{itemize}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[t]
-  \frametitle{Implementation}
-
-  s $=$ current state; @{text "s'"} $=$ next state $=$ @{text "e#s"}\bigskip\bigskip
-  
-
-  When @{text "e"} = \alert{Set th prio}
-  
-  \begin{itemize}
-  \item @{text "RAG s' = RAG s"}
-  \item No precedence needs to be recomputed
-  \end{itemize}
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[t]
-  \frametitle{Implementation}
-
-  s $=$ current state; @{text "s'"} $=$ next state $=$ @{text "e#s"}\bigskip\bigskip
-
-  When @{text "e"} = \alert{Unlock th cs} where there is a thread to take over
-  
-  \begin{itemize}
-  \item @{text "RAG s' = RAG s - {(C cs, T th), (T th', C cs)} \<union> {(C cs, T th')}"}
-  \item we have to recalculate the precedence of the direct descendants
-  \end{itemize}\bigskip
-
-  \pause
-
-  When @{text "e"} =  \alert{Unlock th cs} where no thread takes over
-  
-  \begin{itemize}
-  \item @{text "RAG s' = RAG s - {(C cs, T th)}"}
-  \item no recalculation of precedences
-  \end{itemize}
-
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[t]
-  \frametitle{Implementation}
-
-  s $=$ current state; @{text "s'"} $=$ next state $=$ @{text "e#s"}\bigskip\bigskip
-
-  When @{text "e"} = \alert{Lock th cs} where cs is not locked
-  
-  \begin{itemize}
-  \item @{text "RAG s' = RAG s \<union> {(C cs, T th')}"}
-  \item no recalculation of precedences
-  \end{itemize}\bigskip
-
-  \pause
-
-  When @{text "e"} = \alert{Lock th cs} where cs is locked
-  
-  \begin{itemize}
-  \item @{text "RAG s' = RAG s - {(T th, C cs)}"}
-  \item we have to recalculate the precedence of the descendants
-  \end{itemize}
-
-
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Conclusion}
-  
-  \begin{itemize} \large
-  \item Aims fulfilled \medskip \pause
-  \item Alternative way \pause
-       \begin{itemize} 
-           \item using Isabelle/HOL in OS code development \medskip
-           \item through the Inductive Approach
-       \end{itemize} \pause
-  \item Future research \pause
-       \begin{itemize}
-           \item scheduler in RT-Linux\medskip
-           \item multiprocessor case\medskip
-           \item other ``nails'' ? (networks, \ldots) \medskip \pause
-           \item Refinement to real code and relation between implementations
-        \end{itemize}
-  \end{itemize}
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-text_raw {*
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-  \mode<presentation>{
-  \begin{frame}[c]
-  \frametitle{Questions?}
-  
-  \begin{itemize} \large
-  \item Thank you for listening!
-  \end{itemize}
-  
-  \end{frame}}
-  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
-*}
-
-
-(*<*)
-end
-end
-(*>*)
\ No newline at end of file
--- a/prio/Slides/document/beamerthemeplaincu.sty	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,126 +0,0 @@
-\ProvidesPackage{beamerthemeplaincu}[2003/11/07 ver 0.93]
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-
-% Copyright 2003 by Till Tantau <tantau@cs.tu-berlin.de>.
-%
-% This program can be redistributed and/or modified under the terms
-% of the LaTeX Project Public License Distributed from CTAN
-% archives in directory macros/latex/base/lppl.txt.
-
-\newcommand{\slidecaption}{}
-
-\mode<presentation>
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% comic fonts fonts
-\DeclareFontFamily{T1}{comic}{}%
-\DeclareFontShape{T1}{comic}{m}{n}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{it}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sc}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{b}{n}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{b}{it}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sl}{<->ssub * comic/m/it}{}%
-\DeclareFontShape{T1}{comic}{b}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{b}{sl}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{n}{<->ssub * comic/b/n}{}%
-\DeclareFontShape{T1}{comic}{bx}{it}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{bx}{sl}{<->ssub * comic/b/it}{}%
-%
-\renewcommand{\rmdefault}{comic}%
-\renewcommand{\sfdefault}{comic}%
-\renewcommand{\mathfamilydefault}{cmr}% mathfont should be still the old one
-%
-\DeclareMathVersion{bold}% mathfont needs to be bold
-\DeclareSymbolFont{operators}{OT1}{cmr}{b}{n}%
-\SetSymbolFont{operators}{bold}{OT1}{cmr}{b}{n}%
-\DeclareSymbolFont{letters}{OML}{cmm}{b}{it}%
-\SetSymbolFont{letters}{bold}{OML}{cmm}{b}{it}%
-\DeclareSymbolFont{symbols}{OMS}{cmsy}{b}{n}%
-\SetSymbolFont{symbols}{bold}{OMS}{cmsy}{b}{n}%
-\DeclareSymbolFont{largesymbols}{OMX}{cmex}{b}{n}%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%                              % Title page  
-%\usetitlepagetemplate{
-%  \vbox{}
-%  \vfill
-%  \begin{centering}
-%    \Large\structure{\textrm{\textit{{\inserttitle}}}}
-%    \vskip1em\par
-%    \normalsize\insertauthor\vskip1em\par
-%    {\scriptsize\insertinstitute\par}\par\vskip1em
-%    \insertdate\par\vskip1.5em
-%    \inserttitlegraphic
-%  \end{centering}
-%  \vfill
-%  }
-
-                                % Part page  
-%\usepartpagetemplate{
-%  \begin{centering}
-%    \Large\structure{\textrm{\textit{\partname~\@Roman\c@part}}}
-%    \vskip1em\par
-%    \textrm{\textit{\insertpart}}\par
-%  \end{centering}
-%  }
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Frametitles
-\setbeamerfont{frametitle}{size={\huge}}
-\setbeamerfont{frametitle}{family={\usefont{T1}{ptm}{b}{n}}}
-\setbeamercolor{frametitle}{fg=gray,bg=white}
-
-\setbeamertemplate{frametitle}{%
-\vskip 2mm  % distance from the top margin
-\hskip -3mm % distance from left margin
-\vbox{%
-\begin{minipage}{1.05\textwidth}%
-\centering%
-\insertframetitle%
-\end{minipage}}%
-}
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Foot
-%
-\setbeamertemplate{navigation symbols}{} 
-\usefoottemplate{%
-\vbox{%
-  \tinyline{%
-    \tiny\hfill\textcolor{gray!50}{\slidecaption{} --
-      p.~\insertframenumber/\inserttotalframenumber}}}%
-}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\beamertemplateballitem
-\setlength\leftmargini{2mm}
-\setlength\leftmarginii{0.6cm}
-\setlength\leftmarginiii{1.5cm}
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% blocks
-%\definecolor{cream}{rgb}{1,1,.65}
-\definecolor{cream}{rgb}{1,1,.8}
-\setbeamerfont{block title}{size=\normalsize}
-\setbeamercolor{block title}{fg=black,bg=cream}
-\setbeamercolor{block body}{fg=black,bg=cream}
-
-\setbeamertemplate{blocks}[rounded][shadow=true]
-
-\setbeamercolor{boxcolor}{fg=black,bg=cream}
-
-\mode
-<all>
-
-
-
-
-
-
--- a/prio/Slides/document/mathpartir.sty	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,446 +0,0 @@
-%  Mathpartir --- Math Paragraph for Typesetting Inference Rules
-%
-%  Copyright (C) 2001, 2002, 2003, 2004, 2005 Didier Rémy
-%
-%  Author         : Didier Remy 
-%  Version        : 1.2.0
-%  Bug Reports    : to author
-%  Web Site       : http://pauillac.inria.fr/~remy/latex/
-% 
-%  Mathpartir is free software; you can redistribute it and/or modify
-%  it under the terms of the GNU General Public License as published by
-%  the Free Software Foundation; either version 2, or (at your option)
-%  any later version.
-%  
-%  Mathpartir is distributed in the hope that it will be useful,
-%  but WITHOUT ANY WARRANTY; without even the implied warranty of
-%  MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
-%  GNU General Public License for more details 
-%  (http://pauillac.inria.fr/~remy/license/GPL).
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%  File mathpartir.sty (LaTeX macros)
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\NeedsTeXFormat{LaTeX2e}
-\ProvidesPackage{mathpartir}
-    [2005/12/20 version 1.2.0 Math Paragraph for Typesetting Inference Rules]
-
-%%
-
-%% Identification
-%% Preliminary declarations
-
-\RequirePackage {keyval}
-
-%% Options
-%% More declarations
-
-%% PART I: Typesetting maths in paragraphe mode
-
-%% \newdimen \mpr@tmpdim
-%% Dimens are a precious ressource. Uses seems to be local.
-\let \mpr@tmpdim \@tempdima
-
-% To ensure hevea \hva compatibility, \hva should expands to nothing 
-% in mathpar or in inferrule
-\let \mpr@hva \empty
-
-%% normal paragraph parametters, should rather be taken dynamically
-\def \mpr@savepar {%
-  \edef \MathparNormalpar
-     {\noexpand \lineskiplimit \the\lineskiplimit
-      \noexpand \lineskip \the\lineskip}%
-  }
-
-\def \mpr@rulelineskip {\lineskiplimit=0.3em\lineskip=0.2em plus 0.1em}
-\def \mpr@lesslineskip {\lineskiplimit=0.6em\lineskip=0.5em plus 0.2em}
-\def \mpr@lineskip  {\lineskiplimit=1.2em\lineskip=1.2em plus 0.2em}
-\let \MathparLineskip \mpr@lineskip
-\def \mpr@paroptions {\MathparLineskip}
-\let \mpr@prebindings \relax
-
-\newskip \mpr@andskip \mpr@andskip 2em plus 0.5fil minus 0.5em
-
-\def \mpr@goodbreakand
-   {\hskip -\mpr@andskip  \penalty -1000\hskip \mpr@andskip}
-\def \mpr@and {\hskip \mpr@andskip}
-\def \mpr@andcr {\penalty 50\mpr@and}
-\def \mpr@cr {\penalty -10000\mpr@and}
-\def \mpr@eqno #1{\mpr@andcr #1\hskip 0em plus -1fil \penalty 10}
-
-\def \mpr@bindings {%
-  \let \and \mpr@andcr
-  \let \par \mpr@andcr
-  \let \\\mpr@cr
-  \let \eqno \mpr@eqno
-  \let \hva \mpr@hva
-  } 
-\let \MathparBindings \mpr@bindings
-
-% \@ifundefined {ignorespacesafterend}
-%    {\def \ignorespacesafterend {\aftergroup \ignorespaces}
-
-\newenvironment{mathpar}[1][]
-  {$$\mpr@savepar \parskip 0em \hsize \linewidth \centering
-     \vbox \bgroup \mpr@prebindings \mpr@paroptions #1\ifmmode $\else
-     \noindent $\displaystyle\fi
-     \MathparBindings}
-  {\unskip \ifmmode $\fi\egroup $$\ignorespacesafterend}
-
-\newenvironment{mathparpagebreakable}[1][]
-  {\begingroup 
-   \par
-   \mpr@savepar \parskip 0em \hsize \linewidth \centering
-      \mpr@prebindings \mpr@paroptions #1%
-      \vskip \abovedisplayskip \vskip -\lineskip%
-     \ifmmode  \else  $\displaystyle\fi
-     \MathparBindings
-  }
-  {\unskip
-   \ifmmode $\fi \par\endgroup
-   \vskip \belowdisplayskip
-   \noindent
-  \ignorespacesafterend}
-
-% \def \math@mathpar #1{\setbox0 \hbox {$\displaystyle #1$}\ifnum
-%     \wd0 < \hsize  $$\box0$$\else \bmathpar #1\emathpar \fi}
-
-%%% HOV BOXES
-
-\def \mathvbox@ #1{\hbox \bgroup \mpr@normallineskip 
-  \vbox \bgroup \tabskip 0em \let \\ \cr
-  \halign \bgroup \hfil $##$\hfil\cr #1\crcr \egroup \egroup
-  \egroup}
-
-\def \mathhvbox@ #1{\setbox0 \hbox {\let \\\qquad $#1$}\ifnum \wd0 < \hsize
-      \box0\else \mathvbox {#1}\fi}
-
-
-%% Part II -- operations on lists
-
-\newtoks \mpr@lista
-\newtoks \mpr@listb
-
-\long \def\mpr@cons #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
-{#2}\edef #2{\the \mpr@lista \the \mpr@listb}}
-
-\long \def\mpr@snoc #1\mpr@to#2{\mpr@lista {\\{#1}}\mpr@listb \expandafter
-{#2}\edef #2{\the \mpr@listb\the\mpr@lista}}
-
-\long \def \mpr@concat#1=#2\mpr@to#3{\mpr@lista \expandafter {#2}\mpr@listb
-\expandafter {#3}\edef #1{\the \mpr@listb\the\mpr@lista}}
-
-\def \mpr@head #1\mpr@to #2{\expandafter \mpr@head@ #1\mpr@head@ #1#2}
-\long \def \mpr@head@ #1#2\mpr@head@ #3#4{\def #4{#1}\def#3{#2}}
-
-\def \mpr@flatten #1\mpr@to #2{\expandafter \mpr@flatten@ #1\mpr@flatten@ #1#2}
-\long \def \mpr@flatten@ \\#1\\#2\mpr@flatten@ #3#4{\def #4{#1}\def #3{\\#2}}
-
-\def \mpr@makelist #1\mpr@to #2{\def \mpr@all {#1}%
-   \mpr@lista {\\}\mpr@listb \expandafter {\mpr@all}\edef \mpr@all {\the
-   \mpr@lista \the \mpr@listb \the \mpr@lista}\let #2\empty 
-   \def \mpr@stripof ##1##2\mpr@stripend{\def \mpr@stripped{##2}}\loop
-     \mpr@flatten \mpr@all \mpr@to \mpr@one
-     \expandafter \mpr@snoc \mpr@one \mpr@to #2\expandafter \mpr@stripof
-     \mpr@all \mpr@stripend  
-     \ifx \mpr@stripped \empty \let \mpr@isempty 0\else \let \mpr@isempty 1\fi
-     \ifx 1\mpr@isempty
-   \repeat
-}
-
-\def \mpr@rev #1\mpr@to #2{\let \mpr@tmp \empty
-   \def \\##1{\mpr@cons ##1\mpr@to \mpr@tmp}#1\let #2\mpr@tmp}
-
-%% Part III -- Type inference rules
-
-\newif \if@premisse
-\newbox \mpr@hlist
-\newbox \mpr@vlist
-\newif \ifmpr@center \mpr@centertrue
-\def \mpr@htovlist {%
-   \setbox \mpr@hlist
-      \hbox {\strut
-             \ifmpr@center \hskip -0.5\wd\mpr@hlist\fi
-             \unhbox \mpr@hlist}%
-   \setbox \mpr@vlist
-      \vbox {\if@premisse  \box \mpr@hlist \unvbox \mpr@vlist
-             \else \unvbox \mpr@vlist \box \mpr@hlist
-             \fi}%
-}
-% OLD version
-% \def \mpr@htovlist {%
-%    \setbox \mpr@hlist
-%       \hbox {\strut \hskip -0.5\wd\mpr@hlist \unhbox \mpr@hlist}%
-%    \setbox \mpr@vlist
-%       \vbox {\if@premisse  \box \mpr@hlist \unvbox \mpr@vlist
-%              \else \unvbox \mpr@vlist \box \mpr@hlist
-%              \fi}%
-% }
-
-\def \mpr@item #1{$\displaystyle #1$}
-\def \mpr@sep{2em}
-\def \mpr@blank { }
-\def \mpr@hovbox #1#2{\hbox
-  \bgroup
-  \ifx #1T\@premissetrue
-  \else \ifx #1B\@premissefalse
-  \else
-     \PackageError{mathpartir}
-       {Premisse orientation should either be T or B}
-       {Fatal error in Package}%
-  \fi \fi
-  \def \@test {#2}\ifx \@test \mpr@blank\else
-  \setbox \mpr@hlist \hbox {}%
-  \setbox \mpr@vlist \vbox {}%
-  \if@premisse \let \snoc \mpr@cons \else \let \snoc \mpr@snoc \fi
-  \let \@hvlist \empty \let \@rev \empty
-  \mpr@tmpdim 0em
-  \expandafter \mpr@makelist #2\mpr@to \mpr@flat
-  \if@premisse \mpr@rev \mpr@flat \mpr@to \@rev \else \let \@rev \mpr@flat \fi
-  \def \\##1{%
-     \def \@test {##1}\ifx \@test \empty
-        \mpr@htovlist
-        \mpr@tmpdim 0em %%% last bug fix not extensively checked
-     \else
-      \setbox0 \hbox{\mpr@item {##1}}\relax
-      \advance \mpr@tmpdim by \wd0
-      %\mpr@tmpdim 1.02\mpr@tmpdim
-      \ifnum \mpr@tmpdim < \hsize
-         \ifnum \wd\mpr@hlist > 0
-           \if@premisse
-             \setbox \mpr@hlist 
-                \hbox {\unhbox0 \hskip \mpr@sep \unhbox \mpr@hlist}%
-           \else
-             \setbox \mpr@hlist
-                \hbox {\unhbox \mpr@hlist  \hskip \mpr@sep \unhbox0}%
-           \fi
-         \else 
-         \setbox \mpr@hlist \hbox {\unhbox0}%
-         \fi
-      \else
-         \ifnum \wd \mpr@hlist > 0
-            \mpr@htovlist 
-            \mpr@tmpdim \wd0
-         \fi
-         \setbox \mpr@hlist \hbox {\unhbox0}%
-      \fi
-      \advance \mpr@tmpdim by \mpr@sep
-   \fi
-   }%
-   \@rev
-   \mpr@htovlist
-   \ifmpr@center \hskip \wd\mpr@vlist\fi \box \mpr@vlist
-   \fi
-   \egroup
-}
-
-%%% INFERENCE RULES
-
-\@ifundefined{@@over}{%
-    \let\@@over\over % fallback if amsmath is not loaded
-    \let\@@overwithdelims\overwithdelims
-    \let\@@atop\atop \let\@@atopwithdelims\atopwithdelims
-    \let\@@above\above \let\@@abovewithdelims\abovewithdelims
-  }{}
-
-%% The default
-
-\def \mpr@@fraction #1#2{\hbox {\advance \hsize by -0.5em
-    $\displaystyle {#1\mpr@over #2}$}}
-\def \mpr@@nofraction #1#2{\hbox {\advance \hsize by -0.5em
-    $\displaystyle {#1\@@atop #2}$}}
-
-\let \mpr@fraction \mpr@@fraction
-
-%% A generic solution to arrow
-
-\def \mpr@make@fraction #1#2#3#4#5{\hbox {%
-     \def \mpr@tail{#1}%
-     \def \mpr@body{#2}%
-     \def \mpr@head{#3}%
-     \setbox1=\hbox{$#4$}\setbox2=\hbox{$#5$}%
-     \setbox3=\hbox{$\mkern -3mu\mpr@body\mkern -3mu$}%
-     \setbox3=\hbox{$\mkern -3mu \mpr@body\mkern -3mu$}%
-     \dimen0=\dp1\advance\dimen0 by \ht3\relax\dp1\dimen0\relax
-     \dimen0=\ht2\advance\dimen0 by \dp3\relax\ht2\dimen0\relax
-     \setbox0=\hbox {$\box1 \@@atop \box2$}%
-     \dimen0=\wd0\box0
-     \box0 \hskip -\dimen0\relax
-     \hbox to \dimen0 {$%
-       \mathrel{\mpr@tail}\joinrel
-       \xleaders\hbox{\copy3}\hfil\joinrel\mathrel{\mpr@head}%
-     $}}}
-
-%% Old stuff should be removed in next version
-\def \mpr@@nothing #1#2
-    {$\lower 0.01pt \mpr@@nofraction {#1}{#2}$}
-\def \mpr@@reduce #1#2{\hbox
-    {$\lower 0.01pt \mpr@@fraction {#1}{#2}\mkern -15mu\rightarrow$}}
-\def \mpr@@rewrite #1#2#3{\hbox
-    {$\lower 0.01pt \mpr@@fraction {#2}{#3}\mkern -8mu#1$}}
-\def \mpr@infercenter #1{\vcenter {\mpr@hovbox{T}{#1}}}
-
-\def \mpr@empty {}
-\def \mpr@inferrule
-  {\bgroup
-     \ifnum \linewidth<\hsize \hsize \linewidth\fi
-     \mpr@rulelineskip
-     \let \and \qquad
-     \let \hva \mpr@hva
-     \let \@rulename \mpr@empty
-     \let \@rule@options \mpr@empty
-     \let \mpr@over \@@over
-     \mpr@inferrule@}
-\newcommand {\mpr@inferrule@}[3][]
-  {\everymath={\displaystyle}%       
-   \def \@test {#2}\ifx \empty \@test
-      \setbox0 \hbox {$\vcenter {\mpr@hovbox{B}{#3}}$}%
-   \else 
-   \def \@test {#3}\ifx \empty \@test
-      \setbox0 \hbox {$\vcenter {\mpr@hovbox{T}{#2}}$}%
-   \else
-   \setbox0 \mpr@fraction {\mpr@hovbox{T}{#2}}{\mpr@hovbox{B}{#3}}%
-   \fi \fi
-   \def \@test {#1}\ifx \@test\empty \box0
-   \else \vbox 
-%%% Suggestion de Francois pour les etiquettes longues
-%%%   {\hbox to \wd0 {\RefTirName {#1}\hfil}\box0}\fi
-      {\hbox {\RefTirName {#1}}\box0}\fi
-   \egroup}
-
-\def \mpr@vdotfil #1{\vbox to #1{\leaders \hbox{$\cdot$} \vfil}}
-
-% They are two forms
-% \inferrule [label]{[premisses}{conclusions}
-% or
-% \inferrule* [options]{[premisses}{conclusions}
-%
-% Premisses and conclusions are lists of elements separated by \\
-% Each \\ produces a break, attempting horizontal breaks if possible, 
-% and  vertical breaks if needed. 
-% 
-% An empty element obtained by \\\\ produces a vertical break in all cases. 
-%
-% The former rule is aligned on the fraction bar. 
-% The optional label appears on top of the rule
-% The second form to be used in a derivation tree is aligned on the last
-% line of its conclusion
-% 
-% The second form can be parameterized, using the key=val interface. The
-% folloiwng keys are recognized:
-%       
-%  width                set the width of the rule to val
-%  narrower             set the width of the rule to val\hsize
-%  before               execute val at the beginning/left
-%  lab                  put a label [Val] on top of the rule
-%  lskip                add negative skip on the right
-%  left                 put a left label [Val]
-%  Left                 put a left label [Val],  ignoring its width 
-%  right                put a right label [Val]
-%  Right                put a right label [Val], ignoring its width
-%  leftskip             skip negative space on the left-hand side
-%  rightskip            skip negative space on the right-hand side
-%  vdots                lift the rule by val and fill vertical space with dots
-%  after                execute val at the end/right
-%  
-%  Note that most options must come in this order to avoid strange
-%  typesetting (in particular  leftskip must preceed left and Left and
-%  rightskip must follow Right or right; vdots must come last 
-%  or be only followed by rightskip. 
-%  
-
-%% Keys that make sence in all kinds of rules
-\def \mprset #1{\setkeys{mprset}{#1}}
-\define@key {mprset}{andskip}[]{\mpr@andskip=#1}
-\define@key {mprset}{lineskip}[]{\lineskip=#1}
-\define@key {mprset}{flushleft}[]{\mpr@centerfalse}
-\define@key {mprset}{center}[]{\mpr@centertrue}
-\define@key {mprset}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
-\define@key {mprset}{atop}[]{\let \mpr@fraction \mpr@@nofraction}
-\define@key {mprset}{myfraction}[]{\let \mpr@fraction #1}
-\define@key {mprset}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
-\define@key {mprset}{sep}{\def\mpr@sep{#1}}
-
-\newbox \mpr@right
-\define@key {mpr}{flushleft}[]{\mpr@centerfalse}
-\define@key {mpr}{center}[]{\mpr@centertrue}
-\define@key {mpr}{rewrite}[]{\let \mpr@fraction \mpr@@rewrite}
-\define@key {mpr}{myfraction}[]{\let \mpr@fraction #1}
-\define@key {mpr}{fraction}[]{\def \mpr@fraction {\mpr@make@fraction #1}}
-\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
-     \advance \hsize by -\wd0\box0}
-\define@key {mpr}{width}{\hsize #1}
-\define@key {mpr}{sep}{\def\mpr@sep{#1}}
-\define@key {mpr}{before}{#1}
-\define@key {mpr}{lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
-\define@key {mpr}{Lab}{\let \RefTirName \TirName \def \mpr@rulename {#1}}
-\define@key {mpr}{narrower}{\hsize #1\hsize}
-\define@key {mpr}{leftskip}{\hskip -#1}
-\define@key {mpr}{reduce}[]{\let \mpr@fraction \mpr@@reduce}
-\define@key {mpr}{rightskip}
-  {\setbox \mpr@right \hbox {\unhbox \mpr@right \hskip -#1}}
-\define@key {mpr}{LEFT}{\setbox0 \hbox {$#1$}\relax
-     \advance \hsize by -\wd0\box0}
-\define@key {mpr}{left}{\setbox0 \hbox {$\TirName {#1}\;$}\relax
-     \advance \hsize by -\wd0\box0}
-\define@key {mpr}{Left}{\llap{$\TirName {#1}\;$}}
-\define@key {mpr}{right}
-  {\setbox0 \hbox {$\;\TirName {#1}$}\relax \advance \hsize by -\wd0
-   \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
-\define@key {mpr}{RIGHT}
-  {\setbox0 \hbox {$#1$}\relax \advance \hsize by -\wd0
-   \setbox \mpr@right \hbox {\unhbox \mpr@right \unhbox0}}
-\define@key {mpr}{Right}
-  {\setbox \mpr@right \hbox {\unhbox \mpr@right \rlap {$\;\TirName {#1}$}}}
-\define@key {mpr}{vdots}{\def \mpr@vdots {\@@atop \mpr@vdotfil{#1}}}
-\define@key {mpr}{after}{\edef \mpr@after {\mpr@after #1}}
-
-\newcommand \mpr@inferstar@ [3][]{\setbox0
-  \hbox {\let \mpr@rulename \mpr@empty \let \mpr@vdots \relax
-         \setbox \mpr@right \hbox{}%
-         $\setkeys{mpr}{#1}%
-          \ifx \mpr@rulename \mpr@empty \mpr@inferrule {#2}{#3}\else
-          \mpr@inferrule [{\mpr@rulename}]{#2}{#3}\fi
-          \box \mpr@right \mpr@vdots$}
-  \setbox1 \hbox {\strut}
-  \@tempdima \dp0 \advance \@tempdima by -\dp1
-  \raise \@tempdima \box0}
-
-\def \mpr@infer {\@ifnextchar *{\mpr@inferstar}{\mpr@inferrule}}
-\newcommand \mpr@err@skipargs[3][]{}
-\def \mpr@inferstar*{\ifmmode 
-    \let \@do \mpr@inferstar@
-  \else 
-    \let \@do \mpr@err@skipargs
-    \PackageError {mathpartir}
-      {\string\inferrule* can only be used in math mode}{}%
-  \fi \@do}
-
-
-%%% Exports
-
-% Envirnonment mathpar
-
-\let \inferrule \mpr@infer
-
-% make a short name \infer is not already defined
-\@ifundefined {infer}{\let \infer \mpr@infer}{}
-
-\def \TirNameStyle #1{\small \textsc{#1}}
-\def \tir@name #1{\hbox {\small \TirNameStyle{#1}}}
-\let \TirName \tir@name
-\let \DefTirName \TirName
-\let \RefTirName \TirName
-
-%%% Other Exports
-
-% \let \listcons \mpr@cons
-% \let \listsnoc \mpr@snoc
-% \let \listhead \mpr@head
-% \let \listmake \mpr@makelist
-
-
-
-
-\endinput
--- a/prio/Slides/document/root.beamer.tex	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-\documentclass[14pt,t]{beamer}
-%%%\usepackage{pstricks}
-
-\input{root.tex}
-
-%%% Local Variables:  
-%%% mode: latex
-%%% TeX-master: t
-%%% TeX-command-default: "Slides"
-%%% TeX-view-style: (("." "kghostview --landscape --scale 0.45 --geometry 605x505 %f"))
-%%% End: 
-
--- a/prio/Slides/document/root.notes.tex	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-\documentclass[11pt]{article}
-%%\usepackage{pstricks}
-\usepackage{dina4}
-\usepackage{beamerarticle}
-\usepackage{times}
-\usepackage{hyperref}
-\usepackage{pgf}
-\usepackage{amssymb}
-\setjobnamebeamerversion{root.beamer}
-\input{root.tex}
-
-%%% Local Variables:  
-%%% mode: latex
-%%% TeX-master: t
-%%% TeX-command-default: "Slides"
-%%% TeX-view-style: (("." "kghostview --landscape --scale 0.45 --geometry 605x505 %f"))
-%%% End: 
-
--- a/prio/Slides/document/root.tex	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,147 +0,0 @@
-\usepackage{beamerthemeplaincu}
-%%\usepackage{ulem}
-\usepackage[T1]{fontenc}
-\usepackage{proof}
-\usepackage[latin1]{inputenc}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{mathpartir}
-\usepackage[absolute, overlay]{textpos}
-\usepackage{proof}
-\usepackage{ifthen}
-\usepackage{animate}
-\usepackage{tikz}
-\usepackage{pgf}
-\usetikzlibrary{arrows}
-\usetikzlibrary{automata}
-\usetikzlibrary{shapes}
-\usetikzlibrary{shadows}
-\usetikzlibrary{calc}
-
-% Isabelle configuration
-%%\urlstyle{rm}
-\isabellestyle{rm}
-\renewcommand{\isastyle}{\rm}%
-\renewcommand{\isastyleminor}{\rm}%
-\renewcommand{\isastylescript}{\footnotesize\rm\slshape}%
-\renewcommand{\isatagproof}{}
-\renewcommand{\endisatagproof}{}
-\renewcommand{\isamarkupcmt}[1]{#1}
-
-% Isabelle characters
-\renewcommand{\isacharunderscore}{\_}
-\renewcommand{\isacharbar}{\isamath{\mid}}
-\renewcommand{\isasymiota}{}
-\renewcommand{\isacharbraceleft}{\{}
-\renewcommand{\isacharbraceright}{\}}
-\renewcommand{\isacharless}{$\langle$}
-\renewcommand{\isachargreater}{$\rangle$}
-\renewcommand{\isasymsharp}{\isamath{\#}}
-\renewcommand{\isasymdots}{\isamath{...}}
-\renewcommand{\isasymbullet}{\act}
-
-% mathpatir
-\mprset{sep=1em}
-
-% general math stuff
-\newcommand{\dn}{\stackrel{\mbox{\scriptsize def}}{=}}% for definitions
-\newcommand{\dnn}{\stackrel{\mbox{\Large def}}{=}}
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\emptyset}{\varnothing}% nice round empty set
-\renewcommand{\Gamma}{\varGamma} 
-\DeclareRobustCommand{\flqq}{\mbox{\guillemotleft}}
-\DeclareRobustCommand{\frqq}{\mbox{\guillemotright}}
-\newcommand{\smath}[1]{\textcolor{blue}{\ensuremath{#1}}}
-\newcommand{\fresh}{\mathrel{\#}}
-\newcommand{\act}{{\raisebox{-0.5mm}{\Large$\boldsymbol{\cdot}$}}}% swapping action
-\newcommand{\swap}[2]{(#1\,#2)}% swapping operation
-
-% beamer stuff 
-\renewcommand{\slidecaption}{Salvador, 26.~August 2008}
-
-
-% colours for Isar Code (in article mode everything is black and white)
-\mode<presentation>{
-\definecolor{isacol:brown}{rgb}{.823,.411,.117}
-\definecolor{isacol:lightblue}{rgb}{.274,.509,.705}
-\definecolor{isacol:green}{rgb}{0,.51,0.14}
-\definecolor{isacol:red}{rgb}{.803,0,0}
-\definecolor{isacol:blue}{rgb}{0,0,.803}
-\definecolor{isacol:darkred}{rgb}{.545,0,0}
-\definecolor{isacol:black}{rgb}{0,0,0}}
-\mode<article>{
-\definecolor{isacol:brown}{rgb}{0,0,0}
-\definecolor{isacol:lightblue}{rgb}{0,0,0}
-\definecolor{isacol:green}{rgb}{0,0,0}
-\definecolor{isacol:red}{rgb}{0,0,0}
-\definecolor{isacol:blue}{rgb}{0,0,0}
-\definecolor{isacol:darkred}{rgb}{0,0,0}
-\definecolor{isacol:black}{rgb}{0,0,0}
-}
-
-
-\newcommand{\strong}[1]{{\bfseries {#1}}}
-\newcommand{\bluecmd}[1]{{\color{isacol:lightblue}{\strong{#1}}}}
-\newcommand{\browncmd}[1]{{\color{isacol:brown}{\strong{#1}}}}
-\newcommand{\redcmd}[1]{{\color{isacol:red}{\strong{#1}}}}
-
-\renewcommand{\isakeyword}[1]{%
-\ifthenelse{\equal{#1}{show}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{case}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{assume}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{obtain}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{fix}}{\browncmd{#1}}{%
-\ifthenelse{\equal{#1}{oops}}{\redcmd{#1}}{%
-\ifthenelse{\equal{#1}{thm}}{\redcmd{#1}}{%
-{\bluecmd{#1}}}}}}}}}}%
-
-% inner syntax colour
-\chardef\isachardoublequoteopen=`\"%
-\chardef\isachardoublequoteclose=`\"%
-\chardef\isacharbackquoteopen=`\`%
-\chardef\isacharbackquoteclose=`\`%
-\newenvironment{innersingle}%
-{\isacharbackquoteopen\color{isacol:green}}%
-{\color{isacol:black}\isacharbackquoteclose}
-\newenvironment{innerdouble}%
-{\isachardoublequoteopen\color{isacol:green}}%
-{\color{isacol:black}\isachardoublequoteclose}
-
-%% misc
-\newcommand{\gb}[1]{\textcolor{isacol:green}{#1}}
-\newcommand{\rb}[1]{\textcolor{red}{#1}}
-
-%% animations
-\newcounter{growcnt}
-\newcommand{\grow}[2]
-{\begin{tikzpicture}[baseline=(n.base)]%
-    \node[scale=(0.1 *#1 + 0.001),inner sep=0pt] (n) {#2};
-  \end{tikzpicture}%
-}
-
-%% isatabbing
-%\renewcommand{\isamarkupcmt}[1]%
-%{\ifthenelse{\equal{TABSET}{#1}}{\=}%
-% {\ifthenelse{\equal{TAB}{#1}}{\>}%
-%  {\ifthenelse{\equal{NEWLINE}{#1}}{\\}%
-%   {\ifthenelse{\equal{DOTS}{#1}}{\ldots}{\isastylecmt--- {#1}}}%
-%  }%
-% }%
-%}%
-
-
-\newenvironment{isatabbing}%
-{\renewcommand{\isanewline}{\\}\begin{tabbing}}%
-{\end{tabbing}}
-
-\begin{document}
-\input{session}
-\end{document}
-
-%%% Local Variables:  
-%%% mode: latex
-%%% TeX-master: t
-%%% TeX-command-default: "Slides"
-%%% TeX-view-style: (("." "kghostview --landscape --scale 0.45 --geometry 605x505 %f"))
-%%% End: 
-
Binary file prio/Slides/slides.pdf has changed
--- a/prio/document/beamerthemeplaincu.sty	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,126 +0,0 @@
-\ProvidesPackage{beamerthemeplaincu}[2003/11/07 ver 0.93]
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-
-% Copyright 2003 by Till Tantau <tantau@cs.tu-berlin.de>.
-%
-% This program can be redistributed and/or modified under the terms
-% of the LaTeX Project Public License Distributed from CTAN
-% archives in directory macros/latex/base/lppl.txt.
-
-\newcommand{\slidecaption}{}
-
-\mode<presentation>
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% comic fonts fonts
-\DeclareFontFamily{T1}{comic}{}%
-\DeclareFontShape{T1}{comic}{m}{n}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{it}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sc}{<->s*[.9]comic8t}{}%
-\DeclareFontShape{T1}{comic}{b}{n}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{b}{it}{<->s*[.9]comicbd8t}{}%
-\DeclareFontShape{T1}{comic}{m}{sl}{<->ssub * comic/m/it}{}%
-\DeclareFontShape{T1}{comic}{b}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{b}{sl}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{n}{<->ssub * comic/b/n}{}%
-\DeclareFontShape{T1}{comic}{bx}{it}{<->ssub * comic/b/it}{}%
-\DeclareFontShape{T1}{comic}{bx}{sc}{<->sub * comic/m/sc}{}%
-\DeclareFontShape{T1}{comic}{bx}{sl}{<->ssub * comic/b/it}{}%
-%
-\renewcommand{\rmdefault}{comic}%
-\renewcommand{\sfdefault}{comic}%
-\renewcommand{\mathfamilydefault}{cmr}% mathfont should be still the old one
-%
-\DeclareMathVersion{bold}% mathfont needs to be bold
-\DeclareSymbolFont{operators}{OT1}{cmr}{b}{n}%
-\SetSymbolFont{operators}{bold}{OT1}{cmr}{b}{n}%
-\DeclareSymbolFont{letters}{OML}{cmm}{b}{it}%
-\SetSymbolFont{letters}{bold}{OML}{cmm}{b}{it}%
-\DeclareSymbolFont{symbols}{OMS}{cmsy}{b}{n}%
-\SetSymbolFont{symbols}{bold}{OMS}{cmsy}{b}{n}%
-\DeclareSymbolFont{largesymbols}{OMX}{cmex}{b}{n}%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%                              % Title page  
-%\usetitlepagetemplate{
-%  \vbox{}
-%  \vfill
-%  \begin{centering}
-%    \Large\structure{\textrm{\textit{{\inserttitle}}}}
-%    \vskip1em\par
-%    \normalsize\insertauthor\vskip1em\par
-%    {\scriptsize\insertinstitute\par}\par\vskip1em
-%    \insertdate\par\vskip1.5em
-%    \inserttitlegraphic
-%  \end{centering}
-%  \vfill
-%  }
-
-                                % Part page  
-%\usepartpagetemplate{
-%  \begin{centering}
-%    \Large\structure{\textrm{\textit{\partname~\@Roman\c@part}}}
-%    \vskip1em\par
-%    \textrm{\textit{\insertpart}}\par
-%  \end{centering}
-%  }
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Frametitles
-\setbeamerfont{frametitle}{size={\huge}}
-\setbeamerfont{frametitle}{family={\usefont{T1}{ptm}{b}{n}}}
-\setbeamercolor{frametitle}{fg=gray,bg=white}
-
-\setbeamertemplate{frametitle}{%
-\vskip 2mm  % distance from the top margin
-\hskip -3mm % distance from left margin
-\vbox{%
-\begin{minipage}{1.05\textwidth}%
-\centering%
-\insertframetitle%
-\end{minipage}}%
-}
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Foot
-%
-\setbeamertemplate{navigation symbols}{} 
-\usefoottemplate{%
-\vbox{%
-  \tinyline{%
-    \tiny\hfill\textcolor{gray!50}{\slidecaption{} --
-      p.~\insertframenumber/\inserttotalframenumber}}}%
-}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\beamertemplateballitem
-\setlength\leftmargini{2mm}
-\setlength\leftmarginii{0.6cm}
-\setlength\leftmarginiii{1.5cm}
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% blocks
-%\definecolor{cream}{rgb}{1,1,.65}
-\definecolor{cream}{rgb}{1,1,.8}
-\setbeamerfont{block title}{size=\normalsize}
-\setbeamercolor{block title}{fg=black,bg=cream}
-\setbeamercolor{block body}{fg=black,bg=cream}
-
-\setbeamertemplate{blocks}[rounded][shadow=true]
-
-\setbeamercolor{boxcolor}{fg=black,bg=cream}
-
-\mode
-<all>
-
-
-
-
-
-
--- a/prio/document/llncs.cls	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1189 +0,0 @@
-% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
-% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
-%
-%%
-%% \CharacterTable
-%%  {Upper-case    \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
-%%   Lower-case    \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
-%%   Digits        \0\1\2\3\4\5\6\7\8\9
-%%   Exclamation   \!     Double quote  \"     Hash (number) \#
-%%   Dollar        \$     Percent       \%     Ampersand     \&
-%%   Acute accent  \'     Left paren    \(     Right paren   \)
-%%   Asterisk      \*     Plus          \+     Comma         \,
-%%   Minus         \-     Point         \.     Solidus       \/
-%%   Colon         \:     Semicolon     \;     Less than     \<
-%%   Equals        \=     Greater than  \>     Question mark \?
-%%   Commercial at \@     Left bracket  \[     Backslash     \\
-%%   Right bracket \]     Circumflex    \^     Underscore    \_
-%%   Grave accent  \`     Left brace    \{     Vertical bar  \|
-%%   Right brace   \}     Tilde         \~}
-%%
-\NeedsTeXFormat{LaTeX2e}[1995/12/01]
-\ProvidesClass{llncs}[2002/01/28 v2.13
-^^J LaTeX document class for Lecture Notes in Computer Science]
-% Options
-\let\if@envcntreset\iffalse
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
-\DeclareOption{oribibl}{\let\oribibl=Y}
-\let\if@custvec\iftrue
-\DeclareOption{orivec}{\let\if@custvec\iffalse}
-\let\if@envcntsame\iffalse
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\let\if@envcntsect\iffalse
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\let\if@runhead\iffalse
-\DeclareOption{runningheads}{\let\if@runhead\iftrue}
-
-\let\if@openbib\iffalse
-\DeclareOption{openbib}{\let\if@openbib\iftrue}
-
-% languages
-\let\switcht@@therlang\relax
-\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
-\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
-
-\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
-
-\ProcessOptions
-
-\LoadClass[twoside]{article}
-\RequirePackage{multicol} % needed for the list of participants, index
-
-\setlength{\textwidth}{12.2cm}
-\setlength{\textheight}{19.3cm}
-\renewcommand\@pnumwidth{2em}
-\renewcommand\@tocrmarg{3.5em}
-%
-\def\@dottedtocline#1#2#3#4#5{%
-  \ifnum #1>\c@tocdepth \else
-    \vskip \z@ \@plus.2\p@
-    {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
-               \parfillskip -\rightskip \pretolerance=10000
-     \parindent #2\relax\@afterindenttrue
-     \interlinepenalty\@M
-     \leavevmode
-     \@tempdima #3\relax
-     \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
-     {#4}\nobreak
-     \leaders\hbox{$\m@th
-        \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
-        mu$}\hfill
-     \nobreak
-     \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
-     \par}%
-  \fi}
-%
-\def\switcht@albion{%
-\def\abstractname{Abstract.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{\unskip, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Key words:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}}
-\switcht@albion
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\switcht@francais{%\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{Preuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-
-\renewcommand\small{%
-   \@setfontsize\small\@ixpt{11}%
-   \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
-   \abovedisplayshortskip \z@ \@plus2\p@
-   \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
-   \def\@listi{\leftmargin\leftmargini
-               \parsep 0\p@ \@plus1\p@ \@minus\p@
-               \topsep 8\p@ \@plus2\p@ \@minus4\p@
-               \itemsep0\p@}%
-   \belowdisplayskip \abovedisplayskip
-}
-
-\frenchspacing
-\widowpenalty=10000
-\clubpenalty=10000
-
-\setlength\oddsidemargin   {63\p@}
-\setlength\evensidemargin  {63\p@}
-\setlength\marginparwidth  {90\p@}
-
-\setlength\headsep   {16\p@}
-
-\setlength\footnotesep{7.7\p@}
-\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
-\setlength\intextsep   {8mm\@plus 2\p@ \@minus 2\p@}
-
-\setcounter{secnumdepth}{2}
-
-\newcounter {chapter}
-\renewcommand\thechapter      {\@arabic\c@chapter}
-
-\newif\if@mainmatter \@mainmattertrue
-\newcommand\frontmatter{\cleardoublepage
-            \@mainmatterfalse\pagenumbering{Roman}}
-\newcommand\mainmatter{\cleardoublepage
-       \@mainmattertrue\pagenumbering{arabic}}
-\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
-      \@mainmatterfalse}
-
-\renewcommand\part{\cleardoublepage
-                 \thispagestyle{empty}%
-                 \if@twocolumn
-                     \onecolumn
-                     \@tempswatrue
-                   \else
-                     \@tempswafalse
-                 \fi
-                 \null\vfil
-                 \secdef\@part\@spart}
-
-\def\@part[#1]#2{%
-    \ifnum \c@secnumdepth >-2\relax
-      \refstepcounter{part}%
-      \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
-    \else
-      \addcontentsline{toc}{part}{#1}%
-    \fi
-    \markboth{}{}%
-    {\centering
-     \interlinepenalty \@M
-     \normalfont
-     \ifnum \c@secnumdepth >-2\relax
-       \huge\bfseries \partname~\thepart
-       \par
-       \vskip 20\p@
-     \fi
-     \Huge \bfseries #2\par}%
-    \@endpart}
-\def\@spart#1{%
-    {\centering
-     \interlinepenalty \@M
-     \normalfont
-     \Huge \bfseries #1\par}%
-    \@endpart}
-\def\@endpart{\vfil\newpage
-              \if@twoside
-                \null
-                \thispagestyle{empty}%
-                \newpage
-              \fi
-              \if@tempswa
-                \twocolumn
-              \fi}
-
-\newcommand\chapter{\clearpage
-                    \thispagestyle{empty}%
-                    \global\@topnum\z@
-                    \@afterindentfalse
-                    \secdef\@chapter\@schapter}
-\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
-                       \if@mainmatter
-                         \refstepcounter{chapter}%
-                         \typeout{\@chapapp\space\thechapter.}%
-                         \addcontentsline{toc}{chapter}%
-                                  {\protect\numberline{\thechapter}#1}%
-                       \else
-                         \addcontentsline{toc}{chapter}{#1}%
-                       \fi
-                    \else
-                      \addcontentsline{toc}{chapter}{#1}%
-                    \fi
-                    \chaptermark{#1}%
-                    \addtocontents{lof}{\protect\addvspace{10\p@}}%
-                    \addtocontents{lot}{\protect\addvspace{10\p@}}%
-                    \if@twocolumn
-                      \@topnewpage[\@makechapterhead{#2}]%
-                    \else
-                      \@makechapterhead{#2}%
-                      \@afterheading
-                    \fi}
-\def\@makechapterhead#1{%
-% \vspace*{50\p@}%
-  {\centering
-    \ifnum \c@secnumdepth >\m@ne
-      \if@mainmatter
-        \large\bfseries \@chapapp{} \thechapter
-        \par\nobreak
-        \vskip 20\p@
-      \fi
-    \fi
-    \interlinepenalty\@M
-    \Large \bfseries #1\par\nobreak
-    \vskip 40\p@
-  }}
-\def\@schapter#1{\if@twocolumn
-                   \@topnewpage[\@makeschapterhead{#1}]%
-                 \else
-                   \@makeschapterhead{#1}%
-                   \@afterheading
-                 \fi}
-\def\@makeschapterhead#1{%
-% \vspace*{50\p@}%
-  {\centering
-    \normalfont
-    \interlinepenalty\@M
-    \Large \bfseries  #1\par\nobreak
-    \vskip 40\p@
-  }}
-
-\renewcommand\section{\@startsection{section}{1}{\z@}%
-                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {12\p@ \@plus 4\p@ \@minus 4\p@}%
-                       {\normalfont\large\bfseries\boldmath
-                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
-                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {8\p@ \@plus 4\p@ \@minus 4\p@}%
-                       {\normalfont\normalsize\bfseries\boldmath
-                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
-\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
-                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {-0.5em \@plus -0.22em \@minus -0.1em}%
-                       {\normalfont\normalsize\bfseries\boldmath}}
-\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
-                       {-12\p@ \@plus -4\p@ \@minus -4\p@}%
-                       {-0.5em \@plus -0.22em \@minus -0.1em}%
-                       {\normalfont\normalsize\itshape}}
-\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
-                  \string\subparagraph\space with this class}\vskip0.5cm
-You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
-
-\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
-\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
-\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
-\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
-\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
-\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
-\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
-\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
-\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
-\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
-\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
-
-\let\footnotesize\small
-
-\if@custvec
-\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
-{\mbox{\boldmath$\textstyle#1$}}
-{\mbox{\boldmath$\scriptstyle#1$}}
-{\mbox{\boldmath$\scriptscriptstyle#1$}}}
-\fi
-
-\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
-\penalty50\hskip1em\null\nobreak\hfil\squareforqed
-\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
-
-\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
-\cr\to\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-\gets\cr\to\cr}}}}}
-\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
-<\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
-\noalign{\vskip1.2pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
-\noalign{\vskip1pt}=\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr
-\noalign{\vskip0.9pt}=\cr}}}}}
-\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
-\halign{\hfil
-$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
->\cr\noalign{\vskip-1pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.8pt}<\cr}}}
-{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
->\cr\noalign{\vskip-0.3pt}<\cr}}}}}
-\def\bbbr{{\rm I\!R}} %reelle Zahlen
-\def\bbbm{{\rm I\!M}}
-\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
-\def\bbbf{{\rm I\!F}}
-\def\bbbh{{\rm I\!H}}
-\def\bbbk{{\rm I\!K}}
-\def\bbbp{{\rm I\!P}}
-\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
-{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
-\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
-to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
-0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
-\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
-T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
-to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
-\def\bbbs{{\mathchoice
-{\setbox0=\hbox{$\displaystyle     \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\textstyle        \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
-to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptstyle      \rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
-{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
-to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
-to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
-\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
-{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
-{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
-
-\let\ts\,
-
-\setlength\leftmargini  {17\p@}
-\setlength\leftmargin    {\leftmargini}
-\setlength\leftmarginii  {\leftmargini}
-\setlength\leftmarginiii {\leftmargini}
-\setlength\leftmarginiv  {\leftmargini}
-\setlength  \labelsep  {.5em}
-\setlength  \labelwidth{\leftmargini}
-\addtolength\labelwidth{-\labelsep}
-
-\def\@listI{\leftmargin\leftmargini
-            \parsep 0\p@ \@plus1\p@ \@minus\p@
-            \topsep 8\p@ \@plus2\p@ \@minus4\p@
-            \itemsep0\p@}
-\let\@listi\@listI
-\@listi
-\def\@listii {\leftmargin\leftmarginii
-              \labelwidth\leftmarginii
-              \advance\labelwidth-\labelsep
-              \topsep    0\p@ \@plus2\p@ \@minus\p@}
-\def\@listiii{\leftmargin\leftmarginiii
-              \labelwidth\leftmarginiii
-              \advance\labelwidth-\labelsep
-              \topsep    0\p@ \@plus\p@\@minus\p@
-              \parsep    \z@
-              \partopsep \p@ \@plus\z@ \@minus\p@}
-
-\renewcommand\labelitemi{\normalfont\bfseries --}
-\renewcommand\labelitemii{$\m@th\bullet$}
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
-                                                    {{\contentsname}}}
- \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
- \def\lastand{\ifnum\value{auco}=2\relax
-                 \unskip{} \andname\
-              \else
-                 \unskip \lastandname\
-              \fi}%
- \def\and{\stepcounter{@auth}\relax
-          \ifnum\value{@auth}=\value{auco}%
-             \lastand
-          \else
-             \unskip,
-          \fi}%
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
-   \addvspace{2em plus\p@}%  % space above part line
-   \begingroup
-     \parindent \z@
-     \rightskip \z@ plus 5em
-     \hrule\vskip5pt
-     \large               % same size as for a contribution heading
-     \bfseries\boldmath   % set line in boldface
-     \leavevmode          % TeX command to enter horizontal mode.
-     #1\par
-     \vskip5pt
-     \hrule
-     \vskip1pt
-     \nobreak             % Never break after part entry
-   \endgroup}
-
-\def\@dotsep{2}
-
-\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
-{chapter.\thechapter}\fi}
-
-\def\addnumcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
-                     {\thechapter}#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmark#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
-\def\addcontentsmarkwop#1#2#3{%
-\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
-
-\def\@adcmk[#1]{\ifcase #1 \or
-\def\@gtempa{\addnumcontentsmark}%
-  \or    \def\@gtempa{\addcontentsmark}%
-  \or    \def\@gtempa{\addcontentsmarkwop}%
-  \fi\@gtempa{toc}{chapter}}
-\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
- \else
-      \nobreak
-      \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
-      \@dotsep mu$}\hfill
-      \nobreak\hbox to\@pnumwidth{\hss #2}%
- \fi\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@title#1#2{\addpenalty{-\@highpenalty}
- \addvspace{8pt plus 1pt}
- \@tempdima \z@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \parfillskip -\rightskip \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- #1\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\def\l@author#1#2{\addpenalty{\@highpenalty}
- \@tempdima=\z@ %15\p@
- \begingroup
- \parindent \z@ \rightskip \@tocrmarg
- \advance\rightskip by 0pt plus 2cm
- \pretolerance=10000
- \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
- \textit{#1}\par
- \penalty\@highpenalty \endgroup}
-
-%\setcounter{tocdepth}{0}
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=\z@            % no chapter numbers
-\tocsecnum=15\p@          % section 88. plus 2.222pt
-\tocsubsecnum=23\p@       % subsection 88.8 plus 2.222pt
-\tocsubsubsecnum=27\p@    % subsubsection 88.8.8 plus 1.444pt
-\tocparanum=35\p@         % paragraph 88.8.8.8 plus 1.666pt
-\tocsubparanum=43\p@      % subparagraph 88.8.8.8.8 plus 1.888pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-\renewcommand\listoffigures{%
-    \section*{\listfigurename
-      \@mkboth{\listfigurename}{\listfigurename}}%
-    \@starttoc{lof}%
-    }
-
-\renewcommand\listoftables{%
-    \section*{\listtablename
-      \@mkboth{\listtablename}{\listtablename}}%
-    \@starttoc{lot}%
-    }
-
-\ifx\oribibl\undefined
-\ifx\citeauthoryear\undefined
-\renewenvironment{thebibliography}[1]
-     {\section*{\refname}
-      \def\@biblabel##1{##1.}
-      \small
-      \list{\@biblabel{\@arabic\c@enumiv}}%
-           {\settowidth\labelwidth{\@biblabel{#1}}%
-            \leftmargin\labelwidth
-            \advance\leftmargin\labelsep
-            \if@openbib
-              \advance\leftmargin\bibindent
-              \itemindent -\bibindent
-              \listparindent \itemindent
-              \parsep \z@
-            \fi
-            \usecounter{enumiv}%
-            \let\p@enumiv\@empty
-            \renewcommand\theenumiv{\@arabic\c@enumiv}}%
-      \if@openbib
-        \renewcommand\newblock{\par}%
-      \else
-        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
-      \fi
-      \sloppy\clubpenalty4000\widowpenalty4000%
-      \sfcode`\.=\@m}
-     {\def\@noitemerr
-       {\@latex@warning{Empty `thebibliography' environment}}%
-      \endlist}
-\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
-     {\let\protect\noexpand\immediate
-     \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-\newcount\@tempcntc
-\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
-  \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
-    {\@ifundefined
-       {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
-        ?}\@warning
-       {Citation `\@citeb' on page \thepage \space undefined}}%
-    {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
-     \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
-       \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
-     \else
-      \advance\@tempcntb\@ne
-      \ifnum\@tempcntb=\@tempcntc
-      \else\advance\@tempcntb\m@ne\@citeo
-      \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
-\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
-               \@citea\def\@citea{,\,\hskip\z@skip}%
-               \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
-               {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
-                \def\@citea{--}\fi
-      \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
-\else
-\renewenvironment{thebibliography}[1]
-     {\section*{\refname}
-      \small
-      \list{}%
-           {\settowidth\labelwidth{}%
-            \leftmargin\parindent
-            \itemindent=-\parindent
-            \labelsep=\z@
-            \if@openbib
-              \advance\leftmargin\bibindent
-              \itemindent -\bibindent
-              \listparindent \itemindent
-              \parsep \z@
-            \fi
-            \usecounter{enumiv}%
-            \let\p@enumiv\@empty
-            \renewcommand\theenumiv{}}%
-      \if@openbib
-        \renewcommand\newblock{\par}%
-      \else
-        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
-      \fi
-      \sloppy\clubpenalty4000\widowpenalty4000%
-      \sfcode`\.=\@m}
-     {\def\@noitemerr
-       {\@latex@warning{Empty `thebibliography' environment}}%
-      \endlist}
-      \def\@cite#1{#1}%
-      \def\@lbibitem[#1]#2{\item[]\if@filesw
-        {\def\protect##1{\string ##1\space}\immediate
-      \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
-   \fi
-\else
-\@cons\@openbib@code{\noexpand\small}
-\fi
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
-                \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
-                \noindent\hangindent\wd0\box0}% order index entry
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-\renewenvironment{theindex}
-               {\@mkboth{\indexname}{\indexname}%
-                \thispagestyle{empty}\parindent\z@
-                \parskip\z@ \@plus .3\p@\relax
-                \let\item\par
-                \def\,{\relax\ifmmode\mskip\thinmuskip
-                             \else\hskip0.2em\ignorespaces\fi}%
-                \normalfont\small
-                \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
-                }
-                {\end{multicols}}
-
-\renewcommand\footnoterule{%
-  \kern-3\p@
-  \hrule\@width 2truecm
-  \kern2.6\p@}
-  \newdimen\fnindent
-  \fnindent1em
-\long\def\@makefntext#1{%
-    \parindent \fnindent%
-    \leftskip \fnindent%
-    \noindent
-    \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
-
-\long\def\@makecaption#1#2{%
-  \vskip\abovecaptionskip
-  \sbox\@tempboxa{{\bfseries #1.} #2}%
-  \ifdim \wd\@tempboxa >\hsize
-    {\bfseries #1.} #2\par
-  \else
-    \global \@minipagefalse
-    \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
-  \fi
-  \vskip\belowcaptionskip}
-
-\def\fps@figure{htbp}
-\def\fnum@figure{\figurename\thinspace\thefigure}
-\def \@floatboxreset {%
-        \reset@font
-        \small
-        \@setnobreak
-        \@setminipage
-}
-\def\fps@table{htbp}
-\def\fnum@table{\tablename~\thetable}
-\renewenvironment{table}
-               {\setlength\abovecaptionskip{0\p@}%
-                \setlength\belowcaptionskip{10\p@}%
-                \@float{table}}
-               {\end@float}
-\renewenvironment{table*}
-               {\setlength\abovecaptionskip{0\p@}%
-                \setlength\belowcaptionskip{10\p@}%
-                \@dblfloat{table}}
-               {\end@dblfloat}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
-  ext@#1\endcsname}{#1}{\protect\numberline{\csname
-  the#1\endcsname}{\ignorespaces #2}}\begingroup
-    \@parboxrestore
-    \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
-  \endgroup}
-
-% LaTeX does not provide a command to enter the authors institute
-% addresses. The \institute command is defined here.
-
-\newcounter{@inst}
-\newcounter{@auth}
-\newcounter{auco}
-\newdimen\instindent
-\newbox\authrun
-\newtoks\authorrunning
-\newtoks\tocauthor
-\newbox\titrun
-\newtoks\titlerunning
-\newtoks\toctitle
-
-\def\clearheadinfo{\gdef\@author{No Author Given}%
-                   \gdef\@title{No Title Given}%
-                   \gdef\@subtitle{}%
-                   \gdef\@institute{No Institute Given}%
-                   \gdef\@thanks{}%
-                   \global\titlerunning={}\global\authorrunning={}%
-                   \global\toctitle={}\global\tocauthor={}}
-
-\def\institute#1{\gdef\@institute{#1}}
-
-\def\institutename{\par
- \begingroup
- \parskip=\z@
- \parindent=\z@
- \setcounter{@inst}{1}%
- \def\and{\par\stepcounter{@inst}%
- \noindent$^{\the@inst}$\enspace\ignorespaces}%
- \setbox0=\vbox{\def\thanks##1{}\@institute}%
- \ifnum\c@@inst=1\relax
-   \gdef\fnnstart{0}%
- \else
-   \xdef\fnnstart{\c@@inst}%
-   \setcounter{@inst}{1}%
-   \noindent$^{\the@inst}$\enspace
- \fi
- \ignorespaces
- \@institute\par
- \endgroup}
-
-\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
-   {\star\star\star}\or \dagger\or \ddagger\or
-   \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
-   \or \ddagger\ddagger \else\@ctrerr\fi}}
-
-\def\inst#1{\unskip$^{#1}$}
-\def\fnmsep{\unskip$^,$}
-\def\email#1{{\tt#1}}
-\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
-\@ifpackageloaded{babel}{%
-\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
-\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
-\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
-}{\switcht@@therlang}%
-}
-\def\homedir{\~{ }}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\clearheadinfo
-
-\renewcommand\maketitle{\newpage
-  \refstepcounter{chapter}%
-  \stepcounter{section}%
-  \setcounter{section}{0}%
-  \setcounter{subsection}{0}%
-  \setcounter{figure}{0}
-  \setcounter{table}{0}
-  \setcounter{equation}{0}
-  \setcounter{footnote}{0}%
-  \begingroup
-    \parindent=\z@
-    \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
-    \if@twocolumn
-      \ifnum \col@number=\@ne
-        \@maketitle
-      \else
-        \twocolumn[\@maketitle]%
-      \fi
-    \else
-      \newpage
-      \global\@topnum\z@   % Prevents figures from going at top of page.
-      \@maketitle
-    \fi
-    \thispagestyle{empty}\@thanks
-%
-    \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
-    \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
-    \instindent=\hsize
-    \advance\instindent by-\headlineindent
-%    \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
-%       \addcontentsline{toc}{title}{\the\toctitle}\fi
-    \if@runhead
-       \if!\the\titlerunning!\else
-         \edef\@title{\the\titlerunning}%
-       \fi
-       \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
-       \ifdim\wd\titrun>\instindent
-          \typeout{Title too long for running head. Please supply}%
-          \typeout{a shorter form with \string\titlerunning\space prior to
-                   \string\maketitle}%
-          \global\setbox\titrun=\hbox{\small\rm
-          Title Suppressed Due to Excessive Length}%
-       \fi
-       \xdef\@title{\copy\titrun}%
-    \fi
-%
-    \if!\the\tocauthor!\relax
-      {\def\and{\noexpand\protect\noexpand\and}%
-      \protected@xdef\toc@uthor{\@author}}%
-    \else
-      \def\\{\noexpand\protect\noexpand\newline}%
-      \protected@xdef\scratch{\the\tocauthor}%
-      \protected@xdef\toc@uthor{\scratch}%
-    \fi
-%    \addcontentsline{toc}{author}{\toc@uthor}%
-    \if@runhead
-       \if!\the\authorrunning!
-         \value{@inst}=\value{@auth}%
-         \setcounter{@auth}{1}%
-       \else
-         \edef\@author{\the\authorrunning}%
-       \fi
-       \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
-       \ifdim\wd\authrun>\instindent
-          \typeout{Names of authors too long for running head. Please supply}%
-          \typeout{a shorter form with \string\authorrunning\space prior to
-                   \string\maketitle}%
-          \global\setbox\authrun=\hbox{\small\rm
-          Authors Suppressed Due to Excessive Length}%
-       \fi
-       \xdef\@author{\copy\authrun}%
-       \markboth{\@author}{\@title}%
-     \fi
-  \endgroup
-  \setcounter{footnote}{\fnnstart}%
-  \clearheadinfo}
-%
-\def\@maketitle{\newpage
- \markboth{}{}%
- \def\lastand{\ifnum\value{@inst}=2\relax
-                 \unskip{} \andname\
-              \else
-                 \unskip \lastandname\
-              \fi}%
- \def\and{\stepcounter{@auth}\relax
-          \ifnum\value{@auth}=\value{@inst}%
-             \lastand
-          \else
-             \unskip,
-          \fi}%
- \begin{center}%
- \let\newline\\
- {\Large \bfseries\boldmath
-  \pretolerance=10000
-  \@title \par}\vskip .8cm
-\if!\@subtitle!\else {\large \bfseries\boldmath
-  \vskip -.65cm
-  \pretolerance=10000
-  \@subtitle \par}\vskip .8cm\fi
- \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
- \def\thanks##1{}\@author}%
- \global\value{@inst}=\value{@auth}%
- \global\value{auco}=\value{@auth}%
- \setcounter{@auth}{1}%
-{\lineskip .5em
-\noindent\ignorespaces
-\@author\vskip.35cm}
- {\small\institutename}
- \end{center}%
- }
-
-% definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-%     \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or  \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or  \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\@thmcountersep{}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
-  \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
-   {\@definecounter{#1}\@addtoreset{#1}{#3}%
-   \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
-     \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
-   \expandafter\xdef\csname #1name\endcsname{#2}%
-   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
-                              \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
-   {\@definecounter{#1}%
-   \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
-   \expandafter\xdef\csname #1name\endcsname{#2}%
-   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
-                               \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
-  \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
-  {\expandafter\@ifdefinable\csname #1\endcsname
-  {\global\@namedef{the#1}{\@nameuse{the#2}}%
-  \expandafter\xdef\csname #1name\endcsname{#3}%
-  \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
-  \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
-                    \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
-       the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
-                 \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
-      \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
-   {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
-    \expandafter\xdef\csname #1name\endcsname{#2}%
-    \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
-       {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
-                           \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
-      \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-\if@envcntsect
-   \def\@thmcountersep{.}
-   \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
-\else
-   \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
-   \if@envcntreset
-      \@addtoreset{theorem}{section}
-   \else
-      \@addtoreset{theorem}{chapter}
-   \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % alle Umgebungen wie Theorem.
-   \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % alle Umgebungen mit eigenem Zaehler
-   \if@envcntsect % mit section numeriert
-      \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
-   \else % nicht mit section numeriert
-      \if@envcntreset
-         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
-                                   \@addtoreset{#1}{section}}
-      \else
-         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
-                                   \@addtoreset{#1}{chapter}}%
-      \fi
-   \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
-    \def\@tempa{#1}%
-    \let\@tempd\@elt
-    \def\@elt##1{%
-        \def\@tempb{##1}%
-        \ifx\@tempa\@tempb\else
-            \@addtoreset{##1}{#2}%
-        \fi}%
-    \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
-    \expandafter\def\csname cl@#2\endcsname{}%
-    \@tempc
-    \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
-      \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
-                  \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
-      \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
-      }
-
-\renewenvironment{abstract}{%
-      \list{}{\advance\topsep by0.35cm\relax\small
-      \leftmargin=1cm
-      \labelwidth=\z@
-      \listparindent=\z@
-      \itemindent\listparindent
-      \rightmargin\leftmargin}\item[\hskip\labelsep
-                                    \bfseries\abstractname]}
-    {\endlist}
-
-\newdimen\headlineindent             % dimension for space between
-\headlineindent=1.166cm              % number and text of headings.
-
-\def\ps@headings{\let\@mkboth\@gobbletwo
-   \let\@oddfoot\@empty\let\@evenfoot\@empty
-   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
-                  \leftmark\hfil}
-   \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
-                 \llap{\thepage}}
-   \def\chaptermark##1{}%
-   \def\sectionmark##1{}%
-   \def\subsectionmark##1{}}
-
-\def\ps@titlepage{\let\@mkboth\@gobbletwo
-   \let\@oddfoot\@empty\let\@evenfoot\@empty
-   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
-                  \hfil}
-   \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
-                 \llap{\thepage}}
-   \def\chaptermark##1{}%
-   \def\sectionmark##1{}%
-   \def\subsectionmark##1{}}
-
-\if@runhead\ps@headings\else
-\ps@empty\fi
-
-\setlength\arraycolsep{1.4\p@}
-\setlength\tabcolsep{1.4\p@}
-
-\endinput
-%end of file llncs.cls
--- a/prio/document/root.bib	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,111 +0,0 @@
-@article{OwensReppyTuron09,
-  author = {S.~Owens and J.~Reppy and A.~Turon},
-  title = {{R}egular-{E}xpression {D}erivatives {R}e-{E}xamined},
-  journal = {Journal of Functional Programming},
-  volume = 19,
-  number = {2},
-  year = 2009,
-  pages = {173--190}
-}
-
-
-
-@Unpublished{KraussNipkow11,
-  author = 	 {A.~Kraus and T.~Nipkow},
-  title = 	 {{P}roof {P}earl: {R}egular {E}xpression {E}quivalence and {R}elation {A}lgebra},
-  note = 	 {To appear in Journal of Automated Reasoning},
-  year = 	 {2011}
-}
-
-@Book{Kozen97,
-  author = 	 {D.~Kozen},
-  title = 	 {{A}utomata and {C}omputability},
-  publisher = 	 {Springer Verlag},
-  year = 	 {1997}
-}
-
-
-@incollection{Constable00,
-  author    = {R.~L.~Constable and
-               P.~B.~Jackson and
-               P.~Naumov and
-               J.~C.~Uribe},
-  title     = {{C}onstructively {F}ormalizing {A}utomata {T}heory},
-  booktitle = {Proof, Language, and Interaction},
-  year      = {2000},
-  publisher = {MIT Press},
-  pages     = {213-238}
-}
-
-
-@techreport{Filliatre97,
-  author = {J.-C. Filli\^atre},
-  institution = {LIP - ENS Lyon},
-  number = {97--04},
-  title = {{F}inite {A}utomata {T}heory in {C}oq: 
-           {A} {C}onstructive {P}roof of {K}leene's {T}heorem},
-  type = {Research Report},
-  year = {1997}
-}
-
-@article{OwensSlind08,
-  author    = {S.~Owens and K.~Slind},
-  title     = {{A}dapting {F}unctional {P}rograms to {H}igher {O}rder {L}ogic},
-  journal   = {Higher-Order and Symbolic Computation},
-  volume    = {21},
-  number    = {4},
-  year      = {2008},
-  pages     = {377--409}
-}
-
-@article{Brzozowski64,
- author = {J.~A.~Brzozowski},
- title = {{D}erivatives of {R}egular {E}xpressions},
- journal = {J.~ACM},
- volume = {11},
- issue = {4},
- year = {1964},
- pages = {481--494},
- publisher = {ACM}
-} 
-
-@inproceedings{Nipkow98,
- author={T.~Nipkow},
- title={{V}erified {L}exical {A}nalysis},
- booktitle={Proc.~of the 11th International Conference on Theorem Proving in Higher Order Logics},
- series={LNCS},
- volume=1479,
- pages={1--15},
- year=1998
-}
-
-@inproceedings{BerghoferNipkow00,
-  author={S.~Berghofer and T.~Nipkow},
-  title={{E}xecuting {H}igher {O}rder {L}ogic},
-  booktitle={Proc.~of the International Workshop on Types for Proofs and Programs},
-  year=2002,
-  series={LNCS},
-  volume=2277,
-  pages="24--40"
-}
-
-@book{HopcroftUllman69,
-  author    = {J.~E.~Hopcroft and
-               J.~D.~Ullman},
-  title     = {{F}ormal {L}anguages and {T}heir {R}elation to {A}utomata},
-  publisher = {Addison-Wesley},
-  year      = {1969}
-}
-
-
-@inproceedings{BerghoferReiter09,
-  author    = {S.~Berghofer and
-               M.~Reiter},
-  title     = {{F}ormalizing the {L}ogic-{A}utomaton {C}onnection},
-  booktitle = {Proc.~of the 22nd International
-               Conference on Theorem Proving in Higher Order Logics},
-  year      = {2009},
-  pages     = {147-163},
-  series    = {LNCS},
-  volume    = {5674}
-}
\ No newline at end of file
--- a/prio/document/root.tex	Mon Dec 03 08:16:58 2012 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,73 +0,0 @@
-\documentclass[runningheads]{llncs}
-\usepackage{isabelle}
-\usepackage{isabellesym}
-\usepackage{amsmath}
-\usepackage{amssymb}
-\usepackage{tikz}
-\usepackage{pgf}
-\usetikzlibrary{arrows,automata,decorations,fit,calc}
-\usetikzlibrary{shapes,shapes.arrows,snakes,positioning}
-\usepgflibrary{shapes.misc} % LATEX and plain TEX and pure pgf
-\usetikzlibrary{matrix}
-\usepackage{pdfsetup}
-\usepackage{ot1patch}
-\usepackage{times}
-%%\usepackage{proof}
-%%\usepackage{mathabx}
-\usepackage{stmaryrd}
-
-\titlerunning{Proving the Priority Inheritance Protocol Correct}
-
-
-\urlstyle{rm}
-\isabellestyle{it}
-\renewcommand{\isastyleminor}{\it}%
-\renewcommand{\isastyle}{\normalsize\it}%
-
-
-\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
-\renewcommand{\isasymequiv}{$\dn$}
-\renewcommand{\isasymemptyset}{$\varnothing$}
-\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
-
-\newcommand{\isasymcalL}{\ensuremath{\cal{L}}}
-\newcommand{\isasymbigplus}{\ensuremath{\bigplus}}
-
-\newcommand{\bigplus}{\mbox{\Large\bf$+$}}
-\begin{document}
-
-\title{A Formalisation of the Myhill-Nerode Theorem\\ based on Regular
-  Expressions (Proof Pearl)}
-\author{Chunhan Wu\inst{1} \and Xingyuan Zhang\inst{1} \and Christian Urban\inst{2}}
-\institute{PLA University of Science and Technology, China \and TU Munich, Germany}
-\maketitle
-
-%\mbox{}\\[-10mm]
-\begin{abstract} 
-There are numerous textbooks on regular languages. Nearly all of them
-introduce the subject by describing finite automata and only mentioning on the
-side a connection with regular expressions. Unfortunately, automata are difficult
-to formalise in HOL-based theorem provers. The reason is that
-they need to be represented as graphs, matrices or functions, none of which
-are inductive datatypes. Also convenient operations for disjoint unions of
-graphs and functions are not easily formalisiable in HOL. In contrast, regular
-expressions can be defined conveniently as a datatype and a corresponding
-reasoning infrastructure comes for free. We show in this paper that a central
-result from formal language theory---the Myhill-Nerode theorem---can be
-recreated using only regular expressions.
-
-\end{abstract}
-
-
-\input{session}
-
-%%\mbox{}\\[-10mm]
-\bibliographystyle{plain}
-\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
Binary file prio/paper.pdf has changed