added a bit more text to the paper and separated a theory about Max
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Thu, 22 May 2014 17:40:39 +0100
changeset 35 92f61f6a0fe7
parent 34 313acffe63b6
child 36 af38526275f8
added a bit more text to the paper and separated a theory about Max
CpsG.thy
ExtGG.thy
Journal/Paper.thy
Journal/document/root.tex
Max.thy
Moment.thy
PrioG.thy
PrioGDef.thy
README
journal.pdf
--- a/CpsG.thy	Tue May 20 12:49:21 2014 +0100
+++ b/CpsG.thy	Thu May 22 17:40:39 2014 +0100
@@ -1,6 +1,6 @@
 
 theory CpsG
-imports PrioG 
+imports PrioG Max
 begin
 
 lemma not_thread_holdents:
@@ -23,7 +23,7 @@
         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)
+        by (simp add:RAG_create_unchanged)
       moreover have "th \<notin> threads s" 
       proof -
         from not_in eq_e show ?thesis by simp
@@ -38,13 +38,13 @@
         case True
         with nh eq_e
         show ?thesis 
-          by (auto simp:holdents_test depend_exit_unchanged)
+          by (auto simp:holdents_test RAG_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)
+          by (auto simp:holdents_test RAG_exit_unchanged)
       qed
     next
       case (thread_P thread cs)
@@ -60,7 +60,7 @@
       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)
+        by (unfold step_RAG_p[OF vtp], auto)
       moreover have "holdents s th = {}"
       proof(rule ih)
         from not_in eq_e show "th \<notin> threads s" by simp
@@ -102,7 +102,7 @@
       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)
+        by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
       moreover have "holdents s th = {}"
       proof(rule ih)
         from not_in eq_e show "th \<notin> threads s" by simp
@@ -117,12 +117,12 @@
       from ih [OF this] and eq_e
       show ?thesis 
         apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:depend_set_unchanged)
+        by (simp add:RAG_set_unchanged)
     qed
     next
       case vt_nil
       show ?case
