CpsG.thy
changeset 105 0c89419b4742
parent 90 ed938e2246b9
child 109 4e59c0ce1511
equal deleted inserted replaced
104:43482ab31341 105:0c89419b4742
     1 theory CpsG
     1 theory PIPBasics
     2 imports PIPDefs
     2 imports PIPDefs
     3 begin
     3 begin
       
     4 
       
     5 section {* Generic aulxiliary lemmas *}
     4 
     6 
     5 lemma f_image_eq:
     7 lemma f_image_eq:
     6   assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
     8   assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
     7   shows "f ` A = g ` A"
     9   shows "f ` A = g ` A"
     8 proof
    10 proof
    82   also from assms have "... = ?L"
    84   also from assms have "... = ?L"
    83       by (subst Max.insert, simp+)
    85       by (subst Max.insert, simp+)
    84   finally show ?thesis by simp
    86   finally show ?thesis by simp
    85 qed
    87 qed
    86 
    88 
       
    89 lemma rel_eqI:
       
    90   assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B"
       
    91   and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A"
       
    92   shows "A = B"
       
    93   using assms by auto
       
    94 
       
    95 section {* Lemmas do not depend on trace validity *}
       
    96 
    87 lemma birth_time_lt:  
    97 lemma birth_time_lt:  
    88   assumes "s \<noteq> []"
    98   assumes "s \<noteq> []"
    89   shows "last_set th s < length s"
    99   shows "last_set th s < length s"
    90   using assms
   100   using assms
    91 proof(induct s)
   101 proof(induct s)
   150     thus ?thesis by simp
   160     thus ?thesis by simp
   151   qed
   161   qed
   152   thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
   162   thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
   153 qed
   163 qed
   154 
   164 
       
   165 lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
       
   166   by (unfold s_RAG_def, auto)
       
   167 
       
   168 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
       
   169   by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
       
   170 
       
   171 lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
       
   172   by (unfold s_holding_def wq_def cs_holding_def, simp)
       
   173 
       
   174 lemma children_RAG_alt_def:
       
   175   "children (RAG (s::state)) (Th th) = Cs ` {cs. holding s th cs}"
       
   176   by (unfold s_RAG_def, auto simp:children_def holding_eq)
       
   177 
       
   178 lemma holdents_alt_def:
       
   179   "holdents s th = the_cs ` (children (RAG (s::state)) (Th th))"
       
   180   by (unfold children_RAG_alt_def holdents_def, simp add: image_image)
       
   181 
       
   182 lemma cntCS_alt_def:
       
   183   "cntCS s th = card (children (RAG s) (Th th))"
       
   184   apply (unfold children_RAG_alt_def cntCS_def holdents_def)
       
   185   by (rule card_image[symmetric], auto simp:inj_on_def)
       
   186 
       
   187 lemma runing_ready: 
       
   188   shows "runing s \<subseteq> readys s"
       
   189   unfolding runing_def readys_def
       
   190   by auto 
       
   191 
       
   192 lemma readys_threads:
       
   193   shows "readys s \<subseteq> threads s"
       
   194   unfolding readys_def
       
   195   by auto
       
   196 
       
   197 lemma wq_v_neq [simp]:
       
   198    "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
       
   199   by (auto simp:wq_def Let_def cp_def split:list.splits)
       
   200 
       
   201 lemma runing_head:
       
   202   assumes "th \<in> runing s"
       
   203   and "th \<in> set (wq_fun (schs s) cs)"
       
   204   shows "th = hd (wq_fun (schs s) cs)"
       
   205   using assms
       
   206   by (simp add:runing_def readys_def s_waiting_def wq_def)
       
   207 
       
   208 lemma runing_wqE:
       
   209   assumes "th \<in> runing s"
       
   210   and "th \<in> set (wq s cs)"
       
   211   obtains rest where "wq s cs = th#rest"
       
   212 proof -
       
   213   from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest"
       
   214     by (meson list.set_cases)
       
   215   have "th' = th"
       
   216   proof(rule ccontr)
       
   217     assume "th' \<noteq> th"
       
   218     hence "th \<noteq> hd (wq s cs)" using eq_wq by auto 
       
   219     with assms(2)
       
   220     have "waiting s th cs" 
       
   221       by (unfold s_waiting_def, fold wq_def, auto)
       
   222     with assms show False 
       
   223       by (unfold runing_def readys_def, auto)
       
   224   qed
       
   225   with eq_wq that show ?thesis by metis
       
   226 qed
       
   227 
       
   228 lemma isP_E:
       
   229   assumes "isP e"
       
   230   obtains cs where "e = P (actor e) cs"
       
   231   using assms by (cases e, auto)
       
   232 
       
   233 lemma isV_E:
       
   234   assumes "isV e"
       
   235   obtains cs where "e = V (actor e) cs"
       
   236   using assms by (cases e, auto) 
       
   237 
       
   238 
       
   239 text {*
       
   240   Every thread can only be blocked on one critical resource, 
       
   241   symmetrically, every critical resource can only be held by one thread. 
       
   242   This fact is much more easier according to our definition. 
       
   243 *}
       
   244 lemma held_unique:
       
   245   assumes "holding (s::event list) th1 cs"
       
   246   and "holding s th2 cs"
       
   247   shows "th1 = th2"
       
   248  by (insert assms, unfold s_holding_def, auto)
       
   249 
       
   250 lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
       
   251   apply (induct s, auto)
       
   252   by (case_tac a, auto split:if_splits)
       
   253 
       
   254 lemma last_set_unique: 
       
   255   "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
       
   256           \<Longrightarrow> th1 = th2"
       
   257   apply (induct s, auto)
       
   258   by (case_tac a, auto split:if_splits dest:last_set_lt)
       
   259 
       
   260 lemma preced_unique : 
       
   261   assumes pcd_eq: "preced th1 s = preced th2 s"
       
   262   and th_in1: "th1 \<in> threads s"
       
   263   and th_in2: " th2 \<in> threads s"
       
   264   shows "th1 = th2"
       
   265 proof -
       
   266   from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
       
   267   from last_set_unique [OF this th_in1 th_in2]
       
   268   show ?thesis .
       
   269 qed
       
   270                       
       
   271 lemma preced_linorder: 
       
   272   assumes neq_12: "th1 \<noteq> th2"
       
   273   and th_in1: "th1 \<in> threads s"
       
   274   and th_in2: " th2 \<in> threads s"
       
   275   shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
       
   276 proof -
       
   277   from preced_unique [OF _ th_in1 th_in2] and neq_12 
       
   278   have "preced th1 s \<noteq> preced th2 s" by auto
       
   279   thus ?thesis by auto
       
   280 qed
       
   281 
       
   282 lemma in_RAG_E:
       
   283   assumes "(n1, n2) \<in> RAG (s::state)"
       
   284   obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs"
       
   285       | (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs"
       
   286   using assms[unfolded s_RAG_def, folded waiting_eq holding_eq]
       
   287   by auto
       
   288 
       
   289 lemma count_rec1 [simp]: 
       
   290   assumes "Q e"
       
   291   shows "count Q (e#es) = Suc (count Q es)"
       
   292   using assms
       
   293   by (unfold count_def, auto)
       
   294 
       
   295 lemma count_rec2 [simp]: 
       
   296   assumes "\<not>Q e"
       
   297   shows "count Q (e#es) = (count Q es)"
       
   298   using assms
       
   299   by (unfold count_def, auto)
       
   300 
       
   301 lemma count_rec3 [simp]: 
       
   302   shows "count Q [] =  0"
       
   303   by (unfold count_def, auto)
       
   304 
       
   305 lemma cntP_simp1[simp]:
       
   306   "cntP (P th cs'#s) th = cntP s th + 1"
       
   307   by (unfold cntP_def, simp)
       
   308 
       
   309 lemma cntP_simp2[simp]:
       
   310   assumes "th' \<noteq> th"
       
   311   shows "cntP (P th cs'#s) th' = cntP s th'"
       
   312   using assms
       
   313   by (unfold cntP_def, simp)
       
   314 
       
   315 lemma cntP_simp3[simp]:
       
   316   assumes "\<not> isP e"
       
   317   shows "cntP (e#s) th' = cntP s th'"
       
   318   using assms
       
   319   by (unfold cntP_def, cases e, simp+)
       
   320 
       
   321 lemma cntV_simp1[simp]:
       
   322   "cntV (V th cs'#s) th = cntV s th + 1"
       
   323   by (unfold cntV_def, simp)
       
   324 
       
   325 lemma cntV_simp2[simp]:
       
   326   assumes "th' \<noteq> th"
       
   327   shows "cntV (V th cs'#s) th' = cntV s th'"
       
   328   using assms
       
   329   by (unfold cntV_def, simp)
       
   330 
       
   331 lemma cntV_simp3[simp]:
       
   332   assumes "\<not> isV e"
       
   333   shows "cntV (e#s) th' = cntV s th'"
       
   334   using assms
       
   335   by (unfold cntV_def, cases e, simp+)
       
   336 
       
   337 lemma cntP_diff_inv:
       
   338   assumes "cntP (e#s) th \<noteq> cntP s th"
       
   339   shows "isP e \<and> actor e = th"
       
   340 proof(cases e)
       
   341   case (P th' pty)
       
   342   show ?thesis
       
   343   by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
       
   344         insert assms P, auto simp:cntP_def)
       
   345 qed (insert assms, auto simp:cntP_def)
       
   346   
       
   347 lemma cntV_diff_inv:
       
   348   assumes "cntV (e#s) th \<noteq> cntV s th"
       
   349   shows "isV e \<and> actor e = th"
       
   350 proof(cases e)
       
   351   case (V th' pty)
       
   352   show ?thesis
       
   353   by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
       
   354         insert assms V, auto simp:cntV_def)
       
   355 qed (insert assms, auto simp:cntV_def)
       
   356 
       
   357 lemma eq_dependants: "dependants (wq s) = dependants s"
       
   358   by (simp add: s_dependants_abv wq_def)
       
   359 
       
   360 lemma inj_the_preced: 
       
   361   "inj_on (the_preced s) (threads s)"
       
   362   by (metis inj_onI preced_unique the_preced_def)
       
   363 
       
   364 lemma holding_next_thI:
       
   365   assumes "holding s th cs"
       
   366   and "length (wq s cs) > 1"
       
   367   obtains th' where "next_th s th cs th'"
       
   368 proof -
       
   369   from assms(1)[folded holding_eq, unfolded cs_holding_def]
       
   370   have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" 
       
   371     by (unfold s_holding_def, fold wq_def, auto)
       
   372   then obtain rest where h1: "wq s cs = th#rest" 
       
   373     by (cases "wq s cs", auto)
       
   374   with assms(2) have h2: "rest \<noteq> []" by auto
       
   375   let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
   376   have "next_th s th cs ?th'" using  h1(1) h2 
       
   377     by (unfold next_th_def, auto)
       
   378   from that[OF this] show ?thesis .
       
   379 qed
       
   380 
   155 (* ccc *)
   381 (* ccc *)
   156 
   382 
   157 
   383 section {* Locales used to investigate the execution of PIP *}
       
   384 
       
   385 text {* 
       
   386   The following locale @{text valid_trace} is used to constrain the 
       
   387   trace to be valid. All properties hold for valid traces are 
       
   388   derived under this locale. 
       
   389 *}
   158 locale valid_trace = 
   390 locale valid_trace = 
   159   fixes s
   391   fixes s
   160   assumes vt : "vt s"
   392   assumes vt : "vt s"
   161 
   393 
       
   394 text {* 
       
   395   The following locale @{text valid_trace_e} describes 
       
   396   the valid extension of a valid trace. The event @{text "e"}
       
   397   represents an event in the system, which corresponds 
       
   398   to a one step operation of the PIP protocol. 
       
   399   It is required that @{text "e"} is an event eligible to happen
       
   400   under state @{text "s"}, which is already required to be valid
       
   401   by the parent locale @{text "valid_trace"}.
       
   402 
       
   403   This locale is used to investigate one step execution of PIP, 
       
   404   properties concerning the effects of @{text "e"}'s execution, 
       
   405   for example, how the values of observation functions are changed, 
       
   406   or how desirable properties are kept invariant, are derived
       
   407   under this locale. The state before execution is @{text "s"}, while
       
   408   the state after execution is @{text "e#s"}. Therefore, the lemmas 
       
   409   derived usually relate observations on @{text "e#s"} to those 
       
   410   on @{text "s"}.
       
   411 *}
       
   412 
   162 locale valid_trace_e = valid_trace +
   413 locale valid_trace_e = valid_trace +
   163   fixes e
   414   fixes e
   164   assumes vt_e: "vt (e#s)"
   415   assumes vt_e: "vt (e#s)"
   165 begin
   416 begin
   166 
   417 
       
   418 text {*
       
   419   The following lemma shows that @{text "e"} must be a 
       
   420   eligible event (or a valid step) to be taken under
       
   421   the state represented by @{text "s"}.
       
   422 *}
   167 lemma pip_e: "PIP s e"
   423 lemma pip_e: "PIP s e"
   168   using vt_e by (cases, simp)  
   424   using vt_e by (cases, simp)  
   169 
   425 
   170 end
   426 end
   171 
   427 
       
   428 text {*
       
   429   Because @{term "e#s"} is also a valid trace, properties 
       
   430   derived for valid trace @{term s} also hold on @{term "e#s"}.
       
   431 *}
       
   432 sublocale valid_trace_e < vat_es!: valid_trace "e#s" 
       
   433   using vt_e
       
   434   by (unfold_locales, simp)
       
   435 
       
   436 text {*
       
   437   For each specific event (or operation), there is a sublocale
       
   438   further constraining that the event @{text e} to be that 
       
   439   particular event. 
       
   440 
       
   441   For example, the following 
       
   442   locale @{text "valid_trace_create"} is the sublocale for 
       
   443   event @{term "Create"}:
       
   444 *}
   172 locale valid_trace_create = valid_trace_e + 
   445 locale valid_trace_create = valid_trace_e + 
   173   fixes th prio
   446   fixes th prio
   174   assumes is_create: "e = Create th prio"
   447   assumes is_create: "e = Create th prio"
   175 
   448 
   176 locale valid_trace_exit = valid_trace_e + 
   449 locale valid_trace_exit = valid_trace_e + 
   179 
   452 
   180 locale valid_trace_p = valid_trace_e + 
   453 locale valid_trace_p = valid_trace_e + 
   181   fixes th cs
   454   fixes th cs
   182   assumes is_p: "e = P th cs"
   455   assumes is_p: "e = P th cs"
   183 
   456 
       
   457 text {*
       
   458   locale @{text "valid_trace_p"} is divided further into two 
       
   459   sublocales, namely, @{text "valid_trace_p_h"} 
       
   460   and @{text "valid_trace_p_w"}.
       
   461 *}
       
   462 
       
   463 text {*
       
   464   The following two sublocales @{text "valid_trace_p_h"}
       
   465   and @{text "valid_trace_p_w"} represent two complementary 
       
   466   cases under @{text "valid_trace_p"}, where
       
   467   @{text "valid_trace_p_h"} further constraints that
       
   468   @{text "wq s cs = []"}, which means the waiting queue of 
       
   469   the requested resource @{text "cs"} is empty, in which
       
   470   case,  the requesting thread @{text "th"} 
       
   471   will take hold of @{text "cs"}. 
       
   472 
       
   473   Opposite to @{text "valid_trace_p_h"},
       
   474   @{text "valid_trace_p_w"} constraints that
       
   475   @{text "wq s cs \<noteq> []"}, which means the waiting queue of 
       
   476   the requested resource @{text "cs"} is nonempty, in which
       
   477   case,  the requesting thread @{text "th"} will be blocked
       
   478   on @{text "cs"}: 
       
   479 
       
   480   Peculiar properties will be derived under respective 
       
   481   locales.
       
   482 *}
       
   483 
       
   484 locale valid_trace_p_h = valid_trace_p +
       
   485   assumes we: "wq s cs = []"
       
   486 
       
   487 locale valid_trace_p_w = valid_trace_p +
       
   488   assumes wne: "wq s cs \<noteq> []"
       
   489 begin
       
   490 
       
   491 text {*
       
   492   The following @{text "holder"} designates
       
   493   the holder of @{text "cs"} before the @{text "P"}-operation.
       
   494 *}
       
   495 definition "holder = hd (wq s cs)"
       
   496 
       
   497 text {*
       
   498   The following @{text "waiters"} designates
       
   499   the list of threads waiting for @{text "cs"} 
       
   500   before the @{text "P"}-operation.
       
   501 *}
       
   502 definition "waiters = tl (wq s cs)"
       
   503 end
       
   504 
       
   505 text {* 
       
   506   @{text "valid_trace_v"} is set for the @{term V}-operation.
       
   507 *}
   184 locale valid_trace_v = valid_trace_e + 
   508 locale valid_trace_v = valid_trace_e + 
   185   fixes th cs
   509   fixes th cs
   186   assumes is_v: "e = V th cs"
   510   assumes is_v: "e = V th cs"
   187 begin
   511 begin
       
   512   -- {* The following @{text "rest"} is the tail of 
       
   513         waiting queue of the resource @{text "cs"}
       
   514         to be released by this @{text "V"}-operation.
       
   515      *}
   188   definition "rest = tl (wq s cs)"
   516   definition "rest = tl (wq s cs)"
       
   517 
       
   518   text {*
       
   519     The following @{text "wq'"} is the waiting
       
   520     queue of @{term "cs"}
       
   521     after the @{text "V"}-operation, which
       
   522     is simply a reordering of @{term "rest"}. 
       
   523 
       
   524     The effect of this reordering needs to be 
       
   525     understood by two cases:
       
   526     \begin{enumerate}
       
   527     \item When @{text "rest = []"},
       
   528     the reordering gives rise to an empty list as well, 
       
   529     which means there is no thread holding or waiting 
       
   530     for resource @{term "cs"}, therefore, it is free.
       
   531 
       
   532     \item When @{text "rest \<noteq> []"}, the effect of 
       
   533     this reordering is to arbitrarily 
       
   534     switch one thread in @{term "rest"} to the 
       
   535     head, which, by definition take over the hold
       
   536     of @{term "cs"} and is designated by @{text "taker"}
       
   537     in the following sublocale @{text "valid_trace_v_n"}.
       
   538   *}
   189   definition "wq' = (SOME q. distinct q \<and> set q = set rest)"
   539   definition "wq' = (SOME q. distinct q \<and> set q = set rest)"
   190 end
   540 
       
   541   text {* 
       
   542   The following @{text "rest'"} is the tail of the 
       
   543   waiting queue after the @{text "V"}-operation. 
       
   544   It plays only auxiliary role to ease reasoning. 
       
   545   *}
       
   546   definition "rest' = tl wq'"
       
   547 
       
   548 end
       
   549 
       
   550 text {* 
       
   551   In the following, @{text "valid_trace_v"} is also 
       
   552   divided into two 
       
   553   sublocales: when @{text "rest"} is empty (represented
       
   554   by @{text "valid_trace_v_e"}), which means, there is no thread waiting 
       
   555   for @{text "cs"}, therefore, after the @{text "V"}-operation, 
       
   556   it will become free; otherwise (represented 
       
   557   by @{text "valid_trace_v_n"}), one thread 
       
   558   will be picked from those in @{text "rest"} to take 
       
   559   over @{text "cs"}.
       
   560 *}
       
   561 
       
   562 locale valid_trace_v_e = valid_trace_v +
       
   563   assumes rest_nil: "rest = []"
   191 
   564 
   192 locale valid_trace_v_n = valid_trace_v +
   565 locale valid_trace_v_n = valid_trace_v +
   193   assumes rest_nnl: "rest \<noteq> []"
   566   assumes rest_nnl: "rest \<noteq> []"
   194 
   567 begin
   195 locale valid_trace_v_e = valid_trace_v +
   568 
   196   assumes rest_nil: "rest = []"
   569 text {* 
   197 
   570   The following @{text "taker"} is the thread to 
   198 locale valid_trace_set= valid_trace_e + 
   571   take over @{text "cs"}. 
       
   572 *}
       
   573   definition "taker = hd wq'"
       
   574 
       
   575 end
       
   576 
       
   577 
       
   578 locale valid_trace_set = valid_trace_e + 
   199   fixes th prio
   579   fixes th prio
   200   assumes is_set: "e = Set th prio"
   580   assumes is_set: "e = Set th prio"
   201 
   581 
   202 context valid_trace
   582 context valid_trace
   203 begin
   583 begin
   204 
   584 
       
   585 text {*
       
   586   Induction rule introduced to easy the 
       
   587   derivation of properties for valid trace @{term "s"}.
       
   588   One more premises, namely @{term "valid_trace_e s e"}
       
   589   is added, so that an interpretation of 
       
   590   @{text "valid_trace_e"} can be instantiated 
       
   591   so that all properties derived so far becomes 
       
   592   available in the proof of induction step.
       
   593 
       
   594   You will see its use in the proofs that follows.
       
   595 *}
   205 lemma ind [consumes 0, case_names Nil Cons, induct type]:
   596 lemma ind [consumes 0, case_names Nil Cons, induct type]:
   206   assumes "PP []"
   597   assumes "PP []"
   207      and "(\<And>s e. valid_trace_e s e \<Longrightarrow>
   598      and "(\<And>s e. valid_trace_e s e \<Longrightarrow>
   208                    PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
   599                    PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
   209      shows "PP s"
   600      shows "PP s"
   220   next
   611   next
   221     show "PIP s e" using Step by simp
   612     show "PIP s e" using Step by simp
   222   qed
   613   qed
   223 qed
   614 qed
   224 
   615 
       
   616 text {*
       
   617   The following lemma says that if @{text "s"} is a valid state, so 
       
   618   is its any postfix. Where @{term "monent t s"} is the postfix of 
       
   619   @{term "s"} with length @{term "t"}.
       
   620 *}
   225 lemma  vt_moment: "\<And> t. vt (moment t s)"
   621 lemma  vt_moment: "\<And> t. vt (moment t s)"
   226 proof(induct rule:ind)
   622 proof(induct rule:ind)
   227   case Nil
   623   case Nil
   228   thus ?case by (simp add:vt_nil)
   624   thus ?case by (simp add:vt_nil)
   229 next
   625 next
   244       show ?thesis by simp
   640       show ?thesis by simp
   245     qed
   641     qed
   246     ultimately show ?thesis by simp
   642     ultimately show ?thesis by simp
   247   qed
   643   qed
   248 qed
   644 qed
   249 
   645 end
   250 lemma finite_threads:
   646 
   251   shows "finite (threads s)"
   647 text {*
   252 using vt by (induct) (auto elim: step.cases)
   648   The following locale @{text "valid_moment"} is to inherit the properties 
   253 
   649   derived on any valid state to the prefix of it, with length @{text "i"}.
   254 end
   650 *}
   255 
       
   256 lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
       
   257   by (unfold s_RAG_def, auto)
       
   258 
       
   259 locale valid_moment = valid_trace + 
   651 locale valid_moment = valid_trace + 
   260   fixes i :: nat
   652   fixes i :: nat
   261 
   653 
   262 sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
   654 sublocale valid_moment < vat_moment!: valid_trace "(moment i s)"
   263   by (unfold_locales, insert vt_moment, auto)
   655   by (unfold_locales, insert vt_moment, auto)
   264 
   656 
   265 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
   657 locale valid_moment_e = valid_moment +
   266   by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
   658   assumes less_i: "i < length s"
   267 
   659 begin
   268 lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
   660   definition "next_e  = hd (moment (Suc i) s)"
   269   by (unfold s_holding_def wq_def cs_holding_def, simp)
   661 
   270 
   662   lemma trace_e: 
   271 lemma runing_ready: 
   663     "moment (Suc i) s = next_e#moment i s"
   272   shows "runing s \<subseteq> readys s"
   664    proof -
   273   unfolding runing_def readys_def
   665     from less_i have "Suc i \<le> length s" by auto
   274   by auto 
   666     from moment_plus[OF this, folded next_e_def]
   275 
   667     show ?thesis .
   276 lemma readys_threads:
   668    qed
   277   shows "readys s \<subseteq> threads s"
   669 
   278   unfolding readys_def
   670 end
   279   by auto
   671 
   280 
   672 sublocale valid_moment_e < vat_moment_e!: valid_trace_e "moment i s" "next_e"
   281 lemma wq_v_neq [simp]:
   673   using vt_moment[of "Suc i", unfolded trace_e]
   282    "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
   674   by (unfold_locales, simp)
   283   by (auto simp:wq_def Let_def cp_def split:list.splits)
   675 
   284 
   676 section {* Distinctiveness of waiting queues *}
   285 lemma runing_head:
       
   286   assumes "th \<in> runing s"
       
   287   and "th \<in> set (wq_fun (schs s) cs)"
       
   288   shows "th = hd (wq_fun (schs s) cs)"
       
   289   using assms
       
   290   by (simp add:runing_def readys_def s_waiting_def wq_def)
       
   291 
       
   292 context valid_trace
       
   293 begin
       
   294 
       
   295 lemma runing_wqE:
       
   296   assumes "th \<in> runing s"
       
   297   and "th \<in> set (wq s cs)"
       
   298   obtains rest where "wq s cs = th#rest"
       
   299 proof -
       
   300   from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest"
       
   301     by (meson list.set_cases)
       
   302   have "th' = th"
       
   303   proof(rule ccontr)
       
   304     assume "th' \<noteq> th"
       
   305     hence "th \<noteq> hd (wq s cs)" using eq_wq by auto 
       
   306     with assms(2)
       
   307     have "waiting s th cs" 
       
   308       by (unfold s_waiting_def, fold wq_def, auto)
       
   309     with assms show False 
       
   310       by (unfold runing_def readys_def, auto)
       
   311   qed
       
   312   with eq_wq that show ?thesis by metis
       
   313 qed
       
   314 
       
   315 end
       
   316 
   677 
   317 context valid_trace_create
   678 context valid_trace_create
   318 begin
   679 begin
   319 
   680 
   320 lemma wq_neq_simp [simp]:
   681 lemma wq_kept [simp]:
   321   shows "wq (e#s) cs' = wq s cs'"
   682   shows "wq (e#s) cs' = wq s cs'"
   322     using assms unfolding is_create wq_def
   683     using assms unfolding is_create wq_def
   323   by (auto simp:Let_def)
   684   by (auto simp:Let_def)
   324 
   685 
   325 lemma wq_distinct_kept:
   686 lemma wq_distinct_kept:
   329 end
   690 end
   330 
   691 
   331 context valid_trace_exit
   692 context valid_trace_exit
   332 begin
   693 begin
   333 
   694 
   334 lemma wq_neq_simp [simp]:
   695 lemma wq_kept [simp]:
   335   shows "wq (e#s) cs' = wq s cs'"
   696   shows "wq (e#s) cs' = wq s cs'"
   336     using assms unfolding is_exit wq_def
   697     using assms unfolding is_exit wq_def
   337   by (auto simp:Let_def)
   698   by (auto simp:Let_def)
   338 
   699 
   339 lemma wq_distinct_kept:
   700 lemma wq_distinct_kept:
   340   assumes "distinct (wq s cs')"
   701   assumes "distinct (wq s cs')"
   341   shows "distinct (wq (e#s) cs')"
   702   shows "distinct (wq (e#s) cs')"
   342   using assms by simp
   703   using assms by simp
   343 end
   704 end
   344 
   705 
   345 context valid_trace_p
   706 context valid_trace_p 
   346 begin
   707 begin
   347 
   708 
   348 lemma wq_neq_simp [simp]:
   709 lemma wq_neq_simp [simp]:
   349   assumes "cs' \<noteq> cs"
   710   assumes "cs' \<noteq> cs"
   350   shows "wq (e#s) cs' = wq s cs'"
   711   shows "wq (e#s) cs' = wq s cs'"
   355   shows "th \<in> runing s"
   716   shows "th \<in> runing s"
   356 proof -
   717 proof -
   357   from pip_e[unfolded is_p]
   718   from pip_e[unfolded is_p]
   358   show ?thesis by (cases, simp)
   719   show ?thesis by (cases, simp)
   359 qed
   720 qed
   360 
       
   361 lemma ready_th_s: "th \<in> readys s"
       
   362   using runing_th_s
       
   363   by (unfold runing_def, auto)
       
   364 
       
   365 lemma live_th_s: "th \<in> threads s"
       
   366   using readys_threads ready_th_s by auto
       
   367 
       
   368 lemma live_th_es: "th \<in> threads (e#s)"
       
   369   using live_th_s 
       
   370   by (unfold is_p, simp)
       
   371 
       
   372 lemma th_not_waiting: 
       
   373   "\<not> waiting s th c"
       
   374 proof -
       
   375   have "th \<in> readys s"
       
   376     using runing_ready runing_th_s by blast 
       
   377   thus ?thesis
       
   378     by (unfold readys_def, auto)
       
   379 qed
       
   380 
       
   381 lemma waiting_neq_th: 
       
   382   assumes "waiting s t c"
       
   383   shows "t \<noteq> th"
       
   384   using assms using th_not_waiting by blast 
       
   385 
   721 
   386 lemma th_not_in_wq: 
   722 lemma th_not_in_wq: 
   387   shows "th \<notin> set (wq s cs)"
   723   shows "th \<notin> set (wq s cs)"
   388 proof
   724 proof
   389   assume otherwise: "th \<in> set (wq s cs)"
   725   assume otherwise: "th \<in> set (wq s cs)"
   424   assumes "cs' \<noteq> cs"
   760   assumes "cs' \<noteq> cs"
   425   shows "wq (e#s) cs' = wq s cs'"
   761   shows "wq (e#s) cs' = wq s cs'"
   426     using assms unfolding is_v wq_def
   762     using assms unfolding is_v wq_def
   427   by (auto simp:Let_def)
   763   by (auto simp:Let_def)
   428 
   764 
   429 lemma runing_th_s:
       
   430   shows "th \<in> runing s"
       
   431 proof -
       
   432   from pip_e[unfolded is_v]
       
   433   show ?thesis by (cases, simp)
       
   434 qed
       
   435 
       
   436 lemma th_not_waiting: 
       
   437   "\<not> waiting s th c"
       
   438 proof -
       
   439   have "th \<in> readys s"
       
   440     using runing_ready runing_th_s by blast 
       
   441   thus ?thesis
       
   442     by (unfold readys_def, auto)
       
   443 qed
       
   444 
       
   445 lemma waiting_neq_th: 
       
   446   assumes "waiting s t c"
       
   447   shows "t \<noteq> th"
       
   448   using assms using th_not_waiting by blast 
       
   449 
       
   450 lemma wq_s_cs:
   765 lemma wq_s_cs:
   451   "wq s cs = th#rest"
   766   "wq s cs = th#rest"
   452 proof -
   767 proof -
   453   from pip_e[unfolded is_v]
   768   from pip_e[unfolded is_v]
   454   show ?thesis
   769   show ?thesis
   480 end
   795 end
   481 
   796 
   482 context valid_trace_set
   797 context valid_trace_set
   483 begin
   798 begin
   484 
   799 
   485 lemma wq_neq_simp [simp]:
   800 lemma wq_kept [simp]:
   486   shows "wq (e#s) cs' = wq s cs'"
   801   shows "wq (e#s) cs' = wq s cs'"
   487     using assms unfolding is_set wq_def
   802     using assms unfolding is_set wq_def
   488   by (auto simp:Let_def)
   803   by (auto simp:Let_def)
   489 
   804 
   490 lemma wq_distinct_kept:
   805 lemma wq_distinct_kept:
   494 end
   809 end
   495 
   810 
   496 context valid_trace
   811 context valid_trace
   497 begin
   812 begin
   498 
   813 
   499 lemma actor_inv: 
   814 lemma  finite_threads:
   500   assumes "PIP s e"
   815   shows "finite (threads s)"
   501   and "\<not> isCreate e"
   816   using vt by (induct) (auto elim: step.cases)
   502   shows "actor e \<in> runing s"
   817 
   503   using assms
   818 lemma finite_readys [simp]: "finite (readys s)"
   504   by (induct, auto)
   819   using finite_threads readys_threads rev_finite_subset by blast
   505 
       
   506 lemma isP_E:
       
   507   assumes "isP e"
       
   508   obtains cs where "e = P (actor e) cs"
       
   509   using assms by (cases e, auto)
       
   510 
       
   511 lemma isV_E:
       
   512   assumes "isV e"
       
   513   obtains cs where "e = V (actor e) cs"
       
   514   using assms by (cases e, auto) 
       
   515 
   820 
   516 lemma wq_distinct: "distinct (wq s cs)"
   821 lemma wq_distinct: "distinct (wq s cs)"
   517 proof(induct rule:ind)
   822 proof(induct rule:ind)
   518   case (Cons s e)
   823   case (Cons s e)
   519   interpret vt_e: valid_trace_e s e using Cons by simp
   824   interpret vt_e: valid_trace_e s e using Cons by simp
   544   qed
   849   qed
   545 qed (unfold wq_def Let_def, simp)
   850 qed (unfold wq_def Let_def, simp)
   546 
   851 
   547 end
   852 end
   548 
   853 
       
   854 section {* Waiting queues and threads *}
       
   855 
   549 context valid_trace_e
   856 context valid_trace_e
   550 begin
   857 begin
   551 
   858 
   552 text {*
   859 lemma wq_out_inv: 
   553   The following lemma shows that only the @{text "P"}
   860   assumes s_in: "thread \<in> set (wq s cs)"
   554   operation can add new thread into waiting queues. 
   861   and s_hd: "thread = hd (wq s cs)"
   555   Such kind of lemmas are very obvious, but need to be checked formally.
   862   and s_i: "thread \<noteq> hd (wq (e#s) cs)"
   556   This is a kind of confirmation that our modelling is correct.
   863   shows "e = V thread cs"
   557 *}
   864 proof(cases e)
       
   865 -- {* There are only two non-trivial cases: *}
       
   866   case (V th cs1)
       
   867   show ?thesis
       
   868   proof(cases "cs1 = cs")
       
   869     case True
       
   870     have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] .
       
   871     thus ?thesis
       
   872     proof(cases)
       
   873       case (thread_V)
       
   874       moreover have "th = thread" using thread_V(2) s_hd
       
   875           by (unfold s_holding_def wq_def, simp)
       
   876       ultimately show ?thesis using V True by simp
       
   877     qed
       
   878   qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
       
   879 next
       
   880   case (P th cs1)
       
   881   show ?thesis
       
   882   proof(cases "cs1 = cs")
       
   883     case True
       
   884     with P have "wq (e#s) cs = wq_fun (schs s) cs @ [th]"
       
   885       by (auto simp:wq_def Let_def split:if_splits)
       
   886     with s_i s_hd s_in have False
       
   887       by (metis empty_iff hd_append2 list.set(1) wq_def) 
       
   888     thus ?thesis by simp
       
   889   qed (insert assms P, auto simp:wq_def Let_def split:if_splits)
       
   890 qed (insert assms, auto simp:wq_def Let_def split:if_splits)
   558 
   891 
   559 lemma wq_in_inv: 
   892 lemma wq_in_inv: 
   560   assumes s_ni: "thread \<notin> set (wq s cs)"
   893   assumes s_ni: "thread \<notin> set (wq s cs)"
   561   and s_i: "thread \<in> set (wq (e#s) cs)"
   894   and s_i: "thread \<in> set (wq (e#s) cs)"
   562   shows "e = P thread cs"
   895   shows "e = P thread cs"
   584     qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
   917     qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
   585   qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
   918   qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
   586   thus ?thesis by auto
   919   thus ?thesis by auto
   587 qed (insert assms, auto simp:wq_def Let_def split:if_splits)
   920 qed (insert assms, auto simp:wq_def Let_def split:if_splits)
   588 
   921 
   589 lemma wq_out_inv: 
   922 end
   590   assumes s_in: "thread \<in> set (wq s cs)"
   923 
   591   and s_hd: "thread = hd (wq s cs)"
   924 lemma (in valid_trace_create)
   592   and s_i: "thread \<noteq> hd (wq (e#s) cs)"
   925   th_not_in_threads: "th \<notin> threads s"
   593   shows "e = V thread cs"
   926 proof -
   594 proof(cases e)
   927   from pip_e[unfolded is_create]
   595 -- {* There are only two non-trivial cases: *}
   928   show ?thesis by (cases, simp)
   596   case (V th cs1)
   929 qed
       
   930 
       
   931 lemma (in valid_trace_create)
       
   932   threads_es [simp]: "threads (e#s) = threads s \<union> {th}"
       
   933   by (unfold is_create, simp)
       
   934 
       
   935 lemma (in valid_trace_exit)
       
   936   threads_es [simp]: "threads (e#s) = threads s - {th}"
       
   937   by (unfold is_exit, simp)
       
   938 
       
   939 lemma (in valid_trace_p)
       
   940   threads_es [simp]: "threads (e#s) = threads s"
       
   941   by (unfold is_p, simp)
       
   942 
       
   943 lemma (in valid_trace_v)
       
   944   threads_es [simp]: "threads (e#s) = threads s"
       
   945   by (unfold is_v, simp)
       
   946 
       
   947 lemma (in valid_trace_v)
       
   948   th_not_in_rest[simp]: "th \<notin> set rest"
       
   949 proof
       
   950   assume otherwise: "th \<in> set rest"
       
   951   have "distinct (wq s cs)" by (simp add: wq_distinct)
       
   952   from this[unfolded wq_s_cs] and otherwise
       
   953   show False by auto
       
   954 qed
       
   955 
       
   956 lemma (in valid_trace_v) distinct_rest: "distinct rest"
       
   957   by (simp add: distinct_tl rest_def wq_distinct)
       
   958 
       
   959 lemma (in valid_trace_v)
       
   960   set_wq_es_cs [simp]: "set (wq (e#s) cs) = set (wq s cs) - {th}"
       
   961 proof(unfold wq_es_cs wq'_def, rule someI2)
       
   962   show "distinct rest \<and> set rest = set rest"
       
   963     by (simp add: distinct_rest) 
       
   964 next
       
   965   fix x
       
   966   assume "distinct x \<and> set x = set rest"
       
   967   thus "set x = set (wq s cs) - {th}" 
       
   968       by (unfold wq_s_cs, simp)
       
   969 qed
       
   970 
       
   971 lemma (in valid_trace_exit)
       
   972   th_not_in_wq: "th \<notin> set (wq s cs)"
       
   973 proof -
       
   974   from pip_e[unfolded is_exit]
   597   show ?thesis
   975   show ?thesis
   598   proof(cases "cs1 = cs")
   976   by (cases, unfold holdents_def s_holding_def, fold wq_def, 
   599     case True
   977              auto elim!:runing_wqE)
   600     have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] .
   978 qed
   601     thus ?thesis
   979 
   602     proof(cases)
   980 lemma (in valid_trace) wq_threads: 
   603       case (thread_V)
   981   assumes "th \<in> set (wq s cs)"
   604       moreover have "th = thread" using thread_V(2) s_hd
   982   shows "th \<in> threads s"
   605           by (unfold s_holding_def wq_def, simp)
   983   using assms
   606       ultimately show ?thesis using V True by simp
   984 proof(induct rule:ind)
   607     qed
   985   case (Nil)
   608   qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
   986   thus ?case by (auto simp:wq_def)
   609 next
   987 next
   610   case (P th cs1)
   988   case (Cons s e)
   611   show ?thesis
   989   interpret vt_e: valid_trace_e s e using Cons by simp
   612   proof(cases "cs1 = cs")
   990   show ?case
   613     case True
   991   proof(cases e)
   614     with P have "wq (e#s) cs = wq_fun (schs s) cs @ [th]"
   992     case (Create th' prio')
   615       by (auto simp:wq_def Let_def split:if_splits)
   993     interpret vt: valid_trace_create s e th' prio'
   616     with s_i s_hd s_in have False
   994       using Create by (unfold_locales, simp)
   617       by (metis empty_iff hd_append2 list.set(1) wq_def) 
   995     show ?thesis
   618     thus ?thesis by simp
   996       using Cons.hyps(2) Cons.prems by auto
   619   qed (insert assms P, auto simp:wq_def Let_def split:if_splits)
   997   next
   620 qed (insert assms, auto simp:wq_def Let_def split:if_splits)
   998     case (Exit th')
   621 
   999     interpret vt: valid_trace_exit s e th'
   622 end
  1000       using Exit by (unfold_locales, simp)
   623 
  1001     show ?thesis
       
  1002       using Cons.hyps(2) Cons.prems vt.th_not_in_wq by auto 
       
  1003   next
       
  1004     case (P th' cs')
       
  1005     interpret vt: valid_trace_p s e th' cs'
       
  1006       using P by (unfold_locales, simp)
       
  1007     show ?thesis
       
  1008       using Cons.hyps(2) Cons.prems readys_threads 
       
  1009         runing_ready vt.is_p vt.runing_th_s vt_e.wq_in_inv 
       
  1010         by fastforce 
       
  1011   next
       
  1012     case (V th' cs')
       
  1013     interpret vt: valid_trace_v s e th' cs'
       
  1014       using V by (unfold_locales, simp)
       
  1015     show ?thesis using Cons
       
  1016       using vt.is_v vt.threads_es vt_e.wq_in_inv by blast
       
  1017   next
       
  1018     case (Set th' prio)
       
  1019     interpret vt: valid_trace_set s e th' prio
       
  1020       using Set by (unfold_locales, simp)
       
  1021     show ?thesis using Cons.hyps(2) Cons.prems vt.is_set 
       
  1022         by (auto simp:wq_def Let_def)
       
  1023   qed
       
  1024 qed 
       
  1025 
       
  1026 section {* RAG and threads *}
   624 
  1027 
   625 context valid_trace
  1028 context valid_trace
   626 begin
  1029 begin
   627 
  1030 
   628 
  1031 lemma  dm_RAG_threads:
   629 text {* (* ddd *)
  1032   assumes in_dom: "(Th th) \<in> Domain (RAG s)"
   630   The nature of the work is like this: since it starts from a very simple and basic 
  1033   shows "th \<in> threads s"
   631   model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
  1034 proof -
   632   For instance, the fact 
  1035   from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
   633   that one thread can not be blocked by two critical resources at the same time
  1036   moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
   634   is obvious, because only running threads can make new requests, if one is waiting for 
  1037   ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
   635   a critical resource and get blocked, it can not make another resource request and get 
  1038   hence "th \<in> set (wq s cs)"
   636   blocked the second time (because it is not running). 
  1039     by (unfold s_RAG_def, auto simp:cs_waiting_def)
   637 
  1040   from wq_threads [OF this] show ?thesis .
   638   To derive this fact, one needs to prove by contraction and 
  1041 qed
   639   reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
  1042 
   640   named @{text "p_split"}, which is about status changing along the time axis. It says if 
  1043 lemma rg_RAG_threads: 
   641   a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
  1044   assumes "(Th th) \<in> Range (RAG s)"
   642   but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
  1045   shows "th \<in> threads s"
   643   in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
  1046   using assms
   644   of events leading to it), such that @{text "Q"} switched 
  1047   by (unfold s_RAG_def cs_waiting_def cs_holding_def, 
   645   from being @{text "False"} to @{text "True"} and kept being @{text "True"}
  1048        auto intro:wq_threads)
   646   till the last moment of @{text "s"}.
  1049 
   647 
  1050 lemma RAG_threads:
   648   Suppose a thread @{text "th"} is blocked
  1051   assumes "(Th th) \<in> Field (RAG s)"
   649   on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
  1052   shows "th \<in> threads s"
   650   since no thread is blocked at the very beginning, by applying 
  1053   using assms
   651   @{text "p_split"} to these two blocking facts, there exist 
  1054   by (metis Field_def UnE dm_RAG_threads rg_RAG_threads)
   652   two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
  1055 
   653   @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
  1056 end
   654   and kept on blocked on them respectively ever since.
  1057 
   655  
  1058 section {* The change of @{term RAG} *}
   656   Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
       
   657   However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
       
   658   in blocked state at moment @{text "t2"} and could not
       
   659   make any request and get blocked the second time: Contradiction.
       
   660 *}
       
   661 
       
   662 lemma waiting_unique_pre: (* ddd *)
       
   663   assumes h11: "thread \<in> set (wq s cs1)"
       
   664   and h12: "thread \<noteq> hd (wq s cs1)"
       
   665   assumes h21: "thread \<in> set (wq s cs2)"
       
   666   and h22: "thread \<noteq> hd (wq s cs2)"
       
   667   and neq12: "cs1 \<noteq> cs2"
       
   668   shows "False"
       
   669 proof -
       
   670   let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
       
   671   from h11 and h12 have q1: "?Q cs1 s" by simp
       
   672   from h21 and h22 have q2: "?Q cs2 s" by simp
       
   673   have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
       
   674   have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
       
   675   from p_split [of "?Q cs1", OF q1 nq1]
       
   676   obtain t1 where lt1: "t1 < length s"
       
   677     and np1: "\<not> ?Q cs1 (moment t1 s)"
       
   678     and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto
       
   679   from p_split [of "?Q cs2", OF q2 nq2]
       
   680   obtain t2 where lt2: "t2 < length s"
       
   681     and np2: "\<not> ?Q cs2 (moment t2 s)"
       
   682     and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto
       
   683   { fix s cs
       
   684     assume q: "?Q cs s"
       
   685     have "thread \<notin> runing s"
       
   686     proof
       
   687       assume "thread \<in> runing s"
       
   688       hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> 
       
   689                  thread \<noteq> hd (wq_fun (schs s) cs))"
       
   690         by (unfold runing_def s_waiting_def readys_def, auto)
       
   691       from this[rule_format, of cs] q 
       
   692       show False by (simp add: wq_def) 
       
   693     qed
       
   694   } note q_not_runing = this
       
   695   { fix t1 t2 cs1 cs2
       
   696     assume  lt1: "t1 < length s"
       
   697     and np1: "\<not> ?Q cs1 (moment t1 s)"
       
   698     and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))"
       
   699     and lt2: "t2 < length s"
       
   700     and np2: "\<not> ?Q cs2 (moment t2 s)"
       
   701     and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))"
       
   702     and lt12: "t1 < t2"
       
   703     let ?t3 = "Suc t2"
       
   704     from lt2 have le_t3: "?t3 \<le> length s" by auto
       
   705     from moment_plus [OF this] 
       
   706     obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
       
   707     have "t2 < ?t3" by simp
       
   708     from nn2 [rule_format, OF this] and eq_m
       
   709     have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   710          h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   711     have "vt (e#moment t2 s)"
       
   712     proof -
       
   713       from vt_moment 
       
   714       have "vt (moment ?t3 s)" .
       
   715       with eq_m show ?thesis by simp
       
   716     qed
       
   717     then interpret vt_e: valid_trace_e "moment t2 s" "e"
       
   718         by (unfold_locales, auto, cases, simp)
       
   719     have ?thesis
       
   720     proof -
       
   721       have "thread \<in> runing (moment t2 s)"
       
   722       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   723         case True
       
   724         have "e = V thread cs2"
       
   725         proof -
       
   726           have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
       
   727               using True and np2  by auto 
       
   728           from vt_e.wq_out_inv[OF True this h2]
       
   729           show ?thesis .
       
   730         qed
       
   731         thus ?thesis using vt_e.actor_inv[OF vt_e.pip_e] by auto
       
   732       next
       
   733         case False
       
   734         have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
       
   735         with vt_e.actor_inv[OF vt_e.pip_e]
       
   736         show ?thesis by auto
       
   737       qed
       
   738       moreover have "thread \<notin> runing (moment t2 s)"
       
   739         by (rule q_not_runing[OF nn1[rule_format, OF lt12]])
       
   740       ultimately show ?thesis by simp
       
   741     qed
       
   742   } note lt_case = this
       
   743   show ?thesis
       
   744   proof -
       
   745     { assume "t1 < t2"
       
   746       from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this]
       
   747       have ?thesis .
       
   748     } moreover {
       
   749       assume "t2 < t1"
       
   750       from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this]
       
   751       have ?thesis .
       
   752     } moreover {
       
   753       assume eq_12: "t1 = t2"
       
   754       let ?t3 = "Suc t2"
       
   755       from lt2 have le_t3: "?t3 \<le> length s" by auto
       
   756       from moment_plus [OF this] 
       
   757       obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
       
   758       have lt_2: "t2 < ?t3" by simp
       
   759       from nn2 [rule_format, OF this] and eq_m
       
   760       have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   761            h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   762       from nn1[rule_format, OF lt_2[folded eq_12]] eq_m[folded eq_12]
       
   763       have g1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   764            g2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
       
   765       have "vt (e#moment t2 s)"
       
   766       proof -
       
   767         from vt_moment 
       
   768         have "vt (moment ?t3 s)" .
       
   769         with eq_m show ?thesis by simp
       
   770       qed
       
   771       then interpret vt_e: valid_trace_e "moment t2 s" "e"
       
   772           by (unfold_locales, auto, cases, simp)
       
   773       have "e = V thread cs2 \<or> e = P thread cs2"
       
   774       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   775         case True
       
   776         have "e = V thread cs2"
       
   777         proof -
       
   778           have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
       
   779               using True and np2  by auto 
       
   780           from vt_e.wq_out_inv[OF True this h2]
       
   781           show ?thesis .
       
   782         qed
       
   783         thus ?thesis by auto
       
   784       next
       
   785         case False
       
   786         have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
       
   787         thus ?thesis by auto
       
   788       qed
       
   789       moreover have "e = V thread cs1 \<or> e = P thread cs1"
       
   790       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
       
   791         case True
       
   792         have eq_th: "thread = hd (wq (moment t1 s) cs1)" 
       
   793               using True and np1  by auto 
       
   794         from vt_e.wq_out_inv[folded eq_12, OF True this g2]
       
   795         have "e = V thread cs1" .
       
   796         thus ?thesis by auto
       
   797       next
       
   798         case False
       
   799         have "e = P thread cs1" using vt_e.wq_in_inv[folded eq_12, OF False g1] .
       
   800         thus ?thesis by auto
       
   801       qed
       
   802       ultimately have ?thesis using neq12 by auto
       
   803     } ultimately show ?thesis using nat_neq_iff by blast 
       
   804   qed
       
   805 qed
       
   806 
       
   807 text {*
       
   808   This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
       
   809 *}
       
   810 
       
   811 lemma waiting_unique:
       
   812   assumes "waiting s th cs1"
       
   813   and "waiting s th cs2"
       
   814   shows "cs1 = cs2"
       
   815   using waiting_unique_pre assms
       
   816   unfolding wq_def s_waiting_def
       
   817   by auto
       
   818 
       
   819 end
       
   820 
       
   821 (* not used *)
       
   822 text {*
       
   823   Every thread can only be blocked on one critical resource, 
       
   824   symmetrically, every critical resource can only be held by one thread. 
       
   825   This fact is much more easier according to our definition. 
       
   826 *}
       
   827 lemma held_unique:
       
   828   assumes "holding (s::event list) th1 cs"
       
   829   and "holding s th2 cs"
       
   830   shows "th1 = th2"
       
   831  by (insert assms, unfold s_holding_def, auto)
       
   832 
       
   833 lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
       
   834   apply (induct s, auto)
       
   835   by (case_tac a, auto split:if_splits)
       
   836 
       
   837 lemma last_set_unique: 
       
   838   "\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
       
   839           \<Longrightarrow> th1 = th2"
       
   840   apply (induct s, auto)
       
   841   by (case_tac a, auto split:if_splits dest:last_set_lt)
       
   842 
       
   843 lemma preced_unique : 
       
   844   assumes pcd_eq: "preced th1 s = preced th2 s"
       
   845   and th_in1: "th1 \<in> threads s"
       
   846   and th_in2: " th2 \<in> threads s"
       
   847   shows "th1 = th2"
       
   848 proof -
       
   849   from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
       
   850   from last_set_unique [OF this th_in1 th_in2]
       
   851   show ?thesis .
       
   852 qed
       
   853                       
       
   854 lemma preced_linorder: 
       
   855   assumes neq_12: "th1 \<noteq> th2"
       
   856   and th_in1: "th1 \<in> threads s"
       
   857   and th_in2: " th2 \<in> threads s"
       
   858   shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
       
   859 proof -
       
   860   from preced_unique [OF _ th_in1 th_in2] and neq_12 
       
   861   have "preced th1 s \<noteq> preced th2 s" by auto
       
   862   thus ?thesis by auto
       
   863 qed
       
   864 
  1059 
   865 text {*
  1060 text {*
   866   The following three lemmas show that @{text "RAG"} does not change
  1061   The following three lemmas show that @{text "RAG"} does not change
   867   by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
  1062   by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
   868   events, respectively.
  1063   events, respectively.
   869 *}
  1064 *}
   870 
  1065 
   871 lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
  1066 lemma (in valid_trace_set) RAG_unchanged [simp]: "(RAG (e # s)) = RAG s"
   872 apply (unfold s_RAG_def s_waiting_def wq_def)
  1067    by (unfold is_set s_RAG_def s_waiting_def wq_def, simp add:Let_def)
   873 by (simp add:Let_def)
  1068 
   874 
  1069 lemma (in valid_trace_create) RAG_unchanged [simp]: "(RAG (e # s)) = RAG s"
   875 lemma (in valid_trace_set)
  1070  by (unfold is_create s_RAG_def s_waiting_def wq_def, simp add:Let_def)
   876    RAG_unchanged: "(RAG (e # s)) = RAG s"
  1071 
   877    by (unfold is_set RAG_set_unchanged, simp)
  1072 lemma (in valid_trace_exit) RAG_unchanged[simp]: "(RAG (e # s)) = RAG s"
   878 
  1073   by (unfold is_exit s_RAG_def s_waiting_def wq_def, simp add:Let_def)
   879 lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
       
   880 apply (unfold s_RAG_def s_waiting_def wq_def)
       
   881 by (simp add:Let_def)
       
   882 
       
   883 lemma (in valid_trace_create)
       
   884    RAG_unchanged: "(RAG (e # s)) = RAG s"
       
   885    by (unfold is_create RAG_create_unchanged, simp)
       
   886 
       
   887 lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
       
   888 apply (unfold s_RAG_def s_waiting_def wq_def)
       
   889 by (simp add:Let_def)
       
   890 
       
   891 lemma (in valid_trace_exit)
       
   892    RAG_unchanged: "(RAG (e # s)) = RAG s"
       
   893    by (unfold is_exit RAG_exit_unchanged, simp)
       
   894 
  1074 
   895 context valid_trace_v
  1075 context valid_trace_v
   896 begin
  1076 begin
   897 
       
   898 lemma distinct_rest: "distinct rest"
       
   899   by (simp add: distinct_tl rest_def wq_distinct)
       
   900 
  1077 
   901 lemma holding_cs_eq_th:
  1078 lemma holding_cs_eq_th:
   902   assumes "holding s t cs"
  1079   assumes "holding s t cs"
   903   shows "t = th"
  1080   shows "t = th"
   904 proof -
  1081 proof -
   913 
  1090 
   914 lemma distinct_wq': "distinct wq'"
  1091 lemma distinct_wq': "distinct wq'"
   915   by (metis (mono_tags, lifting) distinct_rest  some_eq_ex wq'_def)
  1092   by (metis (mono_tags, lifting) distinct_rest  some_eq_ex wq'_def)
   916   
  1093   
   917 lemma set_wq': "set wq' = set rest"
  1094 lemma set_wq': "set wq' = set rest"
   918   by (metis (mono_tags, lifting) distinct_rest rest_def 
  1095   by (metis (mono_tags, lifting) distinct_rest some_eq_ex wq'_def)
   919       some_eq_ex wq'_def)
       
   920     
  1096     
   921 lemma th'_in_inv:
  1097 lemma th'_in_inv:
   922   assumes "th' \<in> set wq'"
  1098   assumes "th' \<in> set wq'"
   923   shows "th' \<in> set rest"
  1099   shows "th' \<in> set rest"
   924   using assms set_wq' by simp
  1100   using assms set_wq' by simp
       
  1101 
       
  1102 lemma runing_th_s:
       
  1103   shows "th \<in> runing s"
       
  1104 proof -
       
  1105   from pip_e[unfolded is_v]
       
  1106   show ?thesis by (cases, simp)
       
  1107 qed
   925 
  1108 
   926 lemma neq_t_th: 
  1109 lemma neq_t_th: 
   927   assumes "waiting (e#s) t c"
  1110   assumes "waiting (e#s) t c"
   928   shows "t \<noteq> th"
  1111   shows "t \<noteq> th"
   929 proof
  1112 proof
   943         by (unfold is_v, simp)
  1126         by (unfold is_v, simp)
   944     hence "waiting s t c" using assms 
  1127     hence "waiting s t c" using assms 
   945         by (simp add: cs_waiting_def waiting_eq)
  1128         by (simp add: cs_waiting_def waiting_eq)
   946     hence "t \<notin> readys s" by (unfold readys_def, auto)
  1129     hence "t \<notin> readys s" by (unfold readys_def, auto)
   947     hence "t \<notin> runing s" using runing_ready by auto 
  1130     hence "t \<notin> runing s" using runing_ready by auto 
   948     with runing_th_s[folded otherwise] show ?thesis by auto
  1131     with runing_th_s[folded otherwise] show ?thesis by auto 
   949   qed
  1132   qed
   950 qed
  1133 qed
   951 
  1134 
   952 lemma waiting_esI1:
  1135 lemma waiting_esI1:
   953   assumes "waiting s t c"
  1136   assumes "waiting s t c"
   993 next
  1176 next
   994   fix x
  1177   fix x
   995   assume " distinct x \<and> set x = set rest" 
  1178   assume " distinct x \<and> set x = set rest" 
   996   thus "x \<noteq> []" using rest_nnl by auto
  1179   thus "x \<noteq> []" using rest_nnl by auto
   997 qed 
  1180 qed 
   998 
       
   999 definition "taker = hd wq'"
       
  1000 
       
  1001 definition "rest' = tl wq'"
       
  1002 
  1181 
  1003 lemma eq_wq': "wq' = taker # rest'"
  1182 lemma eq_wq': "wq' = taker # rest'"
  1004   by (simp add: neq_wq' rest'_def taker_def)
  1183   by (simp add: neq_wq' rest'_def taker_def)
  1005 
  1184 
  1006 lemma next_th_taker: 
  1185 lemma next_th_taker: 
  1200   from that[OF False this] show ?thesis .
  1379   from that[OF False this] show ?thesis .
  1201 qed
  1380 qed
  1202 
  1381 
  1203 end 
  1382 end 
  1204 
  1383 
  1205 lemma rel_eqI:
       
  1206   assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B"
       
  1207   and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A"
       
  1208   shows "A = B"
       
  1209   using assms by auto
       
  1210 
       
  1211 lemma in_RAG_E:
       
  1212   assumes "(n1, n2) \<in> RAG (s::state)"
       
  1213   obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs"
       
  1214       | (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs"
       
  1215   using assms[unfolded s_RAG_def, folded waiting_eq holding_eq]
       
  1216   by auto
       
  1217   
  1384   
  1218 context valid_trace_v
  1385 context valid_trace_v
  1219 begin
  1386 begin
  1220 
  1387 
  1221 lemma RAG_es:
  1388 lemma RAG_es:
  1398     qed
  1565     qed
  1399    qed
  1566    qed
  1400  qed
  1567  qed
  1401 qed
  1568 qed
  1402 
  1569 
  1403 end
  1570 lemma 
  1404 
  1571   finite_RAG_kept:
  1405 lemma step_RAG_v: 
  1572   assumes "finite (RAG s)"
  1406 assumes vt:
  1573   shows "finite (RAG (e#s))"
  1407   "vt (V th cs#s)"
  1574 proof(cases "rest = []")
  1408 shows "
  1575   case True
  1409   RAG (V th cs # s) =
  1576   interpret vt: valid_trace_v_e using True
  1410   RAG s - {(Cs cs, Th th)} -
  1577     by (unfold_locales, simp)
  1411   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
  1578   show ?thesis using assms
  1412   {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
  1579     by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
  1413 proof -
       
  1414   interpret vt_v: valid_trace_v s "V th cs"
       
  1415     using assms step_back_vt by (unfold_locales, auto) 
       
  1416   show ?thesis using vt_v.RAG_es .
       
  1417 qed
       
  1418 
       
  1419 lemma (in valid_trace_create)
       
  1420   th_not_in_threads: "th \<notin> threads s"
       
  1421 proof -
       
  1422   from pip_e[unfolded is_create]
       
  1423   show ?thesis by (cases, simp)
       
  1424 qed
       
  1425 
       
  1426 lemma (in valid_trace_create)
       
  1427   threads_es [simp]: "threads (e#s) = threads s \<union> {th}"
       
  1428   by (unfold is_create, simp)
       
  1429 
       
  1430 lemma (in valid_trace_exit)
       
  1431   threads_es [simp]: "threads (e#s) = threads s - {th}"
       
  1432   by (unfold is_exit, simp)
       
  1433 
       
  1434 lemma (in valid_trace_p)
       
  1435   threads_es [simp]: "threads (e#s) = threads s"
       
  1436   by (unfold is_p, simp)
       
  1437 
       
  1438 lemma (in valid_trace_v)
       
  1439   threads_es [simp]: "threads (e#s) = threads s"
       
  1440   by (unfold is_v, simp)
       
  1441 
       
  1442 lemma (in valid_trace_v)
       
  1443   th_not_in_rest[simp]: "th \<notin> set rest"
       
  1444 proof
       
  1445   assume otherwise: "th \<in> set rest"
       
  1446   have "distinct (wq s cs)" by (simp add: wq_distinct)
       
  1447   from this[unfolded wq_s_cs] and otherwise
       
  1448   show False by auto
       
  1449 qed
       
  1450 
       
  1451 lemma (in valid_trace_v)
       
  1452   set_wq_es_cs [simp]: "set (wq (e#s) cs) = set (wq s cs) - {th}"
       
  1453 proof(unfold wq_es_cs wq'_def, rule someI2)
       
  1454   show "distinct rest \<and> set rest = set rest"
       
  1455     by (simp add: distinct_rest)
       
  1456 next
  1580 next
  1457   fix x
  1581   case False
  1458   assume "distinct x \<and> set x = set rest"
  1582   interpret vt: valid_trace_v_n using False
  1459   thus "set x = set (wq s cs) - {th}" 
  1583     by (unfold_locales, simp)
  1460       by (unfold wq_s_cs, simp)
  1584   show ?thesis using assms
  1461 qed
  1585     by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
  1462 
  1586 qed
  1463 lemma (in valid_trace_exit)
  1587 
  1464   th_not_in_wq: "th \<notin> set (wq s cs)"
  1588 end
  1465 proof -
       
  1466   from pip_e[unfolded is_exit]
       
  1467   show ?thesis
       
  1468   by (cases, unfold holdents_def s_holding_def, fold wq_def, 
       
  1469              auto elim!:runing_wqE)
       
  1470 qed
       
  1471 
       
  1472 lemma (in valid_trace) wq_threads: 
       
  1473   assumes "th \<in> set (wq s cs)"
       
  1474   shows "th \<in> threads s"
       
  1475   using assms
       
  1476 proof(induct rule:ind)
       
  1477   case (Nil)
       
  1478   thus ?case by (auto simp:wq_def)
       
  1479 next
       
  1480   case (Cons s e)
       
  1481   interpret vt_e: valid_trace_e s e using Cons by simp
       
  1482   show ?case
       
  1483   proof(cases e)
       
  1484     case (Create th' prio')
       
  1485     interpret vt: valid_trace_create s e th' prio'
       
  1486       using Create by (unfold_locales, simp)
       
  1487     show ?thesis
       
  1488       using Cons.hyps(2) Cons.prems by auto
       
  1489   next
       
  1490     case (Exit th')
       
  1491     interpret vt: valid_trace_exit s e th'
       
  1492       using Exit by (unfold_locales, simp)
       
  1493     show ?thesis
       
  1494       using Cons.hyps(2) Cons.prems vt.th_not_in_wq by auto 
       
  1495   next
       
  1496     case (P th' cs')
       
  1497     interpret vt: valid_trace_p s e th' cs'
       
  1498       using P by (unfold_locales, simp)
       
  1499     show ?thesis
       
  1500       using Cons.hyps(2) Cons.prems readys_threads 
       
  1501         runing_ready vt.is_p vt.runing_th_s vt_e.wq_in_inv 
       
  1502         by fastforce 
       
  1503   next
       
  1504     case (V th' cs')
       
  1505     interpret vt: valid_trace_v s e th' cs'
       
  1506       using V by (unfold_locales, simp)
       
  1507     show ?thesis using Cons
       
  1508       using vt.is_v vt.threads_es vt_e.wq_in_inv by blast
       
  1509   next
       
  1510     case (Set th' prio)
       
  1511     interpret vt: valid_trace_set s e th' prio
       
  1512       using Set by (unfold_locales, simp)
       
  1513     show ?thesis using Cons.hyps(2) Cons.prems vt.is_set 
       
  1514         by (auto simp:wq_def Let_def)
       
  1515   qed
       
  1516 qed 
       
  1517 
       
  1518 context valid_trace
       
  1519 begin
       
  1520 
       
  1521 lemma  dm_RAG_threads:
       
  1522   assumes in_dom: "(Th th) \<in> Domain (RAG s)"
       
  1523   shows "th \<in> threads s"
       
  1524 proof -
       
  1525   from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
       
  1526   moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
       
  1527   ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
       
  1528   hence "th \<in> set (wq s cs)"
       
  1529     by (unfold s_RAG_def, auto simp:cs_waiting_def)
       
  1530   from wq_threads [OF this] show ?thesis .
       
  1531 qed
       
  1532 
       
  1533 lemma rg_RAG_threads: 
       
  1534   assumes "(Th th) \<in> Range (RAG s)"
       
  1535   shows "th \<in> threads s"
       
  1536   using assms
       
  1537   by (unfold s_RAG_def cs_waiting_def cs_holding_def, 
       
  1538        auto intro:wq_threads)
       
  1539 
       
  1540 lemma RAG_threads:
       
  1541   assumes "(Th th) \<in> Field (RAG s)"
       
  1542   shows "th \<in> threads s"
       
  1543   using assms
       
  1544   by (metis Field_def UnE dm_RAG_threads rg_RAG_threads)
       
  1545 
       
  1546 end
       
  1547 
       
  1548 lemma (in valid_trace_v)
       
  1549   preced_es [simp]: "preced th (e#s) = preced th s"
       
  1550   by (unfold is_v preced_def, simp)
       
  1551 
       
  1552 lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s"
       
  1553 proof
       
  1554   fix th'
       
  1555   show "the_preced (V th cs # s) th' = the_preced s th'"
       
  1556     by (unfold the_preced_def preced_def, simp)
       
  1557 qed
       
  1558 
       
  1559 lemma (in valid_trace_v)
       
  1560   the_preced_es: "the_preced (e#s) = the_preced s"
       
  1561   by (unfold is_v preced_def, simp)
       
  1562 
  1589 
  1563 context valid_trace_p
  1590 context valid_trace_p
  1564 begin
  1591 begin
  1565 
       
  1566 lemma not_holding_s_th_cs: "\<not> holding s th cs"
       
  1567 proof
       
  1568   assume otherwise: "holding s th cs"
       
  1569   from pip_e[unfolded is_p]
       
  1570   show False
       
  1571   proof(cases)
       
  1572     case (thread_P)
       
  1573     moreover have "(Cs cs, Th th) \<in> RAG s"
       
  1574       using otherwise cs_holding_def 
       
  1575             holding_eq th_not_in_wq by auto
       
  1576     ultimately show ?thesis by auto
       
  1577   qed
       
  1578 qed
       
  1579 
  1592 
  1580 lemma waiting_kept:
  1593 lemma waiting_kept:
  1581   assumes "waiting s th' cs'"
  1594   assumes "waiting s th' cs'"
  1582   shows "waiting (e#s) th' cs'"
  1595   shows "waiting (e#s) th' cs'"
  1583   using assms
  1596   using assms
  1584   by (metis cs_waiting_def hd_append2 list.sel(1) list.set_intros(2) 
  1597   by (metis cs_waiting_def hd_append2 list.sel(1) list.set_intros(2) 
  1585       rotate1.simps(2) self_append_conv2 set_rotate1 
  1598       rotate1.simps(2) self_append_conv2 set_rotate1 
  1586         th_not_in_wq waiting_eq wq_es_cs wq_neq_simp)
  1599         th_not_in_wq waiting_eq wq_es_cs wq_neq_simp)
  1587   
  1600 
  1588 lemma holding_kept:
  1601 lemma holding_kept:
  1589   assumes "holding s th' cs'"
  1602   assumes "holding s th' cs'"
  1590   shows "holding (e#s) th' cs'"
  1603   shows "holding (e#s) th' cs'"
  1591 proof(cases "cs' = cs")
  1604 proof(cases "cs' = cs")
  1592   case False
  1605   case False
  1600   hence "wq (e#s) cs' = th'#(rest@[th])"
  1613   hence "wq (e#s) cs' = th'#(rest@[th])"
  1601     by (simp add: True wq_es_cs) 
  1614     by (simp add: True wq_es_cs) 
  1602   thus ?thesis
  1615   thus ?thesis
  1603     by (simp add: cs_holding_def holding_eq) 
  1616     by (simp add: cs_holding_def holding_eq) 
  1604 qed
  1617 qed
  1605 
  1618 end 
  1606 end
  1619 
  1607 
  1620 lemma (in valid_trace_p) th_not_waiting: "\<not> waiting s th c"
  1608 locale valid_trace_p_h = valid_trace_p +
  1621 proof -
  1609   assumes we: "wq s cs = []"
  1622   have "th \<in> readys s"
  1610 
  1623     using runing_ready runing_th_s by blast 
  1611 locale valid_trace_p_w = valid_trace_p +
  1624   thus ?thesis
  1612   assumes wne: "wq s cs \<noteq> []"
  1625     by (unfold readys_def, auto)
  1613 begin
  1626 qed
  1614 
       
  1615 definition "holder = hd (wq s cs)"
       
  1616 definition "waiters = tl (wq s cs)"
       
  1617 definition "waiters' = waiters @ [th]"
       
  1618 
       
  1619 lemma wq_s_cs: "wq s cs = holder#waiters"
       
  1620     by (simp add: holder_def waiters_def wne)
       
  1621     
       
  1622 lemma wq_es_cs': "wq (e#s) cs = holder#waiters@[th]"
       
  1623   by (simp add: wq_es_cs wq_s_cs)
       
  1624 
       
  1625 lemma waiting_es_th_cs: "waiting (e#s) th cs"
       
  1626   using cs_waiting_def th_not_in_wq waiting_eq wq_es_cs' wq_s_cs by auto
       
  1627 
       
  1628 lemma RAG_edge: "(Th th, Cs cs) \<in> RAG (e#s)"
       
  1629    by (unfold s_RAG_def, fold waiting_eq, insert waiting_es_th_cs, auto)
       
  1630 
       
  1631 lemma holding_esE:
       
  1632   assumes "holding (e#s) th' cs'"
       
  1633   obtains "holding s th' cs'"
       
  1634   using assms 
       
  1635 proof(cases "cs' = cs")
       
  1636   case False
       
  1637   hence "wq (e#s) cs' = wq s cs'" by simp
       
  1638   with assms show ?thesis
       
  1639     using cs_holding_def holding_eq that by auto 
       
  1640 next
       
  1641   case True
       
  1642   with assms show ?thesis
       
  1643   by (metis cs_holding_def holding_eq list.sel(1) list.set_intros(1) that 
       
  1644         wq_es_cs' wq_s_cs) 
       
  1645 qed
       
  1646 
       
  1647 lemma waiting_esE:
       
  1648   assumes "waiting (e#s) th' cs'"
       
  1649   obtains "th' \<noteq> th" "waiting s th' cs'"
       
  1650      |  "th' = th" "cs' = cs"
       
  1651 proof(cases "waiting s th' cs'")
       
  1652   case True
       
  1653   have "th' \<noteq> th"
       
  1654   proof
       
  1655     assume otherwise: "th' = th"
       
  1656     from True[unfolded this]
       
  1657     show False by (simp add: th_not_waiting) 
       
  1658   qed
       
  1659   from that(1)[OF this True] show ?thesis .
       
  1660 next
       
  1661   case False
       
  1662   hence "th' = th \<and> cs' = cs"
       
  1663       by (metis assms cs_waiting_def holder_def list.sel(1) rotate1.simps(2) 
       
  1664         set_ConsD set_rotate1 waiting_eq wq_es_cs wq_es_cs' wq_neq_simp)
       
  1665   with that(2) show ?thesis by metis
       
  1666 qed
       
  1667 
       
  1668 lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Th th, Cs cs)}" (is "?L = ?R")
       
  1669 proof(rule rel_eqI)
       
  1670   fix n1 n2
       
  1671   assume "(n1, n2) \<in> ?L"
       
  1672   thus "(n1, n2) \<in> ?R" 
       
  1673   proof(cases rule:in_RAG_E)
       
  1674     case (waiting th' cs')
       
  1675     from this(3)
       
  1676     show ?thesis
       
  1677     proof(cases rule:waiting_esE)
       
  1678       case 1
       
  1679       thus ?thesis using waiting(1,2)
       
  1680         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1681     next
       
  1682       case 2
       
  1683       thus ?thesis using waiting(1,2) by auto
       
  1684     qed
       
  1685   next
       
  1686     case (holding th' cs')
       
  1687     from this(3)
       
  1688     show ?thesis
       
  1689     proof(cases rule:holding_esE)
       
  1690       case 1
       
  1691       with holding(1,2)
       
  1692       show ?thesis by (unfold s_RAG_def, fold holding_eq, auto)
       
  1693     qed
       
  1694   qed
       
  1695 next
       
  1696   fix n1 n2
       
  1697   assume "(n1, n2) \<in> ?R"
       
  1698   hence "(n1, n2) \<in> RAG s \<or> (n1 = Th th \<and> n2 = Cs cs)" by auto
       
  1699   thus "(n1, n2) \<in> ?L"
       
  1700   proof
       
  1701     assume "(n1, n2) \<in> RAG s"
       
  1702     thus ?thesis
       
  1703     proof(cases rule:in_RAG_E)
       
  1704       case (waiting th' cs')
       
  1705       from waiting_kept[OF this(3)]
       
  1706       show ?thesis using waiting(1,2)
       
  1707          by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1708     next
       
  1709       case (holding th' cs')
       
  1710       from holding_kept[OF this(3)]
       
  1711       show ?thesis using holding(1,2)
       
  1712          by (unfold s_RAG_def, fold holding_eq, auto)
       
  1713     qed
       
  1714   next
       
  1715     assume "n1 = Th th \<and> n2 = Cs cs"
       
  1716     thus ?thesis using RAG_edge by auto
       
  1717   qed
       
  1718 qed
       
  1719 
       
  1720 end
       
  1721 
  1627 
  1722 context valid_trace_p_h
  1628 context valid_trace_p_h
  1723 begin
  1629 begin
  1724 
  1630 
  1725 lemma wq_es_cs': "wq (e#s) cs = [th]"
  1631 lemma wq_es_cs': "wq (e#s) cs = [th]"
  1813   qed
  1719   qed
  1814 qed
  1720 qed
  1815 
  1721 
  1816 end
  1722 end
  1817 
  1723 
       
  1724 context valid_trace_p_w
       
  1725 begin
       
  1726 
       
  1727 lemma wq_s_cs: "wq s cs = holder#waiters"
       
  1728     by (simp add: holder_def waiters_def wne)
       
  1729     
       
  1730 lemma wq_es_cs': "wq (e#s) cs = holder#waiters@[th]"
       
  1731   by (simp add: wq_es_cs wq_s_cs)
       
  1732 
       
  1733 lemma waiting_es_th_cs: "waiting (e#s) th cs"
       
  1734   using cs_waiting_def th_not_in_wq waiting_eq wq_es_cs' wq_s_cs by auto
       
  1735 
       
  1736 lemma RAG_edge: "(Th th, Cs cs) \<in> RAG (e#s)"
       
  1737    by (unfold s_RAG_def, fold waiting_eq, insert waiting_es_th_cs, auto)
       
  1738 
       
  1739 lemma holding_esE:
       
  1740   assumes "holding (e#s) th' cs'"
       
  1741   obtains "holding s th' cs'"
       
  1742   using assms 
       
  1743 proof(cases "cs' = cs")
       
  1744   case False
       
  1745   hence "wq (e#s) cs' = wq s cs'" by simp
       
  1746   with assms show ?thesis
       
  1747     using cs_holding_def holding_eq that by auto 
       
  1748 next
       
  1749   case True
       
  1750   with assms show ?thesis
       
  1751   by (metis cs_holding_def holding_eq list.sel(1) list.set_intros(1) that 
       
  1752         wq_es_cs' wq_s_cs) 
       
  1753 qed
       
  1754 
       
  1755 lemma waiting_esE:
       
  1756   assumes "waiting (e#s) th' cs'"
       
  1757   obtains "th' \<noteq> th" "waiting s th' cs'"
       
  1758      |  "th' = th" "cs' = cs"
       
  1759 proof(cases "waiting s th' cs'")
       
  1760   case True
       
  1761   have "th' \<noteq> th"
       
  1762   proof
       
  1763     assume otherwise: "th' = th"
       
  1764     from True[unfolded this]
       
  1765     show False by (simp add: th_not_waiting)
       
  1766   qed
       
  1767   from that(1)[OF this True] show ?thesis .
       
  1768 next
       
  1769   case False
       
  1770   hence "th' = th \<and> cs' = cs"
       
  1771       by (metis assms cs_waiting_def holder_def list.sel(1) rotate1.simps(2) 
       
  1772         set_ConsD set_rotate1 waiting_eq wq_es_cs wq_es_cs' wq_neq_simp)
       
  1773   with that(2) show ?thesis by metis
       
  1774 qed
       
  1775 
       
  1776 lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Th th, Cs cs)}" (is "?L = ?R")
       
  1777 proof(rule rel_eqI)
       
  1778   fix n1 n2
       
  1779   assume "(n1, n2) \<in> ?L"
       
  1780   thus "(n1, n2) \<in> ?R" 
       
  1781   proof(cases rule:in_RAG_E)
       
  1782     case (waiting th' cs')
       
  1783     from this(3)
       
  1784     show ?thesis
       
  1785     proof(cases rule:waiting_esE)
       
  1786       case 1
       
  1787       thus ?thesis using waiting(1,2)
       
  1788         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1789     next
       
  1790       case 2
       
  1791       thus ?thesis using waiting(1,2) by auto
       
  1792     qed
       
  1793   next
       
  1794     case (holding th' cs')
       
  1795     from this(3)
       
  1796     show ?thesis
       
  1797     proof(cases rule:holding_esE)
       
  1798       case 1
       
  1799       with holding(1,2)
       
  1800       show ?thesis by (unfold s_RAG_def, fold holding_eq, auto)
       
  1801     qed
       
  1802   qed
       
  1803 next
       
  1804   fix n1 n2
       
  1805   assume "(n1, n2) \<in> ?R"
       
  1806   hence "(n1, n2) \<in> RAG s \<or> (n1 = Th th \<and> n2 = Cs cs)" by auto
       
  1807   thus "(n1, n2) \<in> ?L"
       
  1808   proof
       
  1809     assume "(n1, n2) \<in> RAG s"
       
  1810     thus ?thesis
       
  1811     proof(cases rule:in_RAG_E)
       
  1812       case (waiting th' cs')
       
  1813       from waiting_kept[OF this(3)]
       
  1814       show ?thesis using waiting(1,2)
       
  1815          by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1816     next
       
  1817       case (holding th' cs')
       
  1818       from holding_kept[OF this(3)]
       
  1819       show ?thesis using holding(1,2)
       
  1820          by (unfold s_RAG_def, fold holding_eq, auto)
       
  1821     qed
       
  1822   next
       
  1823     assume "n1 = Th th \<and> n2 = Cs cs"
       
  1824     thus ?thesis using RAG_edge by auto
       
  1825   qed
       
  1826 qed
       
  1827 
       
  1828 end
       
  1829 
  1818 context valid_trace_p
  1830 context valid_trace_p
  1819 begin
  1831 begin
  1820 
  1832 
  1821 lemma RAG_es': "RAG (e # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
  1833 lemma RAG_es: "RAG (e # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
  1822                                                   else RAG s \<union> {(Th th, Cs cs)})"
  1834                                                   else RAG s \<union> {(Th th, Cs cs)})"
  1823 proof(cases "wq s cs = []")
  1835 proof(cases "wq s cs = []")
  1824   case True
  1836   case True
  1825   interpret vt_p: valid_trace_p_h using True
  1837   interpret vt_p: valid_trace_p_h using True
  1826     by (unfold_locales, simp)
  1838     by (unfold_locales, simp)
  1832   show ?thesis by (simp add: vt_p.RAG_es vt_p.wne) 
  1844   show ?thesis by (simp add: vt_p.RAG_es vt_p.wne) 
  1833 qed
  1845 qed
  1834 
  1846 
  1835 end
  1847 end
  1836 
  1848 
       
  1849 section {* Finiteness of RAG *}
       
  1850 
       
  1851 context valid_trace
       
  1852 begin
       
  1853 
       
  1854 lemma finite_RAG:
       
  1855   shows "finite (RAG s)"
       
  1856 proof(induct rule:ind)
       
  1857   case Nil
       
  1858   show ?case 
       
  1859   by (auto simp: s_RAG_def cs_waiting_def 
       
  1860                    cs_holding_def wq_def acyclic_def)
       
  1861 next
       
  1862   case (Cons s e)
       
  1863   interpret vt_e: valid_trace_e s e using Cons by simp
       
  1864   show ?case
       
  1865   proof(cases e)
       
  1866     case (Create th prio)
       
  1867     interpret vt: valid_trace_create s e th prio using Create
       
  1868       by (unfold_locales, simp)
       
  1869     show ?thesis using Cons by simp
       
  1870   next
       
  1871     case (Exit th)
       
  1872     interpret vt: valid_trace_exit s e th using Exit
       
  1873       by (unfold_locales, simp)
       
  1874     show ?thesis using Cons by simp
       
  1875   next
       
  1876     case (P th cs)
       
  1877     interpret vt: valid_trace_p s e th cs using P
       
  1878       by (unfold_locales, simp)
       
  1879     show ?thesis using Cons using vt.RAG_es by auto 
       
  1880   next
       
  1881     case (V th cs)
       
  1882     interpret vt: valid_trace_v s e th cs using V
       
  1883       by (unfold_locales, simp)
       
  1884     show ?thesis using Cons by (simp add: vt.finite_RAG_kept) 
       
  1885   next
       
  1886     case (Set th prio)
       
  1887     interpret vt: valid_trace_set s e th prio using Set
       
  1888       by (unfold_locales, simp)
       
  1889     show ?thesis using Cons by simp
       
  1890   qed
       
  1891 qed
       
  1892 end
       
  1893 
       
  1894 section {* RAG is acyclic *}
       
  1895 
       
  1896 text {* (* ddd *)
       
  1897   The nature of the work is like this: since it starts from a very simple and basic 
       
  1898   model, even intuitively very `basic` and `obvious` properties need to derived from scratch.
       
  1899   For instance, the fact 
       
  1900   that one thread can not be blocked by two critical resources at the same time
       
  1901   is obvious, because only running threads can make new requests, if one is waiting for 
       
  1902   a critical resource and get blocked, it can not make another resource request and get 
       
  1903   blocked the second time (because it is not running). 
       
  1904 
       
  1905   To derive this fact, one needs to prove by contraction and 
       
  1906   reason about time (or @{text "moement"}). The reasoning is based on a generic theorem
       
  1907   named @{text "p_split"}, which is about status changing along the time axis. It says if 
       
  1908   a condition @{text "Q"} is @{text "True"} at a state @{text "s"},
       
  1909   but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} 
       
  1910   in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history 
       
  1911   of events leading to it), such that @{text "Q"} switched 
       
  1912   from being @{text "False"} to @{text "True"} and kept being @{text "True"}
       
  1913   till the last moment of @{text "s"}.
       
  1914 
       
  1915   Suppose a thread @{text "th"} is blocked
       
  1916   on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, 
       
  1917   since no thread is blocked at the very beginning, by applying 
       
  1918   @{text "p_split"} to these two blocking facts, there exist 
       
  1919   two moments @{text "t1"} and @{text "t2"}  in @{text "s"}, such that 
       
  1920   @{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} 
       
  1921   and kept on blocked on them respectively ever since.
       
  1922  
       
  1923   Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}.
       
  1924   However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
       
  1925   in blocked state at moment @{text "t2"} and could not
       
  1926   make any request and get blocked the second time: Contradiction.
       
  1927 *}
       
  1928 
       
  1929 
       
  1930 context valid_trace
       
  1931 begin
       
  1932 
       
  1933 lemma waiting_unique_pre: (* ddd *)
       
  1934   assumes h11: "thread \<in> set (wq s cs1)"
       
  1935   and h12: "thread \<noteq> hd (wq s cs1)"
       
  1936   assumes h21: "thread \<in> set (wq s cs2)"
       
  1937   and h22: "thread \<noteq> hd (wq s cs2)"
       
  1938   and neq12: "cs1 \<noteq> cs2"
       
  1939   shows "False"
       
  1940 proof -
       
  1941   let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
       
  1942   from h11 and h12 have q1: "?Q cs1 s" by simp
       
  1943   from h21 and h22 have q2: "?Q cs2 s" by simp
       
  1944   have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
       
  1945   have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
       
  1946   from p_split [of "?Q cs1", OF q1 nq1]
       
  1947   obtain t1 where lt1: "t1 < length s"
       
  1948     and np1: "\<not> ?Q cs1 (moment t1 s)"
       
  1949     and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto
       
  1950   from p_split [of "?Q cs2", OF q2 nq2]
       
  1951   obtain t2 where lt2: "t2 < length s"
       
  1952     and np2: "\<not> ?Q cs2 (moment t2 s)"
       
  1953     and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto
       
  1954   { fix s cs
       
  1955     assume q: "?Q cs s"
       
  1956     have "thread \<notin> runing s"
       
  1957     proof
       
  1958       assume "thread \<in> runing s"
       
  1959       hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> 
       
  1960                  thread \<noteq> hd (wq_fun (schs s) cs))"
       
  1961         by (unfold runing_def s_waiting_def readys_def, auto)
       
  1962       from this[rule_format, of cs] q 
       
  1963       show False by (simp add: wq_def) 
       
  1964     qed
       
  1965   } note q_not_runing = this
       
  1966   { fix t1 t2 cs1 cs2
       
  1967     assume  lt1: "t1 < length s"
       
  1968     and np1: "\<not> ?Q cs1 (moment t1 s)"
       
  1969     and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))"
       
  1970     and lt2: "t2 < length s"
       
  1971     and np2: "\<not> ?Q cs2 (moment t2 s)"
       
  1972     and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))"
       
  1973     and lt12: "t1 < t2"
       
  1974     let ?t3 = "Suc t2" 
       
  1975     interpret ve2: valid_moment_e _ t2 using lt2
       
  1976      by (unfold_locales, simp)
       
  1977     let ?e = ve2.next_e
       
  1978     have "t2 < ?t3" by simp
       
  1979     from nn2 [rule_format, OF this] and ve2.trace_e
       
  1980     have h1: "thread \<in> set (wq (?e#moment t2 s) cs2)" and
       
  1981          h2: "thread \<noteq> hd (wq (?e#moment t2 s) cs2)" by auto
       
  1982     have ?thesis
       
  1983     proof -
       
  1984       have "thread \<in> runing (moment t2 s)"
       
  1985       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
  1986         case True
       
  1987         have "?e = V thread cs2"
       
  1988         proof -
       
  1989           have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
       
  1990               using True and np2  by auto 
       
  1991           thus ?thesis
       
  1992             using True h2 ve2.vat_moment_e.wq_out_inv by blast 
       
  1993         qed
       
  1994         thus ?thesis
       
  1995           using step.cases ve2.vat_moment_e.pip_e by auto 
       
  1996       next
       
  1997         case False
       
  1998         hence "?e = P thread cs2"
       
  1999           using h1 ve2.vat_moment_e.wq_in_inv by blast 
       
  2000         thus ?thesis
       
  2001           using step.cases ve2.vat_moment_e.pip_e by auto 
       
  2002       qed
       
  2003       moreover have "thread \<notin> runing (moment t2 s)"
       
  2004         by (rule q_not_runing[OF nn1[rule_format, OF lt12]])
       
  2005       ultimately show ?thesis by simp
       
  2006     qed
       
  2007   } note lt_case = this
       
  2008   show ?thesis
       
  2009   proof -
       
  2010     { assume "t1 < t2"
       
  2011       from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this]
       
  2012       have ?thesis .
       
  2013     } moreover {
       
  2014       assume "t2 < t1"
       
  2015       from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this]
       
  2016       have ?thesis .
       
  2017     } moreover { 
       
  2018       assume eq_12: "t1 = t2"
       
  2019       let ?t3 = "Suc t2"
       
  2020       interpret ve2: valid_moment_e _ t2 using lt2
       
  2021         by (unfold_locales, simp)
       
  2022       let ?e = ve2.next_e
       
  2023       have "t2 < ?t3" by simp
       
  2024       from nn2 [rule_format, OF this] and ve2.trace_e
       
  2025       have h1: "thread \<in> set (wq (?e#moment t2 s) cs2)" by auto
       
  2026       have lt_2: "t2 < ?t3" by simp
       
  2027       from nn2 [rule_format, OF this] and ve2.trace_e
       
  2028       have h1: "thread \<in> set (wq (?e#moment t2 s) cs2)" and
       
  2029            h2: "thread \<noteq> hd (wq (?e#moment t2 s) cs2)" by auto
       
  2030       from nn1[rule_format, OF lt_2[folded eq_12], unfolded ve2.trace_e[folded eq_12]] 
       
  2031            eq_12[symmetric]
       
  2032       have g1: "thread \<in> set (wq (?e#moment t1 s) cs1)" and
       
  2033            g2: "thread \<noteq> hd (wq (?e#moment t1 s) cs1)" by auto
       
  2034       have "?e = V thread cs2 \<or> ?e = P thread cs2"
       
  2035         using h1 h2 np2 ve2.vat_moment_e.wq_in_inv 
       
  2036               ve2.vat_moment_e.wq_out_inv by blast
       
  2037       moreover have "?e = V thread cs1 \<or> ?e = P thread cs1"
       
  2038         using eq_12 g1 g2 np1 ve2.vat_moment_e.wq_in_inv 
       
  2039               ve2.vat_moment_e.wq_out_inv by blast
       
  2040       ultimately have ?thesis using neq12 by auto
       
  2041     } ultimately show ?thesis using nat_neq_iff by blast 
       
  2042   qed
       
  2043 qed
       
  2044 
       
  2045 text {*
       
  2046   This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
       
  2047 *}
       
  2048 
       
  2049 lemma waiting_unique:
       
  2050   assumes "waiting s th cs1"
       
  2051   and "waiting s th cs2"
       
  2052   shows "cs1 = cs2"
       
  2053   using waiting_unique_pre assms
       
  2054   unfolding wq_def s_waiting_def
       
  2055   by auto
       
  2056 
       
  2057 end
       
  2058 
       
  2059 lemma (in valid_trace_v)
       
  2060   preced_es [simp]: "preced th (e#s) = preced th s"
       
  2061   by (unfold is_v preced_def, simp)
       
  2062 
       
  2063 lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s"
       
  2064 proof
       
  2065   fix th'
       
  2066   show "the_preced (V th cs # s) th' = the_preced s th'"
       
  2067     by (unfold the_preced_def preced_def, simp)
       
  2068 qed
       
  2069 
       
  2070 
       
  2071 lemma (in valid_trace_v)
       
  2072   the_preced_es: "the_preced (e#s) = the_preced s"
       
  2073   by (unfold is_v preced_def, simp)
       
  2074 
       
  2075 context valid_trace_p
       
  2076 begin
       
  2077 
       
  2078 lemma not_holding_s_th_cs: "\<not> holding s th cs"
       
  2079 proof
       
  2080   assume otherwise: "holding s th cs"
       
  2081   from pip_e[unfolded is_p]
       
  2082   show False
       
  2083   proof(cases)
       
  2084     case (thread_P)
       
  2085     moreover have "(Cs cs, Th th) \<in> RAG s"
       
  2086       using otherwise cs_holding_def 
       
  2087             holding_eq th_not_in_wq by auto
       
  2088     ultimately show ?thesis by auto
       
  2089   qed
       
  2090 qed
       
  2091 
       
  2092 end
       
  2093 
       
  2094 
  1837 lemma (in valid_trace_v_n) finite_waiting_set:
  2095 lemma (in valid_trace_v_n) finite_waiting_set:
  1838   "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
  2096   "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
  1839     by (simp add: waiting_set_eq)
  2097     by (simp add: waiting_set_eq)
  1840 
  2098 
  1841 lemma (in valid_trace_v_n) finite_holding_set:
  2099 lemma (in valid_trace_v_n) finite_holding_set:
  1848 
  2106 
  1849 lemma (in valid_trace_v_e) finite_holding_set:
  2107 lemma (in valid_trace_v_e) finite_holding_set:
  1850   "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
  2108   "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
  1851     by (simp add: holding_set_eq)
  2109     by (simp add: holding_set_eq)
  1852 
  2110 
  1853 context valid_trace_v
       
  1854 begin
       
  1855 
       
  1856 lemma 
       
  1857   finite_RAG_kept:
       
  1858   assumes "finite (RAG s)"
       
  1859   shows "finite (RAG (e#s))"
       
  1860 proof(cases "rest = []")
       
  1861   case True
       
  1862   interpret vt: valid_trace_v_e using True
       
  1863     by (unfold_locales, simp)
       
  1864   show ?thesis using assms
       
  1865     by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
       
  1866 next
       
  1867   case False
       
  1868   interpret vt: valid_trace_v_n using False
       
  1869     by (unfold_locales, simp)
       
  1870   show ?thesis using assms
       
  1871     by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
       
  1872 qed
       
  1873 
       
  1874 end
       
  1875 
  2111 
  1876 context valid_trace_v_e
  2112 context valid_trace_v_e
  1877 begin 
  2113 begin 
  1878 
  2114 
  1879 lemma 
  2115 lemma 
  1914       obtain cs' where h: "(Th taker, Cs cs') \<in> ?A" 
  2150       obtain cs' where h: "(Th taker, Cs cs') \<in> ?A" 
  1915                           "(Th taker, Cs cs') \<in> RAG s"
  2151                           "(Th taker, Cs cs') \<in> RAG s"
  1916         by (unfold s_RAG_def, auto)
  2152         by (unfold s_RAG_def, auto)
  1917       from this(2) have "waiting s taker cs'" 
  2153       from this(2) have "waiting s taker cs'" 
  1918         by (unfold s_RAG_def, fold waiting_eq, auto)
  2154         by (unfold s_RAG_def, fold waiting_eq, auto)
  1919       from waiting_unique[OF this waiting_taker]
  2155       from waiting_unique[OF this waiting_taker] 
  1920       have "cs' = cs" .
  2156       have "cs' = cs" .
  1921       from h(1)[unfolded this] show False by auto
  2157       from h(1)[unfolded this] show False by auto
  1922     qed
  2158     qed
  1923     ultimately show ?thesis by auto
  2159     ultimately show ?thesis by auto
  1924   qed
  2160   qed
  1949       from tranclD[OF this]
  2185       from tranclD[OF this]
  1950       obtain cs' where h: "(Th th, Cs cs') \<in> RAG s"
  2186       obtain cs' where h: "(Th th, Cs cs') \<in> RAG s"
  1951         by (unfold s_RAG_def, auto)
  2187         by (unfold s_RAG_def, auto)
  1952       hence "waiting s th cs'" 
  2188       hence "waiting s th cs'" 
  1953         by (unfold s_RAG_def, fold waiting_eq, auto)
  2189         by (unfold s_RAG_def, fold waiting_eq, auto)
  1954       with th_not_waiting show False by auto
  2190       with th_not_waiting show False by auto 
  1955     qed
  2191     qed
  1956     ultimately show ?thesis by auto
  2192     ultimately show ?thesis by auto
  1957   qed
  2193   qed
  1958   thus ?thesis by (unfold RAG_es, simp)
  2194   thus ?thesis by (unfold RAG_es, simp)
  1959 qed
  2195 qed
  1993 end
  2229 end
  1994 
  2230 
  1995 context valid_trace
  2231 context valid_trace
  1996 begin
  2232 begin
  1997 
  2233 
  1998 lemma finite_RAG:
       
  1999   shows "finite (RAG s)"
       
  2000 proof(induct rule:ind)
       
  2001   case Nil
       
  2002   show ?case 
       
  2003   by (auto simp: s_RAG_def cs_waiting_def 
       
  2004                    cs_holding_def wq_def acyclic_def)
       
  2005 next
       
  2006   case (Cons s e)
       
  2007   interpret vt_e: valid_trace_e s e using Cons by simp
       
  2008   show ?case
       
  2009   proof(cases e)
       
  2010     case (Create th prio)
       
  2011     interpret vt: valid_trace_create s e th prio using Create
       
  2012       by (unfold_locales, simp)
       
  2013     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
       
  2014   next
       
  2015     case (Exit th)
       
  2016     interpret vt: valid_trace_exit s e th using Exit
       
  2017       by (unfold_locales, simp)
       
  2018     show ?thesis using Cons by (simp add: vt.RAG_unchanged)
       
  2019   next
       
  2020     case (P th cs)
       
  2021     interpret vt: valid_trace_p s e th cs using P
       
  2022       by (unfold_locales, simp)
       
  2023     show ?thesis using Cons using vt.RAG_es' by auto 
       
  2024   next
       
  2025     case (V th cs)
       
  2026     interpret vt: valid_trace_v s e th cs using V
       
  2027       by (unfold_locales, simp)
       
  2028     show ?thesis using Cons by (simp add: vt.finite_RAG_kept) 
       
  2029   next
       
  2030     case (Set th prio)
       
  2031     interpret vt: valid_trace_set s e th prio using Set
       
  2032       by (unfold_locales, simp)
       
  2033     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
       
  2034   qed
       
  2035 qed
       
  2036 
       
  2037 lemma acyclic_RAG:
  2234 lemma acyclic_RAG:
  2038   shows "acyclic (RAG s)"
  2235   shows "acyclic (RAG s)"
  2039 proof(induct rule:ind)
  2236 proof(induct rule:ind)
  2040   case Nil
  2237   case Nil
  2041   show ?case 
  2238   show ?case 
  2047   show ?case
  2244   show ?case
  2048   proof(cases e)
  2245   proof(cases e)
  2049     case (Create th prio)
  2246     case (Create th prio)
  2050     interpret vt: valid_trace_create s e th prio using Create
  2247     interpret vt: valid_trace_create s e th prio using Create
  2051       by (unfold_locales, simp)
  2248       by (unfold_locales, simp)
  2052     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
  2249     show ?thesis using Cons by simp 
  2053   next
  2250   next
  2054     case (Exit th)
  2251     case (Exit th)
  2055     interpret vt: valid_trace_exit s e th using Exit
  2252     interpret vt: valid_trace_exit s e th using Exit
  2056       by (unfold_locales, simp)
  2253       by (unfold_locales, simp)
  2057     show ?thesis using Cons by (simp add: vt.RAG_unchanged)
  2254     show ?thesis using Cons by simp
  2058   next
  2255   next
  2059     case (P th cs)
  2256     case (P th cs)
  2060     interpret vt: valid_trace_p s e th cs using P
  2257     interpret vt: valid_trace_p s e th cs using P
  2061       by (unfold_locales, simp)
  2258       by (unfold_locales, simp)
  2062     show ?thesis
  2259     show ?thesis
  2089     qed
  2286     qed
  2090   next
  2287   next
  2091     case (Set th prio)
  2288     case (Set th prio)
  2092     interpret vt: valid_trace_set s e th prio using Set
  2289     interpret vt: valid_trace_set s e th prio using Set
  2093       by (unfold_locales, simp)
  2290       by (unfold_locales, simp)
  2094     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
  2291     show ?thesis using Cons by simp 
  2095   qed
  2292   qed
  2096 qed
  2293 qed
       
  2294 
       
  2295 end
       
  2296 
       
  2297 section {* RAG is single-valued *}
       
  2298 
       
  2299 context valid_trace
       
  2300 begin
       
  2301 
       
  2302 lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
       
  2303   apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
       
  2304   by(auto elim:waiting_unique held_unique)
       
  2305 
       
  2306 lemma sgv_RAG: "single_valued (RAG s)"
       
  2307   using unique_RAG by (auto simp:single_valued_def)
       
  2308 
       
  2309 end
       
  2310 
       
  2311 section {* RAG is well-founded *}
       
  2312 
       
  2313 context valid_trace
       
  2314 begin
  2097 
  2315 
  2098 lemma wf_RAG: "wf (RAG s)"
  2316 lemma wf_RAG: "wf (RAG s)"
  2099 proof(rule finite_acyclic_wf)
  2317 proof(rule finite_acyclic_wf)
  2100   from finite_RAG show "finite (RAG s)" .
  2318   from finite_RAG show "finite (RAG s)" .
  2101 next
  2319 next
  2102   from acyclic_RAG show "acyclic (RAG s)" .
  2320   from acyclic_RAG show "acyclic (RAG s)" .
  2103 qed
  2321 qed
  2104 
  2322 
  2105 lemma sgv_wRAG: "single_valued (wRAG s)"
  2323 lemma wf_RAG_converse: 
  2106   using waiting_unique
  2324   shows "wf ((RAG s)^-1)"
  2107   by (unfold single_valued_def wRAG_def, auto)
  2325 proof(rule finite_acyclic_wf_converse)
  2108 
  2326   from finite_RAG 
  2109 lemma sgv_hRAG: "single_valued (hRAG s)"
  2327   show "finite (RAG s)" .
  2110   using held_unique 
  2328 next
  2111   by (unfold single_valued_def hRAG_def, auto)
  2329   from acyclic_RAG
  2112 
  2330   show "acyclic (RAG s)" .
  2113 lemma sgv_tRAG: "single_valued (tRAG s)"
  2331 qed
  2114   by (unfold tRAG_def, rule single_valued_relcomp, 
  2332 
  2115               insert sgv_wRAG sgv_hRAG, auto)
  2333 end
       
  2334 
       
  2335 section {* RAG forms a forest (or tree) *}
       
  2336 
       
  2337 context valid_trace
       
  2338 begin
       
  2339 
       
  2340 lemma rtree_RAG: "rtree (RAG s)"
       
  2341   using sgv_RAG acyclic_RAG
       
  2342   by (unfold rtree_def rtree_axioms_def sgv_def, auto)
       
  2343 
       
  2344 end
       
  2345 
       
  2346 sublocale valid_trace < rtree_RAG: rtree "RAG s"
       
  2347   using rtree_RAG .
       
  2348 
       
  2349 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
       
  2350 proof -
       
  2351   show "fsubtree (RAG s)"
       
  2352   proof(intro_locales)
       
  2353     show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
       
  2354   next
       
  2355     show "fsubtree_axioms (RAG s)"
       
  2356     proof(unfold fsubtree_axioms_def)
       
  2357       from wf_RAG show "wf (RAG s)" .
       
  2358     qed
       
  2359   qed
       
  2360 qed
       
  2361 
       
  2362 
       
  2363 section {* Derived properties for parts of RAG *}
       
  2364 
       
  2365 context valid_trace
       
  2366 begin
  2116 
  2367 
  2117 lemma acyclic_tRAG: "acyclic (tRAG s)"
  2368 lemma acyclic_tRAG: "acyclic (tRAG s)"
  2118 proof(unfold tRAG_def, rule acyclic_compose)
  2369 proof(unfold tRAG_def, rule acyclic_compose)
  2119   show "acyclic (RAG s)" using acyclic_RAG .
  2370   show "acyclic (RAG s)" using acyclic_RAG .
  2120 next
  2371 next
  2121   show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
  2372   show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
  2122 next
  2373 next
  2123   show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
  2374   show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
  2124 qed
  2375 qed
  2125 
  2376 
  2126 lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
  2377 lemma sgv_wRAG: "single_valued (wRAG s)"
  2127   apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
  2378   using waiting_unique
  2128   by(auto elim:waiting_unique held_unique)
  2379   by (unfold single_valued_def wRAG_def, auto)
  2129 
  2380 
  2130 lemma sgv_RAG: "single_valued (RAG s)"
  2381 lemma sgv_hRAG: "single_valued (hRAG s)"
  2131   using unique_RAG by (auto simp:single_valued_def)
  2382   using held_unique 
  2132 
  2383   by (unfold single_valued_def hRAG_def, auto)
  2133 lemma rtree_RAG: "rtree (RAG s)"
  2384 
  2134   using sgv_RAG acyclic_RAG
  2385 lemma sgv_tRAG: "single_valued (tRAG s)"
  2135   by (unfold rtree_def rtree_axioms_def sgv_def, auto)
  2386   by (unfold tRAG_def, rule single_valued_relcomp, 
  2136 
  2387               insert sgv_wRAG sgv_hRAG, auto)
  2137 end
  2388 
  2138 
  2389 end
  2139 sublocale valid_trace < rtree_RAG: rtree "RAG s"
       
  2140 proof
       
  2141   show "single_valued (RAG s)"
       
  2142   apply (intro_locales)
       
  2143     by (unfold single_valued_def, 
       
  2144         auto intro:unique_RAG)
       
  2145 
       
  2146   show "acyclic (RAG s)"
       
  2147      by (rule acyclic_RAG)
       
  2148 qed
       
  2149 
  2390 
  2150 sublocale valid_trace < rtree_s: rtree "tRAG s"
  2391 sublocale valid_trace < rtree_s: rtree "tRAG s"
  2151 proof(unfold_locales)
  2392 proof(unfold_locales)
  2152   from sgv_tRAG show "single_valued (tRAG s)" .
  2393   from sgv_tRAG show "single_valued (tRAG s)" .
  2153 next
  2394 next
  2154   from acyclic_tRAG show "acyclic (tRAG s)" .
  2395   from acyclic_tRAG show "acyclic (tRAG s)" .
  2155 qed
  2396 qed
  2156 
       
  2157 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
       
  2158 proof -
       
  2159   show "fsubtree (RAG s)"
       
  2160   proof(intro_locales)
       
  2161     show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
       
  2162   next
       
  2163     show "fsubtree_axioms (RAG s)"
       
  2164     proof(unfold fsubtree_axioms_def)
       
  2165       from wf_RAG show "wf (RAG s)" .
       
  2166     qed
       
  2167   qed
       
  2168 qed
       
  2169 
       
  2170 lemma tRAG_alt_def: 
       
  2171   "tRAG s = {(Th th1, Th th2) | th1 th2. 
       
  2172                   \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
       
  2173  by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
       
  2174 
  2397 
  2175 sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
  2398 sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
  2176 proof -
  2399 proof -
  2177   have "fsubtree (tRAG s)"
  2400   have "fsubtree (tRAG s)"
  2178   proof -
  2401   proof -
  2202       by (unfold fsubtree_def fsubtree_axioms_def,auto)
  2425       by (unfold fsubtree_def fsubtree_axioms_def,auto)
  2203   qed
  2426   qed
  2204   from this[folded tRAG_def] show "fsubtree (tRAG s)" .
  2427   from this[folded tRAG_def] show "fsubtree (tRAG s)" .
  2205 qed
  2428 qed
  2206 
  2429 
  2207 
       
  2208 context valid_trace
       
  2209 begin
       
  2210 
       
  2211 lemma finite_subtree_threads:
       
  2212     "finite {th'. Th th' \<in> subtree (RAG s) (Th th)}" (is "finite ?A")
       
  2213 proof -
       
  2214   have "?A = the_thread ` {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
       
  2215         by (auto, insert image_iff, fastforce)
       
  2216   moreover have "finite {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
       
  2217         (is "finite ?B")
       
  2218   proof -
       
  2219      have "?B = (subtree (RAG s) (Th th)) \<inter> {Th th' | th'. True}"
       
  2220       by auto
       
  2221      moreover have "... \<subseteq> (subtree (RAG s) (Th th))" by auto
       
  2222      moreover have "finite ..." by (simp add: fsbtRAGs.finite_subtree) 
       
  2223      ultimately show ?thesis by auto
       
  2224   qed
       
  2225   ultimately show ?thesis by auto
       
  2226 qed
       
  2227 
       
  2228 lemma le_cp:
       
  2229   shows "preced th s \<le> cp s th"
       
  2230   proof(unfold cp_alt_def, rule Max_ge)
       
  2231     show "finite (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})"
       
  2232       by (simp add: finite_subtree_threads)
       
  2233   next
       
  2234     show "preced th s \<in> the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}"
       
  2235       by (simp add: subtree_def the_preced_def)   
       
  2236   qed
       
  2237 
       
  2238 lemma cp_le:
       
  2239   assumes th_in: "th \<in> threads s"
       
  2240   shows "cp s th \<le> Max (the_preced s ` threads s)"
       
  2241 proof(unfold cp_alt_def, rule Max_f_mono)
       
  2242   show "finite (threads s)" by (simp add: finite_threads) 
       
  2243 next
       
  2244   show " {th'. Th th' \<in> subtree (RAG s) (Th th)} \<noteq> {}"
       
  2245     using subtree_def by fastforce
       
  2246 next
       
  2247   show "{th'. Th th' \<in> subtree (RAG s) (Th th)} \<subseteq> threads s"
       
  2248     using assms
       
  2249     by (smt Domain.DomainI dm_RAG_threads mem_Collect_eq 
       
  2250         node.inject(1) rtranclD subsetI subtree_def trancl_domain) 
       
  2251 qed
       
  2252 
       
  2253 lemma max_cp_eq: 
       
  2254   shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)"
       
  2255   (is "?L = ?R")
       
  2256 proof -
       
  2257   have "?L \<le> ?R" 
       
  2258   proof(cases "threads s = {}")
       
  2259     case False
       
  2260     show ?thesis 
       
  2261       by (rule Max.boundedI, 
       
  2262           insert cp_le, 
       
  2263           auto simp:finite_threads False)
       
  2264   qed auto
       
  2265   moreover have "?R \<le> ?L"
       
  2266     by (rule Max_fg_mono, 
       
  2267         simp add: finite_threads,
       
  2268         simp add: le_cp the_preced_def)
       
  2269   ultimately show ?thesis by auto
       
  2270 qed
       
  2271 
       
  2272 lemma wf_RAG_converse: 
       
  2273   shows "wf ((RAG s)^-1)"
       
  2274 proof(rule finite_acyclic_wf_converse)
       
  2275   from finite_RAG 
       
  2276   show "finite (RAG s)" .
       
  2277 next
       
  2278   from acyclic_RAG
       
  2279   show "acyclic (RAG s)" .
       
  2280 qed
       
  2281 
       
  2282 lemma chain_building:
       
  2283   assumes "node \<in> Domain (RAG s)"
       
  2284   obtains th' where "th' \<in> readys s" "(node, Th th') \<in> (RAG s)^+"
       
  2285 proof -
       
  2286   from assms have "node \<in> Range ((RAG s)^-1)" by auto
       
  2287   from wf_base[OF wf_RAG_converse this]
       
  2288   obtain b where h_b: "(b, node) \<in> ((RAG s)\<inverse>)\<^sup>+" "\<forall>c. (c, b) \<notin> (RAG s)\<inverse>" by auto
       
  2289   obtain th' where eq_b: "b = Th th'"
       
  2290   proof(cases b)
       
  2291     case (Cs cs)
       
  2292     from h_b(1)[unfolded trancl_converse] 
       
  2293     have "(node, b) \<in> ((RAG s)\<^sup>+)" by auto
       
  2294     from tranclE[OF this]
       
  2295     obtain n where "(n, b) \<in> RAG s" by auto
       
  2296     from this[unfolded Cs]
       
  2297     obtain th1 where "waiting s th1 cs"
       
  2298       by (unfold s_RAG_def, fold waiting_eq, auto)
       
  2299     from waiting_holding[OF this]
       
  2300     obtain th2 where "holding s th2 cs" .
       
  2301     hence "(Cs cs, Th th2) \<in> RAG s"
       
  2302       by (unfold s_RAG_def, fold holding_eq, auto)
       
  2303     with h_b(2)[unfolded Cs, rule_format]
       
  2304     have False by auto
       
  2305     thus ?thesis by auto
       
  2306   qed auto
       
  2307   have "th' \<in> readys s" 
       
  2308   proof -
       
  2309     from h_b(2)[unfolded eq_b]
       
  2310     have "\<forall>cs. \<not> waiting s th' cs"
       
  2311       by (unfold s_RAG_def, fold waiting_eq, auto)
       
  2312     moreover have "th' \<in> threads s"
       
  2313     proof(rule rg_RAG_threads)
       
  2314       from tranclD[OF h_b(1), unfolded eq_b]
       
  2315       obtain z where "(z, Th th') \<in> (RAG s)" by auto
       
  2316       thus "Th th' \<in> Range (RAG s)" by auto
       
  2317     qed
       
  2318     ultimately show ?thesis by (auto simp:readys_def)
       
  2319   qed
       
  2320   moreover have "(node, Th th') \<in> (RAG s)^+" 
       
  2321     using h_b(1)[unfolded trancl_converse] eq_b by auto
       
  2322   ultimately show ?thesis using that by metis
       
  2323 qed
       
  2324 
       
  2325 text {* \noindent
       
  2326   The following is just an instance of @{text "chain_building"}.
       
  2327 *}
       
  2328 lemma th_chain_to_ready:
       
  2329   assumes th_in: "th \<in> threads s"
       
  2330   shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
       
  2331 proof(cases "th \<in> readys s")
       
  2332   case True
       
  2333   thus ?thesis by auto
       
  2334 next
       
  2335   case False
       
  2336   from False and th_in have "Th th \<in> Domain (RAG s)" 
       
  2337     by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
       
  2338   from chain_building [rule_format, OF this]
       
  2339   show ?thesis by auto
       
  2340 qed
       
  2341 
       
  2342 end
       
  2343 
       
  2344 lemma count_rec1 [simp]: 
       
  2345   assumes "Q e"
       
  2346   shows "count Q (e#es) = Suc (count Q es)"
       
  2347   using assms
       
  2348   by (unfold count_def, auto)
       
  2349 
       
  2350 lemma count_rec2 [simp]: 
       
  2351   assumes "\<not>Q e"
       
  2352   shows "count Q (e#es) = (count Q es)"
       
  2353   using assms
       
  2354   by (unfold count_def, auto)
       
  2355 
       
  2356 lemma count_rec3 [simp]: 
       
  2357   shows "count Q [] =  0"
       
  2358   by (unfold count_def, auto)
       
  2359 
       
  2360 lemma cntP_simp1[simp]:
       
  2361   "cntP (P th cs'#s) th = cntP s th + 1"
       
  2362   by (unfold cntP_def, simp)
       
  2363 
       
  2364 lemma cntP_simp2[simp]:
       
  2365   assumes "th' \<noteq> th"
       
  2366   shows "cntP (P th cs'#s) th' = cntP s th'"
       
  2367   using assms
       
  2368   by (unfold cntP_def, simp)
       
  2369 
       
  2370 lemma cntP_simp3[simp]:
       
  2371   assumes "\<not> isP e"
       
  2372   shows "cntP (e#s) th' = cntP s th'"
       
  2373   using assms
       
  2374   by (unfold cntP_def, cases e, simp+)
       
  2375 
       
  2376 lemma cntV_simp1[simp]:
       
  2377   "cntV (V th cs'#s) th = cntV s th + 1"
       
  2378   by (unfold cntV_def, simp)
       
  2379 
       
  2380 lemma cntV_simp2[simp]:
       
  2381   assumes "th' \<noteq> th"
       
  2382   shows "cntV (V th cs'#s) th' = cntV s th'"
       
  2383   using assms
       
  2384   by (unfold cntV_def, simp)
       
  2385 
       
  2386 lemma cntV_simp3[simp]:
       
  2387   assumes "\<not> isV e"
       
  2388   shows "cntV (e#s) th' = cntV s th'"
       
  2389   using assms
       
  2390   by (unfold cntV_def, cases e, simp+)
       
  2391 
       
  2392 lemma cntP_diff_inv:
       
  2393   assumes "cntP (e#s) th \<noteq> cntP s th"
       
  2394   shows "isP e \<and> actor e = th"
       
  2395 proof(cases e)
       
  2396   case (P th' pty)
       
  2397   show ?thesis
       
  2398   by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
       
  2399         insert assms P, auto simp:cntP_def)
       
  2400 qed (insert assms, auto simp:cntP_def)
       
  2401   
       
  2402 lemma cntV_diff_inv:
       
  2403   assumes "cntV (e#s) th \<noteq> cntV s th"
       
  2404   shows "isV e \<and> actor e = th"
       
  2405 proof(cases e)
       
  2406   case (V th' pty)
       
  2407   show ?thesis
       
  2408   by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
       
  2409         insert assms V, auto simp:cntV_def)
       
  2410 qed (insert assms, auto simp:cntV_def)
       
  2411 
       
  2412 lemma children_RAG_alt_def:
       
  2413   "children (RAG (s::state)) (Th th) = Cs ` {cs. holding s th cs}"
       
  2414   by (unfold s_RAG_def, auto simp:children_def holding_eq)
       
  2415 
       
  2416 lemma holdents_alt_def:
       
  2417   "holdents s th = the_cs ` (children (RAG (s::state)) (Th th))"
       
  2418   by (unfold children_RAG_alt_def holdents_def, simp add: image_image)
       
  2419 
       
  2420 lemma cntCS_alt_def:
       
  2421   "cntCS s th = card (children (RAG s) (Th th))"
       
  2422   apply (unfold children_RAG_alt_def cntCS_def holdents_def)
       
  2423   by (rule card_image[symmetric], auto simp:inj_on_def)
       
  2424 
       
  2425 context valid_trace
       
  2426 begin
       
  2427 
       
  2428 lemma finite_holdents: "finite (holdents s th)"
       
  2429   by (unfold holdents_alt_def, insert fsbtRAGs.finite_children, auto)
       
  2430   
       
  2431 end
       
  2432 
       
  2433 context valid_trace_p_w
       
  2434 begin
       
  2435 
       
  2436 lemma holding_s_holder: "holding s holder cs"
       
  2437   by (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
       
  2438 
       
  2439 lemma holding_es_holder: "holding (e#s) holder cs"
       
  2440   by (unfold s_holding_def, fold wq_def, unfold wq_es_cs wq_s_cs, auto)
       
  2441 
       
  2442 lemma holdents_es:
       
  2443   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") 
       
  2444 proof -
       
  2445   { fix cs'
       
  2446     assume "cs' \<in> ?L"
       
  2447     hence h: "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  2448     have "holding s th' cs'"
       
  2449     proof(cases "cs' = cs")
       
  2450       case True
       
  2451       from held_unique[OF h[unfolded True] holding_es_holder]
       
  2452       have "th' = holder" .
       
  2453       thus ?thesis 
       
  2454         by (unfold True holdents_def, insert holding_s_holder, simp)
       
  2455     next
       
  2456       case False
       
  2457       hence "wq (e#s) cs' = wq s cs'" by simp
       
  2458       from h[unfolded s_holding_def, folded wq_def, unfolded this]
       
  2459       show ?thesis
       
  2460        by (unfold s_holding_def, fold wq_def, auto)
       
  2461     qed 
       
  2462     hence "cs' \<in> ?R" by (auto simp:holdents_def)
       
  2463   } moreover {
       
  2464     fix cs'
       
  2465     assume "cs' \<in> ?R"
       
  2466     hence h: "holding s th' cs'" by (auto simp:holdents_def)
       
  2467     have "holding (e#s) th' cs'"
       
  2468     proof(cases "cs' = cs")
       
  2469       case True
       
  2470       from held_unique[OF h[unfolded True] holding_s_holder]
       
  2471       have "th' = holder" .
       
  2472       thus ?thesis 
       
  2473         by (unfold True holdents_def, insert holding_es_holder, simp)
       
  2474     next
       
  2475       case False
       
  2476       hence "wq s cs' = wq (e#s) cs'" by simp
       
  2477       from h[unfolded s_holding_def, folded wq_def, unfolded this]
       
  2478       show ?thesis
       
  2479        by (unfold s_holding_def, fold wq_def, auto)
       
  2480     qed 
       
  2481     hence "cs' \<in> ?L" by (auto simp:holdents_def)
       
  2482   } ultimately show ?thesis by auto
       
  2483 qed
       
  2484 
       
  2485 lemma cntCS_es_th[simp]: "cntCS (e#s) th' = cntCS s th'"
       
  2486  by (unfold cntCS_def holdents_es, simp)
       
  2487 
       
  2488 lemma th_not_ready_es: 
       
  2489   shows "th \<notin> readys (e#s)"
       
  2490   using waiting_es_th_cs 
       
  2491   by (unfold readys_def, auto)
       
  2492 
       
  2493 end
       
  2494   
       
  2495 context valid_trace_p_h
       
  2496 begin
       
  2497 
       
  2498 lemma th_not_waiting':
       
  2499   "\<not> waiting (e#s) th cs'"
       
  2500 proof(cases "cs' = cs")
       
  2501   case True
       
  2502   show ?thesis
       
  2503     by (unfold True s_waiting_def, fold wq_def, unfold wq_es_cs', auto)
       
  2504 next
       
  2505   case False
       
  2506   from th_not_waiting[of cs', unfolded s_waiting_def, folded wq_def]
       
  2507   show ?thesis
       
  2508     by (unfold s_waiting_def, fold wq_def, insert False, simp)
       
  2509 qed
       
  2510 
       
  2511 lemma ready_th_es: 
       
  2512   shows "th \<in> readys (e#s)"
       
  2513   using th_not_waiting'
       
  2514   by (unfold readys_def, insert live_th_es, auto)
       
  2515 
       
  2516 lemma holdents_es_th:
       
  2517   "holdents (e#s) th = (holdents s th) \<union> {cs}" (is "?L = ?R")
       
  2518 proof -
       
  2519   { fix cs'
       
  2520     assume "cs' \<in> ?L" 
       
  2521     hence "holding (e#s) th cs'"
       
  2522       by (unfold holdents_def, auto)
       
  2523     hence "cs' \<in> ?R"
       
  2524      by (cases rule:holding_esE, auto simp:holdents_def)
       
  2525   } moreover {
       
  2526     fix cs'
       
  2527     assume "cs' \<in> ?R"
       
  2528     hence "holding s th cs' \<or> cs' = cs" 
       
  2529       by (auto simp:holdents_def)
       
  2530     hence "cs' \<in> ?L"
       
  2531     proof
       
  2532       assume "holding s th cs'"
       
  2533       from holding_kept[OF this]
       
  2534       show ?thesis by (auto simp:holdents_def)
       
  2535     next
       
  2536       assume "cs' = cs"
       
  2537       thus ?thesis using holding_es_th_cs
       
  2538         by (unfold holdents_def, auto)
       
  2539     qed
       
  2540   } ultimately show ?thesis by auto
       
  2541 qed
       
  2542 
       
  2543 lemma cntCS_es_th: "cntCS (e#s) th = cntCS s th + 1"
       
  2544 proof -
       
  2545   have "card (holdents s th \<union> {cs}) = card (holdents s th) + 1"
       
  2546   proof(subst card_Un_disjoint)
       
  2547     show "holdents s th \<inter> {cs} = {}"
       
  2548       using not_holding_s_th_cs by (auto simp:holdents_def)
       
  2549   qed (auto simp:finite_holdents)
       
  2550   thus ?thesis
       
  2551    by (unfold cntCS_def holdents_es_th, simp)
       
  2552 qed
       
  2553 
       
  2554 lemma no_holder: 
       
  2555   "\<not> holding s th' cs"
       
  2556 proof
       
  2557   assume otherwise: "holding s th' cs"
       
  2558   from this[unfolded s_holding_def, folded wq_def, unfolded we]
       
  2559   show False by auto
       
  2560 qed
       
  2561 
       
  2562 lemma holdents_es_th':
       
  2563   assumes "th' \<noteq> th"
       
  2564   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  2565 proof -
       
  2566   { fix cs'
       
  2567     assume "cs' \<in> ?L"
       
  2568     hence h_e: "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  2569     have "cs' \<noteq> cs"
       
  2570     proof
       
  2571       assume "cs' = cs"
       
  2572       from held_unique[OF h_e[unfolded this] holding_es_th_cs]
       
  2573       have "th' = th" .
       
  2574       with assms show False by simp
       
  2575     qed
       
  2576     from h_e[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp[OF this]]
       
  2577     have "th' \<in> set (wq s cs') \<and> th' = hd (wq s cs')" .
       
  2578     hence "cs' \<in> ?R" 
       
  2579       by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  2580   } moreover {
       
  2581     fix cs'
       
  2582     assume "cs' \<in> ?R"
       
  2583     hence "holding s th' cs'" by (auto simp:holdents_def)
       
  2584     from holding_kept[OF this]
       
  2585     have "holding (e # s) th' cs'" .
       
  2586     hence "cs' \<in> ?L"
       
  2587       by (unfold holdents_def, auto)
       
  2588   } ultimately show ?thesis by auto
       
  2589 qed
       
  2590 
       
  2591 lemma cntCS_es_th'[simp]: 
       
  2592   assumes "th' \<noteq> th"
       
  2593   shows "cntCS (e#s) th' = cntCS s th'"
       
  2594   by (unfold cntCS_def holdents_es_th'[OF assms], simp)
       
  2595 
       
  2596 end
       
  2597 
       
  2598 context valid_trace_p
       
  2599 begin
       
  2600 
       
  2601 lemma readys_kept1: 
       
  2602   assumes "th' \<noteq> th"
       
  2603   and "th' \<in> readys (e#s)"
       
  2604   shows "th' \<in> readys s"
       
  2605 proof -
       
  2606   { fix cs'
       
  2607     assume wait: "waiting s th' cs'"
       
  2608     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  2609         using assms(2)[unfolded readys_def] by auto
       
  2610     have False
       
  2611     proof(cases "cs' = cs")
       
  2612       case False
       
  2613       with n_wait wait
       
  2614       show ?thesis 
       
  2615         by (unfold s_waiting_def, fold wq_def, auto)
       
  2616     next
       
  2617       case True
       
  2618       show ?thesis
       
  2619       proof(cases "wq s cs = []")
       
  2620         case True
       
  2621         then interpret vt: valid_trace_p_h
       
  2622           by (unfold_locales, simp)
       
  2623         show ?thesis using n_wait wait waiting_kept by auto 
       
  2624       next
       
  2625         case False
       
  2626         then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  2627         show ?thesis using n_wait wait waiting_kept by blast 
       
  2628       qed
       
  2629     qed
       
  2630   } with assms(2) show ?thesis  
       
  2631     by (unfold readys_def, auto)
       
  2632 qed
       
  2633 
       
  2634 lemma readys_kept2: 
       
  2635   assumes "th' \<noteq> th"
       
  2636   and "th' \<in> readys s"
       
  2637   shows "th' \<in> readys (e#s)"
       
  2638 proof -
       
  2639   { fix cs'
       
  2640     assume wait: "waiting (e#s) th' cs'"
       
  2641     have n_wait: "\<not> waiting s th' cs'" 
       
  2642         using assms(2)[unfolded readys_def] by auto
       
  2643     have False
       
  2644     proof(cases "cs' = cs")
       
  2645       case False
       
  2646       with n_wait wait
       
  2647       show ?thesis 
       
  2648         by (unfold s_waiting_def, fold wq_def, auto)
       
  2649     next
       
  2650       case True
       
  2651       show ?thesis
       
  2652       proof(cases "wq s cs = []")
       
  2653         case True
       
  2654         then interpret vt: valid_trace_p_h
       
  2655           by (unfold_locales, simp)
       
  2656         show ?thesis using n_wait vt.waiting_esE wait by blast 
       
  2657       next
       
  2658         case False
       
  2659         then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  2660         show ?thesis using assms(1) n_wait vt.waiting_esE wait by auto 
       
  2661       qed
       
  2662     qed
       
  2663   } with assms(2) show ?thesis  
       
  2664     by (unfold readys_def, auto)
       
  2665 qed
       
  2666 
       
  2667 lemma readys_simp [simp]:
       
  2668   assumes "th' \<noteq> th"
       
  2669   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  2670   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  2671   by metis
       
  2672 
       
  2673 lemma cnp_cnv_cncs_kept: (* ddd *)
       
  2674   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  2675   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  2676 proof(cases "th' = th")
       
  2677   case True
       
  2678   note eq_th' = this
       
  2679   show ?thesis
       
  2680   proof(cases "wq s cs = []")
       
  2681     case True
       
  2682     then interpret vt: valid_trace_p_h by (unfold_locales, simp)
       
  2683     show ?thesis
       
  2684       using assms eq_th' is_p ready_th_s vt.cntCS_es_th vt.ready_th_es pvD_def by auto 
       
  2685   next
       
  2686     case False
       
  2687     then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  2688     show ?thesis
       
  2689       using add.commute add.left_commute assms eq_th' is_p live_th_s 
       
  2690             ready_th_s vt.th_not_ready_es pvD_def
       
  2691       apply (auto)
       
  2692       by (fold is_p, simp)
       
  2693   qed
       
  2694 next
       
  2695   case False
       
  2696   note h_False = False
       
  2697   thus ?thesis
       
  2698   proof(cases "wq s cs = []")
       
  2699     case True
       
  2700     then interpret vt: valid_trace_p_h by (unfold_locales, simp)
       
  2701     show ?thesis using assms
       
  2702       by (insert True h_False pvD_def, auto split:if_splits,unfold is_p, auto)
       
  2703   next
       
  2704     case False
       
  2705     then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  2706     show ?thesis using assms
       
  2707       by (insert False h_False pvD_def, auto split:if_splits,unfold is_p, auto)
       
  2708   qed
       
  2709 qed
       
  2710 
       
  2711 end
       
  2712 
       
  2713 
       
  2714 context valid_trace_v (* ccc *)
       
  2715 begin
       
  2716 
       
  2717 lemma holding_th_cs_s: 
       
  2718   "holding s th cs" 
       
  2719  by  (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
       
  2720 
       
  2721 lemma th_ready_s [simp]: "th \<in> readys s"
       
  2722   using runing_th_s
       
  2723   by (unfold runing_def readys_def, auto)
       
  2724 
       
  2725 lemma th_live_s [simp]: "th \<in> threads s"
       
  2726   using th_ready_s by (unfold readys_def, auto)
       
  2727 
       
  2728 lemma th_ready_es [simp]: "th \<in> readys (e#s)"
       
  2729   using runing_th_s neq_t_th
       
  2730   by (unfold is_v runing_def readys_def, auto)
       
  2731 
       
  2732 lemma th_live_es [simp]: "th \<in> threads (e#s)"
       
  2733   using th_ready_es by (unfold readys_def, auto)
       
  2734 
       
  2735 lemma pvD_th_s[simp]: "pvD s th = 0"
       
  2736   by (unfold pvD_def, simp)
       
  2737 
       
  2738 lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
       
  2739   by (unfold pvD_def, simp)
       
  2740 
       
  2741 lemma cntCS_s_th [simp]: "cntCS s th > 0"
       
  2742 proof -
       
  2743   have "cs \<in> holdents s th" using holding_th_cs_s
       
  2744     by (unfold holdents_def, simp)
       
  2745   moreover have "finite (holdents s th)" using finite_holdents
       
  2746     by simp
       
  2747   ultimately show ?thesis
       
  2748     by (unfold cntCS_def, 
       
  2749         auto intro!:card_gt_0_iff[symmetric, THEN iffD1])
       
  2750 qed
       
  2751 
       
  2752 end
       
  2753 
       
  2754 context valid_trace_v_n
       
  2755 begin
       
  2756 
       
  2757 lemma not_ready_taker_s[simp]: 
       
  2758   "taker \<notin> readys s"
       
  2759   using waiting_taker
       
  2760   by (unfold readys_def, auto)
       
  2761 
       
  2762 lemma taker_live_s [simp]: "taker \<in> threads s"
       
  2763 proof -
       
  2764   have "taker \<in> set wq'" by (simp add: eq_wq') 
       
  2765   from th'_in_inv[OF this]
       
  2766   have "taker \<in> set rest" .
       
  2767   hence "taker \<in> set (wq s cs)" by (simp add: wq_s_cs) 
       
  2768   thus ?thesis using wq_threads by auto 
       
  2769 qed
       
  2770 
       
  2771 lemma taker_live_es [simp]: "taker \<in> threads (e#s)"
       
  2772   using taker_live_s threads_es by blast
       
  2773 
       
  2774 lemma taker_ready_es [simp]:
       
  2775   shows "taker \<in> readys (e#s)"
       
  2776 proof -
       
  2777   { fix cs'
       
  2778     assume "waiting (e#s) taker cs'"
       
  2779     hence False
       
  2780     proof(cases rule:waiting_esE)
       
  2781       case 1
       
  2782       thus ?thesis using waiting_taker waiting_unique by auto 
       
  2783     qed simp
       
  2784   } thus ?thesis by (unfold readys_def, auto)
       
  2785 qed
       
  2786 
       
  2787 lemma neq_taker_th: "taker \<noteq> th"
       
  2788   using th_not_waiting waiting_taker by blast
       
  2789 
       
  2790 lemma not_holding_taker_s_cs:
       
  2791   shows "\<not> holding s taker cs"
       
  2792   using holding_cs_eq_th neq_taker_th by auto
       
  2793 
       
  2794 lemma holdents_es_taker:
       
  2795   "holdents (e#s) taker = holdents s taker \<union> {cs}" (is "?L = ?R")
       
  2796 proof -
       
  2797   { fix cs'
       
  2798     assume "cs' \<in> ?L"
       
  2799     hence "holding (e#s) taker cs'" by (auto simp:holdents_def)
       
  2800     hence "cs' \<in> ?R"
       
  2801     proof(cases rule:holding_esE)
       
  2802       case 2
       
  2803       thus ?thesis by (auto simp:holdents_def)
       
  2804     qed auto
       
  2805   } moreover {
       
  2806     fix cs'
       
  2807     assume "cs' \<in> ?R"
       
  2808     hence "holding s taker cs' \<or> cs' = cs" by (auto simp:holdents_def)
       
  2809     hence "cs' \<in> ?L" 
       
  2810     proof
       
  2811       assume "holding s taker cs'"
       
  2812       hence "holding (e#s) taker cs'" 
       
  2813           using holding_esI2 holding_taker by fastforce 
       
  2814       thus ?thesis by (auto simp:holdents_def)
       
  2815     next
       
  2816       assume "cs' = cs"
       
  2817       with holding_taker
       
  2818       show ?thesis by (auto simp:holdents_def)
       
  2819     qed
       
  2820   } ultimately show ?thesis by auto
       
  2821 qed
       
  2822 
       
  2823 lemma cntCS_es_taker [simp]: "cntCS (e#s) taker = cntCS s taker + 1"
       
  2824 proof -
       
  2825   have "card (holdents s taker \<union> {cs}) = card (holdents s taker) + 1"
       
  2826   proof(subst card_Un_disjoint)
       
  2827     show "holdents s taker \<inter> {cs} = {}"
       
  2828       using not_holding_taker_s_cs by (auto simp:holdents_def)
       
  2829   qed (auto simp:finite_holdents)
       
  2830   thus ?thesis 
       
  2831     by (unfold cntCS_def, insert holdents_es_taker, simp)
       
  2832 qed
       
  2833 
       
  2834 lemma pvD_taker_s[simp]: "pvD s taker = 1"
       
  2835   by (unfold pvD_def, simp)
       
  2836 
       
  2837 lemma pvD_taker_es[simp]: "pvD (e#s) taker = 0"
       
  2838   by (unfold pvD_def, simp)  
       
  2839 
       
  2840 lemma pvD_th_s[simp]: "pvD s th = 0"
       
  2841   by (unfold pvD_def, simp)
       
  2842 
       
  2843 lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
       
  2844   by (unfold pvD_def, simp)
       
  2845 
       
  2846 lemma holdents_es_th:
       
  2847   "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
       
  2848 proof -
       
  2849   { fix cs'
       
  2850     assume "cs' \<in> ?L"
       
  2851     hence "holding (e#s) th cs'" by (auto simp:holdents_def)
       
  2852     hence "cs' \<in> ?R"
       
  2853     proof(cases rule:holding_esE)
       
  2854       case 2
       
  2855       thus ?thesis by (auto simp:holdents_def)
       
  2856     qed (insert neq_taker_th, auto)
       
  2857   } moreover {
       
  2858     fix cs'
       
  2859     assume "cs' \<in> ?R"
       
  2860     hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
       
  2861     from holding_esI2[OF this]
       
  2862     have "cs' \<in> ?L" by (auto simp:holdents_def)
       
  2863   } ultimately show ?thesis by auto
       
  2864 qed
       
  2865 
       
  2866 lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
       
  2867 proof -
       
  2868   have "card (holdents s th - {cs}) = card (holdents s th) - 1"
       
  2869   proof -
       
  2870     have "cs \<in> holdents s th" using holding_th_cs_s
       
  2871       by (auto simp:holdents_def)
       
  2872     moreover have "finite (holdents s th)"
       
  2873         by (simp add: finite_holdents) 
       
  2874     ultimately show ?thesis by auto
       
  2875   qed
       
  2876   thus ?thesis by (unfold cntCS_def holdents_es_th)
       
  2877 qed
       
  2878 
       
  2879 lemma holdents_kept:
       
  2880   assumes "th' \<noteq> taker"
       
  2881   and "th' \<noteq> th"
       
  2882   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  2883 proof -
       
  2884   { fix cs'
       
  2885     assume h: "cs' \<in> ?L"
       
  2886     have "cs' \<in> ?R"
       
  2887     proof(cases "cs' = cs")
       
  2888       case False
       
  2889       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  2890       from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  2891       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  2892       show ?thesis
       
  2893         by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  2894     next
       
  2895       case True
       
  2896       from h[unfolded this]
       
  2897       have "holding (e#s) th' cs" by (auto simp:holdents_def)
       
  2898       from held_unique[OF this holding_taker]
       
  2899       have "th' = taker" .
       
  2900       with assms show ?thesis by auto
       
  2901     qed
       
  2902   } moreover {
       
  2903     fix cs'
       
  2904     assume h: "cs' \<in> ?R"
       
  2905     have "cs' \<in> ?L"
       
  2906     proof(cases "cs' = cs")
       
  2907       case False
       
  2908       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  2909       from h have "holding s th' cs'" by (auto simp:holdents_def)
       
  2910       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  2911       show ?thesis
       
  2912         by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
       
  2913     next
       
  2914       case True
       
  2915       from h[unfolded this]
       
  2916       have "holding s th' cs" by (auto simp:holdents_def)
       
  2917       from held_unique[OF this holding_th_cs_s]
       
  2918       have "th' = th" .
       
  2919       with assms show ?thesis by auto
       
  2920     qed
       
  2921   } ultimately show ?thesis by auto
       
  2922 qed
       
  2923 
       
  2924 lemma cntCS_kept [simp]:
       
  2925   assumes "th' \<noteq> taker"
       
  2926   and "th' \<noteq> th"
       
  2927   shows "cntCS (e#s) th' = cntCS s th'"
       
  2928   by (unfold cntCS_def holdents_kept[OF assms], simp)
       
  2929 
       
  2930 lemma readys_kept1: 
       
  2931   assumes "th' \<noteq> taker"
       
  2932   and "th' \<in> readys (e#s)"
       
  2933   shows "th' \<in> readys s"
       
  2934 proof -
       
  2935   { fix cs'
       
  2936     assume wait: "waiting s th' cs'"
       
  2937     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  2938         using assms(2)[unfolded readys_def] by auto
       
  2939     have False
       
  2940     proof(cases "cs' = cs")
       
  2941       case False
       
  2942       with n_wait wait
       
  2943       show ?thesis 
       
  2944         by (unfold s_waiting_def, fold wq_def, auto)
       
  2945     next
       
  2946       case True
       
  2947       have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
       
  2948         using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
       
  2949       moreover have "\<not> (th' \<in> set rest \<and> th' \<noteq> hd (taker # rest'))" 
       
  2950         using n_wait[unfolded True s_waiting_def, folded wq_def, 
       
  2951                     unfolded wq_es_cs set_wq', unfolded eq_wq'] .
       
  2952       ultimately have "th' = taker" by auto
       
  2953       with assms(1)
       
  2954       show ?thesis by simp
       
  2955     qed
       
  2956   } with assms(2) show ?thesis  
       
  2957     by (unfold readys_def, auto)
       
  2958 qed
       
  2959 
       
  2960 lemma readys_kept2: 
       
  2961   assumes "th' \<noteq> taker"
       
  2962   and "th' \<in> readys s"
       
  2963   shows "th' \<in> readys (e#s)"
       
  2964 proof -
       
  2965   { fix cs'
       
  2966     assume wait: "waiting (e#s) th' cs'"
       
  2967     have n_wait: "\<not> waiting s th' cs'" 
       
  2968         using assms(2)[unfolded readys_def] by auto
       
  2969     have False
       
  2970     proof(cases "cs' = cs")
       
  2971       case False
       
  2972       with n_wait wait
       
  2973       show ?thesis 
       
  2974         by (unfold s_waiting_def, fold wq_def, auto)
       
  2975     next
       
  2976       case True
       
  2977       have "th' \<in> set rest \<and> th' \<noteq> hd (taker # rest')"
       
  2978           using  wait [unfolded True s_waiting_def, folded wq_def, 
       
  2979                     unfolded wq_es_cs set_wq', unfolded eq_wq']  .
       
  2980       moreover have "\<not> (th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest))"
       
  2981           using n_wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
       
  2982       ultimately have "th' = taker" by auto
       
  2983       with assms(1)
       
  2984       show ?thesis by simp
       
  2985     qed
       
  2986   } with assms(2) show ?thesis  
       
  2987     by (unfold readys_def, auto)
       
  2988 qed
       
  2989 
       
  2990 lemma readys_simp [simp]:
       
  2991   assumes "th' \<noteq> taker"
       
  2992   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  2993   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  2994   by metis
       
  2995 
       
  2996 lemma cnp_cnv_cncs_kept:
       
  2997   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  2998   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  2999 proof -
       
  3000   { assume eq_th': "th' = taker"
       
  3001     have ?thesis
       
  3002       apply (unfold eq_th' pvD_taker_es cntCS_es_taker)
       
  3003       by (insert neq_taker_th assms[unfolded eq_th'], unfold is_v, simp)
       
  3004   } moreover {
       
  3005     assume eq_th': "th' = th"
       
  3006     have ?thesis 
       
  3007       apply (unfold eq_th' pvD_th_es cntCS_es_th)
       
  3008       by (insert assms[unfolded eq_th'], unfold is_v, simp)
       
  3009   } moreover {
       
  3010     assume h: "th' \<noteq> taker" "th' \<noteq> th"
       
  3011     have ?thesis using assms
       
  3012       apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
       
  3013       by (fold is_v, unfold pvD_def, simp)
       
  3014   } ultimately show ?thesis by metis
       
  3015 qed
       
  3016 
       
  3017 end
       
  3018 
       
  3019 context valid_trace_v_e
       
  3020 begin
       
  3021 
       
  3022 lemma holdents_es_th:
       
  3023   "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
       
  3024 proof -
       
  3025   { fix cs'
       
  3026     assume "cs' \<in> ?L"
       
  3027     hence "holding (e#s) th cs'" by (auto simp:holdents_def)
       
  3028     hence "cs' \<in> ?R"
       
  3029     proof(cases rule:holding_esE)
       
  3030       case 1
       
  3031       thus ?thesis by (auto simp:holdents_def)
       
  3032     qed 
       
  3033   } moreover {
       
  3034     fix cs'
       
  3035     assume "cs' \<in> ?R"
       
  3036     hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
       
  3037     from holding_esI2[OF this]
       
  3038     have "cs' \<in> ?L" by (auto simp:holdents_def)
       
  3039   } ultimately show ?thesis by auto
       
  3040 qed
       
  3041 
       
  3042 lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
       
  3043 proof -
       
  3044   have "card (holdents s th - {cs}) = card (holdents s th) - 1"
       
  3045   proof -
       
  3046     have "cs \<in> holdents s th" using holding_th_cs_s
       
  3047       by (auto simp:holdents_def)
       
  3048     moreover have "finite (holdents s th)"
       
  3049         by (simp add: finite_holdents) 
       
  3050     ultimately show ?thesis by auto
       
  3051   qed
       
  3052   thus ?thesis by (unfold cntCS_def holdents_es_th)
       
  3053 qed
       
  3054 
       
  3055 lemma holdents_kept:
       
  3056   assumes "th' \<noteq> th"
       
  3057   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3058 proof -
       
  3059   { fix cs'
       
  3060     assume h: "cs' \<in> ?L"
       
  3061     have "cs' \<in> ?R"
       
  3062     proof(cases "cs' = cs")
       
  3063       case False
       
  3064       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3065       from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  3066       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  3067       show ?thesis
       
  3068         by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  3069     next
       
  3070       case True
       
  3071       from h[unfolded this]
       
  3072       have "holding (e#s) th' cs" by (auto simp:holdents_def)
       
  3073       from this[unfolded s_holding_def, folded wq_def, 
       
  3074             unfolded wq_es_cs nil_wq']
       
  3075       show ?thesis by auto
       
  3076     qed
       
  3077   } moreover {
       
  3078     fix cs'
       
  3079     assume h: "cs' \<in> ?R"
       
  3080     have "cs' \<in> ?L"
       
  3081     proof(cases "cs' = cs")
       
  3082       case False
       
  3083       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3084       from h have "holding s th' cs'" by (auto simp:holdents_def)
       
  3085       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  3086       show ?thesis
       
  3087         by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
       
  3088     next
       
  3089       case True
       
  3090       from h[unfolded this]
       
  3091       have "holding s th' cs" by (auto simp:holdents_def)
       
  3092       from held_unique[OF this holding_th_cs_s]
       
  3093       have "th' = th" .
       
  3094       with assms show ?thesis by auto
       
  3095     qed
       
  3096   } ultimately show ?thesis by auto
       
  3097 qed
       
  3098 
       
  3099 lemma cntCS_kept [simp]:
       
  3100   assumes "th' \<noteq> th"
       
  3101   shows "cntCS (e#s) th' = cntCS s th'"
       
  3102   by (unfold cntCS_def holdents_kept[OF assms], simp)
       
  3103 
       
  3104 lemma readys_kept1: 
       
  3105   assumes "th' \<in> readys (e#s)"
       
  3106   shows "th' \<in> readys s"
       
  3107 proof -
       
  3108   { fix cs'
       
  3109     assume wait: "waiting s th' cs'"
       
  3110     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3111         using assms(1)[unfolded readys_def] by auto
       
  3112     have False
       
  3113     proof(cases "cs' = cs")
       
  3114       case False
       
  3115       with n_wait wait
       
  3116       show ?thesis 
       
  3117         by (unfold s_waiting_def, fold wq_def, auto)
       
  3118     next
       
  3119       case True
       
  3120       have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
       
  3121         using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . 
       
  3122       hence "th' \<in> set rest" by auto
       
  3123       with set_wq' have "th' \<in> set wq'" by metis
       
  3124       with nil_wq' show ?thesis by simp
       
  3125     qed
       
  3126   } thus ?thesis using assms
       
  3127     by (unfold readys_def, auto)
       
  3128 qed
       
  3129 
       
  3130 lemma readys_kept2: 
       
  3131   assumes "th' \<in> readys s"
       
  3132   shows "th' \<in> readys (e#s)"
       
  3133 proof -
       
  3134   { fix cs'
       
  3135     assume wait: "waiting (e#s) th' cs'"
       
  3136     have n_wait: "\<not> waiting s th' cs'" 
       
  3137         using assms[unfolded readys_def] by auto
       
  3138     have False
       
  3139     proof(cases "cs' = cs")
       
  3140       case False
       
  3141       with n_wait wait
       
  3142       show ?thesis 
       
  3143         by (unfold s_waiting_def, fold wq_def, auto)
       
  3144     next
       
  3145       case True
       
  3146       have "th' \<in> set [] \<and> th' \<noteq> hd []"
       
  3147         using wait[unfolded True s_waiting_def, folded wq_def, 
       
  3148               unfolded wq_es_cs nil_wq'] .
       
  3149       thus ?thesis by simp
       
  3150     qed
       
  3151   } with assms show ?thesis  
       
  3152     by (unfold readys_def, auto)
       
  3153 qed
       
  3154 
       
  3155 lemma readys_simp [simp]:
       
  3156   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3157   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3158   by metis
       
  3159 
       
  3160 lemma cnp_cnv_cncs_kept:
       
  3161   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3162   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3163 proof -
       
  3164   {
       
  3165     assume eq_th': "th' = th"
       
  3166     have ?thesis 
       
  3167       apply (unfold eq_th' pvD_th_es cntCS_es_th)
       
  3168       by (insert assms[unfolded eq_th'], unfold is_v, simp)
       
  3169   } moreover {
       
  3170     assume h: "th' \<noteq> th"
       
  3171     have ?thesis using assms
       
  3172       apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
       
  3173       by (fold is_v, unfold pvD_def, simp)
       
  3174   } ultimately show ?thesis by metis
       
  3175 qed
       
  3176 
       
  3177 end
       
  3178 
       
  3179 context valid_trace_v
       
  3180 begin
       
  3181 
       
  3182 lemma cnp_cnv_cncs_kept:
       
  3183   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3184   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3185 proof(cases "rest = []")
       
  3186   case True
       
  3187   then interpret vt: valid_trace_v_e by (unfold_locales, simp)
       
  3188   show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
       
  3189 next
       
  3190   case False
       
  3191   then interpret vt: valid_trace_v_n by (unfold_locales, simp)
       
  3192   show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
       
  3193 qed
       
  3194 
       
  3195 end
       
  3196 
       
  3197 context valid_trace_create
       
  3198 begin
       
  3199 
       
  3200 lemma th_not_live_s [simp]: "th \<notin> threads s"
       
  3201 proof -
       
  3202   from pip_e[unfolded is_create]
       
  3203   show ?thesis by (cases, simp)
       
  3204 qed
       
  3205 
       
  3206 lemma th_not_ready_s [simp]: "th \<notin> readys s"
       
  3207   using th_not_live_s by (unfold readys_def, simp)
       
  3208 
       
  3209 lemma th_live_es [simp]: "th \<in> threads (e#s)"
       
  3210   by (unfold is_create, simp)
       
  3211 
       
  3212 lemma not_waiting_th_s [simp]: "\<not> waiting s th cs'"
       
  3213 proof
       
  3214   assume "waiting s th cs'"
       
  3215   from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3216   have "th \<in> set (wq s cs')" by auto
       
  3217   from wq_threads[OF this] have "th \<in> threads s" .
       
  3218   with th_not_live_s show False by simp
       
  3219 qed
       
  3220 
       
  3221 lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
       
  3222 proof
       
  3223   assume "holding s th cs'"
       
  3224   from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
       
  3225   have "th \<in> set (wq s cs')" by auto
       
  3226   from wq_threads[OF this] have "th \<in> threads s" .
       
  3227   with th_not_live_s show False by simp
       
  3228 qed
       
  3229 
       
  3230 lemma not_waiting_th_es [simp]: "\<not> waiting (e#s) th cs'"
       
  3231 proof
       
  3232   assume "waiting (e # s) th cs'"
       
  3233   from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3234   have "th \<in> set (wq s cs')" by auto
       
  3235   from wq_threads[OF this] have "th \<in> threads s" .
       
  3236   with th_not_live_s show False by simp
       
  3237 qed
       
  3238 
       
  3239 lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
       
  3240 proof
       
  3241   assume "holding (e # s) th cs'"
       
  3242   from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
       
  3243   have "th \<in> set (wq s cs')" by auto
       
  3244   from wq_threads[OF this] have "th \<in> threads s" .
       
  3245   with th_not_live_s show False by simp
       
  3246 qed
       
  3247 
       
  3248 lemma ready_th_es [simp]: "th \<in> readys (e#s)"
       
  3249   by (simp add:readys_def)
       
  3250 
       
  3251 lemma holdents_th_s: "holdents s th = {}"
       
  3252   by (unfold holdents_def, auto)
       
  3253 
       
  3254 lemma holdents_th_es: "holdents (e#s) th = {}"
       
  3255   by (unfold holdents_def, auto)
       
  3256 
       
  3257 lemma cntCS_th_s [simp]: "cntCS s th = 0"
       
  3258   by (unfold cntCS_def, simp add:holdents_th_s)
       
  3259 
       
  3260 lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
       
  3261   by (unfold cntCS_def, simp add:holdents_th_es)
       
  3262 
       
  3263 lemma pvD_th_s [simp]: "pvD s th = 0"
       
  3264   by (unfold pvD_def, simp)
       
  3265 
       
  3266 lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
       
  3267   by (unfold pvD_def, simp)
       
  3268 
       
  3269 lemma holdents_kept:
       
  3270   assumes "th' \<noteq> th"
       
  3271   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3272 proof -
       
  3273   { fix cs'
       
  3274     assume h: "cs' \<in> ?L"
       
  3275     hence "cs' \<in> ?R"
       
  3276       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3277              unfold wq_neq_simp, auto)
       
  3278   } moreover {
       
  3279     fix cs'
       
  3280     assume h: "cs' \<in> ?R"
       
  3281     hence "cs' \<in> ?L"
       
  3282       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3283              unfold wq_neq_simp, auto)
       
  3284   } ultimately show ?thesis by auto
       
  3285 qed
       
  3286 
       
  3287 lemma cntCS_kept [simp]:
       
  3288   assumes "th' \<noteq> th"
       
  3289   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  3290   using holdents_kept[OF assms]
       
  3291   by (unfold cntCS_def, simp)
       
  3292 
       
  3293 lemma readys_kept1: 
       
  3294   assumes "th' \<noteq> th"
       
  3295   and "th' \<in> readys (e#s)"
       
  3296   shows "th' \<in> readys s"
       
  3297 proof -
       
  3298   { fix cs'
       
  3299     assume wait: "waiting s th' cs'"
       
  3300     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3301       using assms by (auto simp:readys_def)
       
  3302     from wait[unfolded s_waiting_def, folded wq_def]
       
  3303          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3304     have False by auto
       
  3305   } thus ?thesis using assms
       
  3306     by (unfold readys_def, auto)
       
  3307 qed
       
  3308 
       
  3309 lemma readys_kept2: 
       
  3310   assumes "th' \<noteq> th"
       
  3311   and "th' \<in> readys s"
       
  3312   shows "th' \<in> readys (e#s)"
       
  3313 proof -
       
  3314   { fix cs'
       
  3315     assume wait: "waiting (e#s) th' cs'"
       
  3316     have n_wait: "\<not> waiting s th' cs'"
       
  3317       using assms(2) by (auto simp:readys_def)
       
  3318     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3319          n_wait[unfolded s_waiting_def, folded wq_def]
       
  3320     have False by auto
       
  3321   } with assms show ?thesis  
       
  3322     by (unfold readys_def, auto)
       
  3323 qed
       
  3324 
       
  3325 lemma readys_simp [simp]:
       
  3326   assumes "th' \<noteq> th"
       
  3327   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3328   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3329   by metis
       
  3330 
       
  3331 lemma pvD_kept [simp]:
       
  3332   assumes "th' \<noteq> th"
       
  3333   shows "pvD (e#s) th' = pvD s th'"
       
  3334   using assms
       
  3335   by (unfold pvD_def, simp)
       
  3336 
       
  3337 lemma cnp_cnv_cncs_kept:
       
  3338   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3339   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3340 proof -
       
  3341   {
       
  3342     assume eq_th': "th' = th"
       
  3343     have ?thesis using assms
       
  3344       by (unfold eq_th', simp, unfold is_create, simp)
       
  3345   } moreover {
       
  3346     assume h: "th' \<noteq> th"
       
  3347     hence ?thesis using assms
       
  3348       by (simp, simp add:is_create)
       
  3349   } ultimately show ?thesis by metis
       
  3350 qed
       
  3351 
       
  3352 end
       
  3353 
       
  3354 context valid_trace_exit
       
  3355 begin
       
  3356 
       
  3357 lemma th_live_s [simp]: "th \<in> threads s"
       
  3358 proof -
       
  3359   from pip_e[unfolded is_exit]
       
  3360   show ?thesis
       
  3361   by (cases, unfold runing_def readys_def, simp)
       
  3362 qed
       
  3363 
       
  3364 lemma th_ready_s [simp]: "th \<in> readys s"
       
  3365 proof -
       
  3366   from pip_e[unfolded is_exit]
       
  3367   show ?thesis
       
  3368   by (cases, unfold runing_def, simp)
       
  3369 qed
       
  3370 
       
  3371 lemma th_not_live_es [simp]: "th \<notin> threads (e#s)"
       
  3372   by (unfold is_exit, simp)
       
  3373 
       
  3374 lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
       
  3375 proof -
       
  3376   from pip_e[unfolded is_exit]
       
  3377   show ?thesis 
       
  3378    by (cases, unfold holdents_def, auto)
       
  3379 qed
       
  3380 
       
  3381 lemma cntCS_th_s [simp]: "cntCS s th = 0"
       
  3382 proof -
       
  3383   from pip_e[unfolded is_exit]
       
  3384   show ?thesis 
       
  3385    by (cases, unfold cntCS_def, simp)
       
  3386 qed
       
  3387 
       
  3388 lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
       
  3389 proof
       
  3390   assume "holding (e # s) th cs'"
       
  3391   from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
       
  3392   have "holding s th cs'" 
       
  3393     by (unfold s_holding_def, fold wq_def, auto)
       
  3394   with not_holding_th_s 
       
  3395   show False by simp
       
  3396 qed
       
  3397 
       
  3398 lemma ready_th_es [simp]: "th \<notin> readys (e#s)"
       
  3399   by (simp add:readys_def)
       
  3400 
       
  3401 lemma holdents_th_s: "holdents s th = {}"
       
  3402   by (unfold holdents_def, auto)
       
  3403 
       
  3404 lemma holdents_th_es: "holdents (e#s) th = {}"
       
  3405   by (unfold holdents_def, auto)
       
  3406 
       
  3407 lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
       
  3408   by (unfold cntCS_def, simp add:holdents_th_es)
       
  3409 
       
  3410 lemma pvD_th_s [simp]: "pvD s th = 0"
       
  3411   by (unfold pvD_def, simp)
       
  3412 
       
  3413 lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
       
  3414   by (unfold pvD_def, simp)
       
  3415 
       
  3416 lemma holdents_kept:
       
  3417   assumes "th' \<noteq> th"
       
  3418   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3419 proof -
       
  3420   { fix cs'
       
  3421     assume h: "cs' \<in> ?L"
       
  3422     hence "cs' \<in> ?R"
       
  3423       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3424              unfold wq_neq_simp, auto)
       
  3425   } moreover {
       
  3426     fix cs'
       
  3427     assume h: "cs' \<in> ?R"
       
  3428     hence "cs' \<in> ?L"
       
  3429       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3430              unfold wq_neq_simp, auto)
       
  3431   } ultimately show ?thesis by auto
       
  3432 qed
       
  3433 
       
  3434 lemma cntCS_kept [simp]:
       
  3435   assumes "th' \<noteq> th"
       
  3436   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  3437   using holdents_kept[OF assms]
       
  3438   by (unfold cntCS_def, simp)
       
  3439 
       
  3440 lemma readys_kept1: 
       
  3441   assumes "th' \<noteq> th"
       
  3442   and "th' \<in> readys (e#s)"
       
  3443   shows "th' \<in> readys s"
       
  3444 proof -
       
  3445   { fix cs'
       
  3446     assume wait: "waiting s th' cs'"
       
  3447     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3448       using assms by (auto simp:readys_def)
       
  3449     from wait[unfolded s_waiting_def, folded wq_def]
       
  3450          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3451     have False by auto
       
  3452   } thus ?thesis using assms
       
  3453     by (unfold readys_def, auto)
       
  3454 qed
       
  3455 
       
  3456 lemma readys_kept2: 
       
  3457   assumes "th' \<noteq> th"
       
  3458   and "th' \<in> readys s"
       
  3459   shows "th' \<in> readys (e#s)"
       
  3460 proof -
       
  3461   { fix cs'
       
  3462     assume wait: "waiting (e#s) th' cs'"
       
  3463     have n_wait: "\<not> waiting s th' cs'"
       
  3464       using assms(2) by (auto simp:readys_def)
       
  3465     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3466          n_wait[unfolded s_waiting_def, folded wq_def]
       
  3467     have False by auto
       
  3468   } with assms show ?thesis  
       
  3469     by (unfold readys_def, auto)
       
  3470 qed
       
  3471 
       
  3472 lemma readys_simp [simp]:
       
  3473   assumes "th' \<noteq> th"
       
  3474   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3475   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3476   by metis
       
  3477 
       
  3478 lemma pvD_kept [simp]:
       
  3479   assumes "th' \<noteq> th"
       
  3480   shows "pvD (e#s) th' = pvD s th'"
       
  3481   using assms
       
  3482   by (unfold pvD_def, simp)
       
  3483 
       
  3484 lemma cnp_cnv_cncs_kept:
       
  3485   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3486   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3487 proof -
       
  3488   {
       
  3489     assume eq_th': "th' = th"
       
  3490     have ?thesis using assms
       
  3491       by (unfold eq_th', simp, unfold is_exit, simp)
       
  3492   } moreover {
       
  3493     assume h: "th' \<noteq> th"
       
  3494     hence ?thesis using assms
       
  3495       by (simp, simp add:is_exit)
       
  3496   } ultimately show ?thesis by metis
       
  3497 qed
       
  3498 
       
  3499 end
       
  3500 
       
  3501 context valid_trace_set
       
  3502 begin
       
  3503 
       
  3504 lemma th_live_s [simp]: "th \<in> threads s"
       
  3505 proof -
       
  3506   from pip_e[unfolded is_set]
       
  3507   show ?thesis
       
  3508   by (cases, unfold runing_def readys_def, simp)
       
  3509 qed
       
  3510 
       
  3511 lemma th_ready_s [simp]: "th \<in> readys s"
       
  3512 proof -
       
  3513   from pip_e[unfolded is_set]
       
  3514   show ?thesis
       
  3515   by (cases, unfold runing_def, simp)
       
  3516 qed
       
  3517 
       
  3518 lemma th_not_live_es [simp]: "th \<in> threads (e#s)"
       
  3519   by (unfold is_set, simp)
       
  3520 
       
  3521 
       
  3522 lemma holdents_kept:
       
  3523   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3524 proof -
       
  3525   { fix cs'
       
  3526     assume h: "cs' \<in> ?L"
       
  3527     hence "cs' \<in> ?R"
       
  3528       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3529              unfold wq_neq_simp, auto)
       
  3530   } moreover {
       
  3531     fix cs'
       
  3532     assume h: "cs' \<in> ?R"
       
  3533     hence "cs' \<in> ?L"
       
  3534       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3535              unfold wq_neq_simp, auto)
       
  3536   } ultimately show ?thesis by auto
       
  3537 qed
       
  3538 
       
  3539 lemma cntCS_kept [simp]:
       
  3540   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  3541   using holdents_kept
       
  3542   by (unfold cntCS_def, simp)
       
  3543 
       
  3544 lemma threads_kept[simp]:
       
  3545   "threads (e#s) = threads s"
       
  3546   by (unfold is_set, simp)
       
  3547 
       
  3548 lemma readys_kept1: 
       
  3549   assumes "th' \<in> readys (e#s)"
       
  3550   shows "th' \<in> readys s"
       
  3551 proof -
       
  3552   { fix cs'
       
  3553     assume wait: "waiting s th' cs'"
       
  3554     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3555       using assms by (auto simp:readys_def)
       
  3556     from wait[unfolded s_waiting_def, folded wq_def]
       
  3557          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3558     have False by auto
       
  3559   } moreover have "th' \<in> threads s" 
       
  3560     using assms[unfolded readys_def] by auto
       
  3561   ultimately show ?thesis 
       
  3562     by (unfold readys_def, auto)
       
  3563 qed
       
  3564 
       
  3565 lemma readys_kept2: 
       
  3566   assumes "th' \<in> readys s"
       
  3567   shows "th' \<in> readys (e#s)"
       
  3568 proof -
       
  3569   { fix cs'
       
  3570     assume wait: "waiting (e#s) th' cs'"
       
  3571     have n_wait: "\<not> waiting s th' cs'"
       
  3572       using assms by (auto simp:readys_def)
       
  3573     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3574          n_wait[unfolded s_waiting_def, folded wq_def]
       
  3575     have False by auto
       
  3576   } with assms show ?thesis  
       
  3577     by (unfold readys_def, auto)
       
  3578 qed
       
  3579 
       
  3580 lemma readys_simp [simp]:
       
  3581   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3582   using readys_kept1 readys_kept2
       
  3583   by metis
       
  3584 
       
  3585 lemma pvD_kept [simp]:
       
  3586   shows "pvD (e#s) th' = pvD s th'"
       
  3587   by (unfold pvD_def, simp)
       
  3588 
       
  3589 lemma cnp_cnv_cncs_kept:
       
  3590   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3591   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3592   using assms
       
  3593   by (unfold is_set, simp, fold is_set, simp)
       
  3594 
       
  3595 end
       
  3596 
       
  3597 context valid_trace
       
  3598 begin
       
  3599 
       
  3600 lemma cnp_cnv_cncs: "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3601 proof(induct rule:ind)
       
  3602   case Nil
       
  3603   thus ?case 
       
  3604     by (unfold cntP_def cntV_def pvD_def cntCS_def holdents_def 
       
  3605               s_holding_def, simp)
       
  3606 next
       
  3607   case (Cons s e)
       
  3608   interpret vt_e: valid_trace_e s e using Cons by simp
       
  3609   show ?case
       
  3610   proof(cases e)
       
  3611     case (Create th prio)
       
  3612     interpret vt_create: valid_trace_create s e th prio 
       
  3613       using Create by (unfold_locales, simp)
       
  3614     show ?thesis using Cons by (simp add: vt_create.cnp_cnv_cncs_kept) 
       
  3615   next
       
  3616     case (Exit th)
       
  3617     interpret vt_exit: valid_trace_exit s e th  
       
  3618         using Exit by (unfold_locales, simp)
       
  3619    show ?thesis using Cons by (simp add: vt_exit.cnp_cnv_cncs_kept) 
       
  3620   next
       
  3621     case (P th cs)
       
  3622     interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp)
       
  3623     show ?thesis using Cons by (simp add: vt_p.cnp_cnv_cncs_kept) 
       
  3624   next
       
  3625     case (V th cs)
       
  3626     interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp)
       
  3627     show ?thesis using Cons by (simp add: vt_v.cnp_cnv_cncs_kept) 
       
  3628   next
       
  3629     case (Set th prio)
       
  3630     interpret vt_set: valid_trace_set s e th prio
       
  3631         using Set by (unfold_locales, simp)
       
  3632     show ?thesis using Cons by (simp add: vt_set.cnp_cnv_cncs_kept) 
       
  3633   qed
       
  3634 qed
       
  3635 
       
  3636 lemma not_thread_holdents:
       
  3637   assumes not_in: "th \<notin> threads s" 
       
  3638   shows "holdents s th = {}"
       
  3639 proof -
       
  3640   { fix cs
       
  3641     assume "cs \<in> holdents s th"
       
  3642     hence "holding s th cs" by (auto simp:holdents_def)
       
  3643     from this[unfolded s_holding_def, folded wq_def]
       
  3644     have "th \<in> set (wq s cs)" by auto
       
  3645     with wq_threads have "th \<in> threads s" by auto
       
  3646     with assms
       
  3647     have False by simp
       
  3648   } thus ?thesis by auto
       
  3649 qed
       
  3650 
       
  3651 lemma not_thread_cncs:
       
  3652   assumes not_in: "th \<notin> threads s" 
       
  3653   shows "cntCS s th = 0"
       
  3654   using not_thread_holdents[OF assms]
       
  3655   by (simp add:cntCS_def)
       
  3656 
       
  3657 lemma cnp_cnv_eq:
       
  3658   assumes "th \<notin> threads s"
       
  3659   shows "cntP s th = cntV s th"
       
  3660   using assms cnp_cnv_cncs not_thread_cncs pvD_def
       
  3661   by (auto)
       
  3662 
       
  3663 lemma runing_unique:
       
  3664   assumes runing_1: "th1 \<in> runing s"
       
  3665   and runing_2: "th2 \<in> runing s"
       
  3666   shows "th1 = th2"
       
  3667 proof -
       
  3668   from runing_1 and runing_2 have "cp s th1 = cp s th2"
       
  3669     unfolding runing_def by auto
       
  3670   from this[unfolded cp_alt_def]
       
  3671   have eq_max: 
       
  3672     "Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th1)}) =
       
  3673      Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th2)})" 
       
  3674         (is "Max ?L = Max ?R") .
       
  3675   have "Max ?L \<in> ?L"
       
  3676   proof(rule Max_in)
       
  3677     show "finite ?L" by (simp add: finite_subtree_threads)
       
  3678   next
       
  3679     show "?L \<noteq> {}" using subtree_def by fastforce 
       
  3680   qed
       
  3681   then obtain th1' where 
       
  3682     h_1: "Th th1' \<in> subtree (RAG s) (Th th1)" "the_preced s th1' = Max ?L"
       
  3683     by auto
       
  3684   have "Max ?R \<in> ?R"
       
  3685   proof(rule Max_in)
       
  3686     show "finite ?R" by (simp add: finite_subtree_threads)
       
  3687   next
       
  3688     show "?R \<noteq> {}" using subtree_def by fastforce 
       
  3689   qed
       
  3690   then obtain th2' where 
       
  3691     h_2: "Th th2' \<in> subtree (RAG s) (Th th2)" "the_preced s th2' = Max ?R"
       
  3692     by auto
       
  3693   have "th1' = th2'"
       
  3694   proof(rule preced_unique)
       
  3695     from h_1(1)
       
  3696     show "th1' \<in> threads s"
       
  3697     proof(cases rule:subtreeE)
       
  3698       case 1
       
  3699       hence "th1' = th1" by simp
       
  3700       with runing_1 show ?thesis by (auto simp:runing_def readys_def)
       
  3701     next
       
  3702       case 2
       
  3703       from this(2)
       
  3704       have "(Th th1', Th th1) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
  3705       from tranclD[OF this]
       
  3706       have "(Th th1') \<in> Domain (RAG s)" by auto
       
  3707       from dm_RAG_threads[OF this] show ?thesis .
       
  3708     qed
       
  3709   next
       
  3710     from h_2(1)
       
  3711     show "th2' \<in> threads s"
       
  3712     proof(cases rule:subtreeE)
       
  3713       case 1
       
  3714       hence "th2' = th2" by simp
       
  3715       with runing_2 show ?thesis by (auto simp:runing_def readys_def)
       
  3716     next
       
  3717       case 2
       
  3718       from this(2)
       
  3719       have "(Th th2', Th th2) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
  3720       from tranclD[OF this]
       
  3721       have "(Th th2') \<in> Domain (RAG s)" by auto
       
  3722       from dm_RAG_threads[OF this] show ?thesis .
       
  3723     qed
       
  3724   next
       
  3725     have "the_preced s th1' = the_preced s th2'" 
       
  3726      using eq_max h_1(2) h_2(2) by metis
       
  3727     thus "preced th1' s = preced th2' s" by (simp add:the_preced_def)
       
  3728   qed
       
  3729   from h_1(1)[unfolded this]
       
  3730   have star1: "(Th th2', Th th1) \<in> (RAG s)^*" by (auto simp:subtree_def)
       
  3731   from h_2(1)[unfolded this]
       
  3732   have star2: "(Th th2', Th th2) \<in> (RAG s)^*" by (auto simp:subtree_def)
       
  3733   from star_rpath[OF star1] obtain xs1 
       
  3734     where rp1: "rpath (RAG s) (Th th2') xs1 (Th th1)"
       
  3735     by auto
       
  3736   from star_rpath[OF star2] obtain xs2 
       
  3737     where rp2: "rpath (RAG s) (Th th2') xs2 (Th th2)"
       
  3738     by auto
       
  3739   from rp1 rp2
       
  3740   show ?thesis
       
  3741   proof(cases)
       
  3742     case (less_1 xs')
       
  3743     moreover have "xs' = []"
       
  3744     proof(rule ccontr)
       
  3745       assume otherwise: "xs' \<noteq> []"
       
  3746       from rpath_plus[OF less_1(3) this]
       
  3747       have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+" .
       
  3748       from tranclD[OF this]
       
  3749       obtain cs where "waiting s th1 cs"
       
  3750         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  3751       with runing_1 show False
       
  3752         by (unfold runing_def readys_def, auto)
       
  3753     qed
       
  3754     ultimately have "xs2 = xs1" by simp
       
  3755     from rpath_dest_eq[OF rp1 rp2[unfolded this]]
       
  3756     show ?thesis by simp
       
  3757   next
       
  3758     case (less_2 xs')
       
  3759     moreover have "xs' = []"
       
  3760     proof(rule ccontr)
       
  3761       assume otherwise: "xs' \<noteq> []"
       
  3762       from rpath_plus[OF less_2(3) this]
       
  3763       have "(Th th2, Th th1) \<in> (RAG s)\<^sup>+" .
       
  3764       from tranclD[OF this]
       
  3765       obtain cs where "waiting s th2 cs"
       
  3766         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  3767       with runing_2 show False
       
  3768         by (unfold runing_def readys_def, auto)
       
  3769     qed
       
  3770     ultimately have "xs2 = xs1" by simp
       
  3771     from rpath_dest_eq[OF rp1 rp2[unfolded this]]
       
  3772     show ?thesis by simp
       
  3773   qed
       
  3774 qed
       
  3775 
       
  3776 lemma card_runing: "card (runing s) \<le> 1"
       
  3777 proof(cases "runing s = {}")
       
  3778   case True
       
  3779   thus ?thesis by auto
       
  3780 next
       
  3781   case False
       
  3782   then obtain th where [simp]: "th \<in> runing s" by auto
       
  3783   from runing_unique[OF this]
       
  3784   have "runing s = {th}" by auto
       
  3785   thus ?thesis by auto
       
  3786 qed
       
  3787 
       
  3788 lemma create_pre:
       
  3789   assumes stp: "step s e"
       
  3790   and not_in: "th \<notin> threads s"
       
  3791   and is_in: "th \<in> threads (e#s)"
       
  3792   obtains prio where "e = Create th prio"
       
  3793 proof -
       
  3794   from assms  
       
  3795   show ?thesis
       
  3796   proof(cases)
       
  3797     case (thread_create thread prio)
       
  3798     with is_in not_in have "e = Create th prio" by simp
       
  3799     from that[OF this] show ?thesis .
       
  3800   next
       
  3801     case (thread_exit thread)
       
  3802     with assms show ?thesis by (auto intro!:that)
       
  3803   next
       
  3804     case (thread_P thread)
       
  3805     with assms show ?thesis by (auto intro!:that)
       
  3806   next
       
  3807     case (thread_V thread)
       
  3808     with assms show ?thesis by (auto intro!:that)
       
  3809   next 
       
  3810     case (thread_set thread)
       
  3811     with assms show ?thesis by (auto intro!:that)
       
  3812   qed
       
  3813 qed
       
  3814 
       
  3815 lemma eq_pv_children:
       
  3816   assumes eq_pv: "cntP s th = cntV s th"
       
  3817   shows "children (RAG s) (Th th) = {}"
       
  3818 proof -
       
  3819     from cnp_cnv_cncs and eq_pv
       
  3820     have "cntCS s th = 0" 
       
  3821       by (auto split:if_splits)
       
  3822     from this[unfolded cntCS_def holdents_alt_def]
       
  3823     have card_0: "card (the_cs ` children (RAG s) (Th th)) = 0" .
       
  3824     have "finite (the_cs ` children (RAG s) (Th th))"
       
  3825       by (simp add: fsbtRAGs.finite_children)
       
  3826     from card_0[unfolded card_0_eq[OF this]]
       
  3827     show ?thesis by auto
       
  3828 qed
       
  3829 
       
  3830 lemma eq_pv_holdents:
       
  3831   assumes eq_pv: "cntP s th = cntV s th"
       
  3832   shows "holdents s th = {}"
       
  3833   by (unfold holdents_alt_def eq_pv_children[OF assms], simp)
       
  3834 
       
  3835 lemma eq_pv_subtree:
       
  3836   assumes eq_pv: "cntP s th = cntV s th"
       
  3837   shows "subtree (RAG s) (Th th) = {Th th}"
       
  3838   using eq_pv_children[OF assms]
       
  3839     by (unfold subtree_children, simp)
       
  3840 
       
  3841 end
       
  3842 
       
  3843 lemma cp_gen_alt_def:
       
  3844   "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
       
  3845     by (auto simp:cp_gen_def)
       
  3846 
       
  3847 lemma tRAG_nodeE:
  2430 lemma tRAG_nodeE:
  3848   assumes "(n1, n2) \<in> tRAG s"
  2431   assumes "(n1, n2) \<in> tRAG s"
  3849   obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
  2432   obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
  3850   using assms
  2433   using assms
  3851   by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
  2434   by (auto simp: tRAG_def wRAG_def hRAG_def)
  3852 
  2435 
       
  2436 lemma tRAG_ancestorsE:
       
  2437   assumes "x \<in> ancestors (tRAG s) u"
       
  2438   obtains th where "x = Th th"
       
  2439 proof -
       
  2440   from assms have "(u, x) \<in> (tRAG s)^+" 
       
  2441       by (unfold ancestors_def, auto)
       
  2442   from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
       
  2443   then obtain th where "x = Th th"
       
  2444     by (unfold tRAG_alt_def, auto)
       
  2445   from that[OF this] show ?thesis .
       
  2446 qed
       
  2447                    
  3853 lemma subtree_nodeE:
  2448 lemma subtree_nodeE:
  3854   assumes "n \<in> subtree (tRAG s) (Th th)"
  2449   assumes "n \<in> subtree (tRAG s) (Th th)"
  3855   obtains th1 where "n = Th th1"
  2450   obtains th1 where "n = Th th1"
  3856 proof -
  2451 proof -
  3857   show ?thesis
  2452   show ?thesis
  3971 lemma tRAG_trancl_eq_Th:
  2566 lemma tRAG_trancl_eq_Th:
  3972    "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
  2567    "{Th th' | th'. (Th th', Th th)  \<in> (tRAG s)^+} = 
  3973     {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
  2568     {Th th' | th'. (Th th', Th th)  \<in> (RAG s)^+}"
  3974     using tRAG_trancl_eq by auto
  2569     using tRAG_trancl_eq by auto
  3975 
  2570 
  3976 lemma dependants_alt_def:
       
  3977   "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
       
  3978   by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
       
  3979 
       
  3980 lemma dependants_alt_def1:
       
  3981   "dependants (s::state) th = {th'. (Th th', Th th) \<in> (RAG s)^+}"
       
  3982   using dependants_alt_def tRAG_trancl_eq by auto
       
  3983 
       
  3984 context valid_trace
       
  3985 begin
       
  3986 lemma count_eq_RAG_plus:
       
  3987   assumes "cntP s th = cntV s th"
       
  3988   shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  3989 proof(rule ccontr)
       
  3990     assume otherwise: "{th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} \<noteq> {}"
       
  3991     then obtain th' where "(Th th', Th th) \<in> (RAG s)^+" by auto
       
  3992     from tranclD2[OF this]
       
  3993     obtain z where "z \<in> children (RAG s) (Th th)" 
       
  3994       by (auto simp:children_def)
       
  3995     with eq_pv_children[OF assms]
       
  3996     show False by simp
       
  3997 qed
       
  3998 
       
  3999 lemma eq_pv_dependants:
       
  4000   assumes eq_pv: "cntP s th = cntV s th"
       
  4001   shows "dependants s th = {}"
       
  4002 proof -
       
  4003   from count_eq_RAG_plus[OF assms, folded dependants_alt_def1]
       
  4004   show ?thesis .
       
  4005 qed
       
  4006 
       
  4007 end
       
  4008 
       
  4009 lemma eq_dependants: "dependants (wq s) = dependants s"
       
  4010   by (simp add: s_dependants_abv wq_def)
       
  4011 
       
  4012 context valid_trace
       
  4013 begin
       
  4014 
       
  4015 lemma count_eq_tRAG_plus:
       
  4016   assumes "cntP s th = cntV s th"
       
  4017   shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
       
  4018   using assms eq_pv_dependants dependants_alt_def eq_dependants by auto 
       
  4019 
       
  4020 lemma count_eq_RAG_plus_Th:
       
  4021   assumes "cntP s th = cntV s th"
       
  4022   shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  4023   using count_eq_RAG_plus[OF assms] by auto
       
  4024 
       
  4025 lemma count_eq_tRAG_plus_Th:
       
  4026   assumes "cntP s th = cntV s th"
       
  4027   shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
       
  4028    using count_eq_tRAG_plus[OF assms] by auto
       
  4029 end
       
  4030 
       
  4031 lemma inj_the_preced: 
       
  4032   "inj_on (the_preced s) (threads s)"
       
  4033   by (metis inj_onI preced_unique the_preced_def)
       
  4034 
  2571 
  4035 lemma tRAG_Field:
  2572 lemma tRAG_Field:
  4036   "Field (tRAG s) \<subseteq> Field (RAG s)"
  2573   "Field (tRAG s) \<subseteq> Field (RAG s)"
  4037   by (unfold tRAG_alt_def Field_def, auto)
  2574   by (unfold tRAG_alt_def Field_def, auto)
  4038 
       
  4039 lemma tRAG_ancestorsE:
       
  4040   assumes "x \<in> ancestors (tRAG s) u"
       
  4041   obtains th where "x = Th th"
       
  4042 proof -
       
  4043   from assms have "(u, x) \<in> (tRAG s)^+" 
       
  4044       by (unfold ancestors_def, auto)
       
  4045   from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
       
  4046   then obtain th where "x = Th th"
       
  4047     by (unfold tRAG_alt_def, auto)
       
  4048   from that[OF this] show ?thesis .
       
  4049 qed
       
  4050 
  2575 
  4051 lemma tRAG_mono:
  2576 lemma tRAG_mono:
  4052   assumes "RAG s' \<subseteq> RAG s"
  2577   assumes "RAG s' \<subseteq> RAG s"
  4053   shows "tRAG s' \<subseteq> tRAG s"
  2578   shows "tRAG s' \<subseteq> tRAG s"
  4054   using assms 
  2579   using assms 
  4055   by (unfold tRAG_alt_def, auto)
  2580   by (unfold tRAG_alt_def, auto)
  4056 
       
  4057 lemma holding_next_thI:
       
  4058   assumes "holding s th cs"
       
  4059   and "length (wq s cs) > 1"
       
  4060   obtains th' where "next_th s th cs th'"
       
  4061 proof -
       
  4062   from assms(1)[folded holding_eq, unfolded cs_holding_def]
       
  4063   have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" 
       
  4064     by (unfold s_holding_def, fold wq_def, auto)
       
  4065   then obtain rest where h1: "wq s cs = th#rest" 
       
  4066     by (cases "wq s cs", auto)
       
  4067   with assms(2) have h2: "rest \<noteq> []" by auto
       
  4068   let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
  4069   have "next_th s th cs ?th'" using  h1(1) h2 
       
  4070     by (unfold next_th_def, auto)
       
  4071   from that[OF this] show ?thesis .
       
  4072 qed
       
  4073 
       
  4074 lemma RAG_tRAG_transfer:
       
  4075   assumes "vt s'"
       
  4076   assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
       
  4077   and "(Cs cs, Th th'') \<in> RAG s'"
       
  4078   shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
       
  4079 proof -
       
  4080   interpret vt_s': valid_trace "s'" using assms(1)
       
  4081     by (unfold_locales, simp)
       
  4082   { fix n1 n2
       
  4083     assume "(n1, n2) \<in> ?L"
       
  4084     from this[unfolded tRAG_alt_def]
       
  4085     obtain th1 th2 cs' where 
       
  4086       h: "n1 = Th th1" "n2 = Th th2" 
       
  4087          "(Th th1, Cs cs') \<in> RAG s"
       
  4088          "(Cs cs', Th th2) \<in> RAG s" by auto
       
  4089     from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
       
  4090     from h(3) and assms(2) 
       
  4091     have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
       
  4092           (Th th1, Cs cs') \<in> RAG s'" by auto
       
  4093     hence "(n1, n2) \<in> ?R"
       
  4094     proof
       
  4095       assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
       
  4096       hence eq_th1: "th1 = th" by simp
       
  4097       moreover have "th2 = th''"
       
  4098       proof -
       
  4099         from h1 have "cs' = cs" by simp
       
  4100         from assms(3) cs_in[unfolded this]
       
  4101         show ?thesis using vt_s'.unique_RAG by auto 
       
  4102       qed
       
  4103       ultimately show ?thesis using h(1,2) by auto
       
  4104     next
       
  4105       assume "(Th th1, Cs cs') \<in> RAG s'"
       
  4106       with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
       
  4107         by (unfold tRAG_alt_def, auto)
       
  4108       from this[folded h(1, 2)] show ?thesis by auto
       
  4109     qed
       
  4110   } moreover {
       
  4111     fix n1 n2
       
  4112     assume "(n1, n2) \<in> ?R"
       
  4113     hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
       
  4114     hence "(n1, n2) \<in> ?L" 
       
  4115     proof
       
  4116       assume "(n1, n2) \<in> tRAG s'"
       
  4117       moreover have "... \<subseteq> ?L"
       
  4118       proof(rule tRAG_mono)
       
  4119         show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
       
  4120       qed
       
  4121       ultimately show ?thesis by auto
       
  4122     next
       
  4123       assume eq_n: "(n1, n2) = (Th th, Th th'')"
       
  4124       from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
       
  4125       moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
       
  4126       ultimately show ?thesis 
       
  4127         by (unfold eq_n tRAG_alt_def, auto)
       
  4128     qed
       
  4129   } ultimately show ?thesis by auto
       
  4130 qed
       
  4131 
       
  4132 context valid_trace
       
  4133 begin
       
  4134 
       
  4135 lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
       
  4136 
       
  4137 end
       
  4138 
  2581 
  4139 lemma tRAG_subtree_eq: 
  2582 lemma tRAG_subtree_eq: 
  4140    "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
  2583    "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
  4141    (is "?L = ?R")
  2584    (is "?L = ?R")
  4142 proof -
  2585 proof -
  4162 lemma threads_set_eq: 
  2605 lemma threads_set_eq: 
  4163    "the_thread ` (subtree (tRAG s) (Th th)) = 
  2606    "the_thread ` (subtree (tRAG s) (Th th)) = 
  4164                   {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
  2607                   {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
  4165    by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
  2608    by (auto intro:rev_image_eqI simp:tRAG_subtree_eq)
  4166 
  2609 
  4167 lemma cp_alt_def1: 
       
  4168   "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
       
  4169 proof -
       
  4170   have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
       
  4171        ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
       
  4172        by auto
       
  4173   thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
       
  4174 qed
       
  4175 
       
  4176 lemma cp_gen_def_cond: 
       
  4177   assumes "x = Th th"
       
  4178   shows "cp s th = cp_gen s (Th th)"
       
  4179 by (unfold cp_alt_def1 cp_gen_def, simp)
       
  4180 
       
  4181 lemma cp_gen_over_set:
       
  4182   assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
       
  4183   shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
       
  4184 proof(rule f_image_eq)
       
  4185   fix a
       
  4186   assume "a \<in> A"
       
  4187   from assms[rule_format, OF this]
       
  4188   obtain th where eq_a: "a = Th th" by auto
       
  4189   show "cp_gen s a = (cp s \<circ> the_thread) a"
       
  4190     by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
       
  4191 qed
       
  4192 
       
  4193 context valid_trace
  2610 context valid_trace
  4194 begin
  2611 begin
       
  2612 
       
  2613 lemma RAG_tRAG_transfer:
       
  2614   assumes  "RAG s' = RAG s \<union> {(Th th, Cs cs)}"
       
  2615   and "(Cs cs, Th th'') \<in> RAG s"
       
  2616   shows "tRAG s' = tRAG s \<union> {(Th th, Th th'')}" (is "?L = ?R")
       
  2617 proof -
       
  2618   { fix n1 n2
       
  2619     assume "(n1, n2) \<in> ?L"
       
  2620     from this[unfolded tRAG_alt_def]
       
  2621     obtain th1 th2 cs' where 
       
  2622       h: "n1 = Th th1" "n2 = Th th2" 
       
  2623          "(Th th1, Cs cs') \<in> RAG s'"
       
  2624          "(Cs cs', Th th2) \<in> RAG s'" by auto
       
  2625     from h(4) and assms(1) have cs_in: "(Cs cs', Th th2) \<in> RAG s" by auto
       
  2626     from h(3) and assms(1) 
       
  2627     have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
       
  2628           (Th th1, Cs cs') \<in> RAG s" by auto
       
  2629     hence "(n1, n2) \<in> ?R"
       
  2630     proof
       
  2631       assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
       
  2632       hence eq_th1: "th1 = th" by simp
       
  2633       moreover have "th2 = th''"
       
  2634       proof -
       
  2635         from h1 have "cs' = cs" by simp
       
  2636         from assms(2) cs_in[unfolded this]
       
  2637         show ?thesis using unique_RAG by auto 
       
  2638       qed
       
  2639       ultimately show ?thesis using h(1,2) by auto
       
  2640     next
       
  2641       assume "(Th th1, Cs cs') \<in> RAG s"
       
  2642       with cs_in have "(Th th1, Th th2) \<in> tRAG s"
       
  2643         by (unfold tRAG_alt_def, auto)
       
  2644       from this[folded h(1, 2)] show ?thesis by auto
       
  2645     qed
       
  2646   } moreover {
       
  2647     fix n1 n2
       
  2648     assume "(n1, n2) \<in> ?R"
       
  2649     hence "(n1, n2) \<in>tRAG s \<or> (n1, n2) = (Th th, Th th'')" by auto
       
  2650     hence "(n1, n2) \<in> ?L" 
       
  2651     proof
       
  2652       assume "(n1, n2) \<in> tRAG s"
       
  2653       moreover have "... \<subseteq> ?L"
       
  2654       proof(rule tRAG_mono)
       
  2655         show "RAG s \<subseteq> RAG s'" by (unfold assms(1), auto)
       
  2656       qed
       
  2657       ultimately show ?thesis by auto
       
  2658     next
       
  2659       assume eq_n: "(n1, n2) = (Th th, Th th'')"
       
  2660       from assms(1, 2) have "(Cs cs, Th th'') \<in> RAG s'" by auto
       
  2661       moreover have "(Th th, Cs cs) \<in> RAG s'" using assms(1) by auto
       
  2662       ultimately show ?thesis 
       
  2663         by (unfold eq_n tRAG_alt_def, auto)
       
  2664     qed
       
  2665   } ultimately show ?thesis by auto
       
  2666 qed
  4195 
  2667 
  4196 lemma subtree_tRAG_thread:
  2668 lemma subtree_tRAG_thread:
  4197   assumes "th \<in> threads s"
  2669   assumes "th \<in> threads s"
  4198   shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
  2670   shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
  4199 proof -
  2671 proof -
  4215     qed
  2687     qed
  4216   qed
  2688   qed
  4217   finally show ?thesis .
  2689   finally show ?thesis .
  4218 qed
  2690 qed
  4219 
  2691 
  4220 lemma readys_root:
  2692 lemma dependants_alt_def:
  4221   assumes "th \<in> readys s"
  2693   "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
  4222   shows "root (RAG s) (Th th)"
  2694   by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
  4223 proof -
  2695 
  4224   { fix x
  2696 lemma dependants_alt_def1:
  4225     assume "x \<in> ancestors (RAG s) (Th th)"
  2697   "dependants (s::state) th = {th'. (Th th', Th th) \<in> (RAG s)^+}"
  4226     hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
  2698   using dependants_alt_def tRAG_trancl_eq by auto
  4227     from tranclD[OF this]
  2699 
  4228     obtain z where "(Th th, z) \<in> RAG s" by auto
  2700 end
  4229     with assms(1) have False
  2701 
  4230          apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
  2702 section {* Chain to readys *}
  4231          by (fold wq_def, blast)
       
  4232   } thus ?thesis by (unfold root_def, auto)
       
  4233 qed
       
  4234 
       
  4235 lemma readys_in_no_subtree:
       
  4236   assumes "th \<in> readys s"
       
  4237   and "th' \<noteq> th"
       
  4238   shows "Th th \<notin> subtree (RAG s) (Th th')" 
       
  4239 proof
       
  4240    assume "Th th \<in> subtree (RAG s) (Th th')"
       
  4241    thus False
       
  4242    proof(cases rule:subtreeE)
       
  4243       case 1
       
  4244       with assms show ?thesis by auto
       
  4245    next
       
  4246       case 2
       
  4247       with readys_root[OF assms(1)]
       
  4248       show ?thesis by (auto simp:root_def)
       
  4249    qed
       
  4250 qed
       
  4251 
       
  4252 lemma not_in_thread_isolated:
       
  4253   assumes "th \<notin> threads s"
       
  4254   shows "(Th th) \<notin> Field (RAG s)"
       
  4255 proof
       
  4256   assume "(Th th) \<in> Field (RAG s)"
       
  4257   with dm_RAG_threads and rg_RAG_threads assms
       
  4258   show False by (unfold Field_def, blast)
       
  4259 qed
       
  4260 
       
  4261 end
       
  4262 
       
  4263 definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
       
  4264   where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
       
  4265 
       
  4266 
       
  4267 lemma detached_test:
       
  4268   shows "detached s th = (Th th \<notin> Field (RAG s))"
       
  4269 apply(simp add: detached_def Field_def)
       
  4270 apply(simp add: s_RAG_def)
       
  4271 apply(simp add: s_holding_abv s_waiting_abv)
       
  4272 apply(simp add: Domain_iff Range_iff)
       
  4273 apply(simp add: wq_def)
       
  4274 apply(auto)
       
  4275 done
       
  4276 
  2703 
  4277 context valid_trace
  2704 context valid_trace
  4278 begin
  2705 begin
  4279 
  2706 
  4280 lemma detached_intro:
  2707 lemma chain_building:
  4281   assumes eq_pv: "cntP s th = cntV s th"
  2708   assumes "node \<in> Domain (RAG s)"
  4282   shows "detached s th"
  2709   obtains th' where "th' \<in> readys s" "(node, Th th') \<in> (RAG s)^+"
  4283 proof -
  2710 proof -
  4284   from eq_pv cnp_cnv_cncs
  2711   from assms have "node \<in> Range ((RAG s)^-1)" by auto
  4285   have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:pvD_def)
  2712   from wf_base[OF wf_RAG_converse this]
  4286   thus ?thesis
  2713   obtain b where h_b: "(b, node) \<in> ((RAG s)\<inverse>)\<^sup>+" "\<forall>c. (c, b) \<notin> (RAG s)\<inverse>" by auto
  4287   proof
  2714   obtain th' where eq_b: "b = Th th'"
  4288     assume "th \<notin> threads s"
  2715   proof(cases b)
  4289     with rg_RAG_threads dm_RAG_threads
  2716     case (Cs cs)
  4290     show ?thesis
  2717     from h_b(1)[unfolded trancl_converse] 
  4291       by (auto simp add: detached_def s_RAG_def s_waiting_abv 
  2718     have "(node, b) \<in> ((RAG s)\<^sup>+)" by auto
  4292               s_holding_abv wq_def Domain_iff Range_iff)
  2719     from tranclE[OF this]
  4293   next
  2720     obtain n where "(n, b) \<in> RAG s" by auto
  4294     assume "th \<in> readys s"
  2721     from this[unfolded Cs]
  4295     moreover have "Th th \<notin> Range (RAG s)"
  2722     obtain th1 where "waiting s th1 cs"
  4296     proof -
  2723       by (unfold s_RAG_def, fold waiting_eq, auto)
  4297       from eq_pv_children[OF assms]
  2724     from waiting_holding[OF this]
  4298       have "children (RAG s) (Th th) = {}" .
  2725     obtain th2 where "holding s th2 cs" .
  4299       thus ?thesis
  2726     hence "(Cs cs, Th th2) \<in> RAG s"
  4300       by (unfold children_def, auto)
  2727       by (unfold s_RAG_def, fold holding_eq, auto)
  4301     qed
  2728     with h_b(2)[unfolded Cs, rule_format]
  4302     ultimately show ?thesis
  2729     have False by auto
  4303       by (auto simp add: detached_def s_RAG_def s_waiting_abv 
  2730     thus ?thesis by auto
  4304               s_holding_abv wq_def readys_def)
  2731   qed auto
  4305   qed
  2732   have "th' \<in> readys s" 
  4306 qed
       
  4307 
       
  4308 lemma detached_elim:
       
  4309   assumes dtc: "detached s th"
       
  4310   shows "cntP s th = cntV s th"
       
  4311 proof -
       
  4312   have cncs_z: "cntCS s th = 0"
       
  4313   proof -
  2733   proof -
  4314     from dtc have "holdents s th = {}"
  2734     from h_b(2)[unfolded eq_b]
  4315       unfolding detached_def holdents_test s_RAG_def
  2735     have "\<forall>cs. \<not> waiting s th' cs"
  4316       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
  2736       by (unfold s_RAG_def, fold waiting_eq, auto)
  4317     thus ?thesis by (auto simp:cntCS_def)
  2737     moreover have "th' \<in> threads s"
  4318   qed
  2738     proof(rule rg_RAG_threads)
  4319   show ?thesis
  2739       from tranclD[OF h_b(1), unfolded eq_b]
  4320   proof(cases "th \<in> threads s")
  2740       obtain z where "(z, Th th') \<in> (RAG s)" by auto
  4321     case True
  2741       thus "Th th' \<in> Range (RAG s)" by auto
  4322     with dtc 
  2742     qed
  4323     have "th \<in> readys s"
  2743     ultimately show ?thesis by (auto simp:readys_def)
  4324       by (unfold readys_def detached_def Field_def Domain_def Range_def, 
  2744   qed
  4325            auto simp:waiting_eq s_RAG_def)
  2745   moreover have "(node, Th th') \<in> (RAG s)^+" 
  4326     with cncs_z  show ?thesis using cnp_cnv_cncs by (simp add:pvD_def)
  2746     using h_b(1)[unfolded trancl_converse] eq_b by auto
  4327   next
  2747   ultimately show ?thesis using that by metis
  4328     case False
  2748 qed
  4329     with cncs_z and cnp_cnv_cncs show ?thesis by (simp add:pvD_def)
  2749 
  4330   qed
  2750 text {* \noindent
  4331 qed
  2751   The following is just an instance of @{text "chain_building"}.
  4332 
  2752 *}                    
  4333 lemma detached_eq:
  2753 lemma th_chain_to_ready:
  4334   shows "(detached s th) = (cntP s th = cntV s th)"
  2754   assumes th_in: "th \<in> threads s"
  4335   by (insert vt, auto intro:detached_intro detached_elim)
  2755   shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)"
  4336 
  2756 proof(cases "th \<in> readys s")
  4337 end
       
  4338 
       
  4339 context valid_trace
       
  4340 begin
       
  4341 (* ddd *)
       
  4342 lemma cp_gen_rec:
       
  4343   assumes "x = Th th"
       
  4344   shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
       
  4345 proof(cases "children (tRAG s) x = {}")
       
  4346   case True
  2757   case True
  4347   show ?thesis
  2758   thus ?thesis by auto
  4348     by (unfold True cp_gen_def subtree_children, simp add:assms)
       
  4349 next
  2759 next
  4350   case False
  2760   case False
  4351   hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
  2761   from False and th_in have "Th th \<in> Domain (RAG s)" 
  4352   note fsbttRAGs.finite_subtree[simp]
  2762     by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def)
  4353   have [simp]: "finite (children (tRAG s) x)"
  2763   from chain_building [rule_format, OF this]
  4354      by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
  2764   show ?thesis by auto
  4355             rule children_subtree)
  2765 qed
  4356   { fix r x
  2766 
  4357     have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
  2767 lemma finite_subtree_threads:
  4358   } note this[simp]
  2768     "finite {th'. Th th' \<in> subtree (RAG s) (Th th)}" (is "finite ?A")
  4359   have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
  2769 proof -
       
  2770   have "?A = the_thread ` {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
       
  2771         by (auto, insert image_iff, fastforce)
       
  2772   moreover have "finite {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
       
  2773         (is "finite ?B")
  4360   proof -
  2774   proof -
  4361     from False obtain q where "q \<in> children (tRAG s) x" by blast
  2775      have "?B = (subtree (RAG s) (Th th)) \<inter> {Th th' | th'. True}"
  4362     moreover have "subtree (tRAG s) q \<noteq> {}" by simp
  2776       by auto
  4363     ultimately show ?thesis by blast
  2777      moreover have "... \<subseteq> (subtree (RAG s) (Th th))" by auto
  4364   qed
  2778      moreover have "finite ..." by (simp add: fsbtRAGs.finite_subtree) 
  4365   have h: "Max ((the_preced s \<circ> the_thread) `
  2779      ultimately show ?thesis by auto
  4366                 ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
  2780   qed
  4367         Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
  2781   ultimately show ?thesis by auto
  4368                      (is "?L = ?R")
  2782 qed
  4369   proof -
  2783 
  4370     let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
  2784 lemma runing_unique:
  4371     let "Max (_ \<union> (?h ` ?B))" = ?R
  2785   assumes runing_1: "th1 \<in> runing s"
  4372     let ?L1 = "?f ` \<Union>(?g ` ?B)"
  2786   and runing_2: "th2 \<in> runing s"
  4373     have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
  2787   shows "th1 = th2"
  4374     proof -
  2788 proof -
  4375       have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
  2789   from runing_1 and runing_2 have "cp s th1 = cp s th2"
  4376       also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
  2790     unfolding runing_def by auto
  4377       finally have "Max ?L1 = Max ..." by simp
  2791   from this[unfolded cp_alt_def]
  4378       also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
  2792   have eq_max: 
  4379         by (subst Max_UNION, simp+)
  2793     "Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th1)}) =
  4380       also have "... = Max (cp_gen s ` children (tRAG s) x)"
  2794      Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th2)})" 
  4381           by (unfold image_comp cp_gen_alt_def, simp)
  2795         (is "Max ?L = Max ?R") .
  4382       finally show ?thesis .
  2796   have "Max ?L \<in> ?L"
  4383     qed
  2797   proof(rule Max_in)
  4384     show ?thesis
  2798     show "finite ?L" by (simp add: finite_subtree_threads) 
  4385     proof -
       
  4386       have "?L = Max (?f ` ?A \<union> ?L1)" by simp
       
  4387       also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
       
  4388             by (subst Max_Un, simp+)
       
  4389       also have "... = max (?f x) (Max (?h ` ?B))"
       
  4390         by (unfold eq_Max_L1, simp)
       
  4391       also have "... =?R"
       
  4392         by (rule max_Max_eq, (simp)+, unfold assms, simp)
       
  4393       finally show ?thesis .
       
  4394     qed
       
  4395   qed  thus ?thesis 
       
  4396           by (fold h subtree_children, unfold cp_gen_def, simp) 
       
  4397 qed
       
  4398 
       
  4399 lemma cp_rec:
       
  4400   "cp s th = Max ({the_preced s th} \<union> 
       
  4401                      (cp s o the_thread) ` children (tRAG s) (Th th))"
       
  4402 proof -
       
  4403   have "Th th = Th th" by simp
       
  4404   note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
       
  4405   show ?thesis 
       
  4406   proof -
       
  4407     have "cp_gen s ` children (tRAG s) (Th th) = 
       
  4408                 (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
       
  4409     proof(rule cp_gen_over_set)
       
  4410       show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
       
  4411         by (unfold tRAG_alt_def, auto simp:children_def)
       
  4412     qed
       
  4413     thus ?thesis by (subst (1) h(1), unfold h(2), simp)
       
  4414   qed
       
  4415 qed
       
  4416 
       
  4417 lemma next_th_holding:
       
  4418   assumes nxt: "next_th s th cs th'"
       
  4419   shows "holding (wq s) th cs"
       
  4420 proof -
       
  4421   from nxt[unfolded next_th_def]
       
  4422   obtain rest where h: "wq s cs = th # rest"
       
  4423                        "rest \<noteq> []" 
       
  4424                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
       
  4425   thus ?thesis
       
  4426     by (unfold cs_holding_def, auto)
       
  4427 qed
       
  4428 
       
  4429 lemma next_th_waiting:
       
  4430   assumes nxt: "next_th s th cs th'"
       
  4431   shows "waiting (wq s) th' cs"
       
  4432 proof -
       
  4433   from nxt[unfolded next_th_def]
       
  4434   obtain rest where h: "wq s cs = th # rest"
       
  4435                        "rest \<noteq> []" 
       
  4436                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
       
  4437   from wq_distinct[of cs, unfolded h]
       
  4438   have dst: "distinct (th # rest)" .
       
  4439   have in_rest: "th' \<in> set rest"
       
  4440   proof(unfold h, rule someI2)
       
  4441     show "distinct rest \<and> set rest = set rest" using dst by auto
       
  4442   next
  2799   next
  4443     fix x assume "distinct x \<and> set x = set rest"
  2800     show "?L \<noteq> {}" using subtree_def by fastforce 
  4444     with h(2)
  2801   qed
  4445     show "hd x \<in> set (rest)" by (cases x, auto)
  2802   then obtain th1' where 
  4446   qed
  2803     h_1: "Th th1' \<in> subtree (RAG s) (Th th1)" "the_preced s th1' = Max ?L"
  4447   hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
  2804     by auto
  4448   moreover have "th' \<noteq> hd (wq s cs)"
  2805   have "Max ?R \<in> ?R"
  4449     by (unfold h(1), insert in_rest dst, auto)
  2806   proof(rule Max_in)
  4450   ultimately show ?thesis by (auto simp:cs_waiting_def)
  2807     show "finite ?R" by (simp add: finite_subtree_threads)
  4451 qed
  2808   next
  4452 
  2809     show "?R \<noteq> {}" using subtree_def by fastforce 
  4453 lemma next_th_RAG:
  2810   qed
  4454   assumes nxt: "next_th (s::event list) th cs th'"
  2811   then obtain th2' where 
  4455   shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
  2812     h_2: "Th th2' \<in> subtree (RAG s) (Th th2)" "the_preced s th2' = Max ?R"
  4456   using vt assms next_th_holding next_th_waiting
  2813     by auto
  4457   by (unfold s_RAG_def, simp)
  2814   have "th1' = th2'"
  4458 
  2815   proof(rule preced_unique)
  4459 end
  2816     from h_1(1)
  4460 
  2817     show "th1' \<in> threads s"
  4461 lemma next_th_unique: 
  2818     proof(cases rule:subtreeE)
  4462   assumes nt1: "next_th s th cs th1"
  2819       case 1
  4463   and nt2: "next_th s th cs th2"
  2820       hence "th1' = th1" by simp
  4464   shows "th1 = th2"
  2821       with runing_1 show ?thesis by (auto simp:runing_def readys_def)
  4465 using assms by (unfold next_th_def, auto)
  2822     next
       
  2823       case 2
       
  2824       from this(2)
       
  2825       have "(Th th1', Th th1) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
  2826       from tranclD[OF this]
       
  2827       have "(Th th1') \<in> Domain (RAG s)" by auto
       
  2828       from dm_RAG_threads[OF this] show ?thesis .
       
  2829     qed
       
  2830   next
       
  2831     from h_2(1)
       
  2832     show "th2' \<in> threads s"
       
  2833     proof(cases rule:subtreeE)
       
  2834       case 1
       
  2835       hence "th2' = th2" by simp
       
  2836       with runing_2 show ?thesis by (auto simp:runing_def readys_def)
       
  2837     next
       
  2838       case 2
       
  2839       from this(2)
       
  2840       have "(Th th2', Th th2) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
  2841       from tranclD[OF this]
       
  2842       have "(Th th2') \<in> Domain (RAG s)" by auto
       
  2843       from dm_RAG_threads[OF this] show ?thesis .
       
  2844     qed
       
  2845   next
       
  2846     have "the_preced s th1' = the_preced s th2'" 
       
  2847      using eq_max h_1(2) h_2(2) by metis
       
  2848     thus "preced th1' s = preced th2' s" by (simp add:the_preced_def)
       
  2849   qed
       
  2850   from h_1(1)[unfolded this]
       
  2851   have star1: "(Th th2', Th th1) \<in> (RAG s)^*" by (auto simp:subtree_def)
       
  2852   from h_2(1)[unfolded this]
       
  2853   have star2: "(Th th2', Th th2) \<in> (RAG s)^*" by (auto simp:subtree_def)
       
  2854   from star_rpath[OF star1] obtain xs1 
       
  2855     where rp1: "rpath (RAG s) (Th th2') xs1 (Th th1)"
       
  2856     by auto
       
  2857   from star_rpath[OF star2] obtain xs2 
       
  2858     where rp2: "rpath (RAG s) (Th th2') xs2 (Th th2)"
       
  2859     by auto
       
  2860   from rp1 rp2
       
  2861   show ?thesis
       
  2862   proof(cases)
       
  2863     case (less_1 xs')
       
  2864     moreover have "xs' = []"
       
  2865     proof(rule ccontr)
       
  2866       assume otherwise: "xs' \<noteq> []"
       
  2867       from rpath_plus[OF less_1(3) this]
       
  2868       have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+" .
       
  2869       from tranclD[OF this]
       
  2870       obtain cs where "waiting s th1 cs"
       
  2871         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  2872       with runing_1 show False
       
  2873         by (unfold runing_def readys_def, auto)
       
  2874     qed
       
  2875     ultimately have "xs2 = xs1" by simp
       
  2876     from rpath_dest_eq[OF rp1 rp2[unfolded this]]
       
  2877     show ?thesis by simp
       
  2878   next
       
  2879     case (less_2 xs')
       
  2880     moreover have "xs' = []"
       
  2881     proof(rule ccontr)
       
  2882       assume otherwise: "xs' \<noteq> []"
       
  2883       from rpath_plus[OF less_2(3) this]
       
  2884       have "(Th th2, Th th1) \<in> (RAG s)\<^sup>+" .
       
  2885       from tranclD[OF this]
       
  2886       obtain cs where "waiting s th2 cs"
       
  2887         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  2888       with runing_2 show False
       
  2889         by (unfold runing_def readys_def, auto)
       
  2890     qed
       
  2891     ultimately have "xs2 = xs1" by simp
       
  2892     from rpath_dest_eq[OF rp1 rp2[unfolded this]]
       
  2893     show ?thesis by simp
       
  2894   qed
       
  2895 qed
       
  2896 
       
  2897 lemma card_runing: "card (runing s) \<le> 1"
       
  2898 proof(cases "runing s = {}")
       
  2899   case True
       
  2900   thus ?thesis by auto
       
  2901 next
       
  2902   case False
       
  2903   then obtain th where [simp]: "th \<in> runing s" by auto
       
  2904   from runing_unique[OF this]
       
  2905   have "runing s = {th}" by auto
       
  2906   thus ?thesis by auto
       
  2907 qed
       
  2908 
       
  2909 end
       
  2910 
       
  2911 
       
  2912 section {* Relating @{term cp} and @{term the_preced} and @{term preced} *}
  4466 
  2913 
  4467 context valid_trace
  2914 context valid_trace
  4468 begin
  2915 begin
  4469 
  2916 
  4470 thm th_chain_to_ready
  2917 lemma le_cp:
  4471 
  2918   shows "preced th s \<le> cp s th"
  4472 find_theorems subtree Th RAG
  2919   proof(unfold cp_alt_def, rule Max_ge)
       
  2920     show "finite (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})"
       
  2921       by (simp add: finite_subtree_threads)
       
  2922   next
       
  2923     show "preced th s \<in> the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}"
       
  2924       by (simp add: subtree_def the_preced_def)   
       
  2925   qed
       
  2926 
       
  2927 
       
  2928 lemma cp_le:
       
  2929   assumes th_in: "th \<in> threads s"
       
  2930   shows "cp s th \<le> Max (the_preced s ` threads s)"
       
  2931 proof(unfold cp_alt_def, rule Max_f_mono)
       
  2932   show "finite (threads s)" by (simp add: finite_threads) 
       
  2933 next
       
  2934   show " {th'. Th th' \<in> subtree (RAG s) (Th th)} \<noteq> {}"
       
  2935     using subtree_def by fastforce
       
  2936 next
       
  2937   show "{th'. Th th' \<in> subtree (RAG s) (Th th)} \<subseteq> threads s"
       
  2938     using assms
       
  2939     by (smt Domain.DomainI dm_RAG_threads mem_Collect_eq 
       
  2940         node.inject(1) rtranclD subsetI subtree_def trancl_domain) 
       
  2941 qed
       
  2942 
       
  2943 lemma max_cp_eq: 
       
  2944   shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)"
       
  2945   (is "?L = ?R")
       
  2946 proof -
       
  2947   have "?L \<le> ?R" 
       
  2948   proof(cases "threads s = {}")
       
  2949     case False
       
  2950     show ?thesis 
       
  2951       by (rule Max.boundedI, 
       
  2952           insert cp_le, 
       
  2953           auto simp:finite_threads False)
       
  2954   qed auto
       
  2955   moreover have "?R \<le> ?L"
       
  2956     by (rule Max_fg_mono, 
       
  2957         simp add: finite_threads,
       
  2958         simp add: le_cp the_preced_def)
       
  2959   ultimately show ?thesis by auto
       
  2960 qed
  4473 
  2961 
  4474 lemma threads_alt_def:
  2962 lemma threads_alt_def:
  4475   "(threads s) = (\<Union> th \<in> readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)})"
  2963   "(threads s) = (\<Union> th \<in> readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)})"
  4476     (is "?L = ?R")
  2964     (is "?L = ?R")
  4477 proof -
  2965 proof -
  4498       show ?thesis .
  2986       show ?thesis .
  4499     qed
  2987     qed
  4500   } ultimately show ?thesis by auto
  2988   } ultimately show ?thesis by auto
  4501 qed
  2989 qed
  4502 
  2990 
  4503 lemma finite_readys [simp]: "finite (readys s)"
       
  4504   using finite_threads readys_threads rev_finite_subset by blast
       
  4505 
  2991 
  4506 text {* (* ccc *) \noindent
  2992 text {* (* ccc *) \noindent
  4507   Since the current precedence of the threads in ready queue will always be boosted,
  2993   Since the current precedence of the threads in ready queue will always be boosted,
  4508   there must be one inside it has the maximum precedence of the whole system. 
  2994   there must be one inside it has the maximum precedence of the whole system. 
  4509 *}
  2995 *}
  4542   finally show ?thesis by simp
  3028   finally show ?thesis by simp
  4543 qed (auto simp:threads_alt_def)
  3029 qed (auto simp:threads_alt_def)
  4544 
  3030 
  4545 end
  3031 end
  4546 
  3032 
  4547 end
  3033 section {* Relating @{term cntP}, @{term cntV}, @{term cntCS} and @{term pvD} *}
  4548 
  3034 
       
  3035 context valid_trace_p_w
       
  3036 begin
       
  3037 
       
  3038 lemma holding_s_holder: "holding s holder cs"
       
  3039   by (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
       
  3040 
       
  3041 lemma holding_es_holder: "holding (e#s) holder cs"
       
  3042   by (unfold s_holding_def, fold wq_def, unfold wq_es_cs wq_s_cs, auto)
       
  3043 
       
  3044 lemma holdents_es:
       
  3045   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") 
       
  3046 proof -
       
  3047   { fix cs'
       
  3048     assume "cs' \<in> ?L"
       
  3049     hence h: "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  3050     have "holding s th' cs'"
       
  3051     proof(cases "cs' = cs")
       
  3052       case True
       
  3053       from held_unique[OF h[unfolded True] holding_es_holder]
       
  3054       have "th' = holder" .
       
  3055       thus ?thesis 
       
  3056         by (unfold True holdents_def, insert holding_s_holder, simp)
       
  3057     next
       
  3058       case False
       
  3059       hence "wq (e#s) cs' = wq s cs'" by simp
       
  3060       from h[unfolded s_holding_def, folded wq_def, unfolded this]
       
  3061       show ?thesis
       
  3062        by (unfold s_holding_def, fold wq_def, auto)
       
  3063     qed 
       
  3064     hence "cs' \<in> ?R" by (auto simp:holdents_def)
       
  3065   } moreover {
       
  3066     fix cs'
       
  3067     assume "cs' \<in> ?R"
       
  3068     hence h: "holding s th' cs'" by (auto simp:holdents_def)
       
  3069     have "holding (e#s) th' cs'"
       
  3070     proof(cases "cs' = cs")
       
  3071       case True
       
  3072       from held_unique[OF h[unfolded True] holding_s_holder]
       
  3073       have "th' = holder" .
       
  3074       thus ?thesis 
       
  3075         by (unfold True holdents_def, insert holding_es_holder, simp)
       
  3076     next
       
  3077       case False
       
  3078       hence "wq s cs' = wq (e#s) cs'" by simp
       
  3079       from h[unfolded s_holding_def, folded wq_def, unfolded this]
       
  3080       show ?thesis
       
  3081        by (unfold s_holding_def, fold wq_def, auto)
       
  3082     qed 
       
  3083     hence "cs' \<in> ?L" by (auto simp:holdents_def)
       
  3084   } ultimately show ?thesis by auto
       
  3085 qed
       
  3086 
       
  3087 lemma cntCS_es_th[simp]: "cntCS (e#s) th' = cntCS s th'"
       
  3088  by (unfold cntCS_def holdents_es, simp)
       
  3089 
       
  3090 lemma th_not_ready_es: 
       
  3091   shows "th \<notin> readys (e#s)"
       
  3092   using waiting_es_th_cs 
       
  3093   by (unfold readys_def, auto)
       
  3094 
       
  3095 end
       
  3096   
       
  3097 lemma (in valid_trace) finite_holdents: "finite (holdents s th)"
       
  3098   by (unfold holdents_alt_def, insert fsbtRAGs.finite_children, auto)
       
  3099 
       
  3100 context valid_trace_p 
       
  3101 begin
       
  3102 
       
  3103 lemma ready_th_s: "th \<in> readys s"
       
  3104   using runing_th_s
       
  3105   by (unfold runing_def, auto)
       
  3106 
       
  3107 lemma live_th_s: "th \<in> threads s"
       
  3108   using readys_threads ready_th_s by auto
       
  3109 
       
  3110 lemma live_th_es: "th \<in> threads (e#s)"
       
  3111   using live_th_s 
       
  3112   by (unfold is_p, simp)
       
  3113 
       
  3114 lemma waiting_neq_th: 
       
  3115   assumes "waiting s t c"
       
  3116   shows "t \<noteq> th"
       
  3117   using assms using th_not_waiting by blast 
       
  3118 
       
  3119 end
       
  3120 
       
  3121 context valid_trace_p_h
       
  3122 begin
       
  3123 
       
  3124 lemma th_not_waiting':
       
  3125   "\<not> waiting (e#s) th cs'"
       
  3126 proof(cases "cs' = cs")
       
  3127   case True
       
  3128   show ?thesis
       
  3129     by (unfold True s_waiting_def, fold wq_def, unfold wq_es_cs', auto)
       
  3130 next
       
  3131   case False
       
  3132   from th_not_waiting[of cs', unfolded s_waiting_def, folded wq_def]
       
  3133   show ?thesis
       
  3134     by (unfold s_waiting_def, fold wq_def, insert False, simp)
       
  3135 qed
       
  3136 
       
  3137 lemma ready_th_es: 
       
  3138   shows "th \<in> readys (e#s)"
       
  3139   using th_not_waiting'
       
  3140   by (unfold readys_def, insert live_th_es, auto)
       
  3141 
       
  3142 lemma holdents_es_th:
       
  3143   "holdents (e#s) th = (holdents s th) \<union> {cs}" (is "?L = ?R")
       
  3144 proof -
       
  3145   { fix cs'
       
  3146     assume "cs' \<in> ?L" 
       
  3147     hence "holding (e#s) th cs'"
       
  3148       by (unfold holdents_def, auto)
       
  3149     hence "cs' \<in> ?R"
       
  3150      by (cases rule:holding_esE, auto simp:holdents_def)
       
  3151   } moreover {
       
  3152     fix cs'
       
  3153     assume "cs' \<in> ?R"
       
  3154     hence "holding s th cs' \<or> cs' = cs" 
       
  3155       by (auto simp:holdents_def)
       
  3156     hence "cs' \<in> ?L"
       
  3157     proof
       
  3158       assume "holding s th cs'"
       
  3159       from holding_kept[OF this]
       
  3160       show ?thesis by (auto simp:holdents_def)
       
  3161     next
       
  3162       assume "cs' = cs"
       
  3163       thus ?thesis using holding_es_th_cs
       
  3164         by (unfold holdents_def, auto)
       
  3165     qed
       
  3166   } ultimately show ?thesis by auto
       
  3167 qed
       
  3168 
       
  3169 lemma cntCS_es_th: "cntCS (e#s) th = cntCS s th + 1"
       
  3170 proof -
       
  3171   have "card (holdents s th \<union> {cs}) = card (holdents s th) + 1"
       
  3172   proof(subst card_Un_disjoint)
       
  3173     show "holdents s th \<inter> {cs} = {}"
       
  3174       using not_holding_s_th_cs by (auto simp:holdents_def)
       
  3175   qed (auto simp:finite_holdents)
       
  3176   thus ?thesis
       
  3177    by (unfold cntCS_def holdents_es_th, simp)
       
  3178 qed
       
  3179 
       
  3180 lemma no_holder: 
       
  3181   "\<not> holding s th' cs"
       
  3182 proof
       
  3183   assume otherwise: "holding s th' cs"
       
  3184   from this[unfolded s_holding_def, folded wq_def, unfolded we]
       
  3185   show False by auto
       
  3186 qed
       
  3187 
       
  3188 lemma holdents_es_th':
       
  3189   assumes "th' \<noteq> th"
       
  3190   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3191 proof -
       
  3192   { fix cs'
       
  3193     assume "cs' \<in> ?L"
       
  3194     hence h_e: "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  3195     have "cs' \<noteq> cs"
       
  3196     proof
       
  3197       assume "cs' = cs"
       
  3198       from held_unique[OF h_e[unfolded this] holding_es_th_cs]
       
  3199       have "th' = th" .
       
  3200       with assms show False by simp
       
  3201     qed
       
  3202     from h_e[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp[OF this]]
       
  3203     have "th' \<in> set (wq s cs') \<and> th' = hd (wq s cs')" .
       
  3204     hence "cs' \<in> ?R" 
       
  3205       by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  3206   } moreover {
       
  3207     fix cs'
       
  3208     assume "cs' \<in> ?R"
       
  3209     hence "holding s th' cs'" by (auto simp:holdents_def)
       
  3210     from holding_kept[OF this]
       
  3211     have "holding (e # s) th' cs'" .
       
  3212     hence "cs' \<in> ?L"
       
  3213       by (unfold holdents_def, auto)
       
  3214   } ultimately show ?thesis by auto
       
  3215 qed
       
  3216 
       
  3217 lemma cntCS_es_th'[simp]: 
       
  3218   assumes "th' \<noteq> th"
       
  3219   shows "cntCS (e#s) th' = cntCS s th'"
       
  3220   by (unfold cntCS_def holdents_es_th'[OF assms], simp)
       
  3221 
       
  3222 end
       
  3223 
       
  3224 context valid_trace_p
       
  3225 begin
       
  3226 
       
  3227 lemma readys_kept1: 
       
  3228   assumes "th' \<noteq> th"
       
  3229   and "th' \<in> readys (e#s)"
       
  3230   shows "th' \<in> readys s"
       
  3231 proof -
       
  3232   { fix cs'
       
  3233     assume wait: "waiting s th' cs'"
       
  3234     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3235         using assms(2)[unfolded readys_def] by auto
       
  3236     have False
       
  3237     proof(cases "cs' = cs")
       
  3238       case False
       
  3239       with n_wait wait
       
  3240       show ?thesis 
       
  3241         by (unfold s_waiting_def, fold wq_def, auto)
       
  3242     next
       
  3243       case True
       
  3244       show ?thesis
       
  3245       proof(cases "wq s cs = []")
       
  3246         case True
       
  3247         then interpret vt: valid_trace_p_h
       
  3248           by (unfold_locales, simp)
       
  3249         show ?thesis using n_wait wait waiting_kept by auto 
       
  3250       next
       
  3251         case False
       
  3252         then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  3253         show ?thesis using n_wait wait waiting_kept by blast 
       
  3254       qed
       
  3255     qed
       
  3256   } with assms(2) show ?thesis  
       
  3257     by (unfold readys_def, auto)
       
  3258 qed
       
  3259 
       
  3260 lemma readys_kept2: 
       
  3261   assumes "th' \<noteq> th"
       
  3262   and "th' \<in> readys s"
       
  3263   shows "th' \<in> readys (e#s)"
       
  3264 proof -
       
  3265   { fix cs'
       
  3266     assume wait: "waiting (e#s) th' cs'"
       
  3267     have n_wait: "\<not> waiting s th' cs'" 
       
  3268         using assms(2)[unfolded readys_def] by auto
       
  3269     have False
       
  3270     proof(cases "cs' = cs")
       
  3271       case False
       
  3272       with n_wait wait
       
  3273       show ?thesis 
       
  3274         by (unfold s_waiting_def, fold wq_def, auto)
       
  3275     next
       
  3276       case True
       
  3277       show ?thesis
       
  3278       proof(cases "wq s cs = []")
       
  3279         case True
       
  3280         then interpret vt: valid_trace_p_h
       
  3281           by (unfold_locales, simp)
       
  3282         show ?thesis using n_wait vt.waiting_esE wait by blast 
       
  3283       next
       
  3284         case False
       
  3285         then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  3286         show ?thesis using assms(1) n_wait vt.waiting_esE wait by auto 
       
  3287       qed
       
  3288     qed
       
  3289   } with assms(2) show ?thesis  
       
  3290     by (unfold readys_def, auto)
       
  3291 qed
       
  3292 
       
  3293 lemma readys_simp [simp]:
       
  3294   assumes "th' \<noteq> th"
       
  3295   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3296   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3297   by metis
       
  3298 
       
  3299 lemma cnp_cnv_cncs_kept: (* ddd *)
       
  3300   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3301   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3302 proof(cases "th' = th")
       
  3303   case True
       
  3304   note eq_th' = this
       
  3305   show ?thesis
       
  3306   proof(cases "wq s cs = []")
       
  3307     case True
       
  3308     then interpret vt: valid_trace_p_h by (unfold_locales, simp)
       
  3309     show ?thesis
       
  3310       using assms eq_th' is_p ready_th_s vt.cntCS_es_th vt.ready_th_es pvD_def by auto 
       
  3311   next
       
  3312     case False
       
  3313     then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  3314     show ?thesis
       
  3315       using add.commute add.left_commute assms eq_th' is_p live_th_s 
       
  3316             ready_th_s vt.th_not_ready_es pvD_def
       
  3317       apply (auto)
       
  3318       by (fold is_p, simp)
       
  3319   qed
       
  3320 next
       
  3321   case False
       
  3322   note h_False = False
       
  3323   thus ?thesis
       
  3324   proof(cases "wq s cs = []")
       
  3325     case True
       
  3326     then interpret vt: valid_trace_p_h by (unfold_locales, simp)
       
  3327     show ?thesis using assms
       
  3328       by (insert True h_False pvD_def, auto split:if_splits,unfold is_p, auto)
       
  3329   next
       
  3330     case False
       
  3331     then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  3332     show ?thesis using assms
       
  3333       by (insert False h_False pvD_def, auto split:if_splits,unfold is_p, auto)
       
  3334   qed
       
  3335 qed
       
  3336 
       
  3337 end
       
  3338 
       
  3339 
       
  3340 context valid_trace_v 
       
  3341 begin
       
  3342 
       
  3343 lemma holding_th_cs_s: 
       
  3344   "holding s th cs" 
       
  3345  by  (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
       
  3346 
       
  3347 lemma th_ready_s [simp]: "th \<in> readys s"
       
  3348   using runing_th_s
       
  3349   by (unfold runing_def readys_def, auto)
       
  3350 
       
  3351 lemma th_live_s [simp]: "th \<in> threads s"
       
  3352   using th_ready_s by (unfold readys_def, auto)
       
  3353 
       
  3354 lemma th_ready_es [simp]: "th \<in> readys (e#s)"
       
  3355   using runing_th_s neq_t_th
       
  3356   by (unfold is_v runing_def readys_def, auto)
       
  3357 
       
  3358 lemma th_live_es [simp]: "th \<in> threads (e#s)"
       
  3359   using th_ready_es by (unfold readys_def, auto)
       
  3360 
       
  3361 lemma pvD_th_s[simp]: "pvD s th = 0"
       
  3362   by (unfold pvD_def, simp)
       
  3363 
       
  3364 lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
       
  3365   by (unfold pvD_def, simp)
       
  3366 
       
  3367 lemma cntCS_s_th [simp]: "cntCS s th > 0"
       
  3368 proof -
       
  3369   have "cs \<in> holdents s th" using holding_th_cs_s
       
  3370     by (unfold holdents_def, simp)
       
  3371   moreover have "finite (holdents s th)" using finite_holdents 
       
  3372     by simp
       
  3373   ultimately show ?thesis
       
  3374     by (unfold cntCS_def, 
       
  3375         auto intro!:card_gt_0_iff[symmetric, THEN iffD1])
       
  3376 qed
       
  3377 
       
  3378 end
       
  3379 
       
  3380 context valid_trace_v
       
  3381 begin
       
  3382 
       
  3383 lemma th_not_waiting: 
       
  3384   "\<not> waiting s th c"
       
  3385 proof -
       
  3386   have "th \<in> readys s"
       
  3387     using runing_ready runing_th_s by blast 
       
  3388   thus ?thesis
       
  3389     by (unfold readys_def, auto)
       
  3390 qed
       
  3391 
       
  3392 lemma waiting_neq_th: 
       
  3393   assumes "waiting s t c"
       
  3394   shows "t \<noteq> th"
       
  3395   using assms using th_not_waiting by blast 
       
  3396 
       
  3397 end
       
  3398 
       
  3399 context valid_trace_v_n
       
  3400 begin
       
  3401 
       
  3402 lemma not_ready_taker_s[simp]: 
       
  3403   "taker \<notin> readys s"
       
  3404   using waiting_taker
       
  3405   by (unfold readys_def, auto)
       
  3406 
       
  3407 lemma taker_live_s [simp]: "taker \<in> threads s"
       
  3408 proof -
       
  3409   have "taker \<in> set wq'" by (simp add: eq_wq') 
       
  3410   from th'_in_inv[OF this]
       
  3411   have "taker \<in> set rest" .
       
  3412   hence "taker \<in> set (wq s cs)" by (simp add: wq_s_cs) 
       
  3413   thus ?thesis using wq_threads by auto 
       
  3414 qed
       
  3415 
       
  3416 lemma taker_live_es [simp]: "taker \<in> threads (e#s)"
       
  3417   using taker_live_s threads_es by blast
       
  3418 
       
  3419 lemma taker_ready_es [simp]:
       
  3420   shows "taker \<in> readys (e#s)"
       
  3421 proof -
       
  3422   { fix cs'
       
  3423     assume "waiting (e#s) taker cs'"
       
  3424     hence False
       
  3425     proof(cases rule:waiting_esE)
       
  3426       case 1
       
  3427       thus ?thesis using waiting_taker waiting_unique by auto 
       
  3428     qed simp
       
  3429   } thus ?thesis by (unfold readys_def, auto)
       
  3430 qed
       
  3431 
       
  3432 lemma neq_taker_th: "taker \<noteq> th"
       
  3433   using th_not_waiting waiting_taker by blast 
       
  3434 
       
  3435 lemma not_holding_taker_s_cs:
       
  3436   shows "\<not> holding s taker cs"
       
  3437   using holding_cs_eq_th neq_taker_th by auto
       
  3438 
       
  3439 lemma holdents_es_taker:
       
  3440   "holdents (e#s) taker = holdents s taker \<union> {cs}" (is "?L = ?R")
       
  3441 proof -
       
  3442   { fix cs'
       
  3443     assume "cs' \<in> ?L"
       
  3444     hence "holding (e#s) taker cs'" by (auto simp:holdents_def)
       
  3445     hence "cs' \<in> ?R"
       
  3446     proof(cases rule:holding_esE)
       
  3447       case 2
       
  3448       thus ?thesis by (auto simp:holdents_def)
       
  3449     qed auto
       
  3450   } moreover {
       
  3451     fix cs'
       
  3452     assume "cs' \<in> ?R"
       
  3453     hence "holding s taker cs' \<or> cs' = cs" by (auto simp:holdents_def)
       
  3454     hence "cs' \<in> ?L" 
       
  3455     proof
       
  3456       assume "holding s taker cs'"
       
  3457       hence "holding (e#s) taker cs'" 
       
  3458           using holding_esI2 holding_taker by fastforce 
       
  3459       thus ?thesis by (auto simp:holdents_def)
       
  3460     next
       
  3461       assume "cs' = cs"
       
  3462       with holding_taker
       
  3463       show ?thesis by (auto simp:holdents_def)
       
  3464     qed
       
  3465   } ultimately show ?thesis by auto
       
  3466 qed
       
  3467 
       
  3468 lemma cntCS_es_taker [simp]: "cntCS (e#s) taker = cntCS s taker + 1"
       
  3469 proof -
       
  3470   have "card (holdents s taker \<union> {cs}) = card (holdents s taker) + 1"
       
  3471   proof(subst card_Un_disjoint)
       
  3472     show "holdents s taker \<inter> {cs} = {}"
       
  3473       using not_holding_taker_s_cs by (auto simp:holdents_def)
       
  3474   qed (auto simp:finite_holdents)
       
  3475   thus ?thesis 
       
  3476     by (unfold cntCS_def, insert holdents_es_taker, simp)
       
  3477 qed
       
  3478 
       
  3479 lemma pvD_taker_s[simp]: "pvD s taker = 1"
       
  3480   by (unfold pvD_def, simp)
       
  3481 
       
  3482 lemma pvD_taker_es[simp]: "pvD (e#s) taker = 0"
       
  3483   by (unfold pvD_def, simp)  
       
  3484 
       
  3485 lemma pvD_th_s[simp]: "pvD s th = 0"
       
  3486   by (unfold pvD_def, simp)
       
  3487 
       
  3488 lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
       
  3489   by (unfold pvD_def, simp)
       
  3490 
       
  3491 lemma holdents_es_th:
       
  3492   "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
       
  3493 proof -
       
  3494   { fix cs'
       
  3495     assume "cs' \<in> ?L"
       
  3496     hence "holding (e#s) th cs'" by (auto simp:holdents_def)
       
  3497     hence "cs' \<in> ?R"
       
  3498     proof(cases rule:holding_esE)
       
  3499       case 2
       
  3500       thus ?thesis by (auto simp:holdents_def)
       
  3501     qed (insert neq_taker_th, auto)
       
  3502   } moreover {
       
  3503     fix cs'
       
  3504     assume "cs' \<in> ?R"
       
  3505     hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
       
  3506     from holding_esI2[OF this]
       
  3507     have "cs' \<in> ?L" by (auto simp:holdents_def)
       
  3508   } ultimately show ?thesis by auto
       
  3509 qed
       
  3510 
       
  3511 lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
       
  3512 proof -
       
  3513   have "card (holdents s th - {cs}) = card (holdents s th) - 1"
       
  3514   proof -
       
  3515     have "cs \<in> holdents s th" using holding_th_cs_s
       
  3516       by (auto simp:holdents_def)
       
  3517     moreover have "finite (holdents s th)"
       
  3518         by (simp add: finite_holdents) 
       
  3519     ultimately show ?thesis by auto
       
  3520   qed
       
  3521   thus ?thesis by (unfold cntCS_def holdents_es_th)
       
  3522 qed
       
  3523 
       
  3524 lemma holdents_kept:
       
  3525   assumes "th' \<noteq> taker"
       
  3526   and "th' \<noteq> th"
       
  3527   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3528 proof -
       
  3529   { fix cs'
       
  3530     assume h: "cs' \<in> ?L"
       
  3531     have "cs' \<in> ?R"
       
  3532     proof(cases "cs' = cs")
       
  3533       case False
       
  3534       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3535       from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  3536       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  3537       show ?thesis
       
  3538         by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  3539     next
       
  3540       case True
       
  3541       from h[unfolded this]
       
  3542       have "holding (e#s) th' cs" by (auto simp:holdents_def)
       
  3543       from held_unique[OF this holding_taker]
       
  3544       have "th' = taker" .
       
  3545       with assms show ?thesis by auto
       
  3546     qed
       
  3547   } moreover {
       
  3548     fix cs'
       
  3549     assume h: "cs' \<in> ?R"
       
  3550     have "cs' \<in> ?L"
       
  3551     proof(cases "cs' = cs")
       
  3552       case False
       
  3553       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3554       from h have "holding s th' cs'" by (auto simp:holdents_def)
       
  3555       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  3556       show ?thesis
       
  3557         by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
       
  3558     next
       
  3559       case True
       
  3560       from h[unfolded this]
       
  3561       have "holding s th' cs" by (auto simp:holdents_def)
       
  3562       from held_unique[OF this holding_th_cs_s]
       
  3563       have "th' = th" .
       
  3564       with assms show ?thesis by auto
       
  3565     qed
       
  3566   } ultimately show ?thesis by auto
       
  3567 qed
       
  3568 
       
  3569 lemma cntCS_kept [simp]:
       
  3570   assumes "th' \<noteq> taker"
       
  3571   and "th' \<noteq> th"
       
  3572   shows "cntCS (e#s) th' = cntCS s th'"
       
  3573   by (unfold cntCS_def holdents_kept[OF assms], simp)
       
  3574 
       
  3575 lemma readys_kept1: 
       
  3576   assumes "th' \<noteq> taker"
       
  3577   and "th' \<in> readys (e#s)"
       
  3578   shows "th' \<in> readys s"
       
  3579 proof -
       
  3580   { fix cs'
       
  3581     assume wait: "waiting s th' cs'"
       
  3582     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3583         using assms(2)[unfolded readys_def] by auto
       
  3584     have False
       
  3585     proof(cases "cs' = cs")
       
  3586       case False
       
  3587       with n_wait wait
       
  3588       show ?thesis 
       
  3589         by (unfold s_waiting_def, fold wq_def, auto)
       
  3590     next
       
  3591       case True
       
  3592       have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
       
  3593         using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
       
  3594       moreover have "\<not> (th' \<in> set rest \<and> th' \<noteq> hd (taker # rest'))" 
       
  3595         using n_wait[unfolded True s_waiting_def, folded wq_def, 
       
  3596                     unfolded wq_es_cs set_wq', unfolded eq_wq'] .
       
  3597       ultimately have "th' = taker" by auto
       
  3598       with assms(1)
       
  3599       show ?thesis by simp
       
  3600     qed
       
  3601   } with assms(2) show ?thesis  
       
  3602     by (unfold readys_def, auto)
       
  3603 qed
       
  3604 
       
  3605 lemma readys_kept2: 
       
  3606   assumes "th' \<noteq> taker"
       
  3607   and "th' \<in> readys s"
       
  3608   shows "th' \<in> readys (e#s)"
       
  3609 proof -
       
  3610   { fix cs'
       
  3611     assume wait: "waiting (e#s) th' cs'"
       
  3612     have n_wait: "\<not> waiting s th' cs'" 
       
  3613         using assms(2)[unfolded readys_def] by auto
       
  3614     have False
       
  3615     proof(cases "cs' = cs")
       
  3616       case False
       
  3617       with n_wait wait
       
  3618       show ?thesis 
       
  3619         by (unfold s_waiting_def, fold wq_def, auto)
       
  3620     next
       
  3621       case True
       
  3622       have "th' \<in> set rest \<and> th' \<noteq> hd (taker # rest')"
       
  3623           using  wait [unfolded True s_waiting_def, folded wq_def, 
       
  3624                     unfolded wq_es_cs set_wq', unfolded eq_wq']  .
       
  3625       moreover have "\<not> (th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest))"
       
  3626           using n_wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
       
  3627       ultimately have "th' = taker" by auto
       
  3628       with assms(1)
       
  3629       show ?thesis by simp
       
  3630     qed
       
  3631   } with assms(2) show ?thesis  
       
  3632     by (unfold readys_def, auto)
       
  3633 qed
       
  3634 
       
  3635 lemma readys_simp [simp]:
       
  3636   assumes "th' \<noteq> taker"
       
  3637   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3638   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3639   by metis
       
  3640 
       
  3641 lemma cnp_cnv_cncs_kept:
       
  3642   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3643   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3644 proof -
       
  3645   { assume eq_th': "th' = taker"
       
  3646     have ?thesis
       
  3647       apply (unfold eq_th' pvD_taker_es cntCS_es_taker)
       
  3648       by (insert neq_taker_th assms[unfolded eq_th'], unfold is_v, simp)
       
  3649   } moreover {
       
  3650     assume eq_th': "th' = th"
       
  3651     have ?thesis 
       
  3652       apply (unfold eq_th' pvD_th_es cntCS_es_th)
       
  3653       by (insert assms[unfolded eq_th'], unfold is_v, simp)
       
  3654   } moreover {
       
  3655     assume h: "th' \<noteq> taker" "th' \<noteq> th"
       
  3656     have ?thesis using assms
       
  3657       apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
       
  3658       by (fold is_v, unfold pvD_def, simp)
       
  3659   } ultimately show ?thesis by metis
       
  3660 qed
       
  3661 
       
  3662 end
       
  3663 
       
  3664 context valid_trace_v_e
       
  3665 begin
       
  3666 
       
  3667 lemma holdents_es_th:
       
  3668   "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
       
  3669 proof -
       
  3670   { fix cs'
       
  3671     assume "cs' \<in> ?L"
       
  3672     hence "holding (e#s) th cs'" by (auto simp:holdents_def)
       
  3673     hence "cs' \<in> ?R"
       
  3674     proof(cases rule:holding_esE)
       
  3675       case 1
       
  3676       thus ?thesis by (auto simp:holdents_def)
       
  3677     qed 
       
  3678   } moreover {
       
  3679     fix cs'
       
  3680     assume "cs' \<in> ?R"
       
  3681     hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
       
  3682     from holding_esI2[OF this]
       
  3683     have "cs' \<in> ?L" by (auto simp:holdents_def)
       
  3684   } ultimately show ?thesis by auto
       
  3685 qed
       
  3686 
       
  3687 lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
       
  3688 proof -
       
  3689   have "card (holdents s th - {cs}) = card (holdents s th) - 1"
       
  3690   proof -
       
  3691     have "cs \<in> holdents s th" using holding_th_cs_s
       
  3692       by (auto simp:holdents_def)
       
  3693     moreover have "finite (holdents s th)"
       
  3694         by (simp add: finite_holdents) 
       
  3695     ultimately show ?thesis by auto
       
  3696   qed
       
  3697   thus ?thesis by (unfold cntCS_def holdents_es_th)
       
  3698 qed
       
  3699 
       
  3700 lemma holdents_kept:
       
  3701   assumes "th' \<noteq> th"
       
  3702   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3703 proof -
       
  3704   { fix cs'
       
  3705     assume h: "cs' \<in> ?L"
       
  3706     have "cs' \<in> ?R"
       
  3707     proof(cases "cs' = cs")
       
  3708       case False
       
  3709       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3710       from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  3711       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  3712       show ?thesis
       
  3713         by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  3714     next
       
  3715       case True
       
  3716       from h[unfolded this]
       
  3717       have "holding (e#s) th' cs" by (auto simp:holdents_def)
       
  3718       from this[unfolded s_holding_def, folded wq_def, 
       
  3719             unfolded wq_es_cs nil_wq']
       
  3720       show ?thesis by auto
       
  3721     qed
       
  3722   } moreover {
       
  3723     fix cs'
       
  3724     assume h: "cs' \<in> ?R"
       
  3725     have "cs' \<in> ?L"
       
  3726     proof(cases "cs' = cs")
       
  3727       case False
       
  3728       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3729       from h have "holding s th' cs'" by (auto simp:holdents_def)
       
  3730       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  3731       show ?thesis
       
  3732         by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
       
  3733     next
       
  3734       case True
       
  3735       from h[unfolded this]
       
  3736       have "holding s th' cs" by (auto simp:holdents_def)
       
  3737       from held_unique[OF this holding_th_cs_s]
       
  3738       have "th' = th" .
       
  3739       with assms show ?thesis by auto
       
  3740     qed
       
  3741   } ultimately show ?thesis by auto
       
  3742 qed
       
  3743 
       
  3744 lemma cntCS_kept [simp]:
       
  3745   assumes "th' \<noteq> th"
       
  3746   shows "cntCS (e#s) th' = cntCS s th'"
       
  3747   by (unfold cntCS_def holdents_kept[OF assms], simp)
       
  3748 
       
  3749 lemma readys_kept1: 
       
  3750   assumes "th' \<in> readys (e#s)"
       
  3751   shows "th' \<in> readys s"
       
  3752 proof -
       
  3753   { fix cs'
       
  3754     assume wait: "waiting s th' cs'"
       
  3755     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3756         using assms(1)[unfolded readys_def] by auto
       
  3757     have False
       
  3758     proof(cases "cs' = cs")
       
  3759       case False
       
  3760       with n_wait wait
       
  3761       show ?thesis 
       
  3762         by (unfold s_waiting_def, fold wq_def, auto)
       
  3763     next
       
  3764       case True
       
  3765       have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
       
  3766         using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . 
       
  3767       hence "th' \<in> set rest" by auto
       
  3768       with set_wq' have "th' \<in> set wq'" by metis
       
  3769       with nil_wq' show ?thesis by simp
       
  3770     qed
       
  3771   } thus ?thesis using assms
       
  3772     by (unfold readys_def, auto)
       
  3773 qed
       
  3774 
       
  3775 lemma readys_kept2: 
       
  3776   assumes "th' \<in> readys s"
       
  3777   shows "th' \<in> readys (e#s)"
       
  3778 proof -
       
  3779   { fix cs'
       
  3780     assume wait: "waiting (e#s) th' cs'"
       
  3781     have n_wait: "\<not> waiting s th' cs'" 
       
  3782         using assms[unfolded readys_def] by auto
       
  3783     have False
       
  3784     proof(cases "cs' = cs")
       
  3785       case False
       
  3786       with n_wait wait
       
  3787       show ?thesis 
       
  3788         by (unfold s_waiting_def, fold wq_def, auto)
       
  3789     next
       
  3790       case True
       
  3791       have "th' \<in> set [] \<and> th' \<noteq> hd []"
       
  3792         using wait[unfolded True s_waiting_def, folded wq_def, 
       
  3793               unfolded wq_es_cs nil_wq'] .
       
  3794       thus ?thesis by simp
       
  3795     qed
       
  3796   } with assms show ?thesis  
       
  3797     by (unfold readys_def, auto)
       
  3798 qed
       
  3799 
       
  3800 lemma readys_simp [simp]:
       
  3801   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3802   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3803   by metis
       
  3804 
       
  3805 lemma cnp_cnv_cncs_kept:
       
  3806   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3807   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3808 proof -
       
  3809   {
       
  3810     assume eq_th': "th' = th"
       
  3811     have ?thesis 
       
  3812       apply (unfold eq_th' pvD_th_es cntCS_es_th)
       
  3813       by (insert assms[unfolded eq_th'], unfold is_v, simp)
       
  3814   } moreover {
       
  3815     assume h: "th' \<noteq> th"
       
  3816     have ?thesis using assms
       
  3817       apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
       
  3818       by (fold is_v, unfold pvD_def, simp)
       
  3819   } ultimately show ?thesis by metis
       
  3820 qed
       
  3821 
       
  3822 end
       
  3823 
       
  3824 context valid_trace_v
       
  3825 begin
       
  3826 
       
  3827 lemma cnp_cnv_cncs_kept:
       
  3828   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3829   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3830 proof(cases "rest = []")
       
  3831   case True
       
  3832   then interpret vt: valid_trace_v_e by (unfold_locales, simp)
       
  3833   show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
       
  3834 next
       
  3835   case False
       
  3836   then interpret vt: valid_trace_v_n by (unfold_locales, simp)
       
  3837   show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
       
  3838 qed
       
  3839 
       
  3840 end
       
  3841 
       
  3842 context valid_trace_create
       
  3843 begin
       
  3844 
       
  3845 lemma th_not_live_s [simp]: "th \<notin> threads s"
       
  3846 proof -
       
  3847   from pip_e[unfolded is_create]
       
  3848   show ?thesis by (cases, simp)
       
  3849 qed
       
  3850 
       
  3851 lemma th_not_ready_s [simp]: "th \<notin> readys s"
       
  3852   using th_not_live_s by (unfold readys_def, simp)
       
  3853 
       
  3854 lemma th_live_es [simp]: "th \<in> threads (e#s)"
       
  3855   by (unfold is_create, simp)
       
  3856 
       
  3857 lemma not_waiting_th_s [simp]: "\<not> waiting s th cs'"
       
  3858 proof
       
  3859   assume "waiting s th cs'"
       
  3860   from this[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  3861   have "th \<in> set (wq s cs')" by auto
       
  3862   from wq_threads[OF this] have "th \<in> threads s" .
       
  3863   with th_not_live_s show False by simp
       
  3864 qed
       
  3865 
       
  3866 lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
       
  3867 proof
       
  3868   assume "holding s th cs'"
       
  3869   from this[unfolded s_holding_def, folded wq_def, unfolded wq_kept]
       
  3870   have "th \<in> set (wq s cs')" by auto
       
  3871   from wq_threads[OF this] have "th \<in> threads s" .
       
  3872   with th_not_live_s show False by simp
       
  3873 qed
       
  3874 
       
  3875 lemma not_waiting_th_es [simp]: "\<not> waiting (e#s) th cs'"
       
  3876 proof
       
  3877   assume "waiting (e # s) th cs'"
       
  3878   from this[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  3879   have "th \<in> set (wq s cs')" by auto
       
  3880   from wq_threads[OF this] have "th \<in> threads s" .
       
  3881   with th_not_live_s show False by simp
       
  3882 qed
       
  3883 
       
  3884 lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
       
  3885 proof
       
  3886   assume "holding (e # s) th cs'"
       
  3887   from this[unfolded s_holding_def, folded wq_def, unfolded wq_kept]
       
  3888   have "th \<in> set (wq s cs')" by auto
       
  3889   from wq_threads[OF this] have "th \<in> threads s" .
       
  3890   with th_not_live_s show False by simp
       
  3891 qed
       
  3892 
       
  3893 lemma ready_th_es [simp]: "th \<in> readys (e#s)"
       
  3894   by (simp add:readys_def)
       
  3895 
       
  3896 lemma holdents_th_s: "holdents s th = {}"
       
  3897   by (unfold holdents_def, auto)
       
  3898 
       
  3899 lemma holdents_th_es: "holdents (e#s) th = {}"
       
  3900   by (unfold holdents_def, auto)
       
  3901 
       
  3902 lemma cntCS_th_s [simp]: "cntCS s th = 0"
       
  3903   by (unfold cntCS_def, simp add:holdents_th_s)
       
  3904 
       
  3905 lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
       
  3906   by (unfold cntCS_def, simp add:holdents_th_es)
       
  3907 
       
  3908 lemma pvD_th_s [simp]: "pvD s th = 0"
       
  3909   by (unfold pvD_def, simp)
       
  3910 
       
  3911 lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
       
  3912   by (unfold pvD_def, simp)
       
  3913 
       
  3914 lemma holdents_kept:
       
  3915   assumes "th' \<noteq> th"
       
  3916   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3917 proof -
       
  3918   { fix cs'
       
  3919     assume h: "cs' \<in> ?L"
       
  3920     hence "cs' \<in> ?R"
       
  3921       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3922              unfold wq_kept, auto)
       
  3923   } moreover {
       
  3924     fix cs'
       
  3925     assume h: "cs' \<in> ?R"
       
  3926     hence "cs' \<in> ?L"
       
  3927       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3928              unfold wq_kept, auto)
       
  3929   } ultimately show ?thesis by auto
       
  3930 qed
       
  3931 
       
  3932 lemma cntCS_kept [simp]:
       
  3933   assumes "th' \<noteq> th"
       
  3934   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  3935   using holdents_kept[OF assms]
       
  3936   by (unfold cntCS_def, simp)
       
  3937 
       
  3938 lemma readys_kept1: 
       
  3939   assumes "th' \<noteq> th"
       
  3940   and "th' \<in> readys (e#s)"
       
  3941   shows "th' \<in> readys s"
       
  3942 proof -
       
  3943   { fix cs'
       
  3944     assume wait: "waiting s th' cs'"
       
  3945     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3946       using assms by (auto simp:readys_def)
       
  3947     from wait[unfolded s_waiting_def, folded wq_def]
       
  3948          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  3949     have False by auto
       
  3950   } thus ?thesis using assms
       
  3951     by (unfold readys_def, auto)
       
  3952 qed
       
  3953 
       
  3954 lemma readys_kept2: 
       
  3955   assumes "th' \<noteq> th"
       
  3956   and "th' \<in> readys s"
       
  3957   shows "th' \<in> readys (e#s)"
       
  3958 proof -
       
  3959   { fix cs'
       
  3960     assume wait: "waiting (e#s) th' cs'"
       
  3961     have n_wait: "\<not> waiting s th' cs'"
       
  3962       using assms(2) by (auto simp:readys_def)
       
  3963     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  3964          n_wait[unfolded s_waiting_def, folded wq_def]
       
  3965     have False by auto
       
  3966   } with assms show ?thesis  
       
  3967     by (unfold readys_def, auto)
       
  3968 qed
       
  3969 
       
  3970 lemma readys_simp [simp]:
       
  3971   assumes "th' \<noteq> th"
       
  3972   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3973   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3974   by metis
       
  3975 
       
  3976 lemma pvD_kept [simp]:
       
  3977   assumes "th' \<noteq> th"
       
  3978   shows "pvD (e#s) th' = pvD s th'"
       
  3979   using assms
       
  3980   by (unfold pvD_def, simp)
       
  3981 
       
  3982 lemma cnp_cnv_cncs_kept:
       
  3983   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3984   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3985 proof -
       
  3986   {
       
  3987     assume eq_th': "th' = th"
       
  3988     have ?thesis using assms
       
  3989       by (unfold eq_th', simp, unfold is_create, simp)
       
  3990   } moreover {
       
  3991     assume h: "th' \<noteq> th"
       
  3992     hence ?thesis using assms
       
  3993       by (simp, simp add:is_create)
       
  3994   } ultimately show ?thesis by metis
       
  3995 qed
       
  3996 
       
  3997 end
       
  3998 
       
  3999 context valid_trace_exit
       
  4000 begin
       
  4001 
       
  4002 lemma th_live_s [simp]: "th \<in> threads s"
       
  4003 proof -
       
  4004   from pip_e[unfolded is_exit]
       
  4005   show ?thesis
       
  4006   by (cases, unfold runing_def readys_def, simp)
       
  4007 qed
       
  4008 
       
  4009 lemma th_ready_s [simp]: "th \<in> readys s"
       
  4010 proof -
       
  4011   from pip_e[unfolded is_exit]
       
  4012   show ?thesis
       
  4013   by (cases, unfold runing_def, simp)
       
  4014 qed
       
  4015 
       
  4016 lemma th_not_live_es [simp]: "th \<notin> threads (e#s)"
       
  4017   by (unfold is_exit, simp)
       
  4018 
       
  4019 lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
       
  4020 proof -
       
  4021   from pip_e[unfolded is_exit]
       
  4022   show ?thesis 
       
  4023    by (cases, unfold holdents_def, auto)
       
  4024 qed
       
  4025 
       
  4026 lemma cntCS_th_s [simp]: "cntCS s th = 0"
       
  4027 proof -
       
  4028   from pip_e[unfolded is_exit]
       
  4029   show ?thesis 
       
  4030    by (cases, unfold cntCS_def, simp)
       
  4031 qed
       
  4032 
       
  4033 lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
       
  4034 proof
       
  4035   assume "holding (e # s) th cs'"
       
  4036   from this[unfolded s_holding_def, folded wq_def, unfolded wq_kept]
       
  4037   have "holding s th cs'" 
       
  4038     by (unfold s_holding_def, fold wq_def, auto)
       
  4039   with not_holding_th_s 
       
  4040   show False by simp
       
  4041 qed
       
  4042 
       
  4043 lemma ready_th_es [simp]: "th \<notin> readys (e#s)"
       
  4044   by (simp add:readys_def)
       
  4045 
       
  4046 lemma holdents_th_s: "holdents s th = {}"
       
  4047   by (unfold holdents_def, auto)
       
  4048 
       
  4049 lemma holdents_th_es: "holdents (e#s) th = {}"
       
  4050   by (unfold holdents_def, auto)
       
  4051 
       
  4052 lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
       
  4053   by (unfold cntCS_def, simp add:holdents_th_es)
       
  4054 
       
  4055 lemma pvD_th_s [simp]: "pvD s th = 0"
       
  4056   by (unfold pvD_def, simp)
       
  4057 
       
  4058 lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
       
  4059   by (unfold pvD_def, simp)
       
  4060 
       
  4061 lemma holdents_kept:
       
  4062   assumes "th' \<noteq> th"
       
  4063   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  4064 proof -
       
  4065   { fix cs'
       
  4066     assume h: "cs' \<in> ?L"
       
  4067     hence "cs' \<in> ?R"
       
  4068       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  4069              unfold wq_kept, auto)
       
  4070   } moreover {
       
  4071     fix cs'
       
  4072     assume h: "cs' \<in> ?R"
       
  4073     hence "cs' \<in> ?L"
       
  4074       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  4075              unfold wq_kept, auto)
       
  4076   } ultimately show ?thesis by auto
       
  4077 qed
       
  4078 
       
  4079 lemma cntCS_kept [simp]:
       
  4080   assumes "th' \<noteq> th"
       
  4081   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  4082   using holdents_kept[OF assms]
       
  4083   by (unfold cntCS_def, simp)
       
  4084 
       
  4085 lemma readys_kept1: 
       
  4086   assumes "th' \<noteq> th"
       
  4087   and "th' \<in> readys (e#s)"
       
  4088   shows "th' \<in> readys s"
       
  4089 proof -
       
  4090   { fix cs'
       
  4091     assume wait: "waiting s th' cs'"
       
  4092     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  4093       using assms by (auto simp:readys_def)
       
  4094     from wait[unfolded s_waiting_def, folded wq_def]
       
  4095          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  4096     have False by auto
       
  4097   } thus ?thesis using assms
       
  4098     by (unfold readys_def, auto)
       
  4099 qed
       
  4100 
       
  4101 lemma readys_kept2: 
       
  4102   assumes "th' \<noteq> th"
       
  4103   and "th' \<in> readys s"
       
  4104   shows "th' \<in> readys (e#s)"
       
  4105 proof -
       
  4106   { fix cs'
       
  4107     assume wait: "waiting (e#s) th' cs'"
       
  4108     have n_wait: "\<not> waiting s th' cs'"
       
  4109       using assms(2) by (auto simp:readys_def)
       
  4110     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  4111          n_wait[unfolded s_waiting_def, folded wq_def]
       
  4112     have False by auto
       
  4113   } with assms show ?thesis  
       
  4114     by (unfold readys_def, auto)
       
  4115 qed
       
  4116 
       
  4117 lemma readys_simp [simp]:
       
  4118   assumes "th' \<noteq> th"
       
  4119   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  4120   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  4121   by metis
       
  4122 
       
  4123 lemma pvD_kept [simp]:
       
  4124   assumes "th' \<noteq> th"
       
  4125   shows "pvD (e#s) th' = pvD s th'"
       
  4126   using assms
       
  4127   by (unfold pvD_def, simp)
       
  4128 
       
  4129 lemma cnp_cnv_cncs_kept:
       
  4130   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  4131   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  4132 proof -
       
  4133   {
       
  4134     assume eq_th': "th' = th"
       
  4135     have ?thesis using assms
       
  4136       by (unfold eq_th', simp, unfold is_exit, simp)
       
  4137   } moreover {
       
  4138     assume h: "th' \<noteq> th"
       
  4139     hence ?thesis using assms
       
  4140       by (simp, simp add:is_exit)
       
  4141   } ultimately show ?thesis by metis
       
  4142 qed
       
  4143 
       
  4144 end
       
  4145 
       
  4146 context valid_trace_set
       
  4147 begin
       
  4148 
       
  4149 lemma th_live_s [simp]: "th \<in> threads s"
       
  4150 proof -
       
  4151   from pip_e[unfolded is_set]
       
  4152   show ?thesis
       
  4153   by (cases, unfold runing_def readys_def, simp)
       
  4154 qed
       
  4155 
       
  4156 lemma th_ready_s [simp]: "th \<in> readys s"
       
  4157 proof -
       
  4158   from pip_e[unfolded is_set]
       
  4159   show ?thesis
       
  4160   by (cases, unfold runing_def, simp)
       
  4161 qed
       
  4162 
       
  4163 lemma th_not_live_es [simp]: "th \<in> threads (e#s)"
       
  4164   by (unfold is_set, simp)
       
  4165 
       
  4166 
       
  4167 lemma holdents_kept:
       
  4168   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  4169 proof -
       
  4170   { fix cs'
       
  4171     assume h: "cs' \<in> ?L"
       
  4172     hence "cs' \<in> ?R"
       
  4173       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  4174              unfold wq_kept, auto)
       
  4175   } moreover {
       
  4176     fix cs'
       
  4177     assume h: "cs' \<in> ?R"
       
  4178     hence "cs' \<in> ?L"
       
  4179       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  4180              unfold wq_kept, auto)
       
  4181   } ultimately show ?thesis by auto
       
  4182 qed
       
  4183 
       
  4184 lemma cntCS_kept [simp]:
       
  4185   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  4186   using holdents_kept
       
  4187   by (unfold cntCS_def, simp)
       
  4188 
       
  4189 lemma threads_kept[simp]:
       
  4190   "threads (e#s) = threads s"
       
  4191   by (unfold is_set, simp)
       
  4192 
       
  4193 lemma readys_kept1: 
       
  4194   assumes "th' \<in> readys (e#s)"
       
  4195   shows "th' \<in> readys s"
       
  4196 proof -
       
  4197   { fix cs'
       
  4198     assume wait: "waiting s th' cs'"
       
  4199     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  4200       using assms by (auto simp:readys_def)
       
  4201     from wait[unfolded s_waiting_def, folded wq_def]
       
  4202          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  4203     have False by auto
       
  4204   } moreover have "th' \<in> threads s" 
       
  4205     using assms[unfolded readys_def] by auto
       
  4206   ultimately show ?thesis 
       
  4207     by (unfold readys_def, auto)
       
  4208 qed
       
  4209 
       
  4210 lemma readys_kept2: 
       
  4211   assumes "th' \<in> readys s"
       
  4212   shows "th' \<in> readys (e#s)"
       
  4213 proof -
       
  4214   { fix cs'
       
  4215     assume wait: "waiting (e#s) th' cs'"
       
  4216     have n_wait: "\<not> waiting s th' cs'"
       
  4217       using assms by (auto simp:readys_def)
       
  4218     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept]
       
  4219          n_wait[unfolded s_waiting_def, folded wq_def]
       
  4220     have False by auto
       
  4221   } with assms show ?thesis  
       
  4222     by (unfold readys_def, auto)
       
  4223 qed
       
  4224 
       
  4225 lemma readys_simp [simp]:
       
  4226   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  4227   using readys_kept1 readys_kept2
       
  4228   by metis
       
  4229 
       
  4230 lemma pvD_kept [simp]:
       
  4231   shows "pvD (e#s) th' = pvD s th'"
       
  4232   by (unfold pvD_def, simp)
       
  4233 
       
  4234 lemma cnp_cnv_cncs_kept:
       
  4235   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  4236   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  4237   using assms
       
  4238   by (unfold is_set, simp, fold is_set, simp)
       
  4239 
       
  4240 end
       
  4241 
       
  4242 context valid_trace
       
  4243 begin
       
  4244 
       
  4245 lemma cnp_cnv_cncs: "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  4246 proof(induct rule:ind)
       
  4247   case Nil
       
  4248   thus ?case 
       
  4249     by (unfold cntP_def cntV_def pvD_def cntCS_def holdents_def 
       
  4250               s_holding_def, simp)
       
  4251 next
       
  4252   case (Cons s e)
       
  4253   interpret vt_e: valid_trace_e s e using Cons by simp
       
  4254   show ?case
       
  4255   proof(cases e)
       
  4256     case (Create th prio)
       
  4257     interpret vt_create: valid_trace_create s e th prio 
       
  4258       using Create by (unfold_locales, simp)
       
  4259     show ?thesis using Cons by (simp add: vt_create.cnp_cnv_cncs_kept) 
       
  4260   next
       
  4261     case (Exit th)
       
  4262     interpret vt_exit: valid_trace_exit s e th  
       
  4263         using Exit by (unfold_locales, simp)
       
  4264    show ?thesis using Cons by (simp add: vt_exit.cnp_cnv_cncs_kept) 
       
  4265   next
       
  4266     case (P th cs)
       
  4267     interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp)
       
  4268     show ?thesis using Cons by (simp add: vt_p.cnp_cnv_cncs_kept) 
       
  4269   next
       
  4270     case (V th cs)
       
  4271     interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp)
       
  4272     show ?thesis using Cons by (simp add: vt_v.cnp_cnv_cncs_kept) 
       
  4273   next
       
  4274     case (Set th prio)
       
  4275     interpret vt_set: valid_trace_set s e th prio
       
  4276         using Set by (unfold_locales, simp)
       
  4277     show ?thesis using Cons by (simp add: vt_set.cnp_cnv_cncs_kept) 
       
  4278   qed
       
  4279 qed
       
  4280 
       
  4281 end
       
  4282 
       
  4283 section {* Corollaries of @{thm valid_trace.cnp_cnv_cncs} *}
       
  4284 
       
  4285 context valid_trace
       
  4286 begin
       
  4287 
       
  4288 lemma not_thread_holdents:
       
  4289   assumes not_in: "th \<notin> threads s" 
       
  4290   shows "holdents s th = {}"
       
  4291 proof -
       
  4292   { fix cs
       
  4293     assume "cs \<in> holdents s th"
       
  4294     hence "holding s th cs" by (auto simp:holdents_def)
       
  4295     from this[unfolded s_holding_def, folded wq_def]
       
  4296     have "th \<in> set (wq s cs)" by auto
       
  4297     with wq_threads have "th \<in> threads s" by auto
       
  4298     with assms
       
  4299     have False by simp
       
  4300   } thus ?thesis by auto
       
  4301 qed
       
  4302 
       
  4303 lemma not_thread_cncs:
       
  4304   assumes not_in: "th \<notin> threads s" 
       
  4305   shows "cntCS s th = 0"
       
  4306   using not_thread_holdents[OF assms]
       
  4307   by (simp add:cntCS_def)
       
  4308 
       
  4309 lemma cnp_cnv_eq:
       
  4310   assumes "th \<notin> threads s"
       
  4311   shows "cntP s th = cntV s th"
       
  4312   using assms cnp_cnv_cncs not_thread_cncs pvD_def
       
  4313   by (auto)
       
  4314 
       
  4315 lemma eq_pv_children:
       
  4316   assumes eq_pv: "cntP s th = cntV s th"
       
  4317   shows "children (RAG s) (Th th) = {}"
       
  4318 proof -
       
  4319     from cnp_cnv_cncs and eq_pv
       
  4320     have "cntCS s th = 0" 
       
  4321       by (auto split:if_splits)
       
  4322     from this[unfolded cntCS_def holdents_alt_def]
       
  4323     have card_0: "card (the_cs ` children (RAG s) (Th th)) = 0" .
       
  4324     have "finite (the_cs ` children (RAG s) (Th th))"
       
  4325       by (simp add: fsbtRAGs.finite_children)
       
  4326     from card_0[unfolded card_0_eq[OF this]]
       
  4327     show ?thesis by auto
       
  4328 qed
       
  4329 
       
  4330 lemma eq_pv_holdents:
       
  4331   assumes eq_pv: "cntP s th = cntV s th"
       
  4332   shows "holdents s th = {}"
       
  4333   by (unfold holdents_alt_def eq_pv_children[OF assms], simp)
       
  4334 
       
  4335 lemma eq_pv_subtree:
       
  4336   assumes eq_pv: "cntP s th = cntV s th"
       
  4337   shows "subtree (RAG s) (Th th) = {Th th}"
       
  4338   using eq_pv_children[OF assms]
       
  4339     by (unfold subtree_children, simp)
       
  4340 
       
  4341 lemma count_eq_RAG_plus:
       
  4342   assumes "cntP s th = cntV s th"
       
  4343   shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  4344 proof(rule ccontr)
       
  4345     assume otherwise: "{th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} \<noteq> {}"
       
  4346     then obtain th' where "(Th th', Th th) \<in> (RAG s)^+" by auto
       
  4347     from tranclD2[OF this]
       
  4348     obtain z where "z \<in> children (RAG s) (Th th)" 
       
  4349       by (auto simp:children_def)
       
  4350     with eq_pv_children[OF assms]
       
  4351     show False by simp
       
  4352 qed
       
  4353 
       
  4354 lemma eq_pv_dependants:
       
  4355   assumes eq_pv: "cntP s th = cntV s th"
       
  4356   shows "dependants s th = {}"
       
  4357 proof -
       
  4358   from count_eq_RAG_plus[OF assms, folded dependants_alt_def1]
       
  4359   show ?thesis .
       
  4360 qed
       
  4361 
       
  4362 lemma count_eq_tRAG_plus:
       
  4363   assumes "cntP s th = cntV s th"
       
  4364   shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
       
  4365   using assms eq_pv_dependants dependants_alt_def eq_dependants by auto 
       
  4366 
       
  4367 lemma count_eq_RAG_plus_Th:
       
  4368   assumes "cntP s th = cntV s th"
       
  4369   shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  4370   using count_eq_RAG_plus[OF assms] by auto
       
  4371 
       
  4372 lemma count_eq_tRAG_plus_Th:
       
  4373   assumes "cntP s th = cntV s th"
       
  4374   shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
       
  4375    using count_eq_tRAG_plus[OF assms] by auto
       
  4376 
       
  4377 end
       
  4378 
       
  4379 definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
       
  4380   where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
       
  4381 
       
  4382 lemma detached_test:
       
  4383   shows "detached s th = (Th th \<notin> Field (RAG s))"
       
  4384 apply(simp add: detached_def Field_def)
       
  4385 apply(simp add: s_RAG_def)
       
  4386 apply(simp add: s_holding_abv s_waiting_abv)
       
  4387 apply(simp add: Domain_iff Range_iff)
       
  4388 apply(simp add: wq_def)
       
  4389 apply(auto)
       
  4390 done
       
  4391 
       
  4392 context valid_trace
       
  4393 begin
       
  4394 
       
  4395 lemma detached_intro:
       
  4396   assumes eq_pv: "cntP s th = cntV s th"
       
  4397   shows "detached s th"
       
  4398 proof -
       
  4399   from eq_pv cnp_cnv_cncs
       
  4400   have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:pvD_def)
       
  4401   thus ?thesis
       
  4402   proof
       
  4403     assume "th \<notin> threads s"
       
  4404     with rg_RAG_threads dm_RAG_threads
       
  4405     show ?thesis
       
  4406       by (auto simp add: detached_def s_RAG_def s_waiting_abv 
       
  4407               s_holding_abv wq_def Domain_iff Range_iff)
       
  4408   next
       
  4409     assume "th \<in> readys s"
       
  4410     moreover have "Th th \<notin> Range (RAG s)"
       
  4411     proof -
       
  4412       from eq_pv_children[OF assms]
       
  4413       have "children (RAG s) (Th th) = {}" .
       
  4414       thus ?thesis
       
  4415       by (unfold children_def, auto)
       
  4416     qed
       
  4417     ultimately show ?thesis
       
  4418       by (auto simp add: detached_def s_RAG_def s_waiting_abv 
       
  4419               s_holding_abv wq_def readys_def)
       
  4420   qed
       
  4421 qed
       
  4422 
       
  4423 lemma detached_elim:
       
  4424   assumes dtc: "detached s th"
       
  4425   shows "cntP s th = cntV s th"
       
  4426 proof -
       
  4427   have cncs_z: "cntCS s th = 0"
       
  4428   proof -
       
  4429     from dtc have "holdents s th = {}"
       
  4430       unfolding detached_def holdents_test s_RAG_def
       
  4431       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
       
  4432     thus ?thesis by (auto simp:cntCS_def)
       
  4433   qed
       
  4434   show ?thesis
       
  4435   proof(cases "th \<in> threads s")
       
  4436     case True
       
  4437     with dtc 
       
  4438     have "th \<in> readys s"
       
  4439       by (unfold readys_def detached_def Field_def Domain_def Range_def, 
       
  4440            auto simp:waiting_eq s_RAG_def)
       
  4441     with cncs_z  show ?thesis using cnp_cnv_cncs by (simp add:pvD_def)
       
  4442   next
       
  4443     case False
       
  4444     with cncs_z and cnp_cnv_cncs show ?thesis by (simp add:pvD_def)
       
  4445   qed
       
  4446 qed
       
  4447 
       
  4448 lemma detached_eq:
       
  4449   shows "(detached s th) = (cntP s th = cntV s th)"
       
  4450   by (insert vt, auto intro:detached_intro detached_elim)
       
  4451 
       
  4452 end
       
  4453 
       
  4454 section {* Recursive definition of @{term "cp"} *}
       
  4455 
       
  4456 lemma cp_alt_def1: 
       
  4457   "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
       
  4458 proof -
       
  4459   have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
       
  4460        ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
       
  4461        by auto
       
  4462   thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
       
  4463 qed
       
  4464 
       
  4465 lemma cp_gen_def_cond: 
       
  4466   assumes "x = Th th"
       
  4467   shows "cp s th = cp_gen s (Th th)"
       
  4468 by (unfold cp_alt_def1 cp_gen_def, simp)
       
  4469 
       
  4470 lemma cp_gen_over_set:
       
  4471   assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
       
  4472   shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
       
  4473 proof(rule f_image_eq)
       
  4474   fix a
       
  4475   assume "a \<in> A"
       
  4476   from assms[rule_format, OF this]
       
  4477   obtain th where eq_a: "a = Th th" by auto
       
  4478   show "cp_gen s a = (cp s \<circ> the_thread) a"
       
  4479     by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
       
  4480 qed
       
  4481 
       
  4482 
       
  4483 context valid_trace
       
  4484 begin
       
  4485 (* ddd *)
       
  4486 lemma cp_gen_rec:
       
  4487   assumes "x = Th th"
       
  4488   shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
       
  4489 proof(cases "children (tRAG s) x = {}")
       
  4490   case True
       
  4491   show ?thesis
       
  4492     by (unfold True cp_gen_def subtree_children, simp add:assms)
       
  4493 next
       
  4494   case False
       
  4495   hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
       
  4496   note fsbttRAGs.finite_subtree[simp]
       
  4497   have [simp]: "finite (children (tRAG s) x)"
       
  4498      by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
       
  4499             rule children_subtree)
       
  4500   { fix r x
       
  4501     have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
       
  4502   } note this[simp]
       
  4503   have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
       
  4504   proof -
       
  4505     from False obtain q where "q \<in> children (tRAG s) x" by blast
       
  4506     moreover have "subtree (tRAG s) q \<noteq> {}" by simp
       
  4507     ultimately show ?thesis by blast
       
  4508   qed
       
  4509   have h: "Max ((the_preced s \<circ> the_thread) `
       
  4510                 ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
       
  4511         Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
       
  4512                      (is "?L = ?R")
       
  4513   proof -
       
  4514     let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
       
  4515     let "Max (_ \<union> (?h ` ?B))" = ?R
       
  4516     let ?L1 = "?f ` \<Union>(?g ` ?B)"
       
  4517     have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
       
  4518     proof -
       
  4519       have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
       
  4520       also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
       
  4521       finally have "Max ?L1 = Max ..." by simp
       
  4522       also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
       
  4523         by (subst Max_UNION, simp+)
       
  4524       also have "... = Max (cp_gen s ` children (tRAG s) x)"
       
  4525           by (unfold image_comp cp_gen_alt_def, simp)
       
  4526       finally show ?thesis .
       
  4527     qed
       
  4528     show ?thesis
       
  4529     proof -
       
  4530       have "?L = Max (?f ` ?A \<union> ?L1)" by simp
       
  4531       also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
       
  4532             by (subst Max_Un, simp+)
       
  4533       also have "... = max (?f x) (Max (?h ` ?B))"
       
  4534         by (unfold eq_Max_L1, simp)
       
  4535       also have "... =?R"
       
  4536         by (rule max_Max_eq, (simp)+, unfold assms, simp)
       
  4537       finally show ?thesis .
       
  4538     qed
       
  4539   qed  thus ?thesis 
       
  4540           by (fold h subtree_children, unfold cp_gen_def, simp) 
       
  4541 qed
       
  4542 
       
  4543 lemma cp_rec:
       
  4544   "cp s th = Max ({the_preced s th} \<union> 
       
  4545                      (cp s o the_thread) ` children (tRAG s) (Th th))"
       
  4546 proof -
       
  4547   have "Th th = Th th" by simp
       
  4548   note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
       
  4549   show ?thesis 
       
  4550   proof -
       
  4551     have "cp_gen s ` children (tRAG s) (Th th) = 
       
  4552                 (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
       
  4553     proof(rule cp_gen_over_set)
       
  4554       show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
       
  4555         by (unfold tRAG_alt_def, auto simp:children_def)
       
  4556     qed
       
  4557     thus ?thesis by (subst (1) h(1), unfold h(2), simp)
       
  4558   qed
       
  4559 qed
       
  4560 end
       
  4561 
       
  4562 section {* Other properties useful in Implementation.thy or Correctness.thy *}
       
  4563 
       
  4564 context valid_trace_e 
       
  4565 begin
       
  4566 
       
  4567 lemma actor_inv: 
       
  4568   assumes "\<not> isCreate e"
       
  4569   shows "actor e \<in> runing s"
       
  4570   using pip_e assms 
       
  4571   by (induct, auto)
       
  4572 end
       
  4573 
       
  4574 context valid_trace
       
  4575 begin
       
  4576 
       
  4577 lemma readys_root:
       
  4578   assumes "th \<in> readys s"
       
  4579   shows "root (RAG s) (Th th)"
       
  4580 proof -
       
  4581   { fix x
       
  4582     assume "x \<in> ancestors (RAG s) (Th th)"
       
  4583     hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
  4584     from tranclD[OF this]
       
  4585     obtain z where "(Th th, z) \<in> RAG s" by auto
       
  4586     with assms(1) have False
       
  4587          apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
       
  4588          by (fold wq_def, blast)
       
  4589   } thus ?thesis by (unfold root_def, auto)
       
  4590 qed
       
  4591 
       
  4592 lemma readys_in_no_subtree:
       
  4593   assumes "th \<in> readys s"
       
  4594   and "th' \<noteq> th"
       
  4595   shows "Th th \<notin> subtree (RAG s) (Th th')" 
       
  4596 proof
       
  4597    assume "Th th \<in> subtree (RAG s) (Th th')"
       
  4598    thus False
       
  4599    proof(cases rule:subtreeE)
       
  4600       case 1
       
  4601       with assms show ?thesis by auto
       
  4602    next
       
  4603       case 2
       
  4604       with readys_root[OF assms(1)]
       
  4605       show ?thesis by (auto simp:root_def)
       
  4606    qed
       
  4607 qed
       
  4608 
       
  4609 lemma not_in_thread_isolated:
       
  4610   assumes "th \<notin> threads s"
       
  4611   shows "(Th th) \<notin> Field (RAG s)"
       
  4612 proof
       
  4613   assume "(Th th) \<in> Field (RAG s)"
       
  4614   with dm_RAG_threads and rg_RAG_threads assms
       
  4615   show False by (unfold Field_def, blast)
       
  4616 qed
       
  4617 
       
  4618 lemma next_th_holding:
       
  4619   assumes nxt: "next_th s th cs th'"
       
  4620   shows "holding (wq s) th cs"
       
  4621 proof -
       
  4622   from nxt[unfolded next_th_def]
       
  4623   obtain rest where h: "wq s cs = th # rest"
       
  4624                        "rest \<noteq> []" 
       
  4625                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
       
  4626   thus ?thesis
       
  4627     by (unfold cs_holding_def, auto)
       
  4628 qed
       
  4629 
       
  4630 lemma next_th_waiting:
       
  4631   assumes nxt: "next_th s th cs th'"
       
  4632   shows "waiting (wq s) th' cs"
       
  4633 proof -
       
  4634   from nxt[unfolded next_th_def]
       
  4635   obtain rest where h: "wq s cs = th # rest"
       
  4636                        "rest \<noteq> []" 
       
  4637                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
       
  4638   from wq_distinct[of cs, unfolded h]
       
  4639   have dst: "distinct (th # rest)" .
       
  4640   have in_rest: "th' \<in> set rest"
       
  4641   proof(unfold h, rule someI2)
       
  4642     show "distinct rest \<and> set rest = set rest" using dst by auto
       
  4643   next
       
  4644     fix x assume "distinct x \<and> set x = set rest"
       
  4645     with h(2)
       
  4646     show "hd x \<in> set (rest)" by (cases x, auto)
       
  4647   qed
       
  4648   hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
       
  4649   moreover have "th' \<noteq> hd (wq s cs)"
       
  4650     by (unfold h(1), insert in_rest dst, auto)
       
  4651   ultimately show ?thesis by (auto simp:cs_waiting_def)
       
  4652 qed
       
  4653 
       
  4654 lemma next_th_RAG:
       
  4655   assumes nxt: "next_th (s::event list) th cs th'"
       
  4656   shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
       
  4657   using vt assms next_th_holding next_th_waiting
       
  4658   by (unfold s_RAG_def, simp)
       
  4659 
       
  4660 end
       
  4661 
       
  4662 end