diff -r f7b33c633b96 -r 3d2b59f15f26 PIPDefs.thy --- a/PIPDefs.thy Sun Jan 31 18:15:13 2016 +0800 +++ b/PIPDefs.thy Mon Feb 01 20:56:39 2016 +0800 @@ -657,6 +657,11 @@ by (unfold s_RAG_abv wRAG_def hRAG_def s_waiting_abv s_holding_abv cs_RAG_def, auto) +lemma tRAG_alt_def: + "tRAG s = {(Th th1, Th th2) | th1 th2. + \ cs. (Th th1, Cs cs) \ RAG s \ (Cs cs, Th th2) \ RAG s}" + by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def) + definition "cp_gen s x = Max ((the_preced s \ the_thread) ` subtree (tRAG s) x)"