-      by (auto simp:count_def holdents_test s_depend_def wq_def cs_holding_def)
+      by (auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
   qed
 qed
 
@@ -161,97 +161,33 @@
   shows "th1 = th2"
 using assms by (unfold next_th_def, auto)
 
-lemma wf_depend:
+lemma wf_RAG:
   assumes vt: "vt s"
-  shows "wf (depend s)"
+  shows "wf (RAG s)"
 proof(rule finite_acyclic_wf)
-  from finite_depend[OF vt] show "finite (depend s)" .
+  from finite_RAG[OF vt] show "finite (RAG s)" .
 next
-  from acyclic_depend[OF vt] show "acyclic (depend s)" .
+  from acyclic_RAG[OF vt] show "acyclic (RAG 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}"
+            {(Th th', Th th) | th th'. \<exists>cs. (Th th', Cs cs) \<in> RAG s \<and> (Cs cs, Th th) \<in> RAG 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}"
+  "children s th \<equiv> {th'. \<exists> cs. (Th th', Cs cs) \<in> RAG s \<and> (Cs cs, Th th) \<in> RAG s}"
 unfolding child_def children_def by simp
 
-lemma children_dependants: "children s th \<subseteq> dependants (wq s) th"
-  by (unfold children_def child_def cs_dependants_def, auto simp:eq_depend)
+lemma children_dependants: 
+  "children s th \<subseteq> dependants (wq s) th"
+  unfolding children_def2
+  unfolding cs_dependants_def
+  by (auto simp add: eq_RAG)
 
 lemma child_unique:
   assumes vt: "vt s"
@@ -261,92 +197,92 @@
 using ch1 ch2 
 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]
+  assume h1: "(Th th, Cs cs) \<in> RAG s"
+    and h2: "(Cs cs, Th th1) \<in> RAG s"
+    and h3: "(Th th, Cs csa) \<in> RAG s"
+    and h4: "(Cs csa, Th th2) \<in> RAG s"
+  from unique_RAG[OF vt h1 h3] have "cs = csa" by simp
+  with h4 have "(Cs cs, Th th2) \<in> RAG s" by simp
+  from unique_RAG[OF vt h2 this]
   show "th1 = th2" by simp
 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)^+)"
+lemma RAG_children:
+  assumes h: "(Th th1, Th th2) \<in> (RAG s)^+"
+  shows "th1 \<in> children s th2 \<or> (\<exists> th3. th3 \<in> children s th2 \<and> (Th th1, Th th3) \<in> (RAG 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"
+    assume h1: "(Th th1, c) \<in> (RAG s)\<^sup>+"
+    and h2: "(c, Th th2) \<in> RAG 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>+)"
+      by (case_tac c, auto simp:s_RAG_def)
+    show "th1 \<in> children s th2 \<or> (\<exists>th3. th3 \<in> children s th2 \<and> (Th th1, Th th3) \<in> (RAG 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>+)"
+      assume h3: "(Th th1, ca) \<in> (RAG s)\<^sup>+"
+        and h4: "(ca, c) \<in> RAG s"
+      show "th1 \<in> children s th2 \<or> (\<exists>th3. th3 \<in> children s th2 \<and> (Th th1, Th th3) \<in> (RAG 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)
+          by (case_tac ca, auto simp:s_RAG_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
+        moreover from h3 eq_ca have "(Th th1, Th th3) \<in> (RAG s)\<^sup>+" by simp
         ultimately show ?thesis by auto
       qed
     next
-      assume "(Th th1, c) \<in> depend s"
+      assume "(Th th1, c) \<in> RAG 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"
+    assume "(Th th1, Th th2) \<in> RAG s"
     thus ?thesis
-      by (auto simp:s_depend_def)
+      by (auto simp:s_RAG_def)
   qed
 qed
 
-lemma sub_child: "child s \<subseteq> (depend s)^+"
+lemma sub_child: "child s \<subseteq> (RAG s)^+"
   by (unfold child_def, auto)
 
 lemma wf_child: 
   assumes vt: "vt s"
   shows "wf (child s)"
 apply(rule wf_subset)
-apply(rule wf_trancl[OF wf_depend[OF vt]])
+apply(rule wf_trancl[OF wf_RAG[OF vt]])
 apply(rule sub_child)
 done
 
-lemma depend_child_pre:
+lemma RAG_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")
+  "(Th th, n) \<in> (RAG 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)^+)" .
+  from wf_trancl[OF wf_RAG[OF vt]]
+  have wf: "wf ((RAG 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>+"
+    assume ih[rule_format]: "\<forall>y. (y, Th th') \<in> (RAG s)\<^sup>+ \<longrightarrow>
+               (Th th, y) \<in> (RAG s)\<^sup>+ \<longrightarrow> (\<forall>th'. y = Th th' \<longrightarrow> (Th th, Th th') \<in> (child s)\<^sup>+)"
+    and h: "(Th th, Th th') \<in> (RAG 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>+)" .
+      from RAG_children[OF h]
+      have "th \<in> children s th' \<or> (\<exists>th3. th3 \<in> children s th' \<and> (Th th, Th th3) \<in> (RAG 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>+"
+        assume "\<exists>th3. th3 \<in> children s th' \<and> (Th th, Th th3) \<in> (RAG 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)
+          and th_dp: "(Th th, Th th3) \<in> (RAG s)\<^sup>+" by auto
+        from th3_in have "(Th th3, Th th') \<in> (RAG 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
@@ -354,12 +290,12 @@
   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 RAG_child: "\<lbrakk>vt s; (Th th, Th th') \<in> (RAG s)^+\<rbrakk> \<Longrightarrow> (Th th, Th th') \<in> (child s)^+"
+  by (insert RAG_child_pre, auto)
 
-lemma child_depend_p:
+lemma child_RAG_p:
   assumes "(n1, n2) \<in> (child s)^+"
-  shows "(n1, n2) \<in> (depend s)^+"
+  shows "(n1, n2) \<in> (RAG s)^+"
 proof -
   from assms show ?thesis
   proof(induct)
@@ -368,26 +304,26 @@
   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
+    with sub_child have "(y, z) \<in> (RAG s)^+" by auto
+    moreover have "(n1, y) \<in> (RAG s)^+" by fact
     ultimately show ?case by auto
   qed
 qed
 
-lemma child_depend_eq: 
+lemma child_RAG_eq: 
   assumes vt: "vt s"
-  shows "(Th th1, Th th2) \<in> (child s)^+  \<longleftrightarrow> (Th th1, Th th2) \<in> (depend s)^+"
-  by (auto intro: depend_child[OF vt] child_depend_p)
+  shows "(Th th1, Th th2) \<in> (child s)^+  \<longleftrightarrow> (Th th1, Th th2) \<in> (RAG s)^+"
+  by (auto intro: RAG_child[OF vt] child_RAG_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)^+"
+  and ch3: "(Th th1, Th th2) \<in> (RAG s)^+"
   shows "False"
 proof -
-  from depend_child[OF vt ch3]
+  from RAG_child[OF vt ch3]
   have "(Th th1, Th th2) \<in> (child s)\<^sup>+" .
   thus ?thesis
   proof(rule converse_tranclE)
@@ -412,15 +348,15 @@
   qed
 qed
 
-lemma unique_depend_p:
+lemma unique_RAG_p:
   assumes vt: "vt s"
-  and dp1: "(n, n1) \<in> (depend s)^+"
-  and dp2: "(n, n2) \<in> (depend s)^+"
+  and dp1: "(n, n1) \<in> (RAG s)^+"
+  and dp2: "(n, n2) \<in> (RAG s)^+"
   and neq: "n1 \<noteq> n2"
-  shows "(n1, n2) \<in> (depend s)^+ \<or> (n2, n1) \<in> (depend s)^+"
+  shows "(n1, n2) \<in> (RAG s)^+ \<or> (n2, n1) \<in> (RAG 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
+  from unique_RAG[OF vt]
+  show "\<And>a b c. \<lbrakk>(a, b) \<in> RAG s; (a, c) \<in> RAG s\<rbrakk> \<Longrightarrow> b = c" by auto
 qed
 
 lemma dependants_child_unique:
@@ -433,47 +369,41 @@
 shows "th1 = th2"
 proof -
   { assume neq: "th1 \<noteq> th2"
-    from dp1 have dp1: "(Th th3, Th th1) \<in> (depend s)^+" 
-      by (simp add:s_dependants_def eq_depend)
-    from dp2 have dp2: "(Th th3, Th th2) \<in> (depend s)^+" 
-      by (simp add:s_dependants_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
+    from dp1 have dp1: "(Th th3, Th th1) \<in> (RAG s)^+" 
+      by (simp add:s_dependants_def eq_RAG)
+    from dp2 have dp2: "(Th th3, Th th2) \<in> (RAG s)^+" 
+      by (simp add:s_dependants_def eq_RAG)
+    from unique_RAG_p[OF vt dp1 dp2] and neq
+    have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
     hence False
     proof
-      assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+ "
+      assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ "
       from children_no_dep[OF vt ch1 ch2 this] show ?thesis .
     next
-      assume " (Th th2, Th th1) \<in> (depend s)\<^sup>+"
+      assume " (Th th2, Th th1) \<in> (RAG s)\<^sup>+"
       from children_no_dep[OF vt ch2 ch1 this] show ?thesis .
     qed
   } thus ?thesis by auto
 qed
 
-lemma depend_plus_elim:
+lemma RAG_plus_elim:
   assumes "vt s"
   fixes x
-  assumes "(Th x, Th th) \<in> (depend (wq s))\<^sup>+"
-  shows "\<exists>th'\<in>children s th. x = th' \<or> (Th x, Th th') \<in> (depend (wq s))\<^sup>+"
-  using assms(2)[unfolded eq_depend, folded child_depend_eq[OF `vt s`]]
+  assumes "(Th x, Th th) \<in> (RAG (wq s))\<^sup>+"
+  shows "\<exists>th'\<in>children s th. x = th' \<or> (Th x, Th th') \<in> (RAG (wq s))\<^sup>+"
+  using assms(2)[unfolded eq_RAG, folded child_RAG_eq[OF `vt s`]]
   apply (unfold children_def)
-  by (metis assms(2) children_def depend_children eq_depend)
-
-lemma dependants_expand_pre: 
-  assumes "vt s"
-  shows "dependants (wq s) th = (\<Union> th' \<in> children s th. {th'} \<union> dependants (wq s) th')"
-  apply (unfold cs_dependants_def)
-  apply (auto elim!:depend_plus_elim[OF assms])
-  apply (metis children_def eq_depend mem_Collect_eq set_mp sub_child)
-  apply (unfold children_def, auto)
-  apply (unfold eq_depend, fold  child_depend_eq[OF `vt s`])
-  by (metis trancl.simps)
+  by (metis assms(2) children_def RAG_children eq_RAG)
 
 lemma dependants_expand:
   assumes "vt s"
-  shows "dependants (wq s) th = (\<Union> ((\<lambda> th. {th} \<union> dependants (wq s) th) ` (children s th)))"
-  apply (subst dependants_expand_pre[OF assms])
-  by simp
+  shows "dependants (wq s) th = (children s th) \<union> (\<Union>((dependants (wq s)) ` children s th))"
+apply(simp add: image_def)
+unfolding cs_dependants_def
+apply(auto)
+apply (metis assms RAG_plus_elim mem_Collect_eq)
+apply (metis child_RAG_p children_def eq_RAG mem_Collect_eq r_into_trancl')
+by (metis assms child_RAG_eq children_def eq_RAG mem_Collect_eq trancl.trancl_into_trancl)
 
 lemma finite_children:
   assumes "vt s"
@@ -487,23 +417,11 @@
   using dependants_threads[OF assms] finite_threads[OF assms]
   by (metis rev_finite_subset)
 
-lemma Max_insert:
-  assumes "finite B"
-  and  "B \<noteq> {}"
-  shows "Max ({x} \<union> B) = max x (Max B)"
-  by (metis Max_insert assms insert_is_Un)
-
-lemma dependands_expand2:
-  assumes "vt s"
-  shows "dependants (wq s) th = (children s th) \<union> (\<Union>((dependants (wq s)) ` children s th))"
-  by (subst dependants_expand[OF assms]) (auto)
+abbreviation
+  "preceds s ths \<equiv> {preced th s| th. th \<in> ths}"
 
 abbreviation
-  "preceds s Ths \<equiv> {preced th s| th. th \<in> Ths}"
-
-lemma image_compr:
-  "f ` A = {f x | x. x \<in> A}"
-by auto
+  "cpreceds s ths \<equiv> (cp s) ` ths"
 
 lemma Un_compr:
   "{f th | th. R th \<or> Q th} = ({f th | th. R th} \<union> {f th' | th'. Q th'})"
@@ -531,11 +449,11 @@
   show ?thesis (is "?LHS = ?RHS")
   proof -
     have eq_cp: "cp s = (\<lambda>th. Max (preceds s ({th} \<union> dependants (wq s) th)))"
-      by (simp add: fun_eq_iff cp_eq_cpreced cpreced_def Un_compr image_compr[symmetric])
+      by (simp add: fun_eq_iff cp_eq_cpreced cpreced_def Un_compr image_Collect[symmetric])
   
     have not_emptyness_facts[simp]: 
       "dependants (wq s) th \<noteq> {}" "children s th \<noteq> {}"
-      using False dependands_expand2[OF assms] by(auto simp only: Un_empty)
+      using False dependants_expand[OF assms] by(auto simp only: Un_empty)
 
     have finiteness_facts[simp]:
       "\<And>th. finite (dependants (wq s) th)" "\<And>th. finite (children s th)"
@@ -552,7 +470,7 @@
     (* expanding dependants *)
     also have "\<dots> = max (Max {preced th s}) 
       (Max (preceds s (children s th \<union> \<Union>(dependants (wq s) ` children s th))))"
-      by (subst dependands_expand2[OF `vt s`]) (simp)
+      by (subst dependants_expand[OF `vt s`]) (simp)
 
     (* moving out big Union *)
     also have "\<dots> = max (Max {preced th s})
@@ -602,8 +520,8 @@
     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_dep: "RAG s = RAG s'"
+  by (unfold s_def RAG_set_unchanged, auto)
 
 lemma eq_cp_pre:
   fixes th' 
@@ -613,7 +531,7 @@
   apply (unfold cp_eq_cpreced cpreced_def)
 proof -
   have eq_dp: "\<And> th. dependants (wq s) th = dependants (wq s') th"
-    by (unfold cs_dependants_def, auto simp:eq_dep eq_depend)
+    by (unfold cs_dependants_def, auto simp:eq_dep eq_RAG)
   moreover {
     fix th1 
     assume "th1 \<in> {th'} \<union> dependants (wq s') th'"
@@ -626,7 +544,7 @@
     next
       assume "th1 \<in> dependants (wq s') th'"
       with nd and eq_dp have "th1 \<noteq> th"
-        by (auto simp:eq_depend cs_dependants_def s_dependants_def eq_dep)
+        by (auto simp:eq_RAG cs_dependants_def s_dependants_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> dependants (wq s) th')) = 
@@ -646,12 +564,12 @@
   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_dependants_def, unfold eq_depend, unfold eq_dep, auto)
+  from h have "(Th th, Th th') \<in> (RAG s')\<^sup>+"
+    by (unfold s_dependants_def, unfold eq_RAG, unfold eq_dep, auto)
   from tranclD[OF this]
-  obtain z where "(Th th, z) \<in> depend s'" by auto
+  obtain z where "(Th th, z) \<in> RAG 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)
+    apply (case_tac z, auto simp:readys_def s_waiting_def s_RAG_def s_waiting_def cs_waiting_def)
     by (fold wq_def, blast)
 qed
 
@@ -673,8 +591,8 @@
   shows "cp s th'' = cp s' th''"
 proof -
   from dp2
-  have "(Th th', Th th'') \<in> (depend (wq s))\<^sup>+" by (simp add:s_dependants_def)
-  from depend_child[OF vt_s this[unfolded eq_depend]]
+  have "(Th th', Th th'') \<in> (RAG (wq s))\<^sup>+" by (simp add:s_dependants_def)
+  from RAG_child[OF vt_s this[unfolded eq_RAG]]
   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>
@@ -686,10 +604,10 @@
         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_dependants_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
+      from dp1 have "(Th th, Th th') \<in> (RAG s)^+" by (auto simp:s_dependants_def eq_RAG)
+      moreover from child_RAG_p[OF ch'] and eq_y
+      have "(Th th', Th thy) \<in> (RAG s)^+" by simp
+      ultimately have dp_thy: "(Th th, Th thy) \<in> (RAG s)^+" by auto
       show "cp s th'' = cp s' th''"
         apply (subst cp_rec[OF vt_s])
       proof -
@@ -699,9 +617,9 @@
           proof
             assume "th'' = th"
             with dp_thy y_ch[unfolded eq_y] 
-            have "(Th th, Th th) \<in> (depend s)^+"
+            have "(Th th, Th th) \<in> (RAG s)^+"
               by (auto simp:child_def)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
+            with wf_trancl[OF wf_RAG[OF vt_s]] 
             show False by auto
           qed
         qed
@@ -717,14 +635,14 @@
             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
+              with dp_thy have "(Th th1, Th thy) \<in> (RAG 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> dependants s th1"
             proof
               assume h:"th \<in> dependants s th1"
-              from eq_y dp_thy have "th \<in> dependants s thy" by (auto simp:s_dependants_def eq_depend)
+              from eq_y dp_thy have "th \<in> dependants s thy" by (auto simp:s_dependants_def eq_RAG)
               from dependants_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
@@ -736,7 +654,7 @@
         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)
+          by (unfold children_def child_def s_def RAG_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
@@ -752,9 +670,9 @@
           proof
             assume "th'' = th"
             with dp1 dp'
-            have "(Th th, Th th) \<in> (depend s)^+"
-              by (auto simp:child_def s_dependants_def eq_depend)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
+            have "(Th th, Th th) \<in> (RAG s)^+"
+              by (auto simp:child_def s_dependants_def eq_RAG)
+            with wf_trancl[OF wf_RAG[OF vt_s]] 
             show False by auto
           qed
         qed
@@ -770,8 +688,8 @@
             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_dependants_def eq_depend)
+              with dp1 have "(Th th1, Th th') \<in> (RAG s)^+" 
+                by (auto simp:s_dependants_def eq_RAG)
               from children_no_dep[OF vt_s _ _ this]
               th1_in dp'
               show False by (auto simp:children_def)
@@ -792,7 +710,7 @@
         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)
+          by (unfold children_def child_def s_def RAG_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     
@@ -808,8 +726,8 @@
   shows "cp s th'' = cp s' th''"
 proof -
   from dp
-  have "(Th th, Th th'') \<in> (depend (wq s))\<^sup>+" by (simp add:s_dependants_def)
-  from depend_child[OF vt_s this[unfolded eq_depend]]
+  have "(Th th, Th th'') \<in> (RAG (wq s))\<^sup>+" by (simp add:s_dependants_def)
+  from RAG_child[OF vt_s this[unfolded eq_RAG]]
   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>
@@ -821,8 +739,8 @@
         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
+      from child_RAG_p[OF ch'] and eq_y
+      have dp_thy: "(Th th, Th thy) \<in> (RAG s)^+" by simp
       show "cp s th'' = cp s' th''"
         apply (subst cp_rec[OF vt_s])
       proof -
@@ -832,9 +750,9 @@
           proof
             assume "th'' = th"
             with dp_thy y_ch[unfolded eq_y] 
-            have "(Th th, Th th) \<in> (depend s)^+"
+            have "(Th th, Th th) \<in> (RAG s)^+"
               by (auto simp:child_def)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
+            with wf_trancl[OF wf_RAG[OF vt_s]] 
             show False by auto
           qed
         qed
@@ -850,14 +768,14 @@
             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
+              with dp_thy have "(Th th1, Th thy) \<in> (RAG 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> dependants s th1"
             proof
               assume h:"th \<in> dependants s th1"
-              from eq_y dp_thy have "th \<in> dependants s thy" by (auto simp:s_dependants_def eq_depend)
+              from eq_y dp_thy have "th \<in> dependants s thy" by (auto simp:s_dependants_def eq_RAG)
               from dependants_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
@@ -869,7 +787,7 @@
         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)
+          by (unfold children_def child_def s_def RAG_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
@@ -885,9 +803,9 @@
           proof
             assume "th'' = th"
             with dp dp'
-            have "(Th th, Th th) \<in> (depend s)^+"
-              by (auto simp:child_def s_dependants_def eq_depend)
-            with wf_trancl[OF wf_depend[OF vt_s]] 
+            have "(Th th, Th th) \<in> (RAG s)^+"
+              by (auto simp:child_def s_dependants_def eq_RAG)
+            with wf_trancl[OF wf_RAG[OF vt_s]] 
             show False by auto
           qed
         qed
@@ -906,7 +824,7 @@
               show "th \<notin> dependants s th1"
               proof
                 assume "th \<in> dependants s th1"
-                hence "(Th th, Th th1) \<in> (depend s)^+" by (auto simp:s_dependants_def eq_depend)
+                hence "(Th th, Th th1) \<in> (RAG s)^+" by (auto simp:s_dependants_def eq_RAG)
                 from children_no_dep[OF vt_s _ _ this]
                 and th1_in dp' show False
                   by (auto simp:children_def)
@@ -917,7 +835,7 @@
         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)
+          by (unfold children_def child_def s_def RAG_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     
@@ -998,11 +916,11 @@
 context step_v_cps_nt
 begin
 
-lemma depend_s:
-  "depend s = (depend s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union>
+lemma RAG_s:
+  "RAG s = (RAG 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]
+  from step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
     and nt show ?thesis  by (auto intro:next_th_unique)
 qed
 
@@ -1014,163 +932,163 @@
 proof(auto)
   fix x
   assume "x \<in> dependants (wq s) th''"
-  hence dp: "(Th x, Th th'') \<in> (depend s)^+"
-    by (auto simp:cs_dependants_def eq_depend)
+  hence dp: "(Th x, Th th'') \<in> (RAG s)^+"
+    by (auto simp:cs_dependants_def eq_RAG)
   { fix n
-    have "(n, Th th'') \<in> (depend s)^+ \<Longrightarrow>  (n, Th th'') \<in> (depend s')^+"
+    have "(n, Th th'') \<in> (RAG s)^+ \<Longrightarrow>  (n, Th th'') \<in> (RAG 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
+      assume "(y, Th th'') \<in> RAG s"
+      with RAG_s neq1 neq2
+      have "(y, Th th'') \<in> RAG s'" by auto
+      thus "(y, Th th'') \<in> (RAG 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>+"
+      assume yz: "(y, z) \<in> RAG s"
+        and ztp: "(z, Th th'') \<in> (RAG s)\<^sup>+"
+        and ztp': "(z, Th th'') \<in> (RAG 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] 
+          with yz have dp_yz: "(Cs cs, z) \<in> RAG s" by simp
+          from RAG_s
+          have cst': "(Cs cs, Th th') \<in> RAG s" by simp
+          from unique_RAG[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
+          with ztp have "(Th th', Th th'') \<in> (RAG 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
+          obtain cs' where dp'': "(Th th', Cs cs') \<in> RAG s"
+            by (auto simp:s_RAG_def)
+          with RAG_s have dp': "(Th th', Cs cs') \<in> RAG s'" by auto
+          from dp'' eq_y yz eq_z have "(Cs cs, Cs cs') \<in> (RAG 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']
+            have "(Th th', Cs cs) \<in> RAG s'"
+              by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def)
+            from unique_RAG[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 have "(Cs cs, Cs cs) \<in> (RAG s)^+" by simp
+          moreover note wf_trancl[OF wf_RAG[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
+          with yz have dps: "(Th th', z) \<in> RAG s" by simp
+          with RAG_s have dps': "(Th th', z) \<in> RAG 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]
+            have "(Th th', Cs cs) \<in> RAG s'"
+              by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def)
+            from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] dps' this]
             show ?thesis .
           qed
-          with dps depend_s show False by auto
+          with dps RAG_s show False by auto
         qed
       qed
-      with depend_s yz have "(y, z) \<in> depend s'" by auto
+      with RAG_s yz have "(y, z) \<in> RAG s'" by auto
       with ztp'
-      show "(y, Th th'') \<in> (depend s')\<^sup>+" by auto
+      show "(y, Th th'') \<in> (RAG s')\<^sup>+" by auto
     qed    
   }
   from this[OF dp]
   show "x \<in> dependants (wq s') th''" 
-    by (auto simp:cs_dependants_def eq_depend)
+    by (auto simp:cs_dependants_def eq_RAG)
 next
   fix x
   assume "x \<in> dependants (wq s') th''"
-  hence dp: "(Th x, Th th'') \<in> (depend s')^+"
-    by (auto simp:cs_dependants_def eq_depend)
+  hence dp: "(Th x, Th th'') \<in> (RAG s')^+"
+    by (auto simp:cs_dependants_def eq_RAG)
   { fix n
-    have "(n, Th th'') \<in> (depend s')^+ \<Longrightarrow>  (n, Th th'') \<in> (depend s)^+"
+    have "(n, Th th'') \<in> (RAG s')^+ \<Longrightarrow>  (n, Th th'') \<in> (RAG 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
+      assume "(y, Th th'') \<in> RAG s'"
+      with RAG_s neq1 neq2
+      have "(y, Th th'') \<in> RAG s" by auto
+      thus "(y, Th th'') \<in> (RAG 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>+"
+      assume yz: "(y, z) \<in> RAG s'"
+        and ztp: "(z, Th th'') \<in> (RAG s')\<^sup>+"
+        and ztp': "(z, Th th'') \<in> (RAG 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
+          with yz have dp_yz: "(Cs cs, z) \<in> RAG 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]
+            have "(Cs cs, Th th) \<in> RAG s'"
+              by(cases, auto simp: wq_def s_RAG_def cs_holding_def s_holding_def)
+            from unique_RAG[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
+          obtain u where "(z, u) \<in> RAG 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)
+            by (auto simp:readys_def wq_def s_RAG_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
+          with yz have dps: "(Th th', z) \<in> RAG 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]
+            have "(Th th', Cs cs) \<in> RAG s'"
+              by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def)
+            from unique_RAG[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
+          with ztp have cs_i: "(Cs cs, Th th'') \<in>  (RAG 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'"
+          have cs_th: "(Cs cs, Th th) \<in> RAG s'"
+            by(cases, auto simp: s_RAG_def wq_def cs_holding_def s_holding_def)
+          have "(Cs cs, Th th'') \<notin>  RAG 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]
+            assume "(Cs cs, Th th'') \<in> RAG s'"
+            from unique_RAG[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
+          obtain u where cu: "(Cs cs, u) \<in> RAG s'"  
+            and u_t: "(u, Th th'') \<in> (RAG 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]
+            from unique_RAG[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
+          with u_t have "(Th th, Th th'') \<in> (RAG s')\<^sup>+" by simp
           from converse_tranclE[OF this]
-          obtain v where "(Th th, v) \<in> (depend s')" by auto
+          obtain v where "(Th th, v) \<in> (RAG 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)
+            by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
         qed
       qed
-      with depend_s yz have "(y, z) \<in> depend s" by auto
+      with RAG_s yz have "(y, z) \<in> RAG s" by auto
       with ztp'
-      show "(y, Th th'') \<in> (depend s)\<^sup>+" by auto
+      show "(y, Th th'') \<in> (RAG s)\<^sup>+" by auto
     qed    
   }
   from this[OF dp]
   show "x \<in> dependants (wq s) th''"
-    by (auto simp:cs_dependants_def eq_depend)
+    by (auto simp:cs_dependants_def eq_RAG)
 qed
 
 lemma cp_kept:
@@ -1204,11 +1122,11 @@
 context step_v_cps_nnt
 begin
 
-lemma nw_cs: "(Th th1, Cs cs) \<notin> depend s'"
+lemma nw_cs: "(Th th1, Cs cs) \<notin> RAG s'"
 proof
-  assume "(Th th1, Cs cs) \<in> depend s'"
+  assume "(Th th1, Cs cs) \<in> RAG s'"
   thus "False"
-    apply (auto simp:s_depend_def cs_waiting_def)
+    apply (auto simp:s_RAG_def cs_waiting_def)
   proof -
     assume h1: "th1 \<in> set (wq s' cs)"
       and h2: "th1 \<noteq> hd (wq s' cs)"
@@ -1229,9 +1147,9 @@
   qed
 qed
 
-lemma depend_s: "depend s = depend s' - {(Cs cs, Th th)}"
+lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}"
 proof -
-  from nnt and  step_depend_v[OF vt_s[unfolded s_def], folded s_def]
+  from nnt and  step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
   show ?thesis by auto
 qed
 
@@ -1244,18 +1162,18 @@
   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'"
+      where h1: "(Th th1, Cs cs1) \<in> RAG s'"
+      and h2: "(Cs cs1, Th th2) \<in> RAG 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 h1 have "(Th th1, Cs cs1) \<in> RAG 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
+    with h1 h2 RAG_s have 
+      h1': "(Th th1, Cs cs1) \<in> RAG s" and
+      h2': "(Cs cs1, Th th2) \<in> RAG 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
@@ -1263,18 +1181,18 @@
     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'"
+      where h1: "(Th th1, Cs cs1) \<in> RAG s'"
+      and h2: "(Cs cs1, Th th2) \<in> RAG 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 h1 have "(Th th1, Cs cs1) \<in> RAG 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
+    with h1 h2 RAG_s have 
+      h1': "(Th th1, Cs cs1) \<in> RAG s" and
+      h2': "(Cs cs1, Th th2) \<in> RAG 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
@@ -1290,14 +1208,14 @@
   from assms show ?thesis
   proof(induct)
     case (base y)
-    from base and depend_s 
+    from base and RAG_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'"
+    with RAG_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
@@ -1313,13 +1231,13 @@
   apply (unfold cp_eq_cpreced cpreced_def)
 proof -
   have eq_dp: "\<And> th. dependants (wq s) th = dependants (wq s') th"
-    apply (unfold cs_dependants_def, unfold eq_depend)
+    apply (unfold cs_dependants_def, unfold eq_RAG)
   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>+}"
+    with child_RAG_eq[OF vt_s] child_RAG_eq[OF step_back_vt[OF vt_s[unfolded s_def]]]
+    show "\<And>th. {th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} = {th'. (Th th', Th th) \<in> (RAG s')\<^sup>+}"
       by simp
   qed
   moreover {
@@ -1357,9 +1275,9 @@
 context step_P_cps_e
 begin
 
-lemma depend_s: "depend s = depend s' \<union> {(Cs cs, Th th)}"
+lemma RAG_s: "RAG s = RAG s' \<union> {(Cs cs, Th th)}"
 proof -
-  from ee and  step_depend_p[OF vt_s[unfolded s_def], folded s_def]
+  from ee and  step_RAG_p[OF vt_s[unfolded s_def], folded s_def]
   show ?thesis by auto
 qed
 
@@ -1372,19 +1290,19 @@
   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'"
+      where h1: "(Th th1, Cs cs1) \<in> RAG s'"
+      and h2: "(Cs cs1, Th th2) \<in> RAG 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 h1 have "(Th th1, Cs cs) \<in> RAG s'" by simp
       with ee show False
-        by (auto simp:s_depend_def cs_waiting_def)
+        by (auto simp:s_RAG_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
+    with h1 h2 RAG_s have 
+      h1': "(Th th1, Cs cs1) \<in> RAG s" and
+      h2': "(Cs cs1, Th th2) \<in> RAG 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
@@ -1392,19 +1310,19 @@
     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'"
+      where h1: "(Th th1, Cs cs1) \<in> RAG s'"
+      and h2: "(Cs cs1, Th th2) \<in> RAG 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 h1 have "(Th th1, Cs cs) \<in> RAG s'" by simp
       with ee show False 
-        by (auto simp:s_depend_def cs_waiting_def)
+        by (auto simp:s_RAG_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
+    with h1 h2 RAG_s have 
+      h1': "(Th th1, Cs cs1) \<in> RAG s" and
+      h2': "(Cs cs1, Th th2) \<in> RAG 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
@@ -1420,28 +1338,28 @@
   from assms show ?thesis
   proof(induct)
     case (base y)
-    from base and depend_s
+    from base and RAG_s
     have "(n1, y) \<in> child s'"
       apply (auto simp:child_def)
       proof -
         fix th'
-        assume "(Th th', Cs cs) \<in> depend s'"
+        assume "(Th th', Cs cs) \<in> RAG 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 
+          by (auto simp:s_RAG_def cs_waiting_def)
+        thus "\<exists>cs. (Th th', Cs cs) \<in> RAG s' \<and> (Cs cs, Th th) \<in> RAG 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'"
+    with RAG_s have "(y, z) \<in> child s'"
       apply (auto simp:child_def)
       proof -
         fix th'
-        assume "(Th th', Cs cs) \<in> depend s'"
+        assume "(Th th', Cs cs) \<in> RAG 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 
+          by (auto simp:s_RAG_def cs_waiting_def)
+        thus "\<exists>cs. (Th th', Cs cs) \<in> RAG s' \<and> (Cs cs, Th th) \<in> RAG s'" by auto 
       qed
     moreover have "(n1, y) \<in> (child s')\<^sup>+" by fact
     ultimately show ?case by auto
@@ -1457,13 +1375,13 @@
   apply (unfold cp_eq_cpreced cpreced_def)
 proof -
   have eq_dp: "\<And> th. dependants (wq s) th = dependants (wq s') th"
-    apply (unfold cs_dependants_def, unfold eq_depend)
+    apply (unfold cs_dependants_def, unfold eq_RAG)
   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>+}"
+    with child_RAG_eq[OF vt_s] child_RAG_eq[OF step_back_vt[OF vt_s[unfolded s_def]]]
+    show "\<And>th. {th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} = {th'. (Th th', Th th) \<in> (RAG s')\<^sup>+}"
       by simp
   qed
   moreover {
@@ -1490,9 +1408,9 @@
 context step_P_cps_ne
 begin
 
-lemma depend_s: "depend s = depend s' \<union> {(Th th, Cs cs)}"
+lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
 proof -
-  from step_depend_p[OF vt_s[unfolded s_def]] and ne
+  from step_RAG_p[OF vt_s[unfolded s_def]] and ne
   show ?thesis by (simp add:s_def)
 qed
 
@@ -1502,8 +1420,8 @@
 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"
+    where h1: "(Th th1, Cs cs1) \<in> RAG s"
+    and h2: "(Cs cs1, Th th') \<in> RAG s"
     and eq_y: "y = Th th1"   by (auto simp:child_def)
   have "th1 \<noteq> th"
   proof
@@ -1511,16 +1429,16 @@
     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 h1 h2 RAG_s 
+  have h1': "(Th th1, Cs cs1) \<in> RAG s'" and 
+       h2': "(Cs cs1, Th th') \<in> RAG 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"
+    where h1: "(Th th1, Cs cs1) \<in> RAG s"
+    and h2: "(Cs cs1, Th th2) \<in> RAG s"
     and eq_y: "y = Th th1" and eq_z: "z = Th th2"  by (auto simp:child_def)
   have "th1 \<noteq> th"
   proof
@@ -1530,8 +1448,8 @@
     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 h1 h2 RAG_s have h1': "(Th th1, Cs cs1) \<in> RAG s'"
+                       and h2': "(Cs cs1, Th th2) \<in> RAG 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
@@ -1541,11 +1459,11 @@
   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)
+  with RAG_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)
+  with RAG_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
@@ -1564,34 +1482,34 @@
   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 child_RAG_eq[OF vt_s]
+    have "(Th th, Th th') \<in> (RAG s)\<^sup>+" by simp
     with nd show False 
-      by (simp add:s_dependants_def eq_depend)
+      by (simp add:s_dependants_def eq_RAG)
   qed
   have eq_dp: "dependants (wq s) th' = dependants (wq s') th'"
   proof(auto)
     fix x assume " x \<in> dependants (wq s) th'"
     thus "x \<in> dependants (wq s') th'"
-      apply (auto simp:cs_dependants_def eq_depend)
+      apply (auto simp:cs_dependants_def eq_RAG)
     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
+      assume "(Th x, Th th') \<in> (RAG s)\<^sup>+"
+      with  child_RAG_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
+      with child_RAG_eq[OF step_back_vt[OF vt_s[unfolded s_def]]]
+      show "(Th x, Th th') \<in> (RAG s')\<^sup>+" by simp
     qed
   next
     fix x assume "x \<in> dependants (wq s') th'"
     thus "x \<in> dependants (wq s) th'"
-      apply (auto simp:cs_dependants_def eq_depend)
+      apply (auto simp:cs_dependants_def eq_RAG)
     proof -
-      assume "(Th x, Th th') \<in> (depend s')\<^sup>+"
-      with child_depend_eq[OF step_back_vt[OF vt_s[unfolded s_def]]] 
+      assume "(Th x, Th th') \<in> (RAG s')\<^sup>+"
+      with child_RAG_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
+      with  child_RAG_eq[OF vt_s]
+      show "(Th x, Th th') \<in> (RAG s)\<^sup>+" by simp
     qed
   qed
   moreover {
@@ -1611,8 +1529,8 @@
   shows "cp s th'' = cp s' th''"
 proof -
   from dp2
-  have "(Th th', Th th'') \<in> (depend (wq s))\<^sup>+" by (simp add:s_dependants_def)
-  from depend_child[OF vt_s this[unfolded eq_depend]]
+  have "(Th th', Th th'') \<in> (RAG (wq s))\<^sup>+" by (simp add:s_dependants_def)
+  from RAG_child[OF vt_s this[unfolded eq_RAG]]
   have ch_th': "(Th th', Th th'') \<in> (child s)\<^sup>+" .
   moreover {
     fix n th''
@@ -1625,10 +1543,10 @@
         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_dependants_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
+      from dp1 have "(Th th, Th th') \<in> (RAG s)^+" by (auto simp:s_dependants_def eq_RAG)
+      moreover from child_RAG_p[OF ch'] and eq_y
+      have "(Th th', Th thy) \<in> (RAG s)^+" by simp
+      ultimately have dp_thy: "(Th th, Th thy) \<in> (RAG s)^+" by auto
       show "cp s th'' = cp s' th''"
         apply (subst cp_rec[OF vt_s])
       proof -
@@ -1646,14 +1564,14 @@
             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
+              with dp_thy have "(Th th1, Th thy) \<in> (RAG 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> dependants s th1"
             proof
               assume h:"th \<in> dependants s th1"
-              from eq_y dp_thy have "th \<in> dependants s thy" by (auto simp:s_dependants_def eq_depend)
+              from eq_y dp_thy have "th \<in> dependants s thy" by (auto simp:s_dependants_def eq_RAG)
               from dependants_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
@@ -1665,48 +1583,48 @@
         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)
+          apply (unfold children_def child_def s_def RAG_set_unchanged, simp)
+          apply (fold s_def, auto simp:RAG_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_dependants_def eq_depend)
+            assume "(Cs cs, Th th'') \<in> RAG s'"
+            with RAG_s have cs_th': "(Cs cs, Th th'') \<in> RAG s" by auto
+            from dp1 have "(Th th, Th th') \<in> (RAG s)^+"
+              by (auto simp:s_dependants_def eq_RAG)
             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)
+            obtain cs1 where h1: "(Th th, Cs cs1) \<in> RAG s"
+              and h2: "(Cs cs1 , Th th') \<in> (RAG s)\<^sup>+"
+              by (auto simp:s_RAG_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]
+              from RAG_s have "(Th th, Cs cs) \<in> RAG s" by simp
+              from unique_RAG[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']
+              assume "(Cs cs1, Th th') \<in> RAG s"
+              with eq_cs have "(Cs cs, Th th') \<in> RAG s" by simp
+              from unique_RAG[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']
+              assume "(Cs cs1, y) \<in> RAG s"
+                and ytd: " (y, Th th') \<in> (RAG s)\<^sup>+"
+              with eq_cs have csy: "(Cs cs, y) \<in> RAG s" by simp
+              from unique_RAG[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]
+              with ytd have "(Th th'', Th th') \<in> (RAG s)^+" by simp
+              from RAG_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
+            thus "\<exists>cs. (Th th, Cs cs) \<in> RAG s' \<and> (Cs cs, Th th'') \<in> RAG 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]]])
@@ -1731,8 +1649,8 @@
             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_dependants_def eq_depend)
+              with dp1 have "(Th th1, Th th') \<in> (RAG s)^+" 
+                by (auto simp:s_dependants_def eq_RAG)
               from children_no_dep[OF vt_s _ _ this]
               th1_in dp'
               show False by (auto simp:children_def)
@@ -1753,48 +1671,48 @@
         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)
+          apply (unfold children_def child_def s_def RAG_set_unchanged, simp)
+          apply (fold s_def, auto simp:RAG_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_dependants_def eq_depend)
+            assume "(Cs cs, Th th'') \<in> RAG s'"
+            with RAG_s have cs_th': "(Cs cs, Th th'') \<in> RAG s" by auto
+            from dp1 have "(Th th, Th th') \<in> (RAG s)^+"
+              by (auto simp:s_dependants_def eq_RAG)
             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)
+            obtain cs1 where h1: "(Th th, Cs cs1) \<in> RAG s"
+              and h2: "(Cs cs1 , Th th') \<in> (RAG s)\<^sup>+"
+              by (auto simp:s_RAG_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]
+              from RAG_s have "(Th th, Cs cs) \<in> RAG s" by simp
+              from unique_RAG[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']
+              assume "(Cs cs1, Th th') \<in> RAG s"
+              with eq_cs have "(Cs cs, Th th') \<in> RAG s" by simp
+              from unique_RAG[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']
+              assume "(Cs cs1, y) \<in> RAG s"
+                and ytd: " (y, Th th') \<in> (RAG s)\<^sup>+"
+              with eq_cs have csy: "(Cs cs, y) \<in> RAG s" by simp
+              from unique_RAG[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]
+              with ytd have "(Th th'', Th th') \<in> (RAG s)^+" by simp
+              from RAG_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
+            thus "\<exists>cs. (Th th, Cs cs) \<in> RAG s' \<and> (Cs cs, Th th'') \<in> RAG 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]]])
@@ -1814,8 +1732,8 @@
 context step_create_cps
 begin
 
-lemma eq_dep: "depend s = depend s'"
-  by (unfold s_def depend_create_unchanged, auto)
+lemma eq_dep: "RAG s = RAG s'"
+  by (unfold s_def RAG_create_unchanged, auto)
 
 lemma eq_cp:
   fixes th' 
@@ -1826,11 +1744,11 @@
   have nd: "th \<notin> dependants s th'"
   proof
     assume "th \<in> dependants s th'"
-    hence "(Th th, Th th') \<in> (depend s)^+" by (simp add:s_dependants_def eq_depend)
-    with eq_dep have "(Th th, Th th') \<in> (depend s')^+" by simp
+    hence "(Th th, Th th') \<in> (RAG s)^+" by (simp add:s_dependants_def eq_RAG)
+    with eq_dep have "(Th th, Th th') \<in> (RAG 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]]]
+    obtain y where "(Th th, y) \<in> RAG s'" by auto
+    with dm_RAG_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
@@ -1840,7 +1758,7 @@
     qed
   qed
   have eq_dp: "\<And> th. dependants (wq s) th = dependants (wq s') th"
-    by (unfold cs_dependants_def, auto simp:eq_dep eq_depend)
+    by (unfold cs_dependants_def, auto simp:eq_dep eq_RAG)
   moreover {
     fix th1 
     assume "th1 \<in> {th'} \<union> dependants (wq s') th'"
@@ -1853,7 +1771,7 @@
     next
       assume "th1 \<in> dependants (wq s') th'"
       with nd and eq_dp have "th1 \<noteq> th"
-        by (auto simp:eq_depend cs_dependants_def s_dependants_def eq_dep)
+        by (auto simp:eq_RAG cs_dependants_def s_dependants_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> dependants (wq s) th')) = 
@@ -1874,16 +1792,16 @@
     have "dependants s' th = {}"
     proof -
       { assume "dependants s' th \<noteq> {}"
-        then obtain th' where dp: "(Th th', Th th) \<in> (depend s')^+"
-          by (auto simp:s_dependants_def eq_depend)
+        then obtain th' where dp: "(Th th', Th th) \<in> (RAG s')^+"
+          by (auto simp:s_dependants_def eq_RAG)
         from tranclE[OF this] obtain cs' where 
-          "(Cs cs', Th th) \<in> depend s'" by (auto simp:s_depend_def)
+          "(Cs cs', Th th) \<in> RAG s'" by (auto simp:s_RAG_def)
         with hdn
         have False by (auto simp:holdents_test)
       } thus ?thesis by auto
     qed
     thus ?thesis 
-      by (unfold s_def s_dependants_def eq_depend depend_create_unchanged, simp)
+      by (unfold s_def s_dependants_def eq_RAG RAG_create_unchanged, simp)
   qed
 qed
 
@@ -1902,8 +1820,8 @@
 context step_exit_cps
 begin
 
-lemma eq_dep: "depend s = depend s'"
-  by (unfold s_def depend_exit_unchanged, auto)
+lemma eq_dep: "RAG s = RAG s'"
+  by (unfold s_def RAG_exit_unchanged, auto)
 
 lemma eq_cp:
   fixes th' 
@@ -1914,22 +1832,22 @@
   have nd: "th \<notin> dependants s th'"
   proof
     assume "th \<in> dependants s th'"
-    hence "(Th th, Th th') \<in> (depend s)^+" by (simp add:s_dependants_def eq_depend)
-    with eq_dep have "(Th th, Th th') \<in> (depend s')^+" by simp
+    hence "(Th th, Th th') \<in> (RAG s)^+" by (simp add:s_dependants_def eq_RAG)
+    with eq_dep have "(Th th, Th th') \<in> (RAG 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)
+    obtain cs' where bk: "(Th th, Cs cs') \<in> RAG s'"
+      by (auto simp:s_RAG_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)
+        apply (unfold runing_def readys_def s_waiting_def s_RAG_def)
         by (auto simp:cs_waiting_def wq_def)
     qed
   qed
   have eq_dp: "\<And> th. dependants (wq s) th = dependants (wq s') th"
-    by (unfold cs_dependants_def, auto simp:eq_dep eq_depend)
+    by (unfold cs_dependants_def, auto simp:eq_dep eq_RAG)
   moreover {
     fix th1 
     assume "th1 \<in> {th'} \<union> dependants (wq s') th'"
@@ -1942,7 +1860,7 @@
     next
       assume "th1 \<in> dependants (wq s') th'"
       with nd and eq_dp have "th1 \<noteq> th"
-        by (auto simp:eq_depend cs_dependants_def s_dependants_def eq_dep)
+        by (auto simp:eq_RAG cs_dependants_def s_dependants_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> dependants (wq s) th')) = 
--- a/ExtGG.thy	Tue May 20 12:49:21 2014 +0100
+++ b/ExtGG.thy	Thu May 22 17:40:39 2014 +0100
@@ -978,7 +978,7 @@
   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
+    and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
   have "th' \<in> runing (t@s)"
   proof -
     have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
@@ -1021,7 +1021,7 @@
       next
         from dp
         have "th \<in> dependants (wq (t @ s)) th'" 
-          by (unfold cs_dependants_def, auto simp:eq_depend)
+          by (unfold cs_dependants_def, auto simp:eq_RAG)
         thus "preced th (t @ s) \<in> 
                 (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependants (wq (t @ s)) th')"
           by auto
--- a/Journal/Paper.thy	Tue May 20 12:49:21 2014 +0100
+++ b/Journal/Paper.thy	Thu May 22 17:40:39 2014 +0100
@@ -9,23 +9,22 @@
 
 notation (latex output)
   Cons ("_::_" [78,77] 73) and
+  If  ("(\<^raw:\textrm{>if\<^raw:}> (_)/ \<^raw:\textrm{>then\<^raw:}> (_)/ \<^raw:\textrm{>else\<^raw:}> (_))" 10) and
   vt ("valid'_state") and
   runing ("running") 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
+  preceds ("precs") and
   cpreced ("cprec") and
   cp ("cprec") and
   holdents ("resources") and
   DUMMY  ("\<^raw:\mbox{$\_\!\_$}>")
  
-  
 (*>*)
 
 section {* Introduction *}
@@ -186,7 +185,7 @@
   et al.~\cite{Sha90} to the practically relevant case where critical 
   sections can overlap; see Figure~\ref{overlap} below for an example of 
   this restriction. In the existing literature there is no 
-  proof and also no prove method that cover this generalised case.
+  proof and also no proving method that cover this generalised case.
 
   \begin{figure}
   \begin{center}
@@ -328,6 +327,14 @@
   \end{isabelle}
 
   \noindent
+  We also use the abbreviation 
+
+  \begin{isabelle}\ \ \ \ \ %%%
+  @{abbrev "preceds s ths"}
+  \end{isabelle}
+
+  \noindent
+  for the set of precedences of threads @{text ths} in state @{text s}.
   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 
@@ -390,7 +397,7 @@
   as the union of the sets of waiting and holding edges, namely
 
   \begin{isabelle}\ \ \ \ \ %%%
-  @{thm cs_depend_def}
+  @{thm cs_RAG_def}
   \end{isabelle}
 
 
@@ -478,7 +485,7 @@
 
   \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 
+  While the precedence @{term prec} of any 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
@@ -486,7 +493,8 @@
   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.
+  lowered prematurely. We again introduce an abbreviation for current precedeces of
+  a set of threads, written @{term "cprecs wq s ths"}.
   
   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 
@@ -597,16 +605,18 @@
   \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.
+  Having the scheduler function @{term schs} at our disposal, we can
+  ``lift'', or overload, the notions @{term waiting}, @{term holding},
+  @{term RAG}, @{term dependants} 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}
+  @{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_RAG_abv}      & @{text "\<equiv>"} & @{thm (rhs) s_RAG_abv}\\
+  @{thm (lhs) s_dependants_abv}& @{text "\<equiv>"} & @{thm (rhs) s_dependants_abv}\\
+  @{thm (lhs) cp_def}         & @{text "\<equiv>"} & @{thm (rhs) cp_def}\\  
   \end{tabular}
   \end{isabelle}
 
@@ -625,7 +635,7 @@
   \end{isabelle}
 
   \noindent
-  In the second definition @{term "DUMMY ` DUMMY"} stands for the image of a set under a function.
+  %%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
@@ -791,7 +801,7 @@
   \begin{isabelle}\ \ \ \ \ %%%
   \begin{tabular}{l}
   @{term "th \<in> threads s"}\\
-  @{term "prec th s = Max (cprec s ` threads s)"}\\
+  @{term "prec th s = Max (cprecs s (threads s))"}\\
   @{term "prec th s = (prio, DUMMY)"}
   \end{tabular}
   \end{isabelle}
@@ -836,30 +846,30 @@
   many threads that can block @{text th} in this way and then they 
   need to run with the same 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. There are mainly 
-  two reasons: First, we do not specify what ``running'' the code of a 
-  thread means, for example by giving an operational semantics for
-  machine instructions. Therefore we cannot characterise what are ``good''
+  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. There are mainly two reasons:
+  First, we do not specify what ``running'' the code of a thread
+  means, for example by giving an operational semantics for machine
+  instructions. Therefore we cannot characterise what are ``good''
   programs that contain for every looking request also a corresponding
-  unlocking request for a resource. 
-  %
-  %(HERE)
-  %
-  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 RAG). In this detail, we do not
-  make any progress in comparison with the work by Sha et al.
+  unlocking request for a resource. Second, we would need to specify a
+  kind of global clock that tracks the time how long a thread locks a
+  resource. But this seems difficult, because how do we conveninet
+  distinguish between a thread that ``just'' lock a resource for a
+  very long time and one that locks it forever.  Therefore we decided
+  to leave it out this property and let the programmer assume the
+  responsibility to program threads in such a benign manner (in
+  addition to causing no circularity in the 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
+  It is relatively easy to see that:
 
   \begin{isabelle}\ \ \ \ \ %%%
   \begin{tabular}{@ {}l}
@@ -870,7 +880,7 @@
 
   \noindent
   The second property is by induction of @{term vt}. The next three
-  properties are 
+  properties are: 
 
   \begin{isabelle}\ \ \ \ \ %%%
   \begin{tabular}{@ {}l}
@@ -890,11 +900,11 @@
 
   \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
+  @{text If}~@{thm (prem 1) acyclic_RAG}~@{text "then"}:\\
+  \hspace{5mm}@{thm (concl) acyclic_RAG},
+  @{thm (concl) finite_RAG} and
   @{thm (concl) wf_dep_converse},\\
-  \hspace{5mm}@{text "if"}~@{thm (prem 2) dm_depend_threads}~@{text "then"}~@{thm (concl) dm_depend_threads}
+  \hspace{5mm}@{text "if"}~@{thm (prem 2) dm_RAG_threads}~@{text "then"}~@{thm (concl) dm_RAG_threads}
   and\\
   \hspace{5mm}@{text "if"}~@{thm (prem 2) range_in}~@{text "then"}~@{thm (concl) range_in}.
   \end{tabular}
@@ -1113,7 +1123,7 @@
   
   \noindent
   That means the current precedence of a thread @{text th} can be
-  computed locally by considering only the children of @{text th}. In
+  computed locally by considering only the current precedences of 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
@@ -1121,6 +1131,10 @@
   a standard scheduling operation implemented in most operating
   systems.
 
+  \begin{proof}[of Lemma~\ref{childrenlem}]
+  Test
+  \end{proof}
+
   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
@@ -1242,7 +1256,7 @@
 
 
   \begin{isabelle}\ \ \ \ \ %%%
-  @{thm depend_s}
+  @{thm RAG_s}
   \end{isabelle}
   
   \noindent
@@ -1265,7 +1279,7 @@
 
   \begin{isabelle}\ \ \ \ \ %%%
   \begin{tabular}{@ {}l}
-  @{thm depend_s}\\
+  @{thm RAG_s}\\
   @{thm eq_cp}
   \end{tabular}
   \end{isabelle}
@@ -1284,7 +1298,7 @@
   
   \begin{isabelle}\ \ \ \ \ %%%
   \begin{tabular}{@ {}l}
-  @{thm depend_s}\\
+  @{thm RAG_s}\\
   @{thm eq_cp}
   \end{tabular}
   \end{isabelle}
@@ -1297,7 +1311,7 @@
   
   \begin{isabelle}\ \ \ \ \ %%%
   \begin{tabular}{@ {}l}
-  @{thm depend_s}\\
+  @{thm RAG_s}\\
   @{thm[mode=IfThen] eq_cp}
   \end{tabular}
   \end{isabelle}
--- a/Journal/document/root.tex	Tue May 20 12:49:21 2014 +0100
+++ b/Journal/document/root.tex	Thu May 22 17:40:39 2014 +0100
@@ -53,7 +53,7 @@
 \footnotetext[1]{This paper is a revised, corrected and expanded version of \cite{ZhangUrbanWu12}.
 Compared with that paper we give an actual implementation of our formalised scheduling 
 algorithm in C and the operating system PINTOS. Our implementation follows closely all results
-we proved about optimisations.}
+we proved about optimisations of the Priority Inheritance Protocol.}
 \renewcommand{\thefootnote}{\arabic{footnote}}
 
 \title{Priority Inheritance Protocol Proved Correct}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Max.thy	Thu May 22 17:40:39 2014 +0100
@@ -0,0 +1,78 @@
+theory Max
+imports Main
+begin
+
+section {* Some generic facts about Max *}
+
+
+lemma Max_insert:
+  assumes "finite B"
+  and  "B \<noteq> {}"
+  shows "Max ({x} \<union> B) = max x (Max B)"
+using assms by (simp add: Lattices_Big.Max.insert)
+
+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)"
+using assms
+proof(induct rule: finite_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)"
+        by (metis (full_types) finite_Union h insert.hyps(1) insertCI)
+    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
+
+
+end
\ No newline at end of file
--- a/Moment.thy	Tue May 20 12:49:21 2014 +0100
+++ b/Moment.thy	Thu May 22 17:40:39 2014 +0100
@@ -780,4 +780,4 @@
   with moment_app show ?thesis by auto
 qed
 
-end
\ No newline at end of file
+end
--- a/PrioG.thy	Tue May 20 12:49:21 2014 +0100
+++ b/PrioG.thy	Thu May 22 17:40:39 2014 +0100
@@ -24,7 +24,7 @@
   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>+"
+    assume h1: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
       and h2: "thread \<in> set (wq_fun (schs s) cs)"
       and h3: "thread \<in> runing s"
     show "False" 
@@ -34,8 +34,8 @@
         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)
+      have "(Cs cs, Th thread) \<in> (RAG s)"
+        by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def)
       with h1 show False by auto
     qed
   next
@@ -112,7 +112,7 @@
 qed
 
 lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
-  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (depend s)^+"
+  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (RAG s)^+"
 apply (ind_cases "vt ((P thread cs)#s)")
 apply (ind_cases "step s (P thread cs)")
 by auto
@@ -501,16 +501,16 @@
   qed
 qed
 
-lemma depend_set_unchanged: "(depend (Set th prio # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
+lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
+apply (unfold s_RAG_def s_waiting_def wq_def)
 by (simp add:Let_def)
 
-lemma depend_create_unchanged: "(depend (Create th prio # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
+lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
+apply (unfold s_RAG_def s_waiting_def wq_def)
 by (simp add:Let_def)
 
-lemma depend_exit_unchanged: "(depend (Exit th # s)) = depend s"
-apply (unfold s_depend_def s_waiting_def wq_def)
+lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
+apply (unfold s_RAG_def s_waiting_def wq_def)
 by (simp add:Let_def)
 
 
@@ -773,16 +773,16 @@
   qed
 qed
 
-lemma step_depend_v:
+lemma step_RAG_v:
 fixes th::thread
 assumes vt:
   "vt (V th cs#s)"
 shows "
-  depend (V th cs # s) =
-  depend s - {(Cs cs, Th th)} -
+  RAG (V th cs # s) =
+  RAG s - {(Cs cs, Th th)} -
   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
   {(Cs cs, Th th') |th'.  next_th s th cs th'}"
-  apply (insert vt, unfold s_depend_def) 
+  apply (insert vt, unfold s_RAG_def) 
   apply (auto split:if_splits list.splits simp:Let_def)
   apply (auto elim: step_v_waiting_mono step_v_hold_inv 
               step_v_release step_v_wait_inv
@@ -790,17 +790,17 @@
   apply (erule_tac step_v_not_wait, auto)
   done
 
-lemma step_depend_p:
+lemma step_RAG_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)
+  RAG (P th cs # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
+                                             else RAG s \<union> {(Th th, Cs cs)})"
+  apply(simp only: s_RAG_def wq_def)
   apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def)
   apply(case_tac "csa = cs", auto)
   apply(fold wq_def)
   apply(drule_tac step_back_step)
   apply(ind_cases " step s (P (hd (wq s cs)) cs)")
-  apply(auto simp:s_depend_def wq_def cs_holding_def)
+  apply(auto simp:s_RAG_def wq_def cs_holding_def)
   done
 
 lemma simple_A:
@@ -815,35 +815,35 @@
   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 RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
+  by (unfold s_RAG_def, auto)
 
-lemma acyclic_depend: 
+lemma acyclic_RAG: 
   fixes s
   assumes vt: "vt s"
-  shows "acyclic (depend s)"
+  shows "acyclic (RAG s)"
 proof -
   from vt show ?thesis
   proof(induct)
     case (vt_cons s e)
-    assume ih: "acyclic (depend s)"
+    assume ih: "acyclic (RAG s)"
       and stp: "step s e"
       and vt: "vt s"
     show ?case
     proof(cases e)
       case (Create th prio)
       with ih
-      show ?thesis by (simp add:depend_create_unchanged)
+      show ?thesis by (simp add:RAG_create_unchanged)
     next
       case (Exit th)
-      with ih show ?thesis by (simp add:depend_exit_unchanged)
+      with ih show ?thesis by (simp add:RAG_exit_unchanged)
     next
       case (V th cs)
       from V vt stp have vtt: "vt (V th cs#s)" by auto
-      from step_depend_v [OF this]
+      from step_RAG_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>
+        "RAG (e # s) = 
+            RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
             {(Cs cs, Th th') |th'. next_th s th cs th'}"
         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
       from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
@@ -871,11 +871,11 @@
             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]
+            from RAG_target_th [OF this]
             obtain cs' where eq_x: "x = Cs cs'" by auto
             with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
             hence wt_th': "waiting s ?th' cs'"
-              unfolding s_depend_def s_waiting_def cs_waiting_def wq_def by simp
+              unfolding s_RAG_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]
@@ -945,35 +945,35 @@
   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")
+    from step_RAG_p [OF this] P
+    have "RAG (e # s) = 
+      (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
+      RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
       by simp
     moreover have "acyclic ?R"
     proof(cases "wq s cs = []")
       case True
-      hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
-      have "(Th th, Cs cs) \<notin> (depend s)\<^sup>*"
+      hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
+      have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
       proof
-        assume "(Th th, Cs cs) \<in> (depend s)\<^sup>*"
-        hence "(Th th, Cs cs) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
+        assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
+        hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
         from tranclD2 [OF this]
-        obtain x where "(x, Cs cs) \<in> depend s" by auto
-        with True show False by (auto simp:s_depend_def cs_waiting_def)
+        obtain x where "(x, Cs cs) \<in> RAG s" by auto
+        with True show False by (auto simp:s_RAG_def cs_waiting_def)
       qed
       with acyclic_insert ih eq_r show ?thesis by auto
     next
       case False
-      hence eq_r: "?R =  depend s \<union> {(Th th, Cs cs)}" by simp
-      have "(Cs cs, Th th) \<notin> (depend s)\<^sup>*"
+      hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
+      have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
       proof
-        assume "(Cs cs, Th th) \<in> (depend s)\<^sup>*"
-        hence "(Cs cs, Th th) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
+        assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
+        hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
           moreover from step_back_step [OF vtt] have "step s (P th cs)" .
           ultimately show False
           proof -
-            show " \<lbrakk>(Cs cs, Th th) \<in> (depend s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
+            show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
               by (ind_cases "step s (P th cs)", simp)
           qed
         qed
@@ -983,42 +983,42 @@
     next
       case (Set thread prio)
       with ih
-      thm depend_set_unchanged
-      show ?thesis by (simp add:depend_set_unchanged)
+      thm RAG_set_unchanged
+      show ?thesis by (simp add:RAG_set_unchanged)
     qed
   next
     case vt_nil
-    show "acyclic (depend ([]::state))"
-      by (auto simp: s_depend_def cs_waiting_def 
+    show "acyclic (RAG ([]::state))"
+      by (auto simp: s_RAG_def cs_waiting_def 
                       cs_holding_def wq_def acyclic_def)
   qed
 qed
 
-lemma finite_depend: 
+lemma finite_RAG: 
   fixes s
   assumes vt: "vt s"
-  shows "finite (depend s)"
+  shows "finite (RAG s)"
 proof -
   from vt show ?thesis
   proof(induct)
     case (vt_cons s e)
-    assume ih: "finite (depend s)"
+    assume ih: "finite (RAG s)"
       and stp: "step s e"
       and vt: "vt s"
     show ?case
     proof(cases e)
       case (Create th prio)
       with ih
-      show ?thesis by (simp add:depend_create_unchanged)
+      show ?thesis by (simp add:RAG_create_unchanged)
     next
       case (Exit th)
-      with ih show ?thesis by (simp add:depend_exit_unchanged)
+      with ih show ?thesis by (simp add:RAG_exit_unchanged)
     next
       case (V th cs)
       from V vt stp have vtt: "vt (V th cs#s)" by auto
-      from step_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>
+      from step_RAG_v [OF this]
+      have eq_de: "RAG (e # s) = 
+                   RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
                       {(Cs cs, Th th') |th'. next_th s th cs th'}
 "
         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
@@ -1041,31 +1041,31 @@
     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")
+      from step_RAG_p [OF this] P
+      have "RAG (e # s) = 
+              (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
+                                    RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
         by simp
       moreover have "finite ?R"
       proof(cases "wq s cs = []")
         case True
-        hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
+        hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
         with True and ih show ?thesis by auto
       next
         case False
-        hence "?R = depend s \<union> {(Th th, Cs cs)}" by simp
+        hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
         with False and ih show ?thesis by auto
       qed
       ultimately show ?thesis by auto
     next
       case (Set thread prio)
       with ih
-      show ?thesis by (simp add:depend_set_unchanged)
+      show ?thesis by (simp add:RAG_set_unchanged)
     qed
   next
     case vt_nil
-    show "finite (depend ([]::state))"
-      by (auto simp: s_depend_def cs_waiting_def 
+    show "finite (RAG ([]::state))"
+      by (auto simp: s_RAG_def cs_waiting_def 
                    cs_holding_def wq_def acyclic_def)
   qed
 qed
@@ -1075,20 +1075,20 @@
 lemma wf_dep_converse: 
   fixes s
   assumes vt: "vt s"
-  shows "wf ((depend s)^-1)"
+  shows "wf ((RAG s)^-1)"
 proof(rule finite_acyclic_wf_converse)
-  from finite_depend [OF vt]
-  show "finite (depend s)" .
+  from finite_RAG [OF vt]
+  show "finite (RAG s)" .
 next
-  from acyclic_depend[OF vt]
-  show "acyclic (depend s)" .
+  from acyclic_RAG[OF vt]
+  show "acyclic (RAG 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 th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
+  by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
 
 lemma wq_threads: 
   fixes s cs
@@ -1114,7 +1114,7 @@
         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)
+               s_RAG_def s_holding_def cs_holding_def)
         done
     next
       case (V th' cs')
@@ -1192,8 +1192,8 @@
   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)
+lemma range_in: "\<lbrakk>vt s; (Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
+  apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
   by (auto intro:wq_threads)
 
 lemma readys_v_eq:
@@ -1231,33 +1231,33 @@
 
 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)^+)"
+  shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
 proof -
   from wf_dep_converse [OF vt]
-  have h: "wf ((depend s)\<inverse>)" .
+  have h: "wf ((RAG s)\<inverse>)" .
   show ?thesis
   proof(induct rule:wf_induct [OF h])
     fix x
     assume ih [rule_format]: 
-      "\<forall>y. (y, x) \<in> (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>+)"
+      "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
+           y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
+    show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
     proof
-      assume x_d: "x \<in> Domain (depend s)"
-      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+"
+      assume x_d: "x \<in> Domain (RAG s)"
+      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
       proof(cases x)
         case (Th th)
-        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> 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 x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
+        with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
+        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
+        hence "Cs cs \<in> Domain (RAG s)" by auto
         from ih [OF x_in_r this] obtain th'
-          where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (depend s)\<^sup>+" by auto
-        have "(x, Th th') \<in> (depend s)\<^sup>+" using Th x_in cs_in by auto
+          where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
+        have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
         with th'_ready show ?thesis by auto
       next
         case (Cs cs)
-        from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (depend s)^-1" by (auto simp:s_depend_def)
+        from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
         show ?thesis
         proof(cases "th' \<in> readys s")
           case True
@@ -1265,14 +1265,14 @@
         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)
+          with False have "Th th' \<in> Domain (RAG s)" 
+            by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
           from ih [OF th'_d this]
           obtain th'' where 
             th''_r: "th'' \<in> readys s" and 
-            th''_in: "(Th th', Th th'') \<in> (depend s)\<^sup>+" by auto
+            th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
           from th'_d and th''_in 
-          have "(x, Th th'') \<in> (depend s)\<^sup>+" by auto
+          have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
           with th''_r show ?thesis by auto
         qed
       qed
@@ -1284,14 +1284,14 @@
   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)^+)"
+  shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
 proof(cases "th \<in> readys s")
   case True
   thus ?thesis by auto
 next
   case False
-  from False and th_in have "Th th \<in> Domain (depend s)" 
-    by (auto simp:readys_def s_waiting_def s_depend_def wq_def cs_waiting_def Domain_def)
+  from False and th_in have "Th th \<in> Domain (RAG s)" 
+    by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
   from chain_building [rule_format, OF vt this]
   show ?thesis by auto
 qed
@@ -1305,8 +1305,8 @@
 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)
+lemma unique_RAG: "\<lbrakk>vt s; (n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
+  apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
   by(auto elim:waiting_unique holding_unique)
 
 lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
@@ -1314,34 +1314,34 @@
 
 lemma dchain_unique:
   assumes vt: "vt s"
-  and th1_d: "(n, Th th1) \<in> (depend s)^+"
+  and th1_d: "(n, Th th1) \<in> (RAG s)^+"
   and th1_r: "th1 \<in> readys s"
-  and th2_d: "(n, Th th2) \<in> (depend s)^+"
+  and th2_d: "(n, Th th2) \<in> (RAG s)^+"
   and th2_r: "th2 \<in> readys s"
   shows "th1 = th2"
 proof -
   { assume neq: "th1 \<noteq> th2"
     hence "Th th1 \<noteq> Th th2" by simp
-    from unique_chain [OF _ th1_d th2_d this] and unique_depend [OF vt]
-    have "(Th th1, Th th2) \<in> (depend s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (depend s)\<^sup>+" by auto
+    from unique_chain [OF _ th1_d th2_d this] and unique_RAG [OF vt]
+    have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
     hence "False"
     proof
-      assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+"
+      assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
       from trancl_split [OF this]
-      obtain n where dd: "(Th th1, n) \<in> depend s" by auto
+      obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
       then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
+        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       from dd eq_n have "th1 \<notin> readys s"
-        by (auto simp:readys_def s_depend_def wq_def s_waiting_def cs_waiting_def)
+        by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
       with th1_r show ?thesis by auto
     next
-      assume "(Th th2, Th th1) \<in> (depend s)\<^sup>+"
+      assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
       from trancl_split [OF this]
-      obtain n where dd: "(Th th2, n) \<in> depend s" by auto
+      obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
       then obtain cs where eq_n: "n = Cs cs"
-        by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
+        by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       from dd eq_n have "th2 \<notin> readys s"
-        by (auto simp:readys_def wq_def s_depend_def s_waiting_def cs_waiting_def)
+        by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
       with th2_r show ?thesis by auto
     qed
   } thus ?thesis by auto
@@ -1355,7 +1355,7 @@
   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)
+  unfolding  holdents_test step_RAG_p[OF vt] by (auto)
 qed
 
 lemma step_holdents_p_eq:
@@ -1365,7 +1365,7 @@
   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
+  unfolding  holdents_test step_RAG_p[OF vt] by auto
 qed
 
 
@@ -1375,16 +1375,16 @@
   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>" 
+  from finite_RAG [OF vt]
+  have "finite (RAG s)" .
+  hence "finite (?F `(RAG s))" by simp
+  moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
   proof -
     { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
-      fix x assume "(Cs x, Th th) \<in> depend s"
-      hence "?F (Cs x, Th th) \<in> ?F `(depend s)" by (rule h)
+      fix x assume "(Cs x, Th th) \<in> RAG s"
+      hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
       moreover have "?F (Cs x, Th th) = x" by simp
-      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` depend s" by simp 
+      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
     } thus ?thesis by auto
   qed
   ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
@@ -1397,12 +1397,12 @@
 proof -
   from step_back_step[OF vtv]
   have cs_in: "cs \<in> holdents s thread" 
-    apply (cases, unfold holdents_test s_depend_def, simp)
+    apply (cases, unfold holdents_test s_RAG_def, simp)
     by (unfold cs_holding_def s_holding_def wq_def, auto)
   moreover have cs_not_in: 
     "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
     apply (insert wq_distinct[OF step_back_vt[OF vtv], of cs])
-    apply (unfold holdents_test, unfold step_depend_v[OF vtv],
+    apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
             auto simp:next_th_def)
   proof -
     fix rest
@@ -1425,7 +1425,7 @@
       show "x \<noteq> []" by auto
     qed
     ultimately 
-    show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> depend s"
+    show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
       by auto
   next
     fix rest
@@ -1494,7 +1494,7 @@
         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)
+          by (simp add:RAG_create_unchanged eq_e)
         { assume "th \<noteq> thread"
           with eq_readys eq_e
           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
@@ -1519,7 +1519,7 @@
       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)
+        by (simp add:RAG_exit_unchanged eq_e)
       { assume "th \<noteq> thread"
         with eq_e
         have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
@@ -1544,7 +1544,7 @@
       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>+"
+        and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
       from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
       show ?thesis 
       proof -
@@ -1561,7 +1561,7 @@
             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)
+            by (unfold  step_RAG_p [OF vtp], auto)
           moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
             by (simp add:cntP_def count_def)
           moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
@@ -1582,26 +1582,26 @@
               case True
               with is_runing
               have "th \<in> readys (e#s)"
-                apply (unfold eq_e wq_def, unfold readys_def s_depend_def)
+                apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
                 apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
                 by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
               moreover have "cntCS (e # s) th = 1 + cntCS s th"
               proof -
-                have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> depend s} =
-                  Suc (card {cs. (Cs cs, Th thread) \<in> depend s})" (is "card ?L = Suc (card ?R)")
+                have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
+                  Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
                 proof -
                   have "?L = insert cs ?R" by auto
                   moreover have "card \<dots> = Suc (card (?R - {cs}))" 
                   proof(rule card_insert)
                     from finite_holding [OF vt, of thread]
-                    show " finite {cs. (Cs cs, Th thread) \<in> depend s}"
+                    show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
                       by (unfold holdents_test, simp)
                   qed
                   moreover have "?R - {cs} = ?R"
                   proof -
                     have "cs \<notin> ?R"
                     proof
-                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> depend s}"
+                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
                       with no_dep show False by auto
                     qed
                     thus ?thesis by auto
@@ -1611,7 +1611,7 @@
                 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)
+                  by (unfold step_RAG_p [OF vtp], auto simp:True)
               qed
               moreover from is_runing have "th \<in> readys s"
                 by (simp add:runing_def eq_th)
@@ -1638,7 +1638,7 @@
               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])
+                apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
                 by (auto simp:False)
               moreover note eq_cnp eq_cnv ih[of th]
               moreover from is_runing have "th \<in> readys s"
@@ -1735,13 +1735,13 @@
               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)
+              apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
               proof -
-                have "{csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
-                      {cs. (Cs cs, Th th) \<in> depend s}"
+                have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
+                      {cs. (Cs cs, Th th) \<in> RAG s}"
                 proof -
                   from False eq_wq
-                  have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> depend s"
+                  have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
                     apply (unfold next_th_def, auto)
                   proof -
                     assume ne: "rest \<noteq> []"
@@ -1759,13 +1759,13 @@
                       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"
+                      "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
                       by auto
                   qed    
                   thus ?thesis by auto
                 qed
-                thus "card {csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
-                             card {cs. (Cs cs, Th th) \<in> depend s}" by simp 
+                thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
+                             card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
               qed
             moreover note ih eq_cnp eq_cnv eq_threads
             ultimately show ?thesis by auto
@@ -1796,8 +1796,8 @@
               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
+                  apply (unfold eq_e step_RAG_v[OF vtv], 
+                         auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
                                    Let_def cs_holding_def)
                   by (insert wq_distinct[OF step_back_vt[OF vtv], of cs], auto simp:wq_def)
                 thus ?thesis by (simp add:cntCS_def)
@@ -1862,28 +1862,28 @@
                 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)
+                apply (unfold cntCS_def holdents_test eq_e step_RAG_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})"
+                show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
+                               Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
                   (is "card ?A = Suc (card ?B)")
                 proof -
                   have "?A = insert cs ?B" by auto
                   hence "card ?A = card (insert cs ?B)" by simp
                   also have "\<dots> = Suc (card ?B)"
                   proof(rule card_insert_disjoint)
-                    have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` depend s)" 
+                    have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
                       apply (auto simp:image_def)
                       by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
-                    with finite_depend[OF step_back_vt[OF vtv]]
-                    show "finite {cs. (Cs cs, Th th) \<in> depend s}" by (auto intro:finite_subset)
+                    with finite_RAG[OF step_back_vt[OF vtv]]
+                    show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
                   next
-                    show "cs \<notin> {cs. (Cs cs, Th th) \<in> depend s}"
+                    show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
                     proof
-                      assume "cs \<in> {cs. (Cs cs, Th th) \<in> depend s}"
-                      hence "(Cs cs, Th th) \<in> depend s" by simp
+                      assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
+                      hence "(Cs cs, Th th) \<in> RAG s" by simp
                       with True neq_th eq_wq show False
-                        by (auto simp:next_th_def s_depend_def cs_holding_def)
+                        by (auto simp:next_th_def s_RAG_def cs_holding_def)
                     qed
                   qed
                   finally show ?thesis .
@@ -1905,7 +1905,7 @@
         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)
+          by (simp add:RAG_set_unchanged eq_e)
         from eq_e have eq_readys: "readys (e#s) = readys s" 
           by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
                   auto simp:Let_def)
@@ -1931,7 +1931,7 @@
     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)
+        auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
   qed
 qed
 
@@ -1955,7 +1955,7 @@
         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)
+        by (simp add:RAG_create_unchanged)
       moreover have "th \<notin> threads s" 
       proof -
         from not_in eq_e show ?thesis by simp
@@ -1967,7 +1967,7 @@
       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)
+        by (simp add:RAG_exit_unchanged)
       show ?thesis
       proof(cases "th = thread")
         case True
@@ -1993,7 +1993,7 @@
       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)
