PIPBasics.thy~
changeset 97 c7ba70dc49bd
parent 96 4805c6333fef
parent 93 524bd3caa6b6
child 98 382293d415f3
equal deleted inserted replaced
96:4805c6333fef 97:c7ba70dc49bd
     1 theory PIPBasics
       
     2 imports PIPDefs 
       
     3 begin
       
     4 
       
     5 locale valid_trace = 
       
     6   fixes s
       
     7   assumes vt : "vt s"
       
     8 
       
     9 locale valid_trace_e = valid_trace +
       
    10   fixes e
       
    11   assumes vt_e: "vt (e#s)"
       
    12 begin
       
    13 
       
    14 lemma pip_e: "PIP s e"
       
    15   using vt_e by (cases, simp)  
       
    16 
       
    17 end
       
    18 
       
    19 lemma runing_ready: 
       
    20   shows "runing s \<subseteq> readys s"
       
    21   unfolding runing_def readys_def
       
    22   by auto 
       
    23 
       
    24 lemma readys_threads:
       
    25   shows "readys s \<subseteq> threads s"
       
    26   unfolding readys_def
       
    27   by auto
       
    28 
       
    29 lemma wq_v_neq:
       
    30    "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
       
    31   by (auto simp:wq_def Let_def cp_def split:list.splits)
       
    32 
       
    33 context valid_trace
       
    34 begin
       
    35 
       
    36 lemma actor_inv: 
       
    37   assumes "PIP s e"
       
    38   and "\<not> isCreate e"
       
    39   shows "actor e \<in> runing s"
       
    40   using assms
       
    41   by (induct, auto)
       
    42 
       
    43 lemma ind [consumes 0, case_names Nil Cons, induct type]:
       
    44   assumes "PP []"
       
    45      and "(\<And>s e. valid_trace s \<Longrightarrow> valid_trace (e#s) \<Longrightarrow>
       
    46                    PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
       
    47      shows "PP s"
       
    48 proof(rule vt.induct[OF vt])
       
    49   from assms(1) show "PP []" .
       
    50 next
       
    51   fix s e
       
    52   assume h: "vt s" "PP s" "PIP s e"
       
    53   show "PP (e # s)"
       
    54   proof(cases rule:assms(2))
       
    55     from h(1) show v1: "valid_trace s" by (unfold_locales, simp)
       
    56   next
       
    57     from h(1,3) have "vt (e#s)" by auto
       
    58     thus "valid_trace (e # s)" by (unfold_locales, simp)
       
    59   qed (insert h, auto)
       
    60 qed
       
    61 
       
    62 lemma wq_distinct: "distinct (wq s cs)"
       
    63 proof(rule ind, simp add:wq_def)
       
    64   fix s e
       
    65   assume h1: "step s e"
       
    66   and h2: "distinct (wq s cs)"
       
    67   thus "distinct (wq (e # s) cs)"
       
    68   proof(induct rule:step.induct, auto simp: wq_def Let_def split:list.splits)
       
    69     fix thread s
       
    70     assume h1: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
       
    71       and h2: "thread \<in> set (wq_fun (schs s) cs)"
       
    72       and h3: "thread \<in> runing s"
       
    73     show "False" 
       
    74     proof -
       
    75       from h3 have "\<And> cs. thread \<in>  set (wq_fun (schs s) cs) \<Longrightarrow> 
       
    76                              thread = hd ((wq_fun (schs s) cs))" 
       
    77         by (simp add:runing_def readys_def s_waiting_def wq_def)
       
    78       from this [OF h2] have "thread = hd (wq_fun (schs s) cs)" .
       
    79       with h2
       
    80       have "(Cs cs, Th thread) \<in> (RAG s)"
       
    81         by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def)
       
    82       with h1 show False by auto
       
    83     qed
       
    84   next
       
    85     fix thread s a list
       
    86     assume dst: "distinct list"
       
    87     show "distinct (SOME q. distinct q \<and> set q = set list)"
       
    88     proof(rule someI2)
       
    89       from dst show  "distinct list \<and> set list = set list" by auto
       
    90     next
       
    91       fix q assume "distinct q \<and> set q = set list"
       
    92       thus "distinct q" by auto
       
    93     qed
       
    94   qed
       
    95 qed
       
    96 
       
    97 end
       
    98 
       
    99 
       
   100 context valid_trace_e
       
   101 begin
       
   102 
       
   103 text {*
       
   104   The following lemma shows that only the @{text "P"}
       
   105   operation can add new thread into waiting queues. 
       
   106   Such kind of lemmas are very obvious, but need to be checked formally.
       
   107   This is a kind of confirmation that our modelling is correct.
       
   108 *}
       
   109 
       
   110 lemma block_pre: 
       
   111   assumes s_ni: "thread \<notin>  set (wq s cs)"
       
   112   and s_i: "thread \<in> set (wq (e#s) cs)"
       
   113   shows "e = P thread cs"
       
   114 proof -
       
   115   show ?thesis
       
   116   proof(cases e)
       
   117     case (P th cs)
       
   118     with assms
       
   119     show ?thesis
       
   120       by (auto simp:wq_def Let_def split:if_splits)
       
   121   next
       
   122     case (Create th prio)
       
   123     with assms show ?thesis
       
   124       by (auto simp:wq_def Let_def split:if_splits)
       
   125   next
       
   126     case (Exit th)
       
   127     with assms show ?thesis
       
   128       by (auto simp:wq_def Let_def split:if_splits)
       
   129   next
       
   130     case (Set th prio)
       
   131     with assms show ?thesis
       
   132       by (auto simp:wq_def Let_def split:if_splits)
       
   133   next
       
   134     case (V th cs)
       
   135     with vt_e assms show ?thesis
       
   136       apply (auto simp:wq_def Let_def split:if_splits)
       
   137     proof -
       
   138       fix q qs
       
   139       assume h1: "thread \<notin> set (wq_fun (schs s) cs)"
       
   140         and h2: "q # qs = wq_fun (schs s) cs"
       
   141         and h3: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
       
   142         and vt: "vt (V th cs # s)"
       
   143       from h1 and h2[symmetric] have "thread \<notin> set (q # qs)" by simp
       
   144       moreover have "thread \<in> set qs"
       
   145       proof -
       
   146         have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
       
   147         proof(rule someI2)
       
   148           from wq_distinct [of cs]
       
   149           and h2[symmetric, folded wq_def]
       
   150           show "distinct qs \<and> set qs = set qs" by auto
       
   151         next
       
   152           fix x assume "distinct x \<and> set x = set qs"
       
   153           thus "set x = set qs" by auto
       
   154         qed
       
   155         with h3 show ?thesis by simp
       
   156       qed
       
   157       ultimately show "False" by auto
       
   158       qed
       
   159   qed
       
   160 qed
       
   161 
       
   162 end
       
   163 
       
   164 text {*
       
   165   The following lemmas is also obvious and shallow. It says
       
   166   that only running thread can request for a critical resource 
       
   167   and that the requested resource must be one which is
       
   168   not current held by the thread.
       
   169 *}
       
   170 
       
   171 lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
       
   172   thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (RAG s)^+"
       
   173 apply (ind_cases "vt ((P thread cs)#s)")
       
   174 apply (ind_cases "step s (P thread cs)")
       
   175 by auto
       
   176 
       
   177 lemma abs1:
       
   178   assumes ein: "e \<in> set es"
       
   179   and neq: "hd es \<noteq> hd (es @ [x])"
       
   180   shows "False"
       
   181 proof -
       
   182   from ein have "es \<noteq> []" by auto
       
   183   then obtain e ess where "es = e # ess" by (cases es, auto)
       
   184   with neq show ?thesis by auto
       
   185 qed
       
   186 
       
   187 lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]"
       
   188   by (cases es, auto)
       
   189 
       
   190 inductive_cases evt_cons: "vt (a#s)"
       
   191 
       
   192 context valid_trace_e
       
   193 begin
       
   194 
       
   195 lemma abs2:
       
   196   assumes inq: "thread \<in> set (wq s cs)"
       
   197   and nh: "thread = hd (wq s cs)"
       
   198   and qt: "thread \<noteq> hd (wq (e#s) cs)"
       
   199   and inq': "thread \<in> set (wq (e#s) cs)"
       
   200   shows "False"
       
   201 proof -
       
   202   from vt_e assms show "False"
       
   203     apply (cases e)
       
   204     apply ((simp split:if_splits add:Let_def wq_def)[1])+
       
   205     apply (insert abs1, fast)[1]
       
   206     apply (auto simp:wq_def simp:Let_def split:if_splits list.splits)
       
   207   proof -
       
   208     fix th qs
       
   209     assume vt: "vt (V th cs # s)"
       
   210       and th_in: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
       
   211       and eq_wq: "wq_fun (schs s) cs = thread # qs"
       
   212     show "False"
       
   213     proof -
       
   214       from wq_distinct[of cs]
       
   215         and eq_wq[folded wq_def] have "distinct (thread#qs)" by simp
       
   216       moreover have "thread \<in> set qs"
       
   217       proof -
       
   218         have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
       
   219         proof(rule someI2)
       
   220           from wq_distinct [of cs]
       
   221           and eq_wq [folded wq_def]
       
   222           show "distinct qs \<and> set qs = set qs" by auto
       
   223         next
       
   224           fix x assume "distinct x \<and> set x = set qs"
       
   225           thus "set x = set qs" by auto
       
   226         qed
       
   227         with th_in show ?thesis by auto
       
   228       qed
       
   229       ultimately show ?thesis by auto
       
   230     qed
       
   231   qed
       
   232 qed
       
   233 
       
   234 end
       
   235 
       
   236 context valid_trace
       
   237 begin
       
   238 
       
   239 lemma vt_moment: "\<And> t. vt (moment t s)"
       
   240 proof(induct rule:ind)
       
   241   case Nil
       
   242   thus ?case by (simp add:vt_nil)
       
   243 next
       
   244   case (Cons s e t)
       
   245   show ?case
       
   246   proof(cases "t \<ge> length (e#s)")
       
   247     case True
       
   248     from True have "moment t (e#s) = e#s" by simp
       
   249     thus ?thesis using Cons
       
   250       by (simp add:valid_trace_def)
       
   251   next
       
   252     case False
       
   253     from Cons have "vt (moment t s)" by simp
       
   254     moreover have "moment t (e#s) = moment t s"
       
   255     proof -
       
   256       from False have "t \<le> length s" by simp
       
   257       from moment_app [OF this, of "[e]"] 
       
   258       show ?thesis by simp
       
   259     qed
       
   260     ultimately show ?thesis by simp
       
   261   qed
       
   262 qed
       
   263 
       
   264 (* Wrong:
       
   265     lemma \<lbrakk>thread \<in> set (wq_fun cs1 s); thread \<in> set (wq_fun cs2 s)\<rbrakk> \<Longrightarrow> cs1 = cs2"
       
   266 *)
       
   267 
       
   268 text {* (* ddd *)
       
   269   The nature of the work is like this: since it starts from a very simple and basic 
       
   270   model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
       
   271   For instance, the fact 
       
   272   that one thread can not be blocked by two critical resources at the same time
       
   273   is obvious, because only running threads can make new requests, if one is waiting for 
       
   274   a critical resource and get blocked, it can not make another resource request and get 
       
   275   blocked the second time (because it is not running). 
       
   276 
       
   277   To derive this fact, one needs to prove by contraction and 
       
   278   reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
       
   279   named @{text "p_split"}, which is about status changing along the time axis. It says if 
       
   280   a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
       
   281   but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
       
   282   in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
       
   283   of events leading to it), such that @{text "Q"} switched 
       
   284   from being @{text "False"} to @{text "True"} and kept being @{text "True"}
       
   285   till the last moment of @{text "s"}.
       
   286 
       
   287   Suppose a thread @{text "th"} is blocked
       
   288   on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
       
   289   since no thread is blocked at the very beginning, by applying 
       
   290   @{text "p_split"} to these two blocking facts, there exist 
       
   291   two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
       
   292   @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
       
   293   and kept on blocked on them respectively ever since.
       
   294  
       
   295   Without lose of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
       
   296   However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
       
   297   in blocked state at moment @{text "t2"} and could not
       
   298   make any request and get blocked the second time: Contradiction.
       
   299 *}
       
   300 
       
   301 lemma waiting_unique_pre:
       
   302   assumes h11: "thread \<in> set (wq s cs1)"
       
   303   and h12: "thread \<noteq> hd (wq s cs1)"
       
   304   assumes h21: "thread \<in> set (wq s cs2)"
       
   305   and h22: "thread \<noteq> hd (wq s cs2)"
       
   306   and neq12: "cs1 \<noteq> cs2"
       
   307   shows "False"
       
   308 proof -
       
   309   let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
       
   310   from h11 and h12 have q1: "?Q cs1 s" by simp
       
   311   from h21 and h22 have q2: "?Q cs2 s" by simp
       
   312   have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
       
   313   have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
       
   314   from p_split [of "?Q cs1", OF q1 nq1]
       
   315   obtain t1 where lt1: "t1 < length s"
       
   316     and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and>
       
   317         thread \<noteq> hd (wq (moment t1 s) cs1))"
       
   318     and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and>
       
   319              thread \<noteq> hd (wq (moment i' s) cs1))" by auto
       
   320   from p_split [of "?Q cs2", OF q2 nq2]
       
   321   obtain t2 where lt2: "t2 < length s"
       
   322     and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and>
       
   323         thread \<noteq> hd (wq (moment t2 s) cs2))"
       
   324     and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and>
       
   325              thread \<noteq> hd (wq (moment i' s) cs2))" by auto
       
   326   show ?thesis
       
   327   proof -
       
   328     { 
       
   329       assume lt12: "t1 < t2"
       
   330       let ?t3 = "Suc t2"
       
   331       from lt2 have le_t3: "?t3 \<le> length s" by auto
       
   332       from moment_plus [OF this] 
       
   333       obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
       
   334       have "t2 < ?t3" by simp
       
   335       from nn2 [rule_format, OF this] and eq_m
       
   336       have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   337         h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   338       have "vt (e#moment t2 s)"
       
   339       proof -
       
   340         from vt_moment 
       
   341         have "vt (moment ?t3 s)" .
       
   342         with eq_m show ?thesis by simp
       
   343       qed
       
   344       then interpret vt_e: valid_trace_e "moment t2 s" "e"
       
   345         by (unfold_locales, auto, cases, simp)
       
   346       have ?thesis
       
   347       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   348         case True
       
   349         from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
       
   350           by auto 
       
   351         from vt_e.abs2 [OF True eq_th h2 h1]
       
   352         show ?thesis by auto
       
   353       next
       
   354         case False
       
   355         from vt_e.block_pre[OF False h1]
       
   356         have "e = P thread cs2" .
       
   357         with vt_e.vt_e have "vt ((P thread cs2)# moment t2 s)" by simp
       
   358         from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp
       
   359         with runing_ready have "thread \<in> readys (moment t2 s)" by auto
       
   360         with nn1 [rule_format, OF lt12]
       
   361         show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
       
   362       qed
       
   363     } moreover {
       
   364       assume lt12: "t2 < t1"
       
   365       let ?t3 = "Suc t1"
       
   366       from lt1 have le_t3: "?t3 \<le> length s" by auto
       
   367       from moment_plus [OF this] 
       
   368       obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
       
   369       have lt_t3: "t1 < ?t3" by simp
       
   370       from nn1 [rule_format, OF this] and eq_m
       
   371       have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   372         h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
       
   373       have "vt  (e#moment t1 s)"
       
   374       proof -
       
   375         from vt_moment
       
   376         have "vt (moment ?t3 s)" .
       
   377         with eq_m show ?thesis by simp
       
   378       qed
       
   379       then interpret vt_e: valid_trace_e "moment t1 s" e
       
   380         by (unfold_locales, auto, cases, auto)
       
   381       have ?thesis
       
   382       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
       
   383         case True
       
   384         from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
       
   385           by auto
       
   386         from vt_e.abs2 True eq_th h2 h1
       
   387         show ?thesis by auto
       
   388       next
       
   389         case False
       
   390         from vt_e.block_pre [OF False h1]
       
   391         have "e = P thread cs1" .
       
   392         with vt_e.vt_e have "vt ((P thread cs1)# moment t1 s)" by simp
       
   393         from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp
       
   394         with runing_ready have "thread \<in> readys (moment t1 s)" by auto
       
   395         with nn2 [rule_format, OF lt12]
       
   396         show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
       
   397       qed
       
   398     } moreover {
       
   399       assume eqt12: "t1 = t2"
       
   400       let ?t3 = "Suc t1"
       
   401       from lt1 have le_t3: "?t3 \<le> length s" by auto
       
   402       from moment_plus [OF this] 
       
   403       obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
       
   404       have lt_t3: "t1 < ?t3" by simp
       
   405       from nn1 [rule_format, OF this] and eq_m
       
   406       have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   407         h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
       
   408       have vt_e: "vt (e#moment t1 s)"
       
   409       proof -
       
   410         from vt_moment
       
   411         have "vt (moment ?t3 s)" .
       
   412         with eq_m show ?thesis by simp
       
   413       qed
       
   414       then interpret vt_e: valid_trace_e "moment t1 s" e
       
   415         by (unfold_locales, auto, cases, auto)
       
   416       have ?thesis
       
   417       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
       
   418         case True
       
   419         from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
       
   420           by auto
       
   421         from vt_e.abs2 [OF True eq_th h2 h1]
       
   422         show ?thesis by auto
       
   423       next
       
   424         case False
       
   425         from vt_e.block_pre [OF False h1]
       
   426         have eq_e1: "e = P thread cs1" .
       
   427         have lt_t3: "t1 < ?t3" by simp
       
   428         with eqt12 have "t2 < ?t3" by simp
       
   429         from nn2 [rule_format, OF this] and eq_m and eqt12
       
   430         have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   431           h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   432         show ?thesis
       
   433         proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   434           case True
       
   435           from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
       
   436             by auto
       
   437           from vt_e and eqt12 have "vt (e#moment t2 s)" by simp 
       
   438           then interpret vt_e2: valid_trace_e "moment t2 s" e
       
   439             by (unfold_locales, auto, cases, auto)
       
   440           from vt_e2.abs2 [OF True eq_th h2 h1]
       
   441           show ?thesis .
       
   442         next
       
   443           case False
       
   444           have "vt (e#moment t2 s)"
       
   445           proof -
       
   446             from vt_moment eqt12
       
   447             have "vt (moment (Suc t2) s)" by auto
       
   448             with eq_m eqt12 show ?thesis by simp
       
   449           qed
       
   450           then interpret vt_e2: valid_trace_e "moment t2 s" e
       
   451             by (unfold_locales, auto, cases, auto)
       
   452           from vt_e2.block_pre [OF False h1]
       
   453           have "e = P thread cs2" .
       
   454           with eq_e1 neq12 show ?thesis by auto
       
   455         qed
       
   456       qed
       
   457     } ultimately show ?thesis by arith
       
   458   qed
       
   459 qed
       
   460 
       
   461 text {*
       
   462   This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
       
   463 *}
       
   464 
       
   465 lemma waiting_unique:
       
   466   assumes "waiting s th cs1"
       
   467   and "waiting s th cs2"
       
   468   shows "cs1 = cs2"
       
   469 using waiting_unique_pre assms
       
   470 unfolding wq_def s_waiting_def
       
   471 by auto
       
   472 
       
   473 end
       
   474 
       
   475 (* not used *)
       
   476 text {*
       
   477   Every thread can only be blocked on one critical resource, 
       
   478   symmetrically, every critical resource can only be held by one thread. 
       
   479   This fact is much more easier according to our definition. 
       
   480 *}
       
   481 lemma held_unique:
       
   482   assumes "holding (s::event list) th1 cs"
       
   483   and "holding s th2 cs"
       
   484   shows "th1 = th2"
       
   485  by (insert assms, unfold s_holding_def, auto)
       
   486 
       
   487 
       
   488 lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
       
   489   apply (induct s, auto)
       
   490   by (case_tac a, auto split:if_splits)
       
   491 
       
   492 lemma last_set_unique: 
       
   493   "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
       
   494           \<Longrightarrow> th1 = th2"
       
   495   apply (induct s, auto)
       
   496   by (case_tac a, auto split:if_splits dest:last_set_lt)
       
   497 
       
   498 lemma preced_unique : 
       
   499   assumes pcd_eq: "preced th1 s = preced th2 s"
       
   500   and th_in1: "th1 \<in> threads s"
       
   501   and th_in2: " th2 \<in> threads s"
       
   502   shows "th1 = th2"
       
   503 proof -
       
   504   from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
       
   505   from last_set_unique [OF this th_in1 th_in2]
       
   506   show ?thesis .
       
   507 qed
       
   508 
       
   509 lemma preced_linorder: 
       
   510   assumes neq_12: "th1 \<noteq> th2"
       
   511   and th_in1: "th1 \<in> threads s"
       
   512   and th_in2: " th2 \<in> threads s"
       
   513   shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
       
   514 proof -
       
   515   from preced_unique [OF _ th_in1 th_in2] and neq_12 
       
   516   have "preced th1 s \<noteq> preced th2 s" by auto
       
   517   thus ?thesis by auto
       
   518 qed
       
   519 
       
   520 (* An aux lemma used later *)
       
   521 lemma unique_minus:
       
   522   fixes x y z r
       
   523   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   524   and xy: "(x, y) \<in> r"
       
   525   and xz: "(x, z) \<in> r^+"
       
   526   and neq: "y \<noteq> z"
       
   527   shows "(y, z) \<in> r^+"
       
   528 proof -
       
   529  from xz and neq show ?thesis
       
   530  proof(induct)
       
   531    case (base ya)
       
   532    have "(x, ya) \<in> r" by fact
       
   533    from unique [OF xy this] have "y = ya" .
       
   534    with base show ?case by auto
       
   535  next
       
   536    case (step ya z)
       
   537    show ?case
       
   538    proof(cases "y = ya")
       
   539      case True
       
   540      from step True show ?thesis by simp
       
   541    next
       
   542      case False
       
   543      from step False
       
   544      show ?thesis by auto
       
   545    qed
       
   546  qed
       
   547 qed
       
   548 
       
   549 lemma unique_base:
       
   550   fixes r x y z
       
   551   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   552   and xy: "(x, y) \<in> r"
       
   553   and xz: "(x, z) \<in> r^+"
       
   554   and neq_yz: "y \<noteq> z"
       
   555   shows "(y, z) \<in> r^+"
       
   556 proof -
       
   557   from xz neq_yz show ?thesis
       
   558   proof(induct)
       
   559     case (base ya)
       
   560     from xy unique base show ?case by auto
       
   561   next
       
   562     case (step ya z)
       
   563     show ?case
       
   564     proof(cases "y = ya")
       
   565       case True
       
   566       from True step show ?thesis by auto
       
   567     next
       
   568       case False
       
   569       from False step 
       
   570       have "(y, ya) \<in> r\<^sup>+" by auto
       
   571       with step show ?thesis by auto
       
   572     qed
       
   573   qed
       
   574 qed
       
   575 
       
   576 lemma unique_chain:
       
   577   fixes r x y z
       
   578   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   579   and xy: "(x, y) \<in> r^+"
       
   580   and xz: "(x, z) \<in> r^+"
       
   581   and neq_yz: "y \<noteq> z"
       
   582   shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
       
   583 proof -
       
   584   from xy xz neq_yz show ?thesis
       
   585   proof(induct)
       
   586     case (base y)
       
   587     have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
       
   588     from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
       
   589   next
       
   590     case (step y za)
       
   591     show ?case
       
   592     proof(cases "y = z")
       
   593       case True
       
   594       from True step show ?thesis by auto
       
   595     next
       
   596       case False
       
   597       from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
       
   598       thus ?thesis
       
   599       proof
       
   600         assume "(z, y) \<in> r\<^sup>+"
       
   601         with step have "(z, za) \<in> r\<^sup>+" by auto
       
   602         thus ?thesis by auto
       
   603       next
       
   604         assume h: "(y, z) \<in> r\<^sup>+"
       
   605         from step have yza: "(y, za) \<in> r" by simp
       
   606         from step have "za \<noteq> z" by simp
       
   607         from unique_minus [OF _ yza h this] and unique
       
   608         have "(za, z) \<in> r\<^sup>+" by auto
       
   609         thus ?thesis by auto
       
   610       qed
       
   611     qed
       
   612   qed
       
   613 qed
       
   614 
       
   615 text {*
       
   616   The following three lemmas show that @{text "RAG"} does not change
       
   617   by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
       
   618   events, respectively.
       
   619 *}
       
   620 
       
   621 lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
       
   622 apply (unfold s_RAG_def s_waiting_def wq_def)
       
   623 by (simp add:Let_def)
       
   624 
       
   625 lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
       
   626 apply (unfold s_RAG_def s_waiting_def wq_def)
       
   627 by (simp add:Let_def)
       
   628 
       
   629 lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
       
   630 apply (unfold s_RAG_def s_waiting_def wq_def)
       
   631 by (simp add:Let_def)
       
   632 
       
   633 
       
   634 text {* 
       
   635   The following lemmas are used in the proof of 
       
   636   lemma @{text "step_RAG_v"}, which characterizes how the @{text "RAG"} is changed
       
   637   by @{text "V"}-events. 
       
   638   However, since our model is very concise, such  seemingly obvious lemmas need to be derived from scratch,
       
   639   starting from the model definitions.
       
   640 *}
       
   641 lemma step_v_hold_inv[elim_format]:
       
   642   "\<And>c t. \<lbrakk>vt (V th cs # s); 
       
   643           \<not> holding (wq s) t c; holding (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> 
       
   644             next_th s th cs t \<and> c = cs"
       
   645 proof -
       
   646   fix c t
       
   647   assume vt: "vt (V th cs # s)"
       
   648     and nhd: "\<not> holding (wq s) t c"
       
   649     and hd: "holding (wq (V th cs # s)) t c"
       
   650   show "next_th s th cs t \<and> c = cs"
       
   651   proof(cases "c = cs")
       
   652     case False
       
   653     with nhd hd show ?thesis
       
   654       by (unfold cs_holding_def wq_def, auto simp:Let_def)
       
   655   next
       
   656     case True
       
   657     with step_back_step [OF vt] 
       
   658     have "step s (V th c)" by simp
       
   659     hence "next_th s th cs t"
       
   660     proof(cases)
       
   661       assume "holding s th c"
       
   662       with nhd hd show ?thesis
       
   663         apply (unfold s_holding_def cs_holding_def wq_def next_th_def,
       
   664                auto simp:Let_def split:list.splits if_splits)
       
   665         proof -
       
   666           assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
       
   667           moreover have "\<dots> = set []"
       
   668           proof(rule someI2)
       
   669             show "distinct [] \<and> [] = []" by auto
       
   670           next
       
   671             fix x assume "distinct x \<and> x = []"
       
   672             thus "set x = set []" by auto
       
   673           qed
       
   674           ultimately show False by auto
       
   675         next
       
   676           assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
       
   677           moreover have "\<dots> = set []"
       
   678           proof(rule someI2)
       
   679             show "distinct [] \<and> [] = []" by auto
       
   680           next
       
   681             fix x assume "distinct x \<and> x = []"
       
   682             thus "set x = set []" by auto
       
   683           qed
       
   684           ultimately show False by auto
       
   685         qed
       
   686     qed
       
   687     with True show ?thesis by auto
       
   688   qed
       
   689 qed
       
   690 
       
   691 text {* 
       
   692   The following @{text "step_v_wait_inv"} is also an obvious lemma, which, however, needs to be
       
   693   derived from scratch, which confirms the correctness of the definition of @{text "next_th"}.
       
   694 *}
       
   695 lemma step_v_wait_inv[elim_format]:
       
   696     "\<And>t c. \<lbrakk>vt (V th cs # s); \<not> waiting (wq (V th cs # s)) t c; waiting (wq s) t c
       
   697            \<rbrakk>
       
   698           \<Longrightarrow> (next_th s th cs t \<and> cs = c)"
       
   699 proof -
       
   700   fix t c 
       
   701   assume vt: "vt (V th cs # s)"
       
   702     and nw: "\<not> waiting (wq (V th cs # s)) t c"
       
   703     and wt: "waiting (wq s) t c"
       
   704   from vt interpret vt_v: valid_trace_e s "V th cs" 
       
   705     by  (cases, unfold_locales, simp)
       
   706   show "next_th s th cs t \<and> cs = c"
       
   707   proof(cases "cs = c")
       
   708     case False
       
   709     with nw wt show ?thesis
       
   710       by (auto simp:cs_waiting_def wq_def Let_def)
       
   711   next
       
   712     case True
       
   713     from nw[folded True] wt[folded True]
       
   714     have "next_th s th cs t"
       
   715       apply (unfold next_th_def, auto simp:cs_waiting_def wq_def Let_def split:list.splits)
       
   716     proof -
       
   717       fix a list
       
   718       assume t_in: "t \<in> set list"
       
   719         and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
       
   720         and eq_wq: "wq_fun (schs s) cs = a # list"
       
   721       have " set (SOME q. distinct q \<and> set q = set list) = set list"
       
   722       proof(rule someI2)
       
   723         from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
       
   724         show "distinct list \<and> set list = set list" by auto
       
   725       next
       
   726         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
       
   727           by auto
       
   728       qed
       
   729       with t_ni and t_in show "a = th" by auto
       
   730     next
       
   731       fix a list
       
   732       assume t_in: "t \<in> set list"
       
   733         and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
       
   734         and eq_wq: "wq_fun (schs s) cs = a # list"
       
   735       have " set (SOME q. distinct q \<and> set q = set list) = set list"
       
   736       proof(rule someI2)
       
   737         from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
       
   738         show "distinct list \<and> set list = set list" by auto
       
   739       next
       
   740         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
       
   741           by auto
       
   742       qed
       
   743       with t_ni and t_in show "t = hd (SOME q. distinct q \<and> set q = set list)" by auto
       
   744     next
       
   745       fix a list
       
   746       assume eq_wq: "wq_fun (schs s) cs = a # list"
       
   747       from step_back_step[OF vt]
       
   748       show "a = th"
       
   749       proof(cases)
       
   750         assume "holding s th cs"
       
   751         with eq_wq show ?thesis
       
   752           by (unfold s_holding_def wq_def, auto)
       
   753       qed
       
   754     qed
       
   755     with True show ?thesis by simp
       
   756   qed
       
   757 qed
       
   758 
       
   759 lemma step_v_not_wait[consumes 3]:
       
   760   "\<lbrakk>vt (V th cs # s); next_th s th cs t; waiting (wq (V th cs # s)) t cs\<rbrakk> \<Longrightarrow> False"
       
   761   by (unfold next_th_def cs_waiting_def wq_def, auto simp:Let_def)
       
   762 
       
   763 lemma step_v_release:
       
   764   "\<lbrakk>vt (V th cs # s); holding (wq (V th cs # s)) th cs\<rbrakk> \<Longrightarrow> False"
       
   765 proof -
       
   766   assume vt: "vt (V th cs # s)"
       
   767     and hd: "holding (wq (V th cs # s)) th cs"
       
   768   from vt interpret vt_v: valid_trace_e s "V th cs"
       
   769     by (cases, unfold_locales, simp+)
       
   770   from step_back_step [OF vt] and hd
       
   771   show "False"
       
   772   proof(cases)
       
   773     assume "holding (wq (V th cs # s)) th cs" and "holding s th cs"
       
   774     thus ?thesis
       
   775       apply (unfold s_holding_def wq_def cs_holding_def)
       
   776       apply (auto simp:Let_def split:list.splits)
       
   777     proof -
       
   778       fix list
       
   779       assume eq_wq[folded wq_def]: 
       
   780         "wq_fun (schs s) cs = hd (SOME q. distinct q \<and> set q = set list) # list"
       
   781       and hd_in: "hd (SOME q. distinct q \<and> set q = set list)
       
   782             \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   783       have "set (SOME q. distinct q \<and> set q = set list) = set list"
       
   784       proof(rule someI2)
       
   785         from vt_v.wq_distinct[of cs] and eq_wq
       
   786         show "distinct list \<and> set list = set list" by auto
       
   787       next
       
   788         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
       
   789           by auto
       
   790       qed
       
   791       moreover have "distinct  (hd (SOME q. distinct q \<and> set q = set list) # list)"
       
   792       proof -
       
   793         from vt_v.wq_distinct[of cs] and eq_wq
       
   794         show ?thesis by auto
       
   795       qed
       
   796       moreover note eq_wq and hd_in
       
   797       ultimately show "False" by auto
       
   798     qed
       
   799   qed
       
   800 qed
       
   801 
       
   802 lemma step_v_get_hold:
       
   803   "\<And>th'. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) th' cs; next_th s th cs th'\<rbrakk> \<Longrightarrow> False"
       
   804   apply (unfold cs_holding_def next_th_def wq_def,
       
   805          auto simp:Let_def)
       
   806 proof -
       
   807   fix rest
       
   808   assume vt: "vt (V th cs # s)"
       
   809     and eq_wq[folded wq_def]: " wq_fun (schs s) cs = th # rest"
       
   810     and nrest: "rest \<noteq> []"
       
   811     and ni: "hd (SOME q. distinct q \<and> set q = set rest)
       
   812             \<notin> set (SOME q. distinct q \<and> set q = set rest)"
       
   813   from vt interpret vt_v: valid_trace_e s "V th cs"
       
   814     by (cases, unfold_locales, simp+)
       
   815   have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
   816   proof(rule someI2)
       
   817     from vt_v.wq_distinct[of cs] and eq_wq
       
   818     show "distinct rest \<and> set rest = set rest" by auto
       
   819   next
       
   820     fix x assume "distinct x \<and> set x = set rest"
       
   821     hence "set x = set rest" by auto
       
   822     with nrest
       
   823     show "x \<noteq> []" by (case_tac x, auto)
       
   824   qed
       
   825   with ni show "False" by auto
       
   826 qed
       
   827 
       
   828 lemma step_v_release_inv[elim_format]:
       
   829 "\<And>c t. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> 
       
   830   c = cs \<and> t = th"
       
   831   apply (unfold cs_holding_def wq_def, auto simp:Let_def split:if_splits list.splits)
       
   832   proof -
       
   833     fix a list
       
   834     assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
       
   835     from step_back_step [OF vt] show "a = th"
       
   836     proof(cases)
       
   837       assume "holding s th cs" with eq_wq
       
   838       show ?thesis
       
   839         by (unfold s_holding_def wq_def, auto)
       
   840     qed
       
   841   next
       
   842     fix a list
       
   843     assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
       
   844     from step_back_step [OF vt] show "a = th"
       
   845     proof(cases)
       
   846       assume "holding s th cs" with eq_wq
       
   847       show ?thesis
       
   848         by (unfold s_holding_def wq_def, auto)
       
   849     qed
       
   850   qed
       
   851 
       
   852 lemma step_v_waiting_mono:
       
   853   "\<And>t c. \<lbrakk>vt (V th cs # s); waiting (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> waiting (wq s) t c"
       
   854 proof -
       
   855   fix t c
       
   856   let ?s' = "(V th cs # s)"
       
   857   assume vt: "vt ?s'" 
       
   858     and wt: "waiting (wq ?s') t c"
       
   859   from vt interpret vt_v: valid_trace_e s "V th cs"
       
   860     by (cases, unfold_locales, simp+)
       
   861   show "waiting (wq s) t c"
       
   862   proof(cases "c = cs")
       
   863     case False
       
   864     assume neq_cs: "c \<noteq> cs"
       
   865     hence "waiting (wq ?s') t c = waiting (wq s) t c"
       
   866       by (unfold cs_waiting_def wq_def, auto simp:Let_def)
       
   867     with wt show ?thesis by simp
       
   868   next
       
   869     case True
       
   870     with wt show ?thesis
       
   871       apply (unfold cs_waiting_def wq_def, auto simp:Let_def split:list.splits)
       
   872     proof -
       
   873       fix a list
       
   874       assume not_in: "t \<notin> set list"
       
   875         and is_in: "t \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   876         and eq_wq: "wq_fun (schs s) cs = a # list"
       
   877       have "set (SOME q. distinct q \<and> set q = set list) = set list"
       
   878       proof(rule someI2)
       
   879         from vt_v.wq_distinct [of cs]
       
   880         and eq_wq[folded wq_def]
       
   881         show "distinct list \<and> set list = set list" by auto
       
   882       next
       
   883         fix x assume "distinct x \<and> set x = set list"
       
   884         thus "set x = set list" by auto
       
   885       qed
       
   886       with not_in is_in show "t = a" by auto
       
   887     next
       
   888       fix list
       
   889       assume is_waiting: "waiting (wq (V th cs # s)) t cs"
       
   890       and eq_wq: "wq_fun (schs s) cs = t # list"
       
   891       hence "t \<in> set list"
       
   892         apply (unfold wq_def, auto simp:Let_def cs_waiting_def)
       
   893       proof -
       
   894         assume " t \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   895         moreover have "\<dots> = set list" 
       
   896         proof(rule someI2)
       
   897           from vt_v.wq_distinct [of cs]
       
   898             and eq_wq[folded wq_def]
       
   899           show "distinct list \<and> set list = set list" by auto
       
   900         next
       
   901           fix x assume "distinct x \<and> set x = set list" 
       
   902           thus "set x = set list" by auto
       
   903         qed
       
   904         ultimately show "t \<in> set list" by simp
       
   905       qed
       
   906       with eq_wq and vt_v.wq_distinct [of cs, unfolded wq_def]
       
   907       show False by auto
       
   908     qed
       
   909   qed
       
   910 qed
       
   911 
       
   912 text {* (* ddd *) 
       
   913   The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed
       
   914   with the happening of @{text "V"}-events:
       
   915 *}
       
   916 lemma step_RAG_v:
       
   917 fixes th::thread
       
   918 assumes vt:
       
   919   "vt (V th cs#s)"
       
   920 shows "
       
   921   RAG (V th cs # s) =
       
   922   RAG s - {(Cs cs, Th th)} -
       
   923   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
       
   924   {(Cs cs, Th th') |th'.  next_th s th cs th'}"
       
   925   apply (insert vt, unfold s_RAG_def) 
       
   926   apply (auto split:if_splits list.splits simp:Let_def)
       
   927   apply (auto elim: step_v_waiting_mono step_v_hold_inv 
       
   928               step_v_release step_v_wait_inv
       
   929               step_v_get_hold step_v_release_inv)
       
   930   apply (erule_tac step_v_not_wait, auto)
       
   931   done
       
   932 
       
   933 text {* 
       
   934   The following @{text "step_RAG_p"} lemma charaterizes how @{text "RAG"} is changed
       
   935   with the happening of @{text "P"}-events:
       
   936 *}
       
   937 lemma step_RAG_p:
       
   938   "vt (P th cs#s) \<Longrightarrow>
       
   939   RAG (P th cs # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
       
   940                                              else RAG s \<union> {(Th th, Cs cs)})"
       
   941   apply(simp only: s_RAG_def wq_def)
       
   942   apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def)
       
   943   apply(case_tac "csa = cs", auto)
       
   944   apply(fold wq_def)
       
   945   apply(drule_tac step_back_step)
       
   946   apply(ind_cases " step s (P (hd (wq s cs)) cs)")
       
   947   apply(simp add:s_RAG_def wq_def cs_holding_def)
       
   948   apply(auto)
       
   949   done
       
   950 
       
   951 
       
   952 lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
       
   953   by (unfold s_RAG_def, auto)
       
   954 
       
   955 context valid_trace
       
   956 begin
       
   957 
       
   958 text {*
       
   959   The following lemma shows that @{text "RAG"} is acyclic.
       
   960   The overall structure is by induction on the formation of @{text "vt s"}
       
   961   and then case analysis on event @{text "e"}, where the non-trivial cases 
       
   962   for those for @{text "V"} and @{text "P"} events.
       
   963 *}
       
   964 lemma acyclic_RAG:
       
   965   shows "acyclic (RAG s)"
       
   966 using vt
       
   967 proof(induct)
       
   968   case (vt_cons s e)
       
   969   interpret vt_s: valid_trace s using vt_cons(1)
       
   970     by (unfold_locales, simp)
       
   971   assume ih: "acyclic (RAG s)"
       
   972     and stp: "step s e"
       
   973     and vt: "vt s"
       
   974   show ?case
       
   975   proof(cases e)
       
   976     case (Create th prio)
       
   977     with ih
       
   978     show ?thesis by (simp add:RAG_create_unchanged)
       
   979   next
       
   980     case (Exit th)
       
   981     with ih show ?thesis by (simp add:RAG_exit_unchanged)
       
   982   next
       
   983     case (V th cs)
       
   984     from V vt stp have vtt: "vt (V th cs#s)" by auto
       
   985     from step_RAG_v [OF this]
       
   986     have eq_de: 
       
   987       "RAG (e # s) = 
       
   988       RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
       
   989       {(Cs cs, Th th') |th'. next_th s th cs th'}"
       
   990       (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
       
   991     from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
       
   992     from step_back_step [OF vtt]
       
   993     have "step s (V th cs)" .
       
   994     thus ?thesis
       
   995     proof(cases)
       
   996       assume "holding s th cs"
       
   997       hence th_in: "th \<in> set (wq s cs)" and
       
   998         eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto
       
   999       then obtain rest where
       
  1000         eq_wq: "wq s cs = th#rest"
       
  1001         by (cases "wq s cs", auto)
       
  1002       show ?thesis
       
  1003       proof(cases "rest = []")
       
  1004         case False
       
  1005         let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
  1006         from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
       
  1007           by (unfold next_th_def, auto)
       
  1008         let ?E = "(?A - ?B - ?C)"
       
  1009         have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*"
       
  1010         proof
       
  1011           assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
       
  1012           hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
       
  1013           from tranclD [OF this]
       
  1014           obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
       
  1015           hence th_d: "(Th ?th', x) \<in> ?A" by simp
       
  1016           from RAG_target_th [OF this]
       
  1017           obtain cs' where eq_x: "x = Cs cs'" by auto
       
  1018           with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
       
  1019           hence wt_th': "waiting s ?th' cs'"
       
  1020             unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp
       
  1021           hence "cs' = cs"
       
  1022           proof(rule vt_s.waiting_unique)
       
  1023             from eq_wq vt_s.wq_distinct[of cs]
       
  1024             show "waiting s ?th' cs" 
       
  1025               apply (unfold s_waiting_def wq_def, auto)
       
  1026             proof -
       
  1027               assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1028                 and eq_wq: "wq_fun (schs s) cs = th # rest"
       
  1029               have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1030               proof(rule someI2)
       
  1031                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1032                 show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
       
  1033               next
       
  1034                 fix x assume "distinct x \<and> set x = set rest"
       
  1035                 with False show "x \<noteq> []" by auto
       
  1036               qed
       
  1037               hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
       
  1038                 set (SOME q. distinct q \<and> set q = set rest)" by auto
       
  1039               moreover have "\<dots> = set rest" 
       
  1040               proof(rule someI2)
       
  1041                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1042                 show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
       
  1043               next
       
  1044                 show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1045               qed
       
  1046               moreover note hd_in
       
  1047               ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
       
  1048             next
       
  1049               assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1050                 and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
       
  1051               have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1052               proof(rule someI2)
       
  1053                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1054                 show "distinct rest \<and> set rest = set rest" by auto
       
  1055               next
       
  1056                 fix x assume "distinct x \<and> set x = set rest"
       
  1057                 with False show "x \<noteq> []" by auto
       
  1058               qed
       
  1059               hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
       
  1060                 set (SOME q. distinct q \<and> set q = set rest)" by auto
       
  1061               moreover have "\<dots> = set rest" 
       
  1062               proof(rule someI2)
       
  1063                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1064                 show "distinct rest \<and> set rest = set rest" by auto
       
  1065               next
       
  1066                 show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1067               qed
       
  1068               moreover note hd_in
       
  1069               ultimately show False by auto
       
  1070             qed
       
  1071           qed
       
  1072           with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp
       
  1073           with False
       
  1074           show "False" by (auto simp: next_th_def eq_wq)
       
  1075         qed
       
  1076         with acyclic_insert[symmetric] and ac
       
  1077           and eq_de eq_D show ?thesis by auto
       
  1078       next
       
  1079         case True
       
  1080         with eq_wq
       
  1081         have eq_D: "?D = {}"
       
  1082           by (unfold next_th_def, auto)
       
  1083         with eq_de ac
       
  1084         show ?thesis by auto
       
  1085       qed 
       
  1086     qed
       
  1087   next
       
  1088     case (P th cs)
       
  1089     from P vt stp have vtt: "vt (P th cs#s)" by auto
       
  1090     from step_RAG_p [OF this] P
       
  1091     have "RAG (e # s) = 
       
  1092       (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
       
  1093       RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
       
  1094       by simp
       
  1095     moreover have "acyclic ?R"
       
  1096     proof(cases "wq s cs = []")
       
  1097       case True
       
  1098       hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
       
  1099       have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
       
  1100       proof
       
  1101         assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
       
  1102         hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
       
  1103         from tranclD2 [OF this]
       
  1104         obtain x where "(x, Cs cs) \<in> RAG s" by auto
       
  1105         with True show False by (auto simp:s_RAG_def cs_waiting_def)
       
  1106       qed
       
  1107       with acyclic_insert ih eq_r show ?thesis by auto
       
  1108     next
       
  1109       case False
       
  1110       hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
       
  1111       have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
       
  1112       proof
       
  1113         assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
       
  1114         hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
       
  1115         moreover from step_back_step [OF vtt] have "step s (P th cs)" .
       
  1116         ultimately show False
       
  1117         proof -
       
  1118           show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
       
  1119             by (ind_cases "step s (P th cs)", simp)
       
  1120         qed
       
  1121       qed
       
  1122       with acyclic_insert ih eq_r show ?thesis by auto
       
  1123       qed
       
  1124       ultimately show ?thesis by simp
       
  1125     next
       
  1126       case (Set thread prio)
       
  1127       with ih
       
  1128       thm RAG_set_unchanged
       
  1129       show ?thesis by (simp add:RAG_set_unchanged)
       
  1130     qed
       
  1131   next
       
  1132     case vt_nil
       
  1133     show "acyclic (RAG ([]::state))"
       
  1134       by (auto simp: s_RAG_def cs_waiting_def 
       
  1135         cs_holding_def wq_def acyclic_def)
       
  1136 qed
       
  1137 
       
  1138 
       
  1139 lemma finite_RAG:
       
  1140   shows "finite (RAG s)"
       
  1141 proof -
       
  1142   from vt show ?thesis
       
  1143   proof(induct)
       
  1144     case (vt_cons s e)
       
  1145     interpret vt_s: valid_trace s using vt_cons(1)
       
  1146       by (unfold_locales, simp)
       
  1147     assume ih: "finite (RAG s)"
       
  1148       and stp: "step s e"
       
  1149       and vt: "vt s"
       
  1150     show ?case
       
  1151     proof(cases e)
       
  1152       case (Create th prio)
       
  1153       with ih
       
  1154       show ?thesis by (simp add:RAG_create_unchanged)
       
  1155     next
       
  1156       case (Exit th)
       
  1157       with ih show ?thesis by (simp add:RAG_exit_unchanged)
       
  1158     next
       
  1159       case (V th cs)
       
  1160       from V vt stp have vtt: "vt (V th cs#s)" by auto
       
  1161       from step_RAG_v [OF this]
       
  1162       have eq_de: "RAG (e # s) = 
       
  1163                    RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
       
  1164                       {(Cs cs, Th th') |th'. next_th s th cs th'}
       
  1165 "
       
  1166         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
       
  1167       moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
       
  1168       moreover have "finite ?D"
       
  1169       proof -
       
  1170         have "?D = {} \<or> (\<exists> a. ?D = {a})" 
       
  1171           by (unfold next_th_def, auto)
       
  1172         thus ?thesis
       
  1173         proof
       
  1174           assume h: "?D = {}"
       
  1175           show ?thesis by (unfold h, simp)
       
  1176         next
       
  1177           assume "\<exists> a. ?D = {a}"
       
  1178           thus ?thesis
       
  1179             by (metis finite.simps)
       
  1180         qed
       
  1181       qed
       
  1182       ultimately show ?thesis by simp
       
  1183     next
       
  1184       case (P th cs)
       
  1185       from P vt stp have vtt: "vt (P th cs#s)" by auto
       
  1186       from step_RAG_p [OF this] P
       
  1187       have "RAG (e # s) = 
       
  1188               (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
       
  1189                                     RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
       
  1190         by simp
       
  1191       moreover have "finite ?R"
       
  1192       proof(cases "wq s cs = []")
       
  1193         case True
       
  1194         hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
       
  1195         with True and ih show ?thesis by auto
       
  1196       next
       
  1197         case False
       
  1198         hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
       
  1199         with False and ih show ?thesis by auto
       
  1200       qed
       
  1201       ultimately show ?thesis by auto
       
  1202     next
       
  1203       case (Set thread prio)
       
  1204       with ih
       
  1205       show ?thesis by (simp add:RAG_set_unchanged)
       
  1206     qed
       
  1207   next
       
  1208     case vt_nil
       
  1209     show "finite (RAG ([]::state))"
       
  1210       by (auto simp: s_RAG_def cs_waiting_def 
       
  1211                    cs_holding_def wq_def acyclic_def)
       
  1212   qed
       
  1213 qed
       
  1214 
       
  1215 text {* Several useful lemmas *}
       
  1216 
       
  1217 lemma wf_dep_converse: 
       
  1218   shows "wf ((RAG s)^-1)"
       
  1219 proof(rule finite_acyclic_wf_converse)
       
  1220   from finite_RAG 
       
  1221   show "finite (RAG s)" .
       
  1222 next
       
  1223   from acyclic_RAG
       
  1224   show "acyclic (RAG s)" .
       
  1225 qed
       
  1226 
       
  1227 end
       
  1228 
       
  1229 lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
       
  1230   by (induct l, auto)
       
  1231 
       
  1232 lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
       
  1233   by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1234 
       
  1235 context valid_trace
       
  1236 begin
       
  1237 
       
  1238 lemma wq_threads: 
       
  1239   assumes h: "th \<in> set (wq s cs)"
       
  1240   shows "th \<in> threads s"
       
  1241 proof -
       
  1242  from vt and h show ?thesis
       
  1243   proof(induct arbitrary: th cs)
       
  1244     case (vt_cons s e)
       
  1245     interpret vt_s: valid_trace s
       
  1246       using vt_cons(1) by (unfold_locales, auto)
       
  1247     assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
       
  1248       and stp: "step s e"
       
  1249       and vt: "vt s"
       
  1250       and h: "th \<in> set (wq (e # s) cs)"
       
  1251     show ?case
       
  1252     proof(cases e)
       
  1253       case (Create th' prio)
       
  1254       with ih h show ?thesis
       
  1255         by (auto simp:wq_def Let_def)
       
  1256     next
       
  1257       case (Exit th')
       
  1258       with stp ih h show ?thesis
       
  1259         apply (auto simp:wq_def Let_def)
       
  1260         apply (ind_cases "step s (Exit th')")
       
  1261         apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
       
  1262                s_RAG_def s_holding_def cs_holding_def)
       
  1263         done
       
  1264     next
       
  1265       case (V th' cs')
       
  1266       show ?thesis
       
  1267       proof(cases "cs' = cs")
       
  1268         case False
       
  1269         with h
       
  1270         show ?thesis
       
  1271           apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
       
  1272           by (drule_tac ih, simp)
       
  1273       next
       
  1274         case True
       
  1275         from h
       
  1276         show ?thesis
       
  1277         proof(unfold V wq_def)
       
  1278           assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
       
  1279           show "th \<in> threads (V th' cs' # s)"
       
  1280           proof(cases "cs = cs'")
       
  1281             case False
       
  1282             hence "?l = wq_fun (schs s) cs" by (simp add:Let_def)
       
  1283             with th_in have " th \<in> set (wq s cs)" 
       
  1284               by (fold wq_def, simp)
       
  1285             from ih [OF this] show ?thesis by simp
       
  1286           next
       
  1287             case True
       
  1288             show ?thesis
       
  1289             proof(cases "wq_fun (schs s) cs'")
       
  1290               case Nil
       
  1291               with h V show ?thesis
       
  1292                 apply (auto simp:wq_def Let_def split:if_splits)
       
  1293                 by (fold wq_def, drule_tac ih, simp)
       
  1294             next
       
  1295               case (Cons a rest)
       
  1296               assume eq_wq: "wq_fun (schs s) cs' = a # rest"
       
  1297               with h V show ?thesis
       
  1298                 apply (auto simp:Let_def wq_def split:if_splits)
       
  1299               proof -
       
  1300                 assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1301                 have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
       
  1302                 proof(rule someI2)
       
  1303                   from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def]
       
  1304                   show "distinct rest \<and> set rest = set rest" by auto
       
  1305                 next
       
  1306                   show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
       
  1307                     by auto
       
  1308                 qed
       
  1309                 with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto
       
  1310                 from ih[OF this[folded wq_def]] show "th \<in> threads s" .
       
  1311               next
       
  1312                 assume th_in: "th \<in> set (wq_fun (schs s) cs)"
       
  1313                 from ih[OF this[folded wq_def]]
       
  1314                 show "th \<in> threads s" .
       
  1315               qed
       
  1316             qed
       
  1317           qed
       
  1318         qed
       
  1319       qed
       
  1320     next
       
  1321       case (P th' cs')
       
  1322       from h stp
       
  1323       show ?thesis
       
  1324         apply (unfold P wq_def)
       
  1325         apply (auto simp:Let_def split:if_splits, fold wq_def)
       
  1326         apply (auto intro:ih)
       
  1327         apply(ind_cases "step s (P th' cs')")
       
  1328         by (unfold runing_def readys_def, auto)
       
  1329     next
       
  1330       case (Set thread prio)
       
  1331       with ih h show ?thesis
       
  1332         by (auto simp:wq_def Let_def)
       
  1333     qed
       
  1334   next
       
  1335     case vt_nil
       
  1336     thus ?case by (auto simp:wq_def)
       
  1337   qed
       
  1338 qed
       
  1339 
       
  1340 lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
       
  1341   apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
       
  1342   by (auto intro:wq_threads)
       
  1343 
       
  1344 lemma readys_v_eq:
       
  1345   fixes th thread cs rest
       
  1346   assumes neq_th: "th \<noteq> thread"
       
  1347   and eq_wq: "wq s cs = thread#rest"
       
  1348   and not_in: "th \<notin>  set rest"
       
  1349   shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
       
  1350 proof -
       
  1351   from assms show ?thesis
       
  1352     apply (auto simp:readys_def)
       
  1353     apply(simp add:s_waiting_def[folded wq_def])
       
  1354     apply (erule_tac x = csa in allE)
       
  1355     apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
       
  1356     apply (case_tac "csa = cs", simp)
       
  1357     apply (erule_tac x = cs in allE)
       
  1358     apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits)
       
  1359     apply(auto simp add: wq_def)
       
  1360     apply (auto simp:s_waiting_def wq_def Let_def split:list.splits)
       
  1361     proof -
       
  1362        assume th_nin: "th \<notin> set rest"
       
  1363         and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1364         and eq_wq: "wq_fun (schs s) cs = thread # rest"
       
  1365       have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1366       proof(rule someI2)
       
  1367         from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def]
       
  1368         show "distinct rest \<and> set rest = set rest" by auto
       
  1369       next
       
  1370         show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1371       qed
       
  1372       with th_nin th_in show False by auto
       
  1373     qed
       
  1374 qed
       
  1375 
       
  1376 text {* \noindent
       
  1377   The following lemmas shows that: starting from any node in @{text "RAG"}, 
       
  1378   by chasing out-going edges, it is always possible to reach a node representing a ready
       
  1379   thread. In this lemma, it is the @{text "th'"}.
       
  1380 *}
       
  1381 
       
  1382 lemma chain_building:
       
  1383   shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
       
  1384 proof -
       
  1385   from wf_dep_converse
       
  1386   have h: "wf ((RAG s)\<inverse>)" .
       
  1387   show ?thesis
       
  1388   proof(induct rule:wf_induct [OF h])
       
  1389     fix x
       
  1390     assume ih [rule_format]: 
       
  1391       "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
       
  1392            y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
       
  1393     show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
       
  1394     proof
       
  1395       assume x_d: "x \<in> Domain (RAG s)"
       
  1396       show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
       
  1397       proof(cases x)
       
  1398         case (Th th)
       
  1399         from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
       
  1400         with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
       
  1401         from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
       
  1402         hence "Cs cs \<in> Domain (RAG s)" by auto
       
  1403         from ih [OF x_in_r this] obtain th'
       
  1404           where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
       
  1405         have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
       
  1406         with th'_ready show ?thesis by auto
       
  1407       next
       
  1408         case (Cs cs)
       
  1409         from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
       
  1410         show ?thesis
       
  1411         proof(cases "th' \<in> readys s")
       
  1412           case True
       
  1413           from True and th'_d show ?thesis by auto
       
  1414         next
       
  1415           case False
       
  1416           from th'_d and range_in  have "th' \<in> threads s" by auto
       
  1417           with False have "Th th' \<in> Domain (RAG s)" 
       
  1418             by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
       
  1419           from ih [OF th'_d this]
       
  1420           obtain th'' where 
       
  1421             th''_r: "th'' \<in> readys s" and 
       
  1422             th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
       
  1423           from th'_d and th''_in 
       
  1424           have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
       
  1425           with th''_r show ?thesis by auto
       
  1426         qed
       
  1427       qed
       
  1428     qed
       
  1429   qed
       
  1430 qed
       
  1431 
       
  1432 text {* \noindent
       
  1433   The following is just an instance of @{text "chain_building"}.
       
  1434 *}
       
  1435 lemma th_chain_to_ready:
       
  1436   assumes th_in: "th \<in> threads s"
       
  1437   shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
       
  1438 proof(cases "th \<in> readys s")
       
  1439   case True
       
  1440   thus ?thesis by auto
       
  1441 next
       
  1442   case False
       
  1443   from False and th_in have "Th th \<in> Domain (RAG s)" 
       
  1444     by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
       
  1445   from chain_building [rule_format, OF this]
       
  1446   show ?thesis by auto
       
  1447 qed
       
  1448 
       
  1449 end
       
  1450 
       
  1451 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
       
  1452   by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
       
  1453 
       
  1454 lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
       
  1455   by (unfold s_holding_def wq_def cs_holding_def, simp)
       
  1456 
       
  1457 lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
       
  1458   by (unfold s_holding_def cs_holding_def, auto)
       
  1459 
       
  1460 context valid_trace
       
  1461 begin
       
  1462 
       
  1463 lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
       
  1464   apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
       
  1465   by(auto elim:waiting_unique holding_unique)
       
  1466 
       
  1467 end
       
  1468 
       
  1469 
       
  1470 lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
       
  1471 by (induct rule:trancl_induct, auto)
       
  1472 
       
  1473 context valid_trace
       
  1474 begin
       
  1475 
       
  1476 lemma dchain_unique:
       
  1477   assumes th1_d: "(n, Th th1) \<in> (RAG s)^+"
       
  1478   and th1_r: "th1 \<in> readys s"
       
  1479   and th2_d: "(n, Th th2) \<in> (RAG s)^+"
       
  1480   and th2_r: "th2 \<in> readys s"
       
  1481   shows "th1 = th2"
       
  1482 proof -
       
  1483   { assume neq: "th1 \<noteq> th2"
       
  1484     hence "Th th1 \<noteq> Th th2" by simp
       
  1485     from unique_chain [OF _ th1_d th2_d this] and unique_RAG 
       
  1486     have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
       
  1487     hence "False"
       
  1488     proof
       
  1489       assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
       
  1490       from trancl_split [OF this]
       
  1491       obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
       
  1492       then obtain cs where eq_n: "n = Cs cs"
       
  1493         by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1494       from dd eq_n have "th1 \<notin> readys s"
       
  1495         by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
       
  1496       with th1_r show ?thesis by auto
       
  1497     next
       
  1498       assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
       
  1499       from trancl_split [OF this]
       
  1500       obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
       
  1501       then obtain cs where eq_n: "n = Cs cs"
       
  1502         by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1503       from dd eq_n have "th2 \<notin> readys s"
       
  1504         by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
       
  1505       with th2_r show ?thesis by auto
       
  1506     qed
       
  1507   } thus ?thesis by auto
       
  1508 qed
       
  1509 
       
  1510 end
       
  1511              
       
  1512 
       
  1513 lemma step_holdents_p_add:
       
  1514   fixes th cs s
       
  1515   assumes vt: "vt (P th cs#s)"
       
  1516   and "wq s cs = []"
       
  1517   shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
       
  1518 proof -
       
  1519   from assms show ?thesis
       
  1520   unfolding  holdents_test step_RAG_p[OF vt] by (auto)
       
  1521 qed
       
  1522 
       
  1523 lemma step_holdents_p_eq:
       
  1524   fixes th cs s
       
  1525   assumes vt: "vt (P th cs#s)"
       
  1526   and "wq s cs \<noteq> []"
       
  1527   shows "holdents (P th cs#s) th = holdents s th"
       
  1528 proof -
       
  1529   from assms show ?thesis
       
  1530   unfolding  holdents_test step_RAG_p[OF vt] by auto
       
  1531 qed
       
  1532 
       
  1533 
       
  1534 lemma (in valid_trace) finite_holding :
       
  1535   shows "finite (holdents s th)"
       
  1536 proof -
       
  1537   let ?F = "\<lambda> (x, y). the_cs x"
       
  1538   from finite_RAG 
       
  1539   have "finite (RAG s)" .
       
  1540   hence "finite (?F `(RAG s))" by simp
       
  1541   moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
       
  1542   proof -
       
  1543     { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
       
  1544       fix x assume "(Cs x, Th th) \<in> RAG s"
       
  1545       hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
       
  1546       moreover have "?F (Cs x, Th th) = x" by simp
       
  1547       ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
       
  1548     } thus ?thesis by auto
       
  1549   qed
       
  1550   ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
       
  1551 qed
       
  1552 
       
  1553 lemma cntCS_v_dec: 
       
  1554   fixes s thread cs
       
  1555   assumes vtv: "vt (V thread cs#s)"
       
  1556   shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
       
  1557 proof -
       
  1558   from vtv interpret vt_s: valid_trace s
       
  1559     by (cases, unfold_locales, simp)
       
  1560   from vtv interpret vt_v: valid_trace "V thread cs#s"
       
  1561      by (unfold_locales, simp)
       
  1562   from step_back_step[OF vtv]
       
  1563   have cs_in: "cs \<in> holdents s thread" 
       
  1564     apply (cases, unfold holdents_test s_RAG_def, simp)
       
  1565     by (unfold cs_holding_def s_holding_def wq_def, auto)
       
  1566   moreover have cs_not_in: 
       
  1567     "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
       
  1568     apply (insert vt_s.wq_distinct[of cs])
       
  1569     apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
       
  1570             auto simp:next_th_def)
       
  1571   proof -
       
  1572     fix rest
       
  1573     assume dst: "distinct (rest::thread list)"
       
  1574       and ne: "rest \<noteq> []"
       
  1575     and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1576     moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1577     proof(rule someI2)
       
  1578       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1579     next
       
  1580       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1581     qed
       
  1582     ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1583                      set (SOME q. distinct q \<and> set q = set rest)" by simp
       
  1584     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1585     proof(rule someI2)
       
  1586       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1587     next
       
  1588       fix x assume " distinct x \<and> set x = set rest" with ne
       
  1589       show "x \<noteq> []" by auto
       
  1590     qed
       
  1591     ultimately 
       
  1592     show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
       
  1593       by auto
       
  1594   next
       
  1595     fix rest
       
  1596     assume dst: "distinct (rest::thread list)"
       
  1597       and ne: "rest \<noteq> []"
       
  1598     and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1599     moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1600     proof(rule someI2)
       
  1601       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1602     next
       
  1603       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1604     qed
       
  1605     ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1606                      set (SOME q. distinct q \<and> set q = set rest)" by simp
       
  1607     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1608     proof(rule someI2)
       
  1609       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1610     next
       
  1611       fix x assume " distinct x \<and> set x = set rest" with ne
       
  1612       show "x \<noteq> []" by auto
       
  1613     qed
       
  1614     ultimately show "False" by auto 
       
  1615   qed
       
  1616   ultimately 
       
  1617   have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
       
  1618     by auto
       
  1619   moreover have "card \<dots> = 
       
  1620                     Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
       
  1621   proof(rule card_insert)
       
  1622     from vt_v.finite_holding
       
  1623     show " finite (holdents (V thread cs # s) thread)" .
       
  1624   qed
       
  1625   moreover from cs_not_in 
       
  1626   have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
       
  1627   ultimately show ?thesis by (simp add:cntCS_def)
       
  1628 qed 
       
  1629 
       
  1630 lemma count_rec1 [simp]: 
       
  1631   assumes "Q e"
       
  1632   shows "count Q (e#es) = Suc (count Q es)"
       
  1633   using assms
       
  1634   by (unfold count_def, auto)
       
  1635 
       
  1636 lemma count_rec2 [simp]: 
       
  1637   assumes "\<not>Q e"
       
  1638   shows "count Q (e#es) = (count Q es)"
       
  1639   using assms
       
  1640   by (unfold count_def, auto)
       
  1641 
       
  1642 lemma count_rec3 [simp]: 
       
  1643   shows "count Q [] =  0"
       
  1644   by (unfold count_def, auto)
       
  1645   
       
  1646 lemma cntP_diff_inv:
       
  1647   assumes "cntP (e#s) th \<noteq> cntP s th"
       
  1648   shows "isP e \<and> actor e = th"
       
  1649 proof(cases e)
       
  1650   case (P th' pty)
       
  1651   show ?thesis
       
  1652   by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
       
  1653         insert assms P, auto simp:cntP_def)
       
  1654 qed (insert assms, auto simp:cntP_def)
       
  1655 
       
  1656 lemma isP_E:
       
  1657   assumes "isP e"
       
  1658   obtains cs where "e = P (actor e) cs"
       
  1659   using assms by (cases e, auto)
       
  1660 
       
  1661 lemma isV_E:
       
  1662   assumes "isV e"
       
  1663   obtains cs where "e = V (actor e) cs"
       
  1664   using assms by (cases e, auto) (* ccc *)
       
  1665 
       
  1666 lemma cntV_diff_inv:
       
  1667   assumes "cntV (e#s) th \<noteq> cntV s th"
       
  1668   shows "isV e \<and> actor e = th"
       
  1669 proof(cases e)
       
  1670   case (V th' pty)
       
  1671   show ?thesis
       
  1672   by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
       
  1673         insert assms V, auto simp:cntV_def)
       
  1674 qed (insert assms, auto simp:cntV_def)
       
  1675 
       
  1676 context valid_trace
       
  1677 begin
       
  1678 
       
  1679 text {* (* ddd *) \noindent
       
  1680   The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} 
       
  1681   of one particular thread. 
       
  1682 *} 
       
  1683 
       
  1684 lemma cnp_cnv_cncs:
       
  1685   shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
       
  1686                                        then cntCS s th else cntCS s th + 1)"
       
  1687 proof -
       
  1688   from vt show ?thesis
       
  1689   proof(induct arbitrary:th)
       
  1690     case (vt_cons s e)
       
  1691     interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp)
       
  1692     assume vt: "vt s"
       
  1693     and ih: "\<And>th. cntP s th  = cntV s th +
       
  1694                (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
       
  1695     and stp: "step s e"
       
  1696     from stp show ?case
       
  1697     proof(cases)
       
  1698       case (thread_create thread prio)
       
  1699       assume eq_e: "e = Create thread prio"
       
  1700         and not_in: "thread \<notin> threads s"
       
  1701       show ?thesis
       
  1702       proof -
       
  1703         { fix cs 
       
  1704           assume "thread \<in> set (wq s cs)"
       
  1705           from vt_s.wq_threads [OF this] have "thread \<in> threads s" .
       
  1706           with not_in have "False" by simp
       
  1707         } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
       
  1708           by (auto simp:readys_def threads.simps s_waiting_def 
       
  1709             wq_def cs_waiting_def Let_def)
       
  1710         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
       
  1711         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
       
  1712         have eq_cncs: "cntCS (e#s) th = cntCS s th"
       
  1713           unfolding cntCS_def holdents_test
       
  1714           by (simp add:RAG_create_unchanged eq_e)
       
  1715         { assume "th \<noteq> thread"
       
  1716           with eq_readys eq_e
       
  1717           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
       
  1718                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
       
  1719             by (simp add:threads.simps)
       
  1720           with eq_cnp eq_cnv eq_cncs ih not_in
       
  1721           have ?thesis by simp
       
  1722         } moreover {
       
  1723           assume eq_th: "th = thread"
       
  1724           with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
       
  1725           moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
       
  1726           moreover note eq_cnp eq_cnv eq_cncs
       
  1727           ultimately have ?thesis by auto
       
  1728         } ultimately show ?thesis by blast
       
  1729       qed
       
  1730     next
       
  1731       case (thread_exit thread)
       
  1732       assume eq_e: "e = Exit thread" 
       
  1733       and is_runing: "thread \<in> runing s"
       
  1734       and no_hold: "holdents s thread = {}"
       
  1735       from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
       
  1736       from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
       
  1737       have eq_cncs: "cntCS (e#s) th = cntCS s th"
       
  1738         unfolding cntCS_def holdents_test
       
  1739         by (simp add:RAG_exit_unchanged eq_e)
       
  1740       { assume "th \<noteq> thread"
       
  1741         with eq_e
       
  1742         have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
       
  1743           (th \<in> readys (s) \<or> th \<notin> threads (s))" 
       
  1744           apply (simp add:threads.simps readys_def)
       
  1745           apply (subst s_waiting_def)
       
  1746           apply (simp add:Let_def)
       
  1747           apply (subst s_waiting_def, simp)
       
  1748           done
       
  1749         with eq_cnp eq_cnv eq_cncs ih
       
  1750         have ?thesis by simp
       
  1751       } moreover {
       
  1752         assume eq_th: "th = thread"
       
  1753         with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
       
  1754           by (simp add:runing_def)
       
  1755         moreover from eq_th eq_e have "th \<notin> threads (e#s)"
       
  1756           by simp
       
  1757         moreover note eq_cnp eq_cnv eq_cncs
       
  1758         ultimately have ?thesis by auto
       
  1759       } ultimately show ?thesis by blast
       
  1760     next
       
  1761       case (thread_P thread cs)
       
  1762       assume eq_e: "e = P thread cs"
       
  1763         and is_runing: "thread \<in> runing s"
       
  1764         and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
       
  1765       from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
       
  1766       then interpret vt_p: valid_trace "(P thread cs#s)"
       
  1767         by (unfold_locales, simp)
       
  1768       show ?thesis 
       
  1769       proof -
       
  1770         { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
       
  1771           assume neq_th: "th \<noteq> thread"
       
  1772           with eq_e
       
  1773           have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
       
  1774             apply (simp add:readys_def s_waiting_def wq_def Let_def)
       
  1775             apply (rule_tac hh)
       
  1776              apply (intro iffI allI, clarify)
       
  1777             apply (erule_tac x = csa in allE, auto)
       
  1778             apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
       
  1779             apply (erule_tac x = cs in allE, auto)
       
  1780             by (case_tac "(wq_fun (schs s) cs)", auto)
       
  1781           moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
       
  1782             apply (simp add:cntCS_def holdents_test)
       
  1783             by (unfold  step_RAG_p [OF vtp], auto)
       
  1784           moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
       
  1785             by (simp add:cntP_def count_def)
       
  1786           moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
       
  1787             by (simp add:cntV_def count_def)
       
  1788           moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
       
  1789           moreover note ih [of th] 
       
  1790           ultimately have ?thesis by simp
       
  1791         } moreover {
       
  1792           assume eq_th: "th = thread"
       
  1793           have ?thesis
       
  1794           proof -
       
  1795             from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
       
  1796               by (simp add:cntP_def count_def)
       
  1797             from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
       
  1798               by (simp add:cntV_def count_def)
       
  1799             show ?thesis
       
  1800             proof (cases "wq s cs = []")
       
  1801               case True
       
  1802               with is_runing
       
  1803               have "th \<in> readys (e#s)"
       
  1804                 apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
       
  1805                 apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
       
  1806                 by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
       
  1807               moreover have "cntCS (e # s) th = 1 + cntCS s th"
       
  1808               proof -
       
  1809                 have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
       
  1810                   Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
       
  1811                 proof -
       
  1812                   have "?L = insert cs ?R" by auto
       
  1813                   moreover have "card \<dots> = Suc (card (?R - {cs}))" 
       
  1814                   proof(rule card_insert)
       
  1815                     from vt_s.finite_holding [of thread]
       
  1816                     show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
       
  1817                       by (unfold holdents_test, simp)
       
  1818                   qed
       
  1819                   moreover have "?R - {cs} = ?R"
       
  1820                   proof -
       
  1821                     have "cs \<notin> ?R"
       
  1822                     proof
       
  1823                       assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
       
  1824                       with no_dep show False by auto
       
  1825                     qed
       
  1826                     thus ?thesis by auto
       
  1827                   qed
       
  1828                   ultimately show ?thesis by auto
       
  1829                 qed
       
  1830                 thus ?thesis
       
  1831                   apply (unfold eq_e eq_th cntCS_def)
       
  1832                   apply (simp add: holdents_test)
       
  1833                   by (unfold step_RAG_p [OF vtp], auto simp:True)
       
  1834               qed
       
  1835               moreover from is_runing have "th \<in> readys s"
       
  1836                 by (simp add:runing_def eq_th)
       
  1837               moreover note eq_cnp eq_cnv ih [of th]
       
  1838               ultimately show ?thesis by auto
       
  1839             next
       
  1840               case False
       
  1841               have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
       
  1842                     by (unfold eq_th eq_e wq_def, auto simp:Let_def)
       
  1843               have "th \<notin> readys (e#s)"
       
  1844               proof
       
  1845                 assume "th \<in> readys (e#s)"
       
  1846                 hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
       
  1847                 from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
       
  1848                 hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
       
  1849                   by (simp add:s_waiting_def wq_def)
       
  1850                 moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
       
  1851                 ultimately have "th = hd (wq (e#s) cs)" by blast
       
  1852                 with eq_wq have "th = hd (wq s cs @ [th])" by simp
       
  1853                 hence "th = hd (wq s cs)" using False by auto
       
  1854                 with False eq_wq vt_p.wq_distinct [of cs]
       
  1855                 show False by (fold eq_e, auto)
       
  1856               qed
       
  1857               moreover from is_runing have "th \<in> threads (e#s)" 
       
  1858                 by (unfold eq_e, auto simp:runing_def readys_def eq_th)
       
  1859               moreover have "cntCS (e # s) th = cntCS s th"
       
  1860                 apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
       
  1861                 by (auto simp:False)
       
  1862               moreover note eq_cnp eq_cnv ih[of th]
       
  1863               moreover from is_runing have "th \<in> readys s"
       
  1864                 by (simp add:runing_def eq_th)
       
  1865               ultimately show ?thesis by auto
       
  1866             qed
       
  1867           qed
       
  1868         } ultimately show ?thesis by blast
       
  1869       qed
       
  1870     next
       
  1871       case (thread_V thread cs)
       
  1872       from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto
       
  1873       then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp)
       
  1874       assume eq_e: "e = V thread cs"
       
  1875         and is_runing: "thread \<in> runing s"
       
  1876         and hold: "holding s thread cs"
       
  1877       from hold obtain rest 
       
  1878         where eq_wq: "wq s cs = thread # rest"
       
  1879         by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
       
  1880       have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
       
  1881       have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1882       proof(rule someI2)
       
  1883         from vt_v.wq_distinct[of cs] and eq_wq
       
  1884         show "distinct rest \<and> set rest = set rest"
       
  1885           by (metis distinct.simps(2) vt_s.wq_distinct)
       
  1886       next
       
  1887         show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
       
  1888           by auto
       
  1889       qed
       
  1890       show ?thesis
       
  1891       proof -
       
  1892         { assume eq_th: "th = thread"
       
  1893           from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
       
  1894             by (unfold eq_e, simp add:cntP_def count_def)
       
  1895           moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
       
  1896             by (unfold eq_e, simp add:cntV_def count_def)
       
  1897           moreover from cntCS_v_dec [OF vtv] 
       
  1898           have "cntCS (e # s) thread + 1 = cntCS s thread"
       
  1899             by (simp add:eq_e)
       
  1900           moreover from is_runing have rd_before: "thread \<in> readys s"
       
  1901             by (unfold runing_def, simp)
       
  1902           moreover have "thread \<in> readys (e # s)"
       
  1903           proof -
       
  1904             from is_runing
       
  1905             have "thread \<in> threads (e#s)" 
       
  1906               by (unfold eq_e, auto simp:runing_def readys_def)
       
  1907             moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
       
  1908             proof
       
  1909               fix cs1
       
  1910               { assume eq_cs: "cs1 = cs" 
       
  1911                 have "\<not> waiting (e # s) thread cs1"
       
  1912                 proof -
       
  1913                   from eq_wq
       
  1914                   have "thread \<notin> set (wq (e#s) cs1)"
       
  1915                     apply(unfold eq_e wq_def eq_cs s_holding_def)
       
  1916                     apply (auto simp:Let_def)
       
  1917                   proof -
       
  1918                     assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1919                     with eq_set have "thread \<in> set rest" by simp
       
  1920                     with vt_v.wq_distinct[of cs]
       
  1921                     and eq_wq show False
       
  1922                         by (metis distinct.simps(2) vt_s.wq_distinct)
       
  1923                   qed
       
  1924                   thus ?thesis by (simp add:wq_def s_waiting_def)
       
  1925                 qed
       
  1926               } moreover {
       
  1927                 assume neq_cs: "cs1 \<noteq> cs"
       
  1928                   have "\<not> waiting (e # s) thread cs1" 
       
  1929                   proof -
       
  1930                     from wq_v_neq [OF neq_cs[symmetric]]
       
  1931                     have "wq (V thread cs # s) cs1 = wq s cs1" .
       
  1932                     moreover have "\<not> waiting s thread cs1" 
       
  1933                     proof -
       
  1934                       from runing_ready and is_runing
       
  1935                       have "thread \<in> readys s" by auto
       
  1936                       thus ?thesis by (simp add:readys_def)
       
  1937                     qed
       
  1938                     ultimately show ?thesis 
       
  1939                       by (auto simp:wq_def s_waiting_def eq_e)
       
  1940                   qed
       
  1941               } ultimately show "\<not> waiting (e # s) thread cs1" by blast
       
  1942             qed
       
  1943             ultimately show ?thesis by (simp add:readys_def)
       
  1944           qed
       
  1945           moreover note eq_th ih
       
  1946           ultimately have ?thesis by auto
       
  1947         } moreover {
       
  1948           assume neq_th: "th \<noteq> thread"
       
  1949           from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
       
  1950             by (simp add:cntP_def count_def)
       
  1951           from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
       
  1952             by (simp add:cntV_def count_def)
       
  1953           have ?thesis
       
  1954           proof(cases "th \<in> set rest")
       
  1955             case False
       
  1956             have "(th \<in> readys (e # s)) = (th \<in> readys s)"
       
  1957               apply (insert step_back_vt[OF vtv])
       
  1958               by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq)
       
  1959             moreover have "cntCS (e#s) th = cntCS s th"
       
  1960               apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
       
  1961               proof -
       
  1962                 have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
       
  1963                       {cs. (Cs cs, Th th) \<in> RAG s}"
       
  1964                 proof -
       
  1965                   from False eq_wq
       
  1966                   have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
       
  1967                     apply (unfold next_th_def, auto)
       
  1968                   proof -
       
  1969                     assume ne: "rest \<noteq> []"
       
  1970                       and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1971                       and eq_wq: "wq s cs = thread # rest"
       
  1972                     from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1973                                   set (SOME q. distinct q \<and> set q = set rest)
       
  1974                                   " by simp
       
  1975                     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1976                     proof(rule someI2)
       
  1977                       from vt_s.wq_distinct[ of cs] and eq_wq
       
  1978                       show "distinct rest \<and> set rest = set rest" by auto
       
  1979                     next
       
  1980                       fix x assume "distinct x \<and> set x = set rest"
       
  1981                       with ne show "x \<noteq> []" by auto
       
  1982                     qed
       
  1983                     ultimately show 
       
  1984                       "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
       
  1985                       by auto
       
  1986                   qed    
       
  1987                   thus ?thesis by auto
       
  1988                 qed
       
  1989                 thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
       
  1990                              card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
       
  1991               qed
       
  1992             moreover note ih eq_cnp eq_cnv eq_threads
       
  1993             ultimately show ?thesis by auto
       
  1994           next
       
  1995             case True
       
  1996             assume th_in: "th \<in> set rest"
       
  1997             show ?thesis
       
  1998             proof(cases "next_th s thread cs th")
       
  1999               case False
       
  2000               with eq_wq and th_in have 
       
  2001                 neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest")
       
  2002                 by (auto simp:next_th_def)
       
  2003               have "(th \<in> readys (e # s)) = (th \<in> readys s)"
       
  2004               proof -
       
  2005                 from eq_wq and th_in
       
  2006                 have "\<not> th \<in> readys s"
       
  2007                   apply (auto simp:readys_def s_waiting_def)
       
  2008                   apply (rule_tac x = cs in exI, auto)
       
  2009                   by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def)
       
  2010                 moreover 
       
  2011                 from eq_wq and th_in and neq_hd
       
  2012                 have "\<not> (th \<in> readys (e # s))"
       
  2013                   apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits)
       
  2014                   by (rule_tac x = cs in exI, auto simp:eq_set)
       
  2015                 ultimately show ?thesis by auto
       
  2016               qed
       
  2017               moreover have "cntCS (e#s) th = cntCS s th" 
       
  2018               proof -
       
  2019                 from eq_wq and  th_in and neq_hd
       
  2020                 have "(holdents (e # s) th) = (holdents s th)"
       
  2021                   apply (unfold eq_e step_RAG_v[OF vtv], 
       
  2022                          auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
       
  2023                                    Let_def cs_holding_def)
       
  2024                   by (insert vt_s.wq_distinct[of cs], auto simp:wq_def)
       
  2025                 thus ?thesis by (simp add:cntCS_def)
       
  2026               qed
       
  2027               moreover note ih eq_cnp eq_cnv eq_threads
       
  2028               ultimately show ?thesis by auto
       
  2029             next
       
  2030               case True
       
  2031               let ?rest = " (SOME q. distinct q \<and> set q = set rest)"
       
  2032               let ?t = "hd ?rest"
       
  2033               from True eq_wq th_in neq_th
       
  2034               have "th \<in> readys (e # s)"
       
  2035                 apply (auto simp:eq_e readys_def s_waiting_def wq_def
       
  2036                         Let_def next_th_def)
       
  2037               proof -
       
  2038                 assume eq_wq: "wq_fun (schs s) cs = thread # rest"
       
  2039                   and t_in: "?t \<in> set rest"
       
  2040                 show "?t \<in> threads s"
       
  2041                 proof(rule vt_s.wq_threads)
       
  2042                   from eq_wq and t_in
       
  2043                   show "?t \<in> set (wq s cs)" by (auto simp:wq_def)
       
  2044                 qed
       
  2045               next
       
  2046                 fix csa
       
  2047                 assume eq_wq: "wq_fun (schs s) cs = thread # rest"
       
  2048                   and t_in: "?t \<in> set rest"
       
  2049                   and neq_cs: "csa \<noteq> cs"
       
  2050                   and t_in': "?t \<in>  set (wq_fun (schs s) csa)"
       
  2051                 show "?t = hd (wq_fun (schs s) csa)"
       
  2052                 proof -
       
  2053                   { assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)"
       
  2054                     from vt_s.wq_distinct[of cs] and 
       
  2055                     eq_wq[folded wq_def] and t_in eq_wq
       
  2056                     have "?t \<noteq> thread" by auto
       
  2057                     with eq_wq and t_in
       
  2058                     have w1: "waiting s ?t cs"
       
  2059                       by (auto simp:s_waiting_def wq_def)
       
  2060                     from t_in' neq_hd'
       
  2061                     have w2: "waiting s ?t csa"
       
  2062                       by (auto simp:s_waiting_def wq_def)
       
  2063                     from vt_s.waiting_unique[OF w1 w2]
       
  2064                     and neq_cs have "False" by auto
       
  2065                   } thus ?thesis by auto
       
  2066                 qed
       
  2067               qed
       
  2068               moreover have "cntP s th = cntV s th + cntCS s th + 1"
       
  2069               proof -
       
  2070                 have "th \<notin> readys s" 
       
  2071                 proof -
       
  2072                   from True eq_wq neq_th th_in
       
  2073                   show ?thesis
       
  2074                     apply (unfold readys_def s_waiting_def, auto)
       
  2075                     by (rule_tac x = cs in exI, auto simp add: wq_def)
       
  2076                 qed
       
  2077                 moreover have "th \<in> threads s"
       
  2078                 proof -
       
  2079                   from th_in eq_wq
       
  2080                   have "th \<in> set (wq s cs)" by simp
       
  2081                   from vt_s.wq_threads [OF this] 
       
  2082                   show ?thesis .
       
  2083                 qed
       
  2084                 ultimately show ?thesis using ih by auto
       
  2085               qed
       
  2086               moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
       
  2087                 apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto)
       
  2088               proof -
       
  2089                 show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
       
  2090                                Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
       
  2091                   (is "card ?A = Suc (card ?B)")
       
  2092                 proof -
       
  2093                   have "?A = insert cs ?B" by auto
       
  2094                   hence "card ?A = card (insert cs ?B)" by simp
       
  2095                   also have "\<dots> = Suc (card ?B)"
       
  2096                   proof(rule card_insert_disjoint)
       
  2097                     have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
       
  2098                       apply (auto simp:image_def)
       
  2099                       by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
       
  2100                     with vt_s.finite_RAG
       
  2101                     show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
       
  2102                   next
       
  2103                     show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
       
  2104                     proof
       
  2105                       assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
       
  2106                       hence "(Cs cs, Th th) \<in> RAG s" by simp
       
  2107                       with True neq_th eq_wq show False
       
  2108                         by (auto simp:next_th_def s_RAG_def cs_holding_def)
       
  2109                     qed
       
  2110                   qed
       
  2111                   finally show ?thesis .
       
  2112                 qed
       
  2113               qed
       
  2114               moreover note eq_cnp eq_cnv
       
  2115               ultimately show ?thesis by simp
       
  2116             qed
       
  2117           qed
       
  2118         } ultimately show ?thesis by blast
       
  2119       qed
       
  2120     next
       
  2121       case (thread_set thread prio)
       
  2122       assume eq_e: "e = Set thread prio"
       
  2123         and is_runing: "thread \<in> runing s"
       
  2124       show ?thesis
       
  2125       proof -
       
  2126         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
       
  2127         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
       
  2128         have eq_cncs: "cntCS (e#s) th = cntCS s th"
       
  2129           unfolding cntCS_def holdents_test
       
  2130           by (simp add:RAG_set_unchanged eq_e)
       
  2131         from eq_e have eq_readys: "readys (e#s) = readys s" 
       
  2132           by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
       
  2133                   auto simp:Let_def)
       
  2134         { assume "th \<noteq> thread"
       
  2135           with eq_readys eq_e
       
  2136           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
       
  2137                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
       
  2138             by (simp add:threads.simps)
       
  2139           with eq_cnp eq_cnv eq_cncs ih is_runing
       
  2140           have ?thesis by simp
       
  2141         } moreover {
       
  2142           assume eq_th: "th = thread"
       
  2143           with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
       
  2144             by (unfold runing_def, auto)
       
  2145           moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
       
  2146             by (simp add:runing_def)
       
  2147           moreover note eq_cnp eq_cnv eq_cncs
       
  2148           ultimately have ?thesis by auto
       
  2149         } ultimately show ?thesis by blast
       
  2150       qed   
       
  2151     qed
       
  2152   next
       
  2153     case vt_nil
       
  2154     show ?case 
       
  2155       by (unfold cntP_def cntV_def cntCS_def, 
       
  2156         auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
       
  2157   qed
       
  2158 qed
       
  2159 
       
  2160 lemma not_thread_cncs:
       
  2161   assumes not_in: "th \<notin> threads s" 
       
  2162   shows "cntCS s th = 0"
       
  2163 proof -
       
  2164   from vt not_in show ?thesis
       
  2165   proof(induct arbitrary:th)
       
  2166     case (vt_cons s e th)
       
  2167     interpret vt_s: valid_trace s using vt_cons(1)
       
  2168        by (unfold_locales, simp)
       
  2169     assume vt: "vt s"
       
  2170       and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
       
  2171       and stp: "step s e"
       
  2172       and not_in: "th \<notin> threads (e # s)"
       
  2173     from stp show ?case
       
  2174     proof(cases)
       
  2175       case (thread_create thread prio)
       
  2176       assume eq_e: "e = Create thread prio"
       
  2177         and not_in': "thread \<notin> threads s"
       
  2178       have "cntCS (e # s) th = cntCS s th"
       
  2179         apply (unfold eq_e cntCS_def holdents_test)
       
  2180         by (simp add:RAG_create_unchanged)
       
  2181       moreover have "th \<notin> threads s" 
       
  2182       proof -
       
  2183         from not_in eq_e show ?thesis by simp
       
  2184       qed
       
  2185       moreover note ih ultimately show ?thesis by auto
       
  2186     next
       
  2187       case (thread_exit thread)
       
  2188       assume eq_e: "e = Exit thread"
       
  2189       and nh: "holdents s thread = {}"
       
  2190       have eq_cns: "cntCS (e # s) th = cntCS s th"
       
  2191         apply (unfold eq_e cntCS_def holdents_test)
       
  2192         by (simp add:RAG_exit_unchanged)
       
  2193       show ?thesis
       
  2194       proof(cases "th = thread")
       
  2195         case True
       
  2196         have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
       
  2197         with eq_cns show ?thesis by simp
       
  2198       next
       
  2199         case False
       
  2200         with not_in and eq_e
       
  2201         have "th \<notin> threads s" by simp
       
  2202         from ih[OF this] and eq_cns show ?thesis by simp
       
  2203       qed
       
  2204     next
       
  2205       case (thread_P thread cs)
       
  2206       assume eq_e: "e = P thread cs"
       
  2207       and is_runing: "thread \<in> runing s"
       
  2208       from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto
       
  2209       have neq_th: "th \<noteq> thread" 
       
  2210       proof -
       
  2211         from not_in eq_e have "th \<notin> threads s" by simp
       
  2212         moreover from is_runing have "thread \<in> threads s"
       
  2213           by (simp add:runing_def readys_def)
       
  2214         ultimately show ?thesis by auto
       
  2215       qed
       
  2216       hence "cntCS (e # s) th  = cntCS s th "
       
  2217         apply (unfold cntCS_def holdents_test eq_e)
       
  2218         by (unfold step_RAG_p[OF vtp], auto)
       
  2219       moreover have "cntCS s th = 0"
       
  2220       proof(rule ih)
       
  2221         from not_in eq_e show "th \<notin> threads s" by simp
       
  2222       qed
       
  2223       ultimately show ?thesis by simp
       
  2224     next
       
  2225       case (thread_V thread cs)
       
  2226       assume eq_e: "e = V thread cs"
       
  2227         and is_runing: "thread \<in> runing s"
       
  2228         and hold: "holding s thread cs"
       
  2229       have neq_th: "th \<noteq> thread" 
       
  2230       proof -
       
  2231         from not_in eq_e have "th \<notin> threads s" by simp
       
  2232         moreover from is_runing have "thread \<in> threads s"
       
  2233           by (simp add:runing_def readys_def)
       
  2234         ultimately show ?thesis by auto
       
  2235       qed
       
  2236       from assms thread_V vt stp ih 
       
  2237       have vtv: "vt (V thread cs#s)" by auto
       
  2238       then interpret vt_v: valid_trace "(V thread cs#s)"
       
  2239         by (unfold_locales, simp)
       
  2240       from hold obtain rest 
       
  2241         where eq_wq: "wq s cs = thread # rest"
       
  2242         by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
       
  2243       from not_in eq_e eq_wq
       
  2244       have "\<not> next_th s thread cs th"
       
  2245         apply (auto simp:next_th_def)
       
  2246       proof -
       
  2247         assume ne: "rest \<noteq> []"
       
  2248           and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s")
       
  2249         have "?t \<in> set rest"
       
  2250         proof(rule someI2)
       
  2251           from vt_v.wq_distinct[of cs] and eq_wq
       
  2252           show "distinct rest \<and> set rest = set rest"
       
  2253             by (metis distinct.simps(2) vt_s.wq_distinct) 
       
  2254         next
       
  2255           fix x assume "distinct x \<and> set x = set rest" with ne
       
  2256           show "hd x \<in> set rest" by (cases x, auto)
       
  2257         qed
       
  2258         with eq_wq have "?t \<in> set (wq s cs)" by simp
       
  2259         from vt_s.wq_threads[OF this] and ni
       
  2260         show False
       
  2261           using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` 
       
  2262             ni vt_s.wq_threads by blast 
       
  2263       qed
       
  2264       moreover note neq_th eq_wq
       
  2265       ultimately have "cntCS (e # s) th  = cntCS s th"
       
  2266         by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
       
  2267       moreover have "cntCS s th = 0"
       
  2268       proof(rule ih)
       
  2269         from not_in eq_e show "th \<notin> threads s" by simp
       
  2270       qed
       
  2271       ultimately show ?thesis by simp
       
  2272     next
       
  2273       case (thread_set thread prio)
       
  2274       print_facts
       
  2275       assume eq_e: "e = Set thread prio"
       
  2276         and is_runing: "thread \<in> runing s"
       
  2277       from not_in and eq_e have "th \<notin> threads s" by auto
       
  2278       from ih [OF this] and eq_e
       
  2279       show ?thesis 
       
  2280         apply (unfold eq_e cntCS_def holdents_test)
       
  2281         by (simp add:RAG_set_unchanged)
       
  2282     qed
       
  2283     next
       
  2284       case vt_nil
       
  2285       show ?case
       
  2286       by (unfold cntCS_def, 
       
  2287         auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
       
  2288   qed
       
  2289 qed
       
  2290 
       
  2291 end
       
  2292 
       
  2293 lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
       
  2294   by (auto simp:s_waiting_def cs_waiting_def wq_def)
       
  2295 
       
  2296 context valid_trace
       
  2297 begin
       
  2298 
       
  2299 lemma dm_RAG_threads:
       
  2300   assumes in_dom: "(Th th) \<in> Domain (RAG s)"
       
  2301   shows "th \<in> threads s"
       
  2302 proof -
       
  2303   from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
       
  2304   moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
       
  2305   ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
       
  2306   hence "th \<in> set (wq s cs)"
       
  2307     by (unfold s_RAG_def, auto simp:cs_waiting_def)
       
  2308   from wq_threads [OF this] show ?thesis .
       
  2309 qed
       
  2310 
       
  2311 end
       
  2312 
       
  2313 lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
       
  2314 unfolding cp_def wq_def
       
  2315 apply(induct s rule: schs.induct)
       
  2316 thm cpreced_initial
       
  2317 apply(simp add: Let_def cpreced_initial)
       
  2318 apply(simp add: Let_def)
       
  2319 apply(simp add: Let_def)
       
  2320 apply(simp add: Let_def)
       
  2321 apply(subst (2) schs.simps)
       
  2322 apply(simp add: Let_def)
       
  2323 apply(subst (2) schs.simps)
       
  2324 apply(simp add: Let_def)
       
  2325 done
       
  2326 
       
  2327 context valid_trace
       
  2328 begin
       
  2329 
       
  2330 lemma runing_unique:
       
  2331   assumes runing_1: "th1 \<in> runing s"
       
  2332   and runing_2: "th2 \<in> runing s"
       
  2333   shows "th1 = th2"
       
  2334 proof -
       
  2335   from runing_1 and runing_2 have "cp s th1 = cp s th2"
       
  2336     unfolding runing_def
       
  2337     apply(simp)
       
  2338     done
       
  2339   hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) =
       
  2340                  Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))"
       
  2341     (is "Max (?f ` ?A) = Max (?f ` ?B)")
       
  2342     unfolding cp_eq_cpreced 
       
  2343     unfolding cpreced_def .
       
  2344   obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
       
  2345   proof -
       
  2346     have h1: "finite (?f ` ?A)"
       
  2347     proof -
       
  2348       have "finite ?A" 
       
  2349       proof -
       
  2350         have "finite (dependants (wq s) th1)"
       
  2351         proof-
       
  2352           have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
       
  2353           proof -
       
  2354             let ?F = "\<lambda> (x, y). the_th x"
       
  2355             have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
       
  2356               apply (auto simp:image_def)
       
  2357               by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
       
  2358             moreover have "finite \<dots>"
       
  2359             proof -
       
  2360               from finite_RAG have "finite (RAG s)" .
       
  2361               hence "finite ((RAG (wq s))\<^sup>+)"
       
  2362                 apply (unfold finite_trancl)
       
  2363                 by (auto simp: s_RAG_def cs_RAG_def wq_def)
       
  2364               thus ?thesis by auto
       
  2365             qed
       
  2366             ultimately show ?thesis by (auto intro:finite_subset)
       
  2367           qed
       
  2368           thus ?thesis by (simp add:cs_dependants_def)
       
  2369         qed
       
  2370         thus ?thesis by simp
       
  2371       qed
       
  2372       thus ?thesis by auto
       
  2373     qed
       
  2374     moreover have h2: "(?f ` ?A) \<noteq> {}"
       
  2375     proof -
       
  2376       have "?A \<noteq> {}" by simp
       
  2377       thus ?thesis by simp
       
  2378     qed
       
  2379     from Max_in [OF h1 h2]
       
  2380     have "Max (?f ` ?A) \<in> (?f ` ?A)" .
       
  2381     thus ?thesis 
       
  2382       thm cpreced_def
       
  2383       unfolding cpreced_def[symmetric] 
       
  2384       unfolding cp_eq_cpreced[symmetric] 
       
  2385       unfolding cpreced_def 
       
  2386       using that[intro] by (auto)
       
  2387   qed
       
  2388   obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
       
  2389   proof -
       
  2390     have h1: "finite (?f ` ?B)"
       
  2391     proof -
       
  2392       have "finite ?B" 
       
  2393       proof -
       
  2394         have "finite (dependants (wq s) th2)"
       
  2395         proof-
       
  2396           have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
       
  2397           proof -
       
  2398             let ?F = "\<lambda> (x, y). the_th x"
       
  2399             have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
       
  2400               apply (auto simp:image_def)
       
  2401               by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
       
  2402             moreover have "finite \<dots>"
       
  2403             proof -
       
  2404               from finite_RAG have "finite (RAG s)" .
       
  2405               hence "finite ((RAG (wq s))\<^sup>+)"
       
  2406                 apply (unfold finite_trancl)
       
  2407                 by (auto simp: s_RAG_def cs_RAG_def wq_def)
       
  2408               thus ?thesis by auto
       
  2409             qed
       
  2410             ultimately show ?thesis by (auto intro:finite_subset)
       
  2411           qed
       
  2412           thus ?thesis by (simp add:cs_dependants_def)
       
  2413         qed
       
  2414         thus ?thesis by simp
       
  2415       qed
       
  2416       thus ?thesis by auto
       
  2417     qed
       
  2418     moreover have h2: "(?f ` ?B) \<noteq> {}"
       
  2419     proof -
       
  2420       have "?B \<noteq> {}" by simp
       
  2421       thus ?thesis by simp
       
  2422     qed
       
  2423     from Max_in [OF h1 h2]
       
  2424     have "Max (?f ` ?B) \<in> (?f ` ?B)" .
       
  2425     thus ?thesis by (auto intro:that)
       
  2426   qed
       
  2427   from eq_f_th1 eq_f_th2 eq_max 
       
  2428   have eq_preced: "preced th1' s = preced th2' s" by auto
       
  2429   hence eq_th12: "th1' = th2'"
       
  2430   proof (rule preced_unique)
       
  2431     from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
       
  2432     thus "th1' \<in> threads s"
       
  2433     proof
       
  2434       assume "th1' \<in> dependants (wq s) th1"
       
  2435       hence "(Th th1') \<in> Domain ((RAG s)^+)"
       
  2436         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
       
  2437         by (auto simp:Domain_def)
       
  2438       hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
       
  2439       from dm_RAG_threads[OF this] show ?thesis .
       
  2440     next
       
  2441       assume "th1' = th1"
       
  2442       with runing_1 show ?thesis
       
  2443         by (unfold runing_def readys_def, auto)
       
  2444     qed
       
  2445   next
       
  2446     from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
       
  2447     thus "th2' \<in> threads s"
       
  2448     proof
       
  2449       assume "th2' \<in> dependants (wq s) th2"
       
  2450       hence "(Th th2') \<in> Domain ((RAG s)^+)"
       
  2451         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
       
  2452         by (auto simp:Domain_def)
       
  2453       hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
       
  2454       from dm_RAG_threads[OF this] show ?thesis .
       
  2455     next
       
  2456       assume "th2' = th2"
       
  2457       with runing_2 show ?thesis
       
  2458         by (unfold runing_def readys_def, auto)
       
  2459     qed
       
  2460   qed
       
  2461   from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp
       
  2462   thus ?thesis
       
  2463   proof
       
  2464     assume eq_th': "th1' = th1"
       
  2465     from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
       
  2466     thus ?thesis
       
  2467     proof
       
  2468       assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
       
  2469     next
       
  2470       assume "th2' \<in> dependants (wq s) th2"
       
  2471       with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
       
  2472       hence "(Th th1, Th th2) \<in> (RAG s)^+"
       
  2473         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       
  2474       hence "Th th1 \<in> Domain ((RAG s)^+)" 
       
  2475         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
       
  2476         by (auto simp:Domain_def)
       
  2477       hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
       
  2478       then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
       
  2479       from RAG_target_th [OF this]
       
  2480       obtain cs' where "n = Cs cs'" by auto
       
  2481       with d have "(Th th1, Cs cs') \<in> RAG s" by simp
       
  2482       with runing_1 have "False"
       
  2483         apply (unfold runing_def readys_def s_RAG_def)
       
  2484         by (auto simp:eq_waiting)
       
  2485       thus ?thesis by simp
       
  2486     qed
       
  2487   next
       
  2488     assume th1'_in: "th1' \<in> dependants (wq s) th1"
       
  2489     from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
       
  2490     thus ?thesis 
       
  2491     proof
       
  2492       assume "th2' = th2"
       
  2493       with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
       
  2494       hence "(Th th2, Th th1) \<in> (RAG s)^+"
       
  2495         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       
  2496       hence "Th th2 \<in> Domain ((RAG s)^+)" 
       
  2497         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
       
  2498         by (auto simp:Domain_def)
       
  2499       hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
       
  2500       then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
       
  2501       from RAG_target_th [OF this]
       
  2502       obtain cs' where "n = Cs cs'" by auto
       
  2503       with d have "(Th th2, Cs cs') \<in> RAG s" by simp
       
  2504       with runing_2 have "False"
       
  2505         apply (unfold runing_def readys_def s_RAG_def)
       
  2506         by (auto simp:eq_waiting)
       
  2507       thus ?thesis by simp
       
  2508     next
       
  2509       assume "th2' \<in> dependants (wq s) th2"
       
  2510       with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
       
  2511       hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
       
  2512         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       
  2513       from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
       
  2514         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       
  2515       show ?thesis
       
  2516       proof(rule dchain_unique[OF h1 _ h2, symmetric])
       
  2517         from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
       
  2518         from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
       
  2519       qed
       
  2520     qed
       
  2521   qed
       
  2522 qed
       
  2523 
       
  2524 
       
  2525 lemma "card (runing s) \<le> 1"
       
  2526 apply(subgoal_tac "finite (runing s)")
       
  2527 prefer 2
       
  2528 apply (metis finite_nat_set_iff_bounded lessI runing_unique)
       
  2529 apply(rule ccontr)
       
  2530 apply(simp)
       
  2531 apply(case_tac "Suc (Suc 0) \<le> card (runing s)")
       
  2532 apply(subst (asm) card_le_Suc_iff)
       
  2533 apply(simp)
       
  2534 apply(auto)[1]
       
  2535 apply (metis insertCI runing_unique)
       
  2536 apply(auto) 
       
  2537 done
       
  2538 
       
  2539 end
       
  2540 
       
  2541 
       
  2542 lemma create_pre:
       
  2543   assumes stp: "step s e"
       
  2544   and not_in: "th \<notin> threads s"
       
  2545   and is_in: "th \<in> threads (e#s)"
       
  2546   obtains prio where "e = Create th prio"
       
  2547 proof -
       
  2548   from assms  
       
  2549   show ?thesis
       
  2550   proof(cases)
       
  2551     case (thread_create thread prio)
       
  2552     with is_in not_in have "e = Create th prio" by simp
       
  2553     from that[OF this] show ?thesis .
       
  2554   next
       
  2555     case (thread_exit thread)
       
  2556     with assms show ?thesis by (auto intro!:that)
       
  2557   next
       
  2558     case (thread_P thread)
       
  2559     with assms show ?thesis by (auto intro!:that)
       
  2560   next
       
  2561     case (thread_V thread)
       
  2562     with assms show ?thesis by (auto intro!:that)
       
  2563   next 
       
  2564     case (thread_set thread)
       
  2565     with assms show ?thesis by (auto intro!:that)
       
  2566   qed
       
  2567 qed
       
  2568 
       
  2569 lemma length_down_to_in: 
       
  2570   assumes le_ij: "i \<le> j"
       
  2571     and le_js: "j \<le> length s"
       
  2572   shows "length (down_to j i s) = j - i"
       
  2573 proof -
       
  2574   have "length (down_to j i s) = length (from_to i j (rev s))"
       
  2575     by (unfold down_to_def, auto)
       
  2576   also have "\<dots> = j - i"
       
  2577   proof(rule length_from_to_in[OF le_ij])
       
  2578     from le_js show "j \<le> length (rev s)" by simp
       
  2579   qed
       
  2580   finally show ?thesis .
       
  2581 qed
       
  2582 
       
  2583 
       
  2584 lemma moment_head: 
       
  2585   assumes le_it: "Suc i \<le> length t"
       
  2586   obtains e where "moment (Suc i) t = e#moment i t"
       
  2587 proof -
       
  2588   have "i \<le> Suc i" by simp
       
  2589   from length_down_to_in [OF this le_it]
       
  2590   have "length (down_to (Suc i) i t) = 1" by auto
       
  2591   then obtain e where "down_to (Suc i) i t = [e]"
       
  2592     apply (cases "(down_to (Suc i) i t)") by auto
       
  2593   moreover have "down_to (Suc i) 0 t = down_to (Suc i) i t @ down_to i 0 t"
       
  2594     by (rule down_to_conc[symmetric], auto)
       
  2595   ultimately have eq_me: "moment (Suc i) t = e#(moment i t)"
       
  2596     by (auto simp:down_to_moment)
       
  2597   from that [OF this] show ?thesis .
       
  2598 qed
       
  2599 
       
  2600 context valid_trace
       
  2601 begin
       
  2602 
       
  2603 lemma cnp_cnv_eq:
       
  2604   assumes "th \<notin> threads s"
       
  2605   shows "cntP s th = cntV s th"
       
  2606   using assms
       
  2607   using cnp_cnv_cncs not_thread_cncs by auto
       
  2608 
       
  2609 end
       
  2610 
       
  2611 
       
  2612 lemma eq_RAG: 
       
  2613   "RAG (wq s) = RAG s"
       
  2614 by (unfold cs_RAG_def s_RAG_def, auto)
       
  2615 
       
  2616 context valid_trace
       
  2617 begin
       
  2618 
       
  2619 lemma count_eq_dependants:
       
  2620   assumes eq_pv: "cntP s th = cntV s th"
       
  2621   shows "dependants (wq s) th = {}"
       
  2622 proof -
       
  2623   from cnp_cnv_cncs and eq_pv
       
  2624   have "cntCS s th = 0" 
       
  2625     by (auto split:if_splits)
       
  2626   moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
       
  2627   proof -
       
  2628     from finite_holding[of th] show ?thesis
       
  2629       by (simp add:holdents_test)
       
  2630   qed
       
  2631   ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
       
  2632     by (unfold cntCS_def holdents_test cs_dependants_def, auto)
       
  2633   show ?thesis
       
  2634   proof(unfold cs_dependants_def)
       
  2635     { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
       
  2636       then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
       
  2637       hence "False"
       
  2638       proof(cases)
       
  2639         assume "(Th th', Th th) \<in> RAG (wq s)"
       
  2640         thus "False" by (auto simp:cs_RAG_def)
       
  2641       next
       
  2642         fix c
       
  2643         assume "(c, Th th) \<in> RAG (wq s)"
       
  2644         with h and eq_RAG show "False"
       
  2645           by (cases c, auto simp:cs_RAG_def)
       
  2646       qed
       
  2647     } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
       
  2648   qed
       
  2649 qed
       
  2650 
       
  2651 lemma dependants_threads:
       
  2652   shows "dependants (wq s) th \<subseteq> threads s"
       
  2653 proof
       
  2654   { fix th th'
       
  2655     assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
       
  2656     have "Th th \<in> Domain (RAG s)"
       
  2657     proof -
       
  2658       from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
       
  2659       hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
       
  2660       with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
       
  2661       thus ?thesis using eq_RAG by simp
       
  2662     qed
       
  2663     from dm_RAG_threads[OF this]
       
  2664     have "th \<in> threads s" .
       
  2665   } note hh = this
       
  2666   fix th1 
       
  2667   assume "th1 \<in> dependants (wq s) th"
       
  2668   hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
       
  2669     by (unfold cs_dependants_def, simp)
       
  2670   from hh [OF this] show "th1 \<in> threads s" .
       
  2671 qed
       
  2672 
       
  2673 lemma finite_threads:
       
  2674   shows "finite (threads s)"
       
  2675 using vt by (induct) (auto elim: step.cases)
       
  2676 
       
  2677 end
       
  2678 
       
  2679 lemma Max_f_mono:
       
  2680   assumes seq: "A \<subseteq> B"
       
  2681   and np: "A \<noteq> {}"
       
  2682   and fnt: "finite B"
       
  2683   shows "Max (f ` A) \<le> Max (f ` B)"
       
  2684 proof(rule Max_mono)
       
  2685   from seq show "f ` A \<subseteq> f ` B" by auto
       
  2686 next
       
  2687   from np show "f ` A \<noteq> {}" by auto
       
  2688 next
       
  2689   from fnt and seq show "finite (f ` B)" by auto
       
  2690 qed
       
  2691 
       
  2692 context valid_trace
       
  2693 begin
       
  2694 
       
  2695 lemma cp_le:
       
  2696   assumes th_in: "th \<in> threads s"
       
  2697   shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
       
  2698 proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def)
       
  2699   show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
       
  2700          \<le> Max ((\<lambda>th. preced th s) ` threads s)"
       
  2701     (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
       
  2702   proof(rule Max_f_mono)
       
  2703     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp
       
  2704   next
       
  2705     from finite_threads
       
  2706     show "finite (threads s)" .
       
  2707   next
       
  2708     from th_in
       
  2709     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
       
  2710       apply (auto simp:Domain_def)
       
  2711       apply (rule_tac dm_RAG_threads)
       
  2712       apply (unfold trancl_domain [of "RAG s", symmetric])
       
  2713       by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def)
       
  2714   qed
       
  2715 qed
       
  2716 
       
  2717 lemma le_cp:
       
  2718   shows "preced th s \<le> cp s th"
       
  2719 proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
       
  2720   show "Prc (priority th s) (last_set th s)
       
  2721     \<le> Max (insert (Prc (priority th s) (last_set th s))
       
  2722             ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))"
       
  2723     (is "?l \<le> Max (insert ?l ?A)")
       
  2724   proof(cases "?A = {}")
       
  2725     case False
       
  2726     have "finite ?A" (is "finite (?f ` ?B)")
       
  2727     proof -
       
  2728       have "finite ?B" 
       
  2729       proof-
       
  2730         have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
       
  2731         proof -
       
  2732           let ?F = "\<lambda> (x, y). the_th x"
       
  2733           have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
       
  2734             apply (auto simp:image_def)
       
  2735             by (rule_tac x = "(Th x, Th th)" in bexI, auto)
       
  2736           moreover have "finite \<dots>"
       
  2737           proof -
       
  2738             from finite_RAG have "finite (RAG s)" .
       
  2739             hence "finite ((RAG (wq s))\<^sup>+)"
       
  2740               apply (unfold finite_trancl)
       
  2741               by (auto simp: s_RAG_def cs_RAG_def wq_def)
       
  2742             thus ?thesis by auto
       
  2743           qed
       
  2744           ultimately show ?thesis by (auto intro:finite_subset)
       
  2745         qed
       
  2746         thus ?thesis by (simp add:cs_dependants_def)
       
  2747       qed
       
  2748       thus ?thesis by simp
       
  2749     qed
       
  2750     from Max_insert [OF this False, of ?l] show ?thesis by auto
       
  2751   next
       
  2752     case True
       
  2753     thus ?thesis by auto
       
  2754   qed
       
  2755 qed
       
  2756 
       
  2757 lemma max_cp_eq: 
       
  2758   shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)"
       
  2759   (is "?l = ?r")
       
  2760 proof(cases "threads s = {}")
       
  2761   case True
       
  2762   thus ?thesis by auto
       
  2763 next
       
  2764   case False
       
  2765   have "?l \<in> ((cp s) ` threads s)"
       
  2766   proof(rule Max_in)
       
  2767     from finite_threads
       
  2768     show "finite (cp s ` threads s)" by auto
       
  2769   next
       
  2770     from False show "cp s ` threads s \<noteq> {}" by auto
       
  2771   qed
       
  2772   then obtain th 
       
  2773     where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto
       
  2774   have "\<dots> \<le> ?r" by (rule cp_le[OF th_in])
       
  2775   moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th")
       
  2776   proof -
       
  2777     have "?r \<in> (?f ` ?A)"
       
  2778     proof(rule Max_in)
       
  2779       from finite_threads
       
  2780       show " finite ((\<lambda>th. preced th s) ` threads s)" by auto
       
  2781     next
       
  2782       from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
       
  2783     qed
       
  2784     then obtain th' where 
       
  2785       th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto
       
  2786     from le_cp [of th']  eq_r
       
  2787     have "?r \<le> cp s th'" by auto
       
  2788     moreover have "\<dots> \<le> cp s th"
       
  2789     proof(fold eq_l)
       
  2790       show " cp s th' \<le> Max (cp s ` threads s)"
       
  2791       proof(rule Max_ge)
       
  2792         from th_in' show "cp s th' \<in> cp s ` threads s"
       
  2793           by auto
       
  2794       next
       
  2795         from finite_threads
       
  2796         show "finite (cp s ` threads s)" by auto
       
  2797       qed
       
  2798     qed
       
  2799     ultimately show ?thesis by auto
       
  2800   qed
       
  2801   ultimately show ?thesis using eq_l by auto
       
  2802 qed
       
  2803 
       
  2804 lemma max_cp_readys_threads_pre:
       
  2805   assumes np: "threads s \<noteq> {}"
       
  2806   shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
       
  2807 proof(unfold max_cp_eq)
       
  2808   show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
       
  2809   proof -
       
  2810     let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
       
  2811     let ?f = "(\<lambda>th. preced th s)"
       
  2812     have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
       
  2813     proof(rule Max_in)
       
  2814       from finite_threads show "finite (?f ` threads s)" by simp
       
  2815     next
       
  2816       from np show "?f ` threads s \<noteq> {}" by simp
       
  2817     qed
       
  2818     then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
       
  2819       by (auto simp:Image_def)
       
  2820     from th_chain_to_ready [OF tm_in]
       
  2821     have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
       
  2822     thus ?thesis
       
  2823     proof
       
  2824       assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
       
  2825       then obtain th' where th'_in: "th' \<in> readys s" 
       
  2826         and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
       
  2827       have "cp s th' = ?f tm"
       
  2828       proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
       
  2829         from dependants_threads finite_threads
       
  2830         show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
       
  2831           by (auto intro:finite_subset)
       
  2832       next
       
  2833         fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
       
  2834         from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
       
  2835         moreover have "p \<le> \<dots>"
       
  2836         proof(rule Max_ge)
       
  2837           from finite_threads
       
  2838           show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2839         next
       
  2840           from p_in and th'_in and dependants_threads[of th']
       
  2841           show "p \<in> (\<lambda>th. preced th s) ` threads s"
       
  2842             by (auto simp:readys_def)
       
  2843         qed
       
  2844         ultimately show "p \<le> preced tm s" by auto
       
  2845       next
       
  2846         show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
       
  2847         proof -
       
  2848           from tm_chain
       
  2849           have "tm \<in> dependants (wq s) th'"
       
  2850             by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
       
  2851           thus ?thesis by auto
       
  2852         qed
       
  2853       qed
       
  2854       with tm_max
       
  2855       have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2856       show ?thesis
       
  2857       proof (fold h, rule Max_eqI)
       
  2858         fix q 
       
  2859         assume "q \<in> cp s ` readys s"
       
  2860         then obtain th1 where th1_in: "th1 \<in> readys s"
       
  2861           and eq_q: "q = cp s th1" by auto
       
  2862         show "q \<le> cp s th'"
       
  2863           apply (unfold h eq_q)
       
  2864           apply (unfold cp_eq_cpreced cpreced_def)
       
  2865           apply (rule Max_mono)
       
  2866         proof -
       
  2867           from dependants_threads [of th1] th1_in
       
  2868           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
       
  2869                  (\<lambda>th. preced th s) ` threads s"
       
  2870             by (auto simp:readys_def)
       
  2871         next
       
  2872           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
       
  2873         next
       
  2874           from finite_threads 
       
  2875           show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2876         qed
       
  2877       next
       
  2878         from finite_threads
       
  2879         show "finite (cp s ` readys s)" by (auto simp:readys_def)
       
  2880       next
       
  2881         from th'_in
       
  2882         show "cp s th' \<in> cp s ` readys s" by simp
       
  2883       qed
       
  2884     next
       
  2885       assume tm_ready: "tm \<in> readys s"
       
  2886       show ?thesis
       
  2887       proof(fold tm_max)
       
  2888         have cp_eq_p: "cp s tm = preced tm s"
       
  2889         proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
       
  2890           fix y 
       
  2891           assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
       
  2892           show "y \<le> preced tm s"
       
  2893           proof -
       
  2894             { fix y'
       
  2895               assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
       
  2896               have "y' \<le> preced tm s"
       
  2897               proof(unfold tm_max, rule Max_ge)
       
  2898                 from hy' dependants_threads[of tm]
       
  2899                 show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
       
  2900               next
       
  2901                 from finite_threads
       
  2902                 show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2903               qed
       
  2904             } with hy show ?thesis by auto
       
  2905           qed
       
  2906         next
       
  2907           from dependants_threads[of tm] finite_threads
       
  2908           show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
       
  2909             by (auto intro:finite_subset)
       
  2910         next
       
  2911           show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
       
  2912             by simp
       
  2913         qed 
       
  2914         moreover have "Max (cp s ` readys s) = cp s tm"
       
  2915         proof(rule Max_eqI)
       
  2916           from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
       
  2917         next
       
  2918           from finite_threads
       
  2919           show "finite (cp s ` readys s)" by (auto simp:readys_def)
       
  2920         next
       
  2921           fix y assume "y \<in> cp s ` readys s"
       
  2922           then obtain th1 where th1_readys: "th1 \<in> readys s"
       
  2923             and h: "y = cp s th1" by auto
       
  2924           show "y \<le> cp s tm"
       
  2925             apply(unfold cp_eq_p h)
       
  2926             apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
       
  2927           proof -
       
  2928             from finite_threads
       
  2929             show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2930           next
       
  2931             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
       
  2932               by simp
       
  2933           next
       
  2934             from dependants_threads[of th1] th1_readys
       
  2935             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
       
  2936                     \<subseteq> (\<lambda>th. preced th s) ` threads s"
       
  2937               by (auto simp:readys_def)
       
  2938           qed
       
  2939         qed
       
  2940         ultimately show " Max (cp s ` readys s) = preced tm s" by simp
       
  2941       qed 
       
  2942     qed
       
  2943   qed
       
  2944 qed
       
  2945 
       
  2946 text {* (* ccc *) \noindent
       
  2947   Since the current precedence of the threads in ready queue will always be boosted,
       
  2948   there must be one inside it has the maximum precedence of the whole system. 
       
  2949 *}
       
  2950 lemma max_cp_readys_threads:
       
  2951   shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
       
  2952 proof(cases "threads s = {}")
       
  2953   case True
       
  2954   thus ?thesis 
       
  2955     by (auto simp:readys_def)
       
  2956 next
       
  2957   case False
       
  2958   show ?thesis by (rule max_cp_readys_threads_pre[OF False])
       
  2959 qed
       
  2960 
       
  2961 end
       
  2962 
       
  2963 lemma eq_holding: "holding (wq s) th cs = holding s th cs"
       
  2964   apply (unfold s_holding_def cs_holding_def wq_def, simp)
       
  2965   done
       
  2966 
       
  2967 lemma f_image_eq:
       
  2968   assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
       
  2969   shows "f ` A = g ` A"
       
  2970 proof
       
  2971   show "f ` A \<subseteq> g ` A"
       
  2972     by(rule image_subsetI, auto intro:h)
       
  2973 next
       
  2974   show "g ` A \<subseteq> f ` A"
       
  2975    by (rule image_subsetI, auto intro:h[symmetric])
       
  2976 qed
       
  2977 
       
  2978 
       
  2979 definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
       
  2980   where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
       
  2981 
       
  2982 
       
  2983 lemma detached_test:
       
  2984   shows "detached s th = (Th th \<notin> Field (RAG s))"
       
  2985 apply(simp add: detached_def Field_def)
       
  2986 apply(simp add: s_RAG_def)
       
  2987 apply(simp add: s_holding_abv s_waiting_abv)
       
  2988 apply(simp add: Domain_iff Range_iff)
       
  2989 apply(simp add: wq_def)
       
  2990 apply(auto)
       
  2991 done
       
  2992 
       
  2993 context valid_trace
       
  2994 begin
       
  2995 
       
  2996 lemma detached_intro:
       
  2997   assumes eq_pv: "cntP s th = cntV s th"
       
  2998   shows "detached s th"
       
  2999 proof -
       
  3000  from cnp_cnv_cncs
       
  3001   have eq_cnt: "cntP s th =
       
  3002     cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
       
  3003   hence cncs_zero: "cntCS s th = 0"
       
  3004     by (auto simp:eq_pv split:if_splits)
       
  3005   with eq_cnt
       
  3006   have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
       
  3007   thus ?thesis
       
  3008   proof
       
  3009     assume "th \<notin> threads s"
       
  3010     with range_in dm_RAG_threads
       
  3011     show ?thesis
       
  3012       by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
       
  3013   next
       
  3014     assume "th \<in> readys s"
       
  3015     moreover have "Th th \<notin> Range (RAG s)"
       
  3016     proof -
       
  3017       from card_0_eq [OF finite_holding] and cncs_zero
       
  3018       have "holdents s th = {}"
       
  3019         by (simp add:cntCS_def)
       
  3020       thus ?thesis
       
  3021         apply(auto simp:holdents_test)
       
  3022         apply(case_tac a)
       
  3023         apply(auto simp:holdents_test s_RAG_def)
       
  3024         done
       
  3025     qed
       
  3026     ultimately show ?thesis
       
  3027       by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
       
  3028   qed
       
  3029 qed
       
  3030 
       
  3031 lemma detached_elim:
       
  3032   assumes dtc: "detached s th"
       
  3033   shows "cntP s th = cntV s th"
       
  3034 proof -
       
  3035   from cnp_cnv_cncs
       
  3036   have eq_pv: " cntP s th =
       
  3037     cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
       
  3038   have cncs_z: "cntCS s th = 0"
       
  3039   proof -
       
  3040     from dtc have "holdents s th = {}"
       
  3041       unfolding detached_def holdents_test s_RAG_def
       
  3042       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
       
  3043     thus ?thesis by (auto simp:cntCS_def)
       
  3044   qed
       
  3045   show ?thesis
       
  3046   proof(cases "th \<in> threads s")
       
  3047     case True
       
  3048     with dtc 
       
  3049     have "th \<in> readys s"
       
  3050       by (unfold readys_def detached_def Field_def Domain_def Range_def, 
       
  3051            auto simp:eq_waiting s_RAG_def)
       
  3052     with cncs_z and eq_pv show ?thesis by simp
       
  3053   next
       
  3054     case False
       
  3055     with cncs_z and eq_pv show ?thesis by simp
       
  3056   qed
       
  3057 qed
       
  3058 
       
  3059 lemma detached_eq:
       
  3060   shows "(detached s th) = (cntP s th = cntV s th)"
       
  3061   by (insert vt, auto intro:detached_intro detached_elim)
       
  3062 
       
  3063 end
       
  3064 
       
  3065 text {* 
       
  3066   The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived
       
  3067   from the concise and miniature model of PIP given in PrioGDef.thy.
       
  3068 *}
       
  3069 
       
  3070 lemma eq_dependants: "dependants (wq s) = dependants s"
       
  3071   by (simp add: s_dependants_abv wq_def)
       
  3072 
       
  3073 lemma next_th_unique: 
       
  3074   assumes nt1: "next_th s th cs th1"
       
  3075   and nt2: "next_th s th cs th2"
       
  3076   shows "th1 = th2"
       
  3077 using assms by (unfold next_th_def, auto)
       
  3078 
       
  3079 lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
       
  3080   apply (induct s, simp)
       
  3081 proof -
       
  3082   fix a s
       
  3083   assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
       
  3084     and eq_as: "a # s \<noteq> []"
       
  3085   show "last_set th (a # s) < length (a # s)"
       
  3086   proof(cases "s \<noteq> []")
       
  3087     case False
       
  3088     from False show ?thesis
       
  3089       by (cases a, auto simp:last_set.simps)
       
  3090   next
       
  3091     case True
       
  3092     from ih [OF True] show ?thesis
       
  3093       by (cases a, auto simp:last_set.simps)
       
  3094   qed
       
  3095 qed
       
  3096 
       
  3097 lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
       
  3098   by (induct s, auto simp:threads.simps)
       
  3099 
       
  3100 lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
       
  3101   apply (drule_tac th_in_ne)
       
  3102   by (unfold preced_def, auto intro: birth_time_lt)
       
  3103 
       
  3104 lemma inj_the_preced: 
       
  3105   "inj_on (the_preced s) (threads s)"
       
  3106   by (metis inj_onI preced_unique the_preced_def)
       
  3107 
       
  3108 lemma tRAG_alt_def: 
       
  3109   "tRAG s = {(Th th1, Th th2) | th1 th2. 
       
  3110                   \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
       
  3111  by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
       
  3112 
       
  3113 lemma tRAG_Field:
       
  3114   "Field (tRAG s) \<subseteq> Field (RAG s)"
       
  3115   by (unfold tRAG_alt_def Field_def, auto)
       
  3116 
       
  3117 lemma tRAG_ancestorsE:
       
  3118   assumes "x \<in> ancestors (tRAG s) u"
       
  3119   obtains th where "x = Th th"
       
  3120 proof -
       
  3121   from assms have "(u, x) \<in> (tRAG s)^+" 
       
  3122       by (unfold ancestors_def, auto)
       
  3123   from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
       
  3124   then obtain th where "x = Th th"
       
  3125     by (unfold tRAG_alt_def, auto)
       
  3126   from that[OF this] show ?thesis .
       
  3127 qed
       
  3128 
       
  3129 lemma tRAG_mono:
       
  3130   assumes "RAG s' \<subseteq> RAG s"
       
  3131   shows "tRAG s' \<subseteq> tRAG s"
       
  3132   using assms 
       
  3133   by (unfold tRAG_alt_def, auto)
       
  3134 
       
  3135 lemma holding_next_thI:
       
  3136   assumes "holding s th cs"
       
  3137   and "length (wq s cs) > 1"
       
  3138   obtains th' where "next_th s th cs th'"
       
  3139 proof -
       
  3140   from assms(1)[folded eq_holding, unfolded cs_holding_def]
       
  3141   have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" .
       
  3142   then obtain rest where h1: "wq s cs = th#rest" 
       
  3143     by (cases "wq s cs", auto)
       
  3144   with assms(2) have h2: "rest \<noteq> []" by auto
       
  3145   let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
  3146   have "next_th s th cs ?th'" using  h1(1) h2 
       
  3147     by (unfold next_th_def, auto)
       
  3148   from that[OF this] show ?thesis .
       
  3149 qed
       
  3150 
       
  3151 lemma RAG_tRAG_transfer:
       
  3152   assumes "vt s'"
       
  3153   assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
       
  3154   and "(Cs cs, Th th'') \<in> RAG s'"
       
  3155   shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
       
  3156 proof -
       
  3157   interpret vt_s': valid_trace "s'" using assms(1)
       
  3158     by (unfold_locales, simp)
       
  3159   interpret rtree: rtree "RAG s'"
       
  3160   proof
       
  3161   show "single_valued (RAG s')"
       
  3162   apply (intro_locales)
       
  3163     by (unfold single_valued_def, 
       
  3164         auto intro:vt_s'.unique_RAG)
       
  3165 
       
  3166   show "acyclic (RAG s')"
       
  3167      by (rule vt_s'.acyclic_RAG)
       
  3168   qed
       
  3169   { fix n1 n2
       
  3170     assume "(n1, n2) \<in> ?L"
       
  3171     from this[unfolded tRAG_alt_def]
       
  3172     obtain th1 th2 cs' where 
       
  3173       h: "n1 = Th th1" "n2 = Th th2" 
       
  3174          "(Th th1, Cs cs') \<in> RAG s"
       
  3175          "(Cs cs', Th th2) \<in> RAG s" by auto
       
  3176     from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
       
  3177     from h(3) and assms(2) 
       
  3178     have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
       
  3179           (Th th1, Cs cs') \<in> RAG s'" by auto
       
  3180     hence "(n1, n2) \<in> ?R"
       
  3181     proof
       
  3182       assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
       
  3183       hence eq_th1: "th1 = th" by simp
       
  3184       moreover have "th2 = th''"
       
  3185       proof -
       
  3186         from h1 have "cs' = cs" by simp
       
  3187         from assms(3) cs_in[unfolded this] rtree.sgv
       
  3188         show ?thesis
       
  3189           by (unfold single_valued_def, auto)
       
  3190       qed
       
  3191       ultimately show ?thesis using h(1,2) by auto
       
  3192     next
       
  3193       assume "(Th th1, Cs cs') \<in> RAG s'"
       
  3194       with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
       
  3195         by (unfold tRAG_alt_def, auto)
       
  3196       from this[folded h(1, 2)] show ?thesis by auto
       
  3197     qed
       
  3198   } moreover {
       
  3199     fix n1 n2
       
  3200     assume "(n1, n2) \<in> ?R"
       
  3201     hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
       
  3202     hence "(n1, n2) \<in> ?L" 
       
  3203     proof
       
  3204       assume "(n1, n2) \<in> tRAG s'"
       
  3205       moreover have "... \<subseteq> ?L"
       
  3206       proof(rule tRAG_mono)
       
  3207         show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
       
  3208       qed
       
  3209       ultimately show ?thesis by auto
       
  3210     next
       
  3211       assume eq_n: "(n1, n2) = (Th th, Th th'')"
       
  3212       from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
       
  3213       moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
       
  3214       ultimately show ?thesis 
       
  3215         by (unfold eq_n tRAG_alt_def, auto)
       
  3216     qed
       
  3217   } ultimately show ?thesis by auto
       
  3218 qed
       
  3219 
       
  3220 context valid_trace
       
  3221 begin
       
  3222 
       
  3223 lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
       
  3224 
       
  3225 end
       
  3226 
       
  3227 lemma cp_alt_def:
       
  3228   "cp s th =  
       
  3229            Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})"
       
  3230 proof -
       
  3231   have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) =
       
  3232         Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" 
       
  3233           (is "Max (_ ` ?L) = Max (_ ` ?R)")
       
  3234   proof -
       
  3235     have "?L = ?R" 
       
  3236     by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def)
       
  3237     thus ?thesis by simp
       
  3238   qed
       
  3239   thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
       
  3240 qed
       
  3241 
       
  3242 lemma cp_gen_alt_def:
       
  3243   "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
       
  3244     by (auto simp:cp_gen_def)
       
  3245 
       
  3246 lemma tRAG_nodeE:
       
  3247   assumes "(n1, n2) \<in> tRAG s"
       
  3248   obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
       
  3249   using assms
       
  3250   by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
       
  3251 
       
  3252 lemma subtree_nodeE:
       
  3253   assumes "n \<in> subtree (tRAG s) (Th th)"
       
  3254   obtains th1 where "n = Th th1"
       
  3255 proof -
       
  3256   show ?thesis
       
  3257   proof(rule subtreeE[OF assms])
       
  3258     assume "n = Th th"
       
  3259     from that[OF this] show ?thesis .
       
  3260   next
       
  3261     assume "Th th \<in> ancestors (tRAG s) n"
       
  3262     hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
       
  3263     hence "\<exists> th1. n = Th th1"
       
  3264     proof(induct)
       
  3265       case (base y)
       
  3266       from tRAG_nodeE[OF this] show ?case by metis
       
  3267     next
       
  3268       case (step y z)
       
  3269       thus ?case by auto
       
  3270     qed
       
  3271     with that show ?thesis by auto
       
  3272   qed
       
  3273 qed
       
  3274 
       
  3275 lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*"
       
  3276 proof -
       
  3277   have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" 
       
  3278     by (rule rtrancl_mono, auto simp:RAG_split)
       
  3279   also have "... \<subseteq> ((RAG s)^*)^*"
       
  3280     by (rule rtrancl_mono, auto)
       
  3281   also have "... = (RAG s)^*" by simp
       
  3282   finally show ?thesis by (unfold tRAG_def, simp)
       
  3283 qed
       
  3284 
       
  3285 lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
       
  3286 proof -
       
  3287   { fix a
       
  3288     assume "a \<in> subtree (tRAG s) x"
       
  3289     hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
       
  3290     with tRAG_star_RAG[of s]
       
  3291     have "(a, x) \<in> (RAG s)^*" by auto
       
  3292     hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
       
  3293   } thus ?thesis by auto
       
  3294 qed
       
  3295 
       
  3296 lemma tRAG_trancl_eq:
       
  3297    "{th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
       
  3298     {th'. (Th th', Th th)  \<in> (RAG s)^+}"
       
  3299    (is "?L = ?R")
       
  3300 proof -
       
  3301   { fix th'
       
  3302     assume "th' \<in> ?L"
       
  3303     hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
       
  3304     from tranclD[OF this]
       
  3305     obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
       
  3306     from tRAG_subtree_RAG[of s] and this(2)
       
  3307     have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
       
  3308     moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
       
  3309     ultimately have "th' \<in> ?R"  by auto 
       
  3310   } moreover 
       
  3311   { fix th'
       
  3312     assume "th' \<in> ?R"
       
  3313     hence "(Th th', Th th) \<in> (RAG s)^+" by (auto)
       
  3314     from plus_rpath[OF this]
       
  3315     obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto
       
  3316     hence "(Th th', Th th) \<in> (tRAG s)^+"
       
  3317     proof(induct xs arbitrary:th' th rule:length_induct)
       
  3318       case (1 xs th' th)
       
  3319       then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto)
       
  3320       show ?case
       
  3321       proof(cases "xs1")
       
  3322         case Nil
       
  3323         from 1(2)[unfolded Cons1 Nil]
       
  3324         have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
       
  3325         hence "(Th th', x1) \<in> (RAG s)" by (cases, simp)
       
  3326         then obtain cs where "x1 = Cs cs" 
       
  3327               by (unfold s_RAG_def, auto)
       
  3328         from rpath_nnl_lastE[OF rp[unfolded this]]
       
  3329         show ?thesis by auto
       
  3330       next
       
  3331         case (Cons x2 xs2)
       
  3332         from 1(2)[unfolded Cons1[unfolded this]]
       
  3333         have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
       
  3334         from rpath_edges_on[OF this]
       
  3335         have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
       
  3336         have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
       
  3337             by (simp add: edges_on_unfold)
       
  3338         with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
       
  3339         then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
       
  3340         have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
       
  3341             by (simp add: edges_on_unfold)
       
  3342         from this eds
       
  3343         have rg2: "(x1, x2) \<in> RAG s" by auto
       
  3344         from this[unfolded eq_x1] 
       
  3345         obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
       
  3346         from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
       
  3347         have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto)
       
  3348         from rp have "rpath (RAG s) x2 xs2 (Th th)"
       
  3349            by  (elim rpath_ConsE, simp)
       
  3350         from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
       
  3351         show ?thesis
       
  3352         proof(cases "xs2 = []")
       
  3353           case True
       
  3354           from rpath_nilE[OF rp'[unfolded this]]
       
  3355           have "th1 = th" by auto
       
  3356           from rt1[unfolded this] show ?thesis by auto
       
  3357         next
       
  3358           case False
       
  3359           from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons]
       
  3360           have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp
       
  3361           with rt1 show ?thesis by auto
       
  3362         qed
       
  3363       qed
       
  3364     qed
       
  3365     hence "th' \<in> ?L" by auto
       
  3366   } ultimately show ?thesis by blast
       
  3367 qed
       
  3368 
       
  3369 lemma tRAG_trancl_eq_Th:
       
  3370    "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
       
  3371     {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
       
  3372     using tRAG_trancl_eq by auto
       
  3373 
       
  3374 lemma dependants_alt_def:
       
  3375   "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
       
  3376   by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
       
  3377   
       
  3378 context valid_trace
       
  3379 begin
       
  3380 
       
  3381 lemma count_eq_tRAG_plus:
       
  3382   assumes "cntP s th = cntV s th"
       
  3383   shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
       
  3384   using assms count_eq_dependants dependants_alt_def eq_dependants by auto 
       
  3385 
       
  3386 lemma count_eq_RAG_plus:
       
  3387   assumes "cntP s th = cntV s th"
       
  3388   shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  3389   using assms count_eq_dependants cs_dependants_def eq_RAG by auto
       
  3390 
       
  3391 lemma count_eq_RAG_plus_Th:
       
  3392   assumes "cntP s th = cntV s th"
       
  3393   shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  3394   using count_eq_RAG_plus[OF assms] by auto
       
  3395 
       
  3396 lemma count_eq_tRAG_plus_Th:
       
  3397   assumes "cntP s th = cntV s th"
       
  3398   shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
       
  3399    using count_eq_tRAG_plus[OF assms] by auto
       
  3400 
       
  3401 end
       
  3402 
       
  3403 lemma tRAG_subtree_eq: 
       
  3404    "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
       
  3405    (is "?L = ?R")
       
  3406 proof -
       
  3407   { fix n 
       
  3408     assume h: "n \<in> ?L"
       
  3409     hence "n \<in> ?R"
       
  3410     by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) 
       
  3411   } moreover {
       
  3412     fix n
       
  3413     assume "n \<in> ?R"
       
  3414     then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*"
       
  3415       by (auto simp:subtree_def)
       
  3416     from rtranclD[OF this(2)]
       
  3417     have "n \<in> ?L"
       
  3418     proof
       
  3419       assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+"
       
  3420       with h have "n \<in> {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}" by auto
       
  3421       thus ?thesis using subtree_def tRAG_trancl_eq by fastforce
       
  3422     qed (insert h, auto simp:subtree_def)
       
  3423   } ultimately show ?thesis by auto
       
  3424 qed
       
  3425 
       
  3426 lemma threads_set_eq: 
       
  3427    "the_thread ` (subtree (tRAG s) (Th th)) = 
       
  3428                   {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
       
  3429    by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
       
  3430 
       
  3431 lemma cp_alt_def1: 
       
  3432   "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
       
  3433 proof -
       
  3434   have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
       
  3435        ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
       
  3436        by auto
       
  3437   thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
       
  3438 qed
       
  3439 
       
  3440 lemma cp_gen_def_cond: 
       
  3441   assumes "x = Th th"
       
  3442   shows "cp s th = cp_gen s (Th th)"
       
  3443 by (unfold cp_alt_def1 cp_gen_def, simp)
       
  3444 
       
  3445 lemma cp_gen_over_set:
       
  3446   assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
       
  3447   shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
       
  3448 proof(rule f_image_eq)
       
  3449   fix a
       
  3450   assume "a \<in> A"
       
  3451   from assms[rule_format, OF this]
       
  3452   obtain th where eq_a: "a = Th th" by auto
       
  3453   show "cp_gen s a = (cp s \<circ> the_thread) a"
       
  3454     by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
       
  3455 qed
       
  3456 
       
  3457 
       
  3458 context valid_trace
       
  3459 begin
       
  3460 
       
  3461 lemma RAG_threads:
       
  3462   assumes "(Th th) \<in> Field (RAG s)"
       
  3463   shows "th \<in> threads s"
       
  3464   using assms
       
  3465   by (metis Field_def UnE dm_RAG_threads range_in vt)
       
  3466 
       
  3467 lemma subtree_tRAG_thread:
       
  3468   assumes "th \<in> threads s"
       
  3469   shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
       
  3470 proof -
       
  3471   have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
       
  3472     by (unfold tRAG_subtree_eq, simp)
       
  3473   also have "... \<subseteq> ?R"
       
  3474   proof
       
  3475     fix x
       
  3476     assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}"
       
  3477     then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto
       
  3478     from this(2)
       
  3479     show "x \<in> ?R"
       
  3480     proof(cases rule:subtreeE)
       
  3481       case 1
       
  3482       thus ?thesis by (simp add: assms h(1)) 
       
  3483     next
       
  3484       case 2
       
  3485       thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) 
       
  3486     qed
       
  3487   qed
       
  3488   finally show ?thesis .
       
  3489 qed
       
  3490 
       
  3491 lemma readys_root:
       
  3492   assumes "th \<in> readys s"
       
  3493   shows "root (RAG s) (Th th)"
       
  3494 proof -
       
  3495   { fix x
       
  3496     assume "x \<in> ancestors (RAG s) (Th th)"
       
  3497     hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
  3498     from tranclD[OF this]
       
  3499     obtain z where "(Th th, z) \<in> RAG s" by auto
       
  3500     with assms(1) have False
       
  3501          apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
       
  3502          by (fold wq_def, blast)
       
  3503   } thus ?thesis by (unfold root_def, auto)
       
  3504 qed
       
  3505 
       
  3506 lemma readys_in_no_subtree:
       
  3507   assumes "th \<in> readys s"
       
  3508   and "th' \<noteq> th"
       
  3509   shows "Th th \<notin> subtree (RAG s) (Th th')" 
       
  3510 proof
       
  3511    assume "Th th \<in> subtree (RAG s) (Th th')"
       
  3512    thus False
       
  3513    proof(cases rule:subtreeE)
       
  3514       case 1
       
  3515       with assms show ?thesis by auto
       
  3516    next
       
  3517       case 2
       
  3518       with readys_root[OF assms(1)]
       
  3519       show ?thesis by (auto simp:root_def)
       
  3520    qed
       
  3521 qed
       
  3522 
       
  3523 lemma not_in_thread_isolated:
       
  3524   assumes "th \<notin> threads s"
       
  3525   shows "(Th th) \<notin> Field (RAG s)"
       
  3526 proof
       
  3527   assume "(Th th) \<in> Field (RAG s)"
       
  3528   with dm_RAG_threads and range_in assms
       
  3529   show False by (unfold Field_def, blast)
       
  3530 qed
       
  3531 
       
  3532 lemma wf_RAG: "wf (RAG s)"
       
  3533 proof(rule finite_acyclic_wf)
       
  3534   from finite_RAG show "finite (RAG s)" .
       
  3535 next
       
  3536   from acyclic_RAG show "acyclic (RAG s)" .
       
  3537 qed
       
  3538 
       
  3539 lemma sgv_wRAG: "single_valued (wRAG s)"
       
  3540   using waiting_unique
       
  3541   by (unfold single_valued_def wRAG_def, auto)
       
  3542 
       
  3543 lemma sgv_hRAG: "single_valued (hRAG s)"
       
  3544   using holding_unique 
       
  3545   by (unfold single_valued_def hRAG_def, auto)
       
  3546 
       
  3547 lemma sgv_tRAG: "single_valued (tRAG s)"
       
  3548   by (unfold tRAG_def, rule single_valued_relcomp, 
       
  3549               insert sgv_wRAG sgv_hRAG, auto)
       
  3550 
       
  3551 lemma acyclic_tRAG: "acyclic (tRAG s)"
       
  3552 proof(unfold tRAG_def, rule acyclic_compose)
       
  3553   show "acyclic (RAG s)" using acyclic_RAG .
       
  3554 next
       
  3555   show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
  3556 next
       
  3557   show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
  3558 qed
       
  3559 
       
  3560 lemma sgv_RAG: "single_valued (RAG s)"
       
  3561   using unique_RAG by (auto simp:single_valued_def)
       
  3562 
       
  3563 lemma rtree_RAG: "rtree (RAG s)"
       
  3564   using sgv_RAG acyclic_RAG
       
  3565   by (unfold rtree_def rtree_axioms_def sgv_def, auto)
       
  3566 
       
  3567 end
       
  3568 
       
  3569 sublocale valid_trace < rtree_RAG: rtree "RAG s"
       
  3570 proof
       
  3571   show "single_valued (RAG s)"
       
  3572   apply (intro_locales)
       
  3573     by (unfold single_valued_def, 
       
  3574         auto intro:unique_RAG)
       
  3575 
       
  3576   show "acyclic (RAG s)"
       
  3577      by (rule acyclic_RAG)
       
  3578 qed
       
  3579 
       
  3580 sublocale valid_trace < rtree_s: rtree "tRAG s"
       
  3581 proof(unfold_locales)
       
  3582   from sgv_tRAG show "single_valued (tRAG s)" .
       
  3583 next
       
  3584   from acyclic_tRAG show "acyclic (tRAG s)" .
       
  3585 qed
       
  3586 
       
  3587 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
       
  3588 proof -
       
  3589   show "fsubtree (RAG s)"
       
  3590   proof(intro_locales)
       
  3591     show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
       
  3592   next
       
  3593     show "fsubtree_axioms (RAG s)"
       
  3594     proof(unfold fsubtree_axioms_def)
       
  3595       from wf_RAG show "wf (RAG s)" .
       
  3596     qed
       
  3597   qed
       
  3598 qed
       
  3599 
       
  3600 sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
       
  3601 proof -
       
  3602   have "fsubtree (tRAG s)"
       
  3603   proof -
       
  3604     have "fbranch (tRAG s)"
       
  3605     proof(unfold tRAG_def, rule fbranch_compose)
       
  3606         show "fbranch (wRAG s)"
       
  3607         proof(rule finite_fbranchI)
       
  3608            from finite_RAG show "finite (wRAG s)"
       
  3609            by (unfold RAG_split, auto)
       
  3610         qed
       
  3611     next
       
  3612         show "fbranch (hRAG s)"
       
  3613         proof(rule finite_fbranchI)
       
  3614            from finite_RAG 
       
  3615            show "finite (hRAG s)" by (unfold RAG_split, auto)
       
  3616         qed
       
  3617     qed
       
  3618     moreover have "wf (tRAG s)"
       
  3619     proof(rule wf_subset)
       
  3620       show "wf (RAG s O RAG s)" using wf_RAG
       
  3621         by (fold wf_comp_self, simp)
       
  3622     next
       
  3623       show "tRAG s \<subseteq> (RAG s O RAG s)"
       
  3624         by (unfold tRAG_alt_def, auto)
       
  3625     qed
       
  3626     ultimately show ?thesis
       
  3627       by (unfold fsubtree_def fsubtree_axioms_def,auto)
       
  3628   qed
       
  3629   from this[folded tRAG_def] show "fsubtree (tRAG s)" .
       
  3630 qed
       
  3631 
       
  3632 lemma Max_UNION: 
       
  3633   assumes "finite A"
       
  3634   and "A \<noteq> {}"
       
  3635   and "\<forall> M \<in> f ` A. finite M"
       
  3636   and "\<forall> M \<in> f ` A. M \<noteq> {}"
       
  3637   shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
       
  3638   using assms[simp]
       
  3639 proof -
       
  3640   have "?L = Max (\<Union>(f ` A))"
       
  3641     by (fold Union_image_eq, simp)
       
  3642   also have "... = ?R"
       
  3643     by (subst Max_Union, simp+)
       
  3644   finally show ?thesis .
       
  3645 qed
       
  3646 
       
  3647 lemma max_Max_eq:
       
  3648   assumes "finite A"
       
  3649     and "A \<noteq> {}"
       
  3650     and "x = y"
       
  3651   shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
       
  3652 proof -
       
  3653   have "?R = Max (insert y A)" by simp
       
  3654   also from assms have "... = ?L"
       
  3655       by (subst Max.insert, simp+)
       
  3656   finally show ?thesis by simp
       
  3657 qed
       
  3658 
       
  3659 context valid_trace
       
  3660 begin
       
  3661 
       
  3662 (* ddd *)
       
  3663 lemma cp_gen_rec:
       
  3664   assumes "x = Th th"
       
  3665   shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
       
  3666 proof(cases "children (tRAG s) x = {}")
       
  3667   case True
       
  3668   show ?thesis
       
  3669     by (unfold True cp_gen_def subtree_children, simp add:assms)
       
  3670 next
       
  3671   case False
       
  3672   hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
       
  3673   note fsbttRAGs.finite_subtree[simp]
       
  3674   have [simp]: "finite (children (tRAG s) x)"
       
  3675      by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
       
  3676             rule children_subtree)
       
  3677   { fix r x
       
  3678     have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
       
  3679   } note this[simp]
       
  3680   have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
       
  3681   proof -
       
  3682     from False obtain q where "q \<in> children (tRAG s) x" by blast
       
  3683     moreover have "subtree (tRAG s) q \<noteq> {}" by simp
       
  3684     ultimately show ?thesis by blast
       
  3685   qed
       
  3686   have h: "Max ((the_preced s \<circ> the_thread) `
       
  3687                 ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
       
  3688         Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
       
  3689                      (is "?L = ?R")
       
  3690   proof -
       
  3691     let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
       
  3692     let "Max (_ \<union> (?h ` ?B))" = ?R
       
  3693     let ?L1 = "?f ` \<Union>(?g ` ?B)"
       
  3694     have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
       
  3695     proof -
       
  3696       have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
       
  3697       also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
       
  3698       finally have "Max ?L1 = Max ..." by simp
       
  3699       also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
       
  3700         by (subst Max_UNION, simp+)
       
  3701       also have "... = Max (cp_gen s ` children (tRAG s) x)"
       
  3702           by (unfold image_comp cp_gen_alt_def, simp)
       
  3703       finally show ?thesis .
       
  3704     qed
       
  3705     show ?thesis
       
  3706     proof -
       
  3707       have "?L = Max (?f ` ?A \<union> ?L1)" by simp
       
  3708       also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
       
  3709             by (subst Max_Un, simp+)
       
  3710       also have "... = max (?f x) (Max (?h ` ?B))"
       
  3711         by (unfold eq_Max_L1, simp)
       
  3712       also have "... =?R"
       
  3713         by (rule max_Max_eq, (simp)+, unfold assms, simp)
       
  3714       finally show ?thesis .
       
  3715     qed
       
  3716   qed  thus ?thesis 
       
  3717           by (fold h subtree_children, unfold cp_gen_def, simp) 
       
  3718 qed
       
  3719 
       
  3720 lemma cp_rec:
       
  3721   "cp s th = Max ({the_preced s th} \<union> 
       
  3722                      (cp s o the_thread) ` children (tRAG s) (Th th))"
       
  3723 proof -
       
  3724   have "Th th = Th th" by simp
       
  3725   note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
       
  3726   show ?thesis 
       
  3727   proof -
       
  3728     have "cp_gen s ` children (tRAG s) (Th th) = 
       
  3729                 (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
       
  3730     proof(rule cp_gen_over_set)
       
  3731       show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
       
  3732         by (unfold tRAG_alt_def, auto simp:children_def)
       
  3733     qed
       
  3734     thus ?thesis by (subst (1) h(1), unfold h(2), simp)
       
  3735   qed
       
  3736 qed
       
  3737 
       
  3738 end
       
  3739 
       
  3740 (* keep *)
       
  3741 lemma next_th_holding:
       
  3742   assumes vt: "vt s"
       
  3743   and nxt: "next_th s th cs th'"
       
  3744   shows "holding (wq s) th cs"
       
  3745 proof -
       
  3746   from nxt[unfolded next_th_def]
       
  3747   obtain rest where h: "wq s cs = th # rest"
       
  3748                        "rest \<noteq> []" 
       
  3749                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
       
  3750   thus ?thesis
       
  3751     by (unfold cs_holding_def, auto)
       
  3752 qed
       
  3753 
       
  3754 context valid_trace
       
  3755 begin
       
  3756 
       
  3757 lemma next_th_waiting:
       
  3758   assumes nxt: "next_th s th cs th'"
       
  3759   shows "waiting (wq s) th' cs"
       
  3760 proof -
       
  3761   from nxt[unfolded next_th_def]
       
  3762   obtain rest where h: "wq s cs = th # rest"
       
  3763                        "rest \<noteq> []" 
       
  3764                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
       
  3765   from wq_distinct[of cs, unfolded h]
       
  3766   have dst: "distinct (th # rest)" .
       
  3767   have in_rest: "th' \<in> set rest"
       
  3768   proof(unfold h, rule someI2)
       
  3769     show "distinct rest \<and> set rest = set rest" using dst by auto
       
  3770   next
       
  3771     fix x assume "distinct x \<and> set x = set rest"
       
  3772     with h(2)
       
  3773     show "hd x \<in> set (rest)" by (cases x, auto)
       
  3774   qed
       
  3775   hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
       
  3776   moreover have "th' \<noteq> hd (wq s cs)"
       
  3777     by (unfold h(1), insert in_rest dst, auto)
       
  3778   ultimately show ?thesis by (auto simp:cs_waiting_def)
       
  3779 qed
       
  3780 
       
  3781 lemma next_th_RAG:
       
  3782   assumes nxt: "next_th (s::event list) th cs th'"
       
  3783   shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
       
  3784   using vt assms next_th_holding next_th_waiting
       
  3785   by (unfold s_RAG_def, simp)
       
  3786 
       
  3787 end
       
  3788 
       
  3789 
       
  3790 -- {* A useless definition *}
       
  3791 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
       
  3792 where "cps s = {(th, cp s th) | th . th \<in> threads s}"
       
  3793 
       
  3794 end