+        by (unfold step_RAG_p[OF vtp], auto)
       moreover have "cntCS s th = 0"
       proof(rule ih)
         from not_in eq_e show "th \<notin> threads s" by simp
@@ -2035,7 +2035,7 @@
       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)
+        by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
       moreover have "cntCS s th = 0"
       proof(rule ih)
         from not_in eq_e show "th \<notin> threads s" by simp
@@ -2050,30 +2050,30 @@
       from ih [OF this] and eq_e
       show ?thesis 
         apply (unfold eq_e cntCS_def holdents_test)
-        by (simp add:depend_set_unchanged)
+        by (simp add:RAG_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)
+        auto simp:count_def holdents_test s_RAG_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:
+lemma dm_RAG_threads:
   fixes th s
   assumes vt: "vt s"
-  and in_dom: "(Th th) \<in> Domain (depend s)"
+  and in_dom: "(Th th) \<in> Domain (RAG 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
+  from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
+  moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
+  ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
   hence "th \<in> set (wq s cs)"
-    by (unfold s_depend_def, auto simp:cs_waiting_def)
+    by (unfold s_RAG_def, auto simp:cs_waiting_def)
   from wq_threads [OF vt this] show ?thesis .
 qed
 
@@ -2112,18 +2112,18 @@
       proof -
         have "finite (dependants (wq s) th1)"
         proof-
-          have "finite {th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+}"
+          have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
           proof -
             let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
+            have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
               apply (auto simp:image_def)
               by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
             moreover have "finite \<dots>"
             proof -
-              from finite_depend[OF vt] have "finite (depend s)" .
-              hence "finite ((depend (wq s))\<^sup>+)"
+              from finite_RAG[OF vt] have "finite (RAG s)" .
+              hence "finite ((RAG (wq s))\<^sup>+)"
                 apply (unfold finite_trancl)
-                by (auto simp: s_depend_def cs_depend_def wq_def)
+                by (auto simp: s_RAG_def cs_RAG_def wq_def)
               thus ?thesis by auto
             qed
             ultimately show ?thesis by (auto intro:finite_subset)
@@ -2151,18 +2151,18 @@
       proof -
         have "finite (dependants (wq s) th2)"
         proof-
-          have "finite {th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+}"
+          have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
           proof -
             let ?F = "\<lambda> (x, y). the_th x"
-            have "{th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
+            have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
               apply (auto simp:image_def)
               by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
             moreover have "finite \<dots>"
             proof -
-              from finite_depend[OF vt] have "finite (depend s)" .
-              hence "finite ((depend (wq s))\<^sup>+)"
+              from finite_RAG[OF vt] have "finite (RAG s)" .
+              hence "finite ((RAG (wq s))\<^sup>+)"
                 apply (unfold finite_trancl)
-                by (auto simp: s_depend_def cs_depend_def wq_def)
+                by (auto simp: s_RAG_def cs_RAG_def wq_def)
               thus ?thesis by auto
             qed
             ultimately show ?thesis by (auto intro:finite_subset)
@@ -2190,11 +2190,11 @@
     thus "th1' \<in> threads s"
     proof
       assume "th1' \<in> dependants (wq s) th1"
-      hence "(Th th1') \<in> Domain ((depend s)^+)"
-        apply (unfold cs_dependants_def cs_depend_def s_depend_def)
+      hence "(Th th1') \<in> Domain ((RAG s)^+)"
+        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
         by (auto simp:Domain_def)
-      hence "(Th th1') \<in> Domain (depend s)" by (simp add:trancl_domain)
-      from dm_depend_threads[OF vt this] show ?thesis .
+      hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
+      from dm_RAG_threads[OF vt this] show ?thesis .
     next
       assume "th1' = th1"
       with runing_1 show ?thesis
@@ -2205,11 +2205,11 @@
     thus "th2' \<in> threads s"
     proof
       assume "th2' \<in> dependants (wq s) th2"
-      hence "(Th th2') \<in> Domain ((depend s)^+)"
-        apply (unfold cs_dependants_def cs_depend_def s_depend_def)
+      hence "(Th th2') \<in> Domain ((RAG s)^+)"
+        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
         by (auto simp:Domain_def)
-      hence "(Th th2') \<in> Domain (depend s)" by (simp add:trancl_domain)
-      from dm_depend_threads[OF vt this] show ?thesis .
+      hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
+      from dm_RAG_threads[OF vt this] show ?thesis .
     next
       assume "th2' = th2"
       with runing_2 show ?thesis
@@ -2227,18 +2227,18 @@
     next
       assume "th2' \<in> dependants (wq s) th2"
       with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
-      hence "(Th th1, Th th2) \<in> (depend s)^+"
-        by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
-      hence "Th th1 \<in> Domain ((depend s)^+)" 
-        apply (unfold cs_dependants_def cs_depend_def s_depend_def)
+      hence "(Th th1, Th th2) \<in> (RAG s)^+"
+        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
+      hence "Th th1 \<in> Domain ((RAG s)^+)" 
+        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
         by (auto simp:Domain_def)
-      hence "Th th1 \<in> Domain (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]
+      hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
+      then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
+      from RAG_target_th [OF this]
       obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th1, Cs cs') \<in> depend s" by simp
+      with d have "(Th th1, Cs cs') \<in> RAG s" by simp
       with runing_1 have "False"
-        apply (unfold runing_def readys_def s_depend_def)
+        apply (unfold runing_def readys_def s_RAG_def)
         by (auto simp:eq_waiting)
       thus ?thesis by simp
     qed
@@ -2249,27 +2249,27 @@
     proof
       assume "th2' = th2"
       with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
-      hence "(Th th2, Th th1) \<in> (depend s)^+"
-        by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
-      hence "Th th2 \<in> Domain ((depend s)^+)" 
-        apply (unfold cs_dependants_def cs_depend_def s_depend_def)
+      hence "(Th th2, Th th1) \<in> (RAG s)^+"
+        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
+      hence "Th th2 \<in> Domain ((RAG s)^+)" 
+        apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
         by (auto simp:Domain_def)
-      hence "Th th2 \<in> Domain (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]
+      hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
+      then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
+      from RAG_target_th [OF this]
       obtain cs' where "n = Cs cs'" by auto
-      with d have "(Th th2, Cs cs') \<in> depend s" by simp
+      with d have "(Th th2, Cs cs') \<in> RAG s" by simp
       with runing_2 have "False"
-        apply (unfold runing_def readys_def s_depend_def)
+        apply (unfold runing_def readys_def s_RAG_def)
         by (auto simp:eq_waiting)
       thus ?thesis by simp
     next
       assume "th2' \<in> dependants (wq s) th2"
       with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
-      hence h1: "(Th th1', Th th2) \<in> (depend s)^+"
-        by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
-      from th1'_in have h2: "(Th th1', Th th1) \<in> (depend s)^+"
-        by (unfold cs_dependants_def s_depend_def cs_depend_def, simp)
+      hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
+        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
+      from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
+        by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       show ?thesis
       proof(rule dchain_unique[OF vt h1 _ h2, symmetric])
         from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
@@ -2411,9 +2411,9 @@
   qed
 qed
 
-lemma eq_depend: 
-  "depend (wq s) = depend s"
-by (unfold cs_depend_def s_depend_def, auto)
+lemma eq_RAG: 
+  "RAG (wq s) = RAG s"
+by (unfold cs_RAG_def s_RAG_def, auto)
 
 lemma count_eq_dependants:
   assumes vt: "vt s"
@@ -2423,28 +2423,28 @@
   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}"
+  moreover have "finite {cs. (Cs cs, Th th) \<in> RAG 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} = {}"
+  ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
     by (unfold cntCS_def holdents_test cs_dependants_def, auto)
   show ?thesis
   proof(unfold cs_dependants_def)
-    { assume "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}"
-      then obtain th' where "(Th th', Th th) \<in> (depend (wq s))\<^sup>+" by auto
+    { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
+      then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
       hence "False"
       proof(cases)
-        assume "(Th th', Th th) \<in> depend (wq s)"
-        thus "False" by (auto simp:cs_depend_def)
+        assume "(Th th', Th th) \<in> RAG (wq s)"
+        thus "False" by (auto simp:cs_RAG_def)
       next
         fix c
-        assume "(c, Th th) \<in> depend (wq s)"
-        with h and eq_depend show "False"
-          by (cases c, auto simp:cs_depend_def)
+        assume "(c, Th th) \<in> RAG (wq s)"
+        with h and eq_RAG show "False"
+          by (cases c, auto simp:cs_RAG_def)
       qed
-    } thus "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} = {}" by auto
+    } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
   qed
 qed
 
@@ -2454,20 +2454,20 @@
   shows "dependants (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)"
+    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
+    have "Th th \<in> Domain (RAG s)"
     proof -
-      from h obtain th' where "(Th th, Th th') \<in> (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
+      from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
+      hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
+      with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
+      thus ?thesis using eq_RAG by simp
     qed
-    from dm_depend_threads[OF vt this]
+    from dm_RAG_threads[OF vt this]
     have "th \<in> threads s" .
   } note hh = this
   fix th1 
   assume "th1 \<in> dependants (wq s) th"
-  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (depend (wq s))\<^sup>+}"
+  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
     by (unfold cs_dependants_def, simp)
   from hh [OF this] show "th1 \<in> threads s" .
 qed
@@ -2496,21 +2496,21 @@
   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_dependants_def)
-  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}))
+  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
          \<le> Max ((\<lambda>th. preced th s) ` threads s)"
     (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
   proof(rule Max_f_mono)
-    show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}" by simp
+    show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (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"
+    show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
       apply (auto simp:Domain_def)
-      apply (rule_tac dm_depend_threads[OF vt])
-      apply (unfold trancl_domain [of "depend s", symmetric])
-      by (unfold cs_depend_def s_depend_def, auto simp:Domain_def)
+      apply (rule_tac dm_RAG_threads[OF vt])
+      apply (unfold trancl_domain [of "RAG s", symmetric])
+      by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def)
   qed
 qed
 
@@ -2528,18 +2528,18 @@
     proof -
       have "finite ?B" 
       proof-
-        have "finite {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}"
+        have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
         proof -
           let ?F = "\<lambda> (x, y). the_th x"
-          have "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
+          have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
             apply (auto simp:image_def)
             by (rule_tac x = "(Th x, Th th)" in bexI, auto)
           moreover have "finite \<dots>"
           proof -
-            from finite_depend[OF vt] have "finite (depend s)" .
-            hence "finite ((depend (wq s))\<^sup>+)"
+            from finite_RAG[OF vt] have "finite (RAG s)" .
+            hence "finite ((RAG (wq s))\<^sup>+)"
               apply (unfold finite_trancl)
-              by (auto simp: s_depend_def cs_depend_def wq_def)
+              by (auto simp: s_RAG_def cs_RAG_def wq_def)
             thus ?thesis by auto
           qed
           ultimately show ?thesis by (auto intro:finite_subset)
@@ -2621,12 +2621,12 @@
     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>+)" .
+    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
     thus ?thesis
     proof
-      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+ "
+      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
       then obtain th' where th'_in: "th' \<in> readys s" 
-        and tm_chain:"(Th tm, Th th') \<in> (depend s)\<^sup>+" by auto
+        and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
       have "cp s th' = ?f tm"
       proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
         from dependants_threads[OF vt] finite_threads[OF vt]
@@ -2650,7 +2650,7 @@
         proof -
           from tm_chain
           have "tm \<in> dependants (wq s) th'"
-            by (unfold cs_dependants_def s_depend_def cs_depend_def, auto)
+            by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
           thus ?thesis by auto
         qed
       qed
@@ -2780,9 +2780,9 @@
 
 
 lemma detached_test:
-  shows "detached s th = (Th th \<notin> Field (depend s))"
+  shows "detached s th = (Th th \<notin> Field (RAG s))"
 apply(simp add: detached_def Field_def)
-apply(simp add: s_depend_def)
+apply(simp add: s_RAG_def)
 apply(simp add: s_holding_abv s_waiting_abv)
 apply(simp add: Domain_iff Range_iff)
 apply(simp add: wq_def)
@@ -2805,12 +2805,12 @@
   thus ?thesis
   proof
     assume "th \<notin> threads s"
-    with range_in[OF vt] dm_depend_threads[OF vt]
+    with range_in[OF vt] dm_RAG_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)
+      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
   next
     assume "th \<in> readys s"
-    moreover have "Th th \<notin> Range (depend s)"
+    moreover have "Th th \<notin> Range (RAG s)"
     proof -
       from card_0_eq [OF finite_holding [OF vt]] and cncs_zero
       have "holdents s th = {}"
@@ -2818,11 +2818,11 @@
       thus ?thesis
         apply(auto simp:holdents_test)
         apply(case_tac a)
-        apply(auto simp:holdents_test s_depend_def)
+        apply(auto simp:holdents_test s_RAG_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)
+      by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
   qed
 qed
 
@@ -2838,7 +2838,7 @@
   have cncs_z: "cntCS s th = 0"
   proof -
     from dtc have "holdents s th = {}"
-      unfolding detached_def holdents_test s_depend_def
+      unfolding detached_def holdents_test s_RAG_def
       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
     thus ?thesis by (auto simp:cntCS_def)
   qed
@@ -2848,7 +2848,7 @@
     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)
+           auto simp:eq_waiting s_RAG_def)
     with cncs_z and eq_pv show ?thesis by simp
   next
     case False
@@ -2862,4 +2862,4 @@
   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
+end
--- a/PrioGDef.thy	Tue May 20 12:49:21 2014 +0100
+++ b/PrioGDef.thy	Thu May 22 17:40:39 2014 +0100
@@ -121,7 +121,7 @@
 consts 
   holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool" 
   waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
-  depend :: "'b \<Rightarrow> (node \<times> node) set"
+  RAG :: "'b \<Rightarrow> (node \<times> node) set"
   dependants :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
 
 text {*
@@ -153,21 +153,21 @@
   "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 
+  @{text "RAG 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>
+  cs_RAG_def: 
+  "RAG (wq::cs \<Rightarrow> thread list) \<equiv>
       {(Th th, Cs cs) | th cs. waiting wq th cs} \<union> {(Cs cs, Th th) | cs th. holding wq th cs}"
   -- {* 
   \begin{minipage}{0.9\textwidth}
-  The following @{text "dependants wq th"} represents the set of threads which are depending on
-  thread @{text "th"} in Resource Allocation Graph @{text "depend wq"}:
+  The following @{text "dependants wq th"} represents the set of threads which are RAGing on
+  thread @{text "th"} in Resource Allocation Graph @{text "RAG wq"}:
   \end{minipage}
   *}
   cs_dependants_def: 
-  "dependants (wq::cs \<Rightarrow> thread list) th \<equiv> {th' . (Th th', Th th) \<in> (depend wq)^+}"
+  "dependants (wq::cs \<Rightarrow> thread list) th \<equiv> {th' . (Th th', Th th) \<in> (RAG wq)^+}"
 
 
 text {*
@@ -243,7 +243,7 @@
       \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 
+        function. The RAGency of precedence function on waiting queue function is the reason to 
         put them in the same record so that they can evolve together.
   \end{enumerate}
   \end{minipage}
@@ -269,7 +269,7 @@
 lemma cpreced_initial: 
   "cpreced (\<lambda> cs. []) [] = (\<lambda>_. (Prc 0 0))"
 apply(simp add: cpreced_def)
-apply(simp add: cs_dependants_def cs_depend_def cs_waiting_def cs_holding_def)
+apply(simp add: cs_dependants_def cs_RAG_def cs_waiting_def cs_holding_def)
 apply(simp add: preced_def)
 done
 
@@ -305,9 +305,9 @@
   where "cp s \<equiv> cprec_fun (schs s)"
 
 text {* \noindent
-  Functions @{text "holding"}, @{text "waiting"}, @{text "depend"} and 
+  Functions @{text "holding"}, @{text "waiting"}, @{text "RAG"} and 
   @{text "dependants"} still have the 
-  same meaning, but redefined so that they no longer depend on the 
+  same meaning, but redefined so that they no longer RAG on the 
   fictitious {\em waiting queue function}
   @{text "wq"}, but on system state @{text "s"}.
   *}
@@ -316,8 +316,8 @@
   "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_RAG_abv: 
+  "RAG (s::state) \<equiv> RAG (wq_fun (schs s))"
   s_dependants_abv: 
   "dependants (s::state) \<equiv> dependants (wq_fun (schs s))"
 
@@ -334,14 +334,14 @@
   "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) =
+lemma s_RAG_def: 
+  "RAG (s::state) =
     {(Th th, Cs cs) | th cs. waiting (wq s) th cs} \<union> {(Cs cs, Th th) | cs th. holding (wq s) th cs}"
-  by (auto simp:s_depend_abv wq_def cs_depend_def)
+  by (auto simp:s_RAG_abv wq_def cs_RAG_def)
 
 lemma
   s_dependants_def: 
-  "dependants (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (depend (wq s))^+}"
+  "dependants (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (RAG (wq s))^+}"
   by (auto simp:s_dependants_abv wq_def cs_dependants_def)
 
 text {*
@@ -366,9 +366,9 @@
   where "holdents s th \<equiv> {cs . holding s th cs}"
 
 lemma holdents_test: 
-  "holdents s th = {cs . (Cs cs, Th th) \<in> depend s}"
+  "holdents s th = {cs . (Cs cs, Th th) \<in> RAG s}"
 unfolding holdents_def
-unfolding s_depend_def
+unfolding s_RAG_def
 unfolding s_holding_abv
 unfolding wq_def
 by (simp)
@@ -403,7 +403,7 @@
   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> 
+  thread_P: "\<lbrakk>thread \<in> runing s;  (Cs cs, Th thread)  \<notin> (RAG s)^+\<rbrakk> \<Longrightarrow> 
                                                                 step s (P thread cs)" |
   -- {*
   \begin{minipage}{0.9\textwidth}
--- a/README	Tue May 20 12:49:21 2014 +0100
+++ b/README	Thu May 22 17:40:39 2014 +0100
@@ -1,6 +1,7 @@
 Theories:
 =========
 
+ Max.thy                 Some generic facts about Max.
  Precedence_ord.thy      A theory of precedences.
  Moment.thy              The notion of moment.
  PrioGDef.thy            The formal definition of the PIP-model.
@@ -8,6 +9,7 @@
  ExtGG.thy               The correctness proof of the PIP-model.
  CpsG.thy                Properties interesting for an implementation.
 
+
 The repository can be checked using Isabelle 2013-2.
 
   isabelle build -d . PIP
Binary file journal.pdf has changed