|    129   The following lemma shows that @{term "th"} is not in the  |    101   The following lemma shows that @{term "th"} is not in the  | 
|    130   sub-tree of any other thread.  |    102   sub-tree of any other thread.  | 
|    131 *} |    103 *} | 
|    132 lemma th_in_no_subtree: |    104 lemma th_in_no_subtree: | 
|    133   assumes "th' \<noteq> th" |    105   assumes "th' \<noteq> th" | 
|    134   shows "Th th \<notin> subtree (RAG s') (Th th')" |    106   shows "Th th \<notin> subtree (RAG s) (Th th')" | 
|    135 proof - |    107 proof - | 
|    136   have "th \<in> readys s'" |    108   from readys_in_no_subtree[OF th_ready_s assms(1)] | 
|    137   proof - |         | 
|    138     from step_back_step [OF vt_s[unfolded s_def]] |         | 
|    139     have "step s' (Set th prio)" . |         | 
|    140     hence "th \<in> runing s'" by (cases, simp) |         | 
|    141     thus ?thesis by (simp add:readys_def runing_def) |         | 
|    142   qed |         | 
|    143   from vat_s'.readys_in_no_subtree[OF this assms(1)] |         | 
|    144   show ?thesis by blast |    109   show ?thesis by blast | 
|    145 qed |    110 qed | 
|    146  |    111  | 
|    147 text {*  |    112 text {*  | 
|    148   By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"},  |    113   By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"},  | 
|    149   it is obvious that the change of priority only affects the @{text "cp"}-value  |    114   it is obvious that the change of priority only affects the @{text "cp"}-value  | 
|    150   of the initiating thread @{text "th"}. |    115   of the initiating thread @{text "th"}. | 
|    151 *} |    116 *} | 
|    152 lemma eq_cp: |    117 lemma eq_cp: | 
|    153   assumes "th' \<noteq> th" |    118   assumes "th' \<noteq> th" | 
|    154   shows "cp s th' = cp s' th'" |    119   shows "cp (e#s) th' = cp s th'" | 
|    155   by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) |    120   by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) | 
|    156  |    121  | 
|    157 end |    122 end | 
|    158  |    123  | 
|    159 section {* The @{term V} operation *} |    124 section {* The @{term V} operation *} | 
|    160  |    125  | 
|    161 text {* |    126 text {* | 
|    162   The following @{text "step_v_cps"} is the locale for @{text "V"}-operation. |    127   The following @{text "step_v_cps"} is the locale for @{text "V"}-operation. | 
|    163 *} |    128 *} | 
|    164  |    129  | 
|    165 locale step_v_cps = |    130  | 
|    166   -- {* @{text "th"} is the initiating thread *} |    131 context valid_trace_v | 
|    167   -- {* @{text "cs"} is the critical resource release by the @{text "V"}-operation *} |         | 
|    168   fixes s' th cs s    -- {* @{text "s'"} is the state before operation*} |         | 
|    169   defines s_def : "s \<equiv> (V th cs#s')" -- {* @{text "s"} is the state after operation*} |         | 
|    170   -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *} |         | 
|    171   assumes vt_s: "vt s" |         | 
|    172  |         | 
|    173 sublocale step_v_cps < vat_s : valid_trace "s" |         | 
|    174 proof |         | 
|    175   from vt_s show "vt s" . |         | 
|    176 qed |         | 
|    177  |         | 
|    178 sublocale step_v_cps < vat_s' : valid_trace "s'" |         | 
|    179 proof |         | 
|    180   from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . |         | 
|    181 qed |         | 
|    182  |         | 
|    183 context step_v_cps |         | 
|    184 begin |    132 begin | 
|    185  |    133  | 
|    186 lemma ready_th_s': "th \<in> readys s'" |    134 lemma ancestors_th: "ancestors (RAG s) (Th th) = {}" | 
|    187   using step_back_step[OF vt_s[unfolded s_def]] |    135 proof - | 
|    188   by (cases, simp add:runing_def) |    136   from readys_root[OF th_ready_s] | 
|    189  |         | 
|    190 lemma ancestors_th: "ancestors (RAG s') (Th th) = {}" |         | 
|    191 proof - |         | 
|    192   from vat_s'.readys_root[OF ready_th_s'] |         | 
|    193   show ?thesis |    137   show ?thesis | 
|    194   by (unfold root_def, simp) |    138   by (unfold root_def, simp) | 
|    195 qed |    139 qed | 
|    196  |    140  | 
|    197 lemma holding_th: "holding s' th cs" |         | 
|    198 proof - |         | 
|    199   from vt_s[unfolded s_def] |         | 
|    200   have " PIP s' (V th cs)" by (cases, simp) |         | 
|    201   thus ?thesis by (cases, auto) |         | 
|    202 qed |         | 
|    203  |         | 
|    204 lemma edge_of_th: |    141 lemma edge_of_th: | 
|    205     "(Cs cs, Th th) \<in> RAG s'"  |    142     "(Cs cs, Th th) \<in> RAG s"  | 
|    206 proof - |    143 proof - | 
|    207  from holding_th |    144  from holding_th_cs_s | 
|    208  show ?thesis  |    145  show ?thesis  | 
|    209     by (unfold s_RAG_def holding_eq, auto) |    146     by (unfold s_RAG_def holding_eq, auto) | 
|    210 qed |    147 qed | 
|    211  |    148  | 
|    212 lemma ancestors_cs:  |    149 lemma ancestors_cs:  | 
|    213   "ancestors (RAG s') (Cs cs) = {Th th}" |    150   "ancestors (RAG s) (Cs cs) = {Th th}" | 
|    214 proof - |    151 proof - | 
|    215   have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th)  \<union>  {Th th}" |    152   have "ancestors (RAG s) (Cs cs) = ancestors (RAG s) (Th th)  \<union>  {Th th}" | 
|    216   proof(rule vat_s'.rtree_RAG.ancestors_accum) |    153    by (rule rtree_RAG.ancestors_accum[OF edge_of_th]) | 
|    217     from vt_s[unfolded s_def] |         | 
|    218     have " PIP s' (V th cs)" by (cases, simp) |         | 
|    219     thus "(Cs cs, Th th) \<in> RAG s'"  |         | 
|    220     proof(cases) |         | 
|    221       assume "holding s' th cs" |         | 
|    222       from this[unfolded holding_eq] |         | 
|    223       show ?thesis by (unfold s_RAG_def, auto) |         | 
|    224     qed |         | 
|    225   qed |         | 
|    226   from this[unfolded ancestors_th] show ?thesis by simp |    154   from this[unfolded ancestors_th] show ?thesis by simp | 
|    227 qed |    155 qed | 
|    228  |         | 
|    229 lemma preced_kept: "the_preced s = the_preced s'" |         | 
|    230   by (auto simp: s_def the_preced_def preced_def) |         | 
|    231  |    156  | 
|    232 end |    157 end | 
|    233  |    158  | 
|    234 text {* |    159 text {* | 
|    235   The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation,  |    160   The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation,  | 
|    236   which represents the case when there is another thread @{text "th'"} |    161   which represents the case when there is another thread @{text "th'"} | 
|    237   to take over the critical resource released by the initiating thread @{text "th"}. |    162   to take over the critical resource released by the initiating thread @{text "th"}. | 
|    238 *} |    163 *} | 
|    239 locale step_v_cps_nt = step_v_cps + |    164  | 
|    240   fixes th' |    165 context valid_trace_v_n | 
|    241   -- {* @{text "th'"} is assumed to take over @{text "cs"} *} |         | 
|    242   assumes nt: "next_th s' th cs th'"  |         | 
|    243  |         | 
|    244 context step_v_cps_nt |         | 
|    245 begin |    166 begin | 
|    246  |    167  | 
|    247 text {* |    168 lemma sub_RAGs':  | 
|    248   Lemma @{text "RAG_s"} confirms the change of RAG: |    169   "{(Cs cs, Th th), (Th taker, Cs cs)} \<subseteq> RAG s" | 
|    249   two edges removed and one added, as shown by the following diagram. |    170      using next_th_RAG[OF next_th_taker]  . | 
|    250 *} |         | 
|    251  |         | 
|    252 (* |         | 
|    253   RAG before the V-operation |         | 
|    254     th1 ----| |         | 
|    255             | |         | 
|    256     th' ----| |         | 
|    257             |----> cs -----| |         | 
|    258     th2 ----|              | |         | 
|    259             |              | |         | 
|    260     th3 ----|              | |         | 
|    261                            |------> th |         | 
|    262     th4 ----|              | |         | 
|    263             |              | |         | 
|    264     th5 ----|              | |         | 
|    265             |----> cs'-----| |         | 
|    266     th6 ----| |         | 
|    267             | |         | 
|    268     th7 ----| |         | 
|    269  |         | 
|    270  RAG after the V-operation |         | 
|    271     th1 ----| |         | 
|    272             | |         | 
|    273             |----> cs ----> th' |         | 
|    274     th2 ----|               |         | 
|    275             |               |         | 
|    276     th3 ----|               |         | 
|    277                             |         | 
|    278     th4 ----|               |         | 
|    279             |               |         | 
|    280     th5 ----|               |         | 
|    281             |----> cs'----> th |         | 
|    282     th6 ----| |         | 
|    283             | |         | 
|    284     th7 ----| |         | 
|    285 *) |         | 
|    286  |         | 
|    287 lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'" |         | 
|    288                 using next_th_RAG[OF nt]  . |         | 
|    289  |    171  | 
|    290 lemma ancestors_th':  |    172 lemma ancestors_th':  | 
|    291   "ancestors (RAG s') (Th th') = {Th th, Cs cs}"  |    173   "ancestors (RAG s) (Th taker) = {Th th, Cs cs}"  | 
|    292 proof - |    174 proof - | 
|    293   have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}" |    175   have "ancestors (RAG s) (Th taker) = ancestors (RAG s) (Cs cs) \<union> {Cs cs}" | 
|    294   proof(rule  vat_s'.rtree_RAG.ancestors_accum) |    176   proof(rule  rtree_RAG.ancestors_accum) | 
|    295     from sub_RAGs' show "(Th th', Cs cs) \<in> RAG s'" by auto |    177     from sub_RAGs' show "(Th taker, Cs cs) \<in> RAG s" by auto | 
|    296   qed |    178   qed | 
|    297   thus ?thesis using ancestors_th ancestors_cs by auto |    179   thus ?thesis using ancestors_th ancestors_cs by auto | 
|    298 qed |    180 qed | 
|    299  |    181  | 
|    300 lemma RAG_s: |    182 lemma RAG_s: | 
|    301   "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union> |    183   "RAG (e#s) = (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) \<union> | 
|    302                                          {(Cs cs, Th th')}" |    184                                          {(Cs cs, Th taker)}" | 
|    303 proof - |    185  by (unfold RAG_es waiting_set_eq holding_set_eq, auto) | 
|    304   from step_RAG_v[OF vt_s[unfolded s_def], folded s_def] |         | 
|    305     and nt show ?thesis  by (auto intro:next_th_unique) |         | 
|    306 qed |         | 
|    307  |    186  | 
|    308 lemma subtree_kept: (* ddd *) |    187 lemma subtree_kept: (* ddd *) | 
|    309   assumes "th1 \<notin> {th, th'}" |    188   assumes "th1 \<notin> {th, taker}" | 
|    310   shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R") |    189   shows "subtree (RAG (e#s)) (Th th1) =  | 
|    311 proof - |    190                      subtree (RAG s) (Th th1)" (is "_ = ?R") | 
|    312   let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})" |    191 proof - | 
|    313   let ?RAG'' = "?RAG' \<union> {(Cs cs, Th th')}" |    192   let ?RAG' = "(RAG s - {(Cs cs, Th th), (Th taker, Cs cs)})" | 
|         |    193   let ?RAG'' = "?RAG' \<union> {(Cs cs, Th taker)}" | 
|    314   have "subtree ?RAG' (Th th1) = ?R"  |    194   have "subtree ?RAG' (Th th1) = ?R"  | 
|    315   proof(rule subset_del_subtree_outside) |    195   proof(rule subset_del_subtree_outside) | 
|    316     show "Range {(Cs cs, Th th), (Th th', Cs cs)} \<inter> subtree (RAG s') (Th th1) = {}" |    196     show "Range {(Cs cs, Th th), (Th taker, Cs cs)} \<inter> subtree (RAG s) (Th th1) = {}" | 
|    317     proof - |    197     proof - | 
|    318       have "(Th th) \<notin> subtree (RAG s') (Th th1)" |    198       have "(Th th) \<notin> subtree (RAG s) (Th th1)" | 
|    319       proof(rule subtree_refute) |    199       proof(rule subtree_refute) | 
|    320         show "Th th1 \<notin> ancestors (RAG s') (Th th)" |    200         show "Th th1 \<notin> ancestors (RAG s) (Th th)" | 
|    321           by (unfold ancestors_th, simp) |    201           by (unfold ancestors_th, simp) | 
|    322       next |    202       next | 
|    323         from assms show "Th th1 \<noteq> Th th" by simp |    203         from assms show "Th th1 \<noteq> Th th" by simp | 
|    324       qed |    204       qed | 
|    325       moreover have "(Cs cs) \<notin>  subtree (RAG s') (Th th1)" |    205       moreover have "(Cs cs) \<notin>  subtree (RAG s) (Th th1)" | 
|    326       proof(rule subtree_refute) |    206       proof(rule subtree_refute) | 
|    327         show "Th th1 \<notin> ancestors (RAG s') (Cs cs)" |    207         show "Th th1 \<notin> ancestors (RAG s) (Cs cs)" | 
|    328           by (unfold ancestors_cs, insert assms, auto) |    208           by (unfold ancestors_cs, insert assms, auto) | 
|    329       qed simp |    209       qed simp | 
|    330       ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s') (Th th1) = {}" by auto |    210       ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s) (Th th1) = {}" by auto | 
|    331       thus ?thesis by simp |    211       thus ?thesis by simp | 
|    332      qed |    212      qed | 
|    333   qed |    213   qed | 
|    334   moreover have "subtree ?RAG'' (Th th1) =  subtree ?RAG' (Th th1)" |    214   moreover have "subtree ?RAG'' (Th th1) =  subtree ?RAG' (Th th1)" | 
|    335   proof(rule subtree_insert_next) |    215   proof(rule subtree_insert_next) | 
|    336     show "Th th' \<notin> subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)" |    216     show "Th taker \<notin> subtree (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th th1)" | 
|    337     proof(rule subtree_refute) |    217     proof(rule subtree_refute) | 
|    338       show "Th th1 \<notin> ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')" |    218       show "Th th1 \<notin> ancestors (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th taker)" | 
|    339             (is "_ \<notin> ?R") |    219             (is "_ \<notin> ?R") | 
|    340       proof - |    220       proof - | 
|    341           have "?R \<subseteq> ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto) |    221           have "?R \<subseteq> ancestors (RAG s) (Th taker)" by (rule ancestors_mono, auto) | 
|    342           moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp |    222           moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp | 
|    343           ultimately show ?thesis by auto |    223           ultimately show ?thesis by auto | 
|    344       qed |    224       qed | 
|    345     next |    225     next | 
|    346       from assms show "Th th1 \<noteq> Th th'" by simp |    226       from assms show "Th th1 \<noteq> Th taker" by simp | 
|    347     qed |    227     qed | 
|    348   qed |    228   qed | 
|    349   ultimately show ?thesis by (unfold RAG_s, simp) |    229   ultimately show ?thesis by (unfold RAG_s, simp) | 
|    350 qed |    230 qed | 
|    351  |    231  | 
|    352 lemma cp_kept: |    232 lemma cp_kept: | 
|    353   assumes "th1 \<notin> {th, th'}" |    233   assumes "th1 \<notin> {th, taker}" | 
|    354   shows "cp s th1 = cp s' th1" |    234   shows "cp (e#s) th1 = cp s th1" | 
|    355     by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) |    235     by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp) | 
|    356  |    236  | 
|    357 end |    237 end | 
|    358  |    238  | 
|    359 locale step_v_cps_nnt = step_v_cps + |    239  | 
|    360   assumes nnt: "\<And> th'. (\<not> next_th s' th cs th')" |    240 context valid_trace_v_e | 
|    361  |         | 
|    362 context step_v_cps_nnt |         | 
|    363 begin |    241 begin | 
|    364  |    242  | 
|    365 lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}" |    243 find_theorems RAG s e | 
|    366 proof - |    244  | 
|    367   from nnt and  step_RAG_v[OF vt_s[unfolded s_def], folded s_def] |    245 lemma RAG_s: "RAG (e#s) = RAG s - {(Cs cs, Th th)}" | 
|    368   show ?thesis by auto |    246   by (unfold RAG_es waiting_set_eq holding_set_eq, simp) | 
|    369 qed |         | 
|    370  |    247  | 
|    371 lemma subtree_kept: |    248 lemma subtree_kept: | 
|    372   assumes "th1 \<noteq> th" |    249   assumes "th1 \<noteq> th" | 
|    373   shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" |    250   shows "subtree (RAG (e#s)) (Th th1) = subtree (RAG s) (Th th1)" | 
|    374 proof(unfold RAG_s, rule subset_del_subtree_outside) |    251 proof(unfold RAG_s, rule subset_del_subtree_outside) | 
|    375   show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s') (Th th1) = {}" |    252   show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s) (Th th1) = {}" | 
|    376   proof - |    253   proof - | 
|    377     have "(Th th) \<notin> subtree (RAG s') (Th th1)" |    254     have "(Th th) \<notin> subtree (RAG s) (Th th1)" | 
|    378     proof(rule subtree_refute) |    255     proof(rule subtree_refute) | 
|    379       show "Th th1 \<notin> ancestors (RAG s') (Th th)" |    256       show "Th th1 \<notin> ancestors (RAG s) (Th th)" | 
|    380           by (unfold ancestors_th, simp) |    257           by (unfold ancestors_th, simp) | 
|    381     next |    258     next | 
|    382       from assms show "Th th1 \<noteq> Th th" by simp |    259       from assms show "Th th1 \<noteq> Th th" by simp | 
|    383     qed |    260     qed | 
|    384     thus ?thesis by auto |    261     thus ?thesis by auto | 
|    385   qed |    262   qed | 
|    386 qed |    263 qed | 
|    387  |    264  | 
|    388 lemma cp_kept_1: |    265 lemma cp_kept_1: | 
|    389   assumes "th1 \<noteq> th" |    266   assumes "th1 \<noteq> th" | 
|    390   shows "cp s th1 = cp s' th1" |    267   shows "cp (e#s) th1 = cp s th1" | 
|    391     by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) |    268     by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp) | 
|    392  |    269  | 
|    393 lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}" |    270 lemma subtree_cs: "subtree (RAG s) (Cs cs) = {Cs cs}" | 
|    394 proof - |    271 proof - | 
|    395   { fix n |    272   { fix n | 
|    396     have "(Cs cs) \<notin> ancestors (RAG s') n" |    273     have "(Cs cs) \<notin> ancestors (RAG s) n" | 
|    397     proof |    274     proof | 
|    398       assume "Cs cs \<in> ancestors (RAG s') n" |    275       assume "Cs cs \<in> ancestors (RAG s) n" | 
|    399       hence "(n, Cs cs) \<in> (RAG s')^+" by (auto simp:ancestors_def) |    276       hence "(n, Cs cs) \<in> (RAG s)^+" by (auto simp:ancestors_def) | 
|    400       from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s'" by auto |    277       from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s" by auto | 
|    401       then obtain th' where "nn = Th th'" |    278       then obtain th' where "nn = Th th'" | 
|    402         by (unfold s_RAG_def, auto) |    279         by (unfold s_RAG_def, auto) | 
|    403       from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s'" . |    280       from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s" . | 
|    404       from this[unfolded s_RAG_def] |    281       from this[unfolded s_RAG_def] | 
|    405       have "waiting (wq s') th' cs" by auto |    282       have "waiting (wq s) th' cs" by auto | 
|    406       from this[unfolded cs_waiting_def] |    283       from this[unfolded cs_waiting_def] | 
|    407       have "1 < length (wq s' cs)" |    284       have "1 < length (wq s cs)" | 
|    408           by (cases "wq s' cs", auto) |    285           by (cases "wq s cs", auto) | 
|    409       from holding_next_thI[OF holding_th this] |    286       from holding_next_thI[OF holding_th_cs_s this] | 
|    410       obtain th' where "next_th s' th cs th'" by auto |    287       obtain th' where "next_th s th cs th'" by auto | 
|    411       with nnt show False by auto |    288       thus False using no_taker by blast | 
|    412     qed |    289     qed | 
|    413   } note h = this |    290   } note h = this | 
|    414   {  fix n |    291   {  fix n | 
|    415      assume "n \<in> subtree (RAG s') (Cs cs)" |    292      assume "n \<in> subtree (RAG s) (Cs cs)" | 
|    416      hence "n = (Cs cs)" |    293      hence "n = (Cs cs)" | 
|    417      by (elim subtreeE, insert h, auto) |    294      by (elim subtreeE, insert h, auto) | 
|    418   } moreover have "(Cs cs) \<in> subtree (RAG s') (Cs cs)" |    295   } moreover have "(Cs cs) \<in> subtree (RAG s) (Cs cs)" | 
|    419       by (auto simp:subtree_def) |    296       by (auto simp:subtree_def) | 
|    420   ultimately show ?thesis by auto  |    297   ultimately show ?thesis by auto  | 
|    421 qed |    298 qed | 
|    422  |    299  | 
|    423 lemma subtree_th:  |    300 lemma subtree_th:  | 
|    424   "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}" |    301   "subtree (RAG (e#s)) (Th th) = subtree (RAG s) (Th th) - {Cs cs}" | 
|    425 proof(unfold RAG_s, fold subtree_cs, rule vat_s'.rtree_RAG.subtree_del_inside) |    302 proof(unfold RAG_s, fold subtree_cs, rule rtree_RAG.subtree_del_inside) | 
|    426   from edge_of_th |    303   from edge_of_th | 
|    427   show "(Cs cs, Th th) \<in> edges_in (RAG s') (Th th)" |    304   show "(Cs cs, Th th) \<in> edges_in (RAG s) (Th th)" | 
|    428     by (unfold edges_in_def, auto simp:subtree_def) |    305     by (unfold edges_in_def, auto simp:subtree_def) | 
|    429 qed |    306 qed | 
|    430  |    307  | 
|    431 lemma cp_kept_2:  |    308 lemma cp_kept_2:  | 
|    432   shows "cp s th = cp s' th"  |    309   shows "cp (e#s) th = cp s th"  | 
|    433  by (unfold cp_alt_def subtree_th preced_kept, auto) |    310  by (unfold cp_alt_def subtree_th the_preced_es, auto) | 
|    434  |    311  | 
|    435 lemma eq_cp: |    312 lemma eq_cp: | 
|    436   shows "cp s th' = cp s' th'" |    313   shows "cp (e#s) th' = cp s th'" | 
|    437   using cp_kept_1 cp_kept_2 |    314   using cp_kept_1 cp_kept_2 | 
|    438   by (cases "th' = th", auto) |    315   by (cases "th' = th", auto) | 
|         |    316  | 
|    439 end |    317 end | 
|    440  |    318  | 
|    441  |    319  | 
|    442 locale step_P_cps = |         | 
|    443   fixes s' th cs s  |         | 
|    444   defines s_def : "s \<equiv> (P th cs#s')" |         | 
|    445   assumes vt_s: "vt s" |         | 
|    446  |         | 
|    447 sublocale step_P_cps < vat_s : valid_trace "s" |         | 
|    448 proof |         | 
|    449   from vt_s show "vt s" . |         | 
|    450 qed |         | 
|    451  |         | 
|    452 section {* The @{term P} operation *} |    320 section {* The @{term P} operation *} | 
|    453  |    321  | 
|    454 sublocale step_P_cps < vat_s' : valid_trace "s'" |    322 context valid_trace_p | 
|    455 proof |         | 
|    456   from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . |         | 
|    457 qed |         | 
|    458  |         | 
|    459 context step_P_cps |         | 
|    460 begin |    323 begin | 
|    461  |    324  | 
|    462 lemma readys_th: "th \<in> readys s'" |    325 lemma root_th: "root (RAG s) (Th th)" | 
|    463 proof - |    326   by (simp add: ready_th_s readys_root) | 
|    464     from step_back_step [OF vt_s[unfolded s_def]] |         | 
|    465     have "PIP s' (P th cs)" . |         | 
|    466     hence "th \<in> runing s'" by (cases, simp) |         | 
|    467     thus ?thesis by (simp add:readys_def runing_def) |         | 
|    468 qed |         | 
|    469  |         | 
|    470 lemma root_th: "root (RAG s') (Th th)" |         | 
|    471   using readys_root[OF readys_th] . |         | 
|    472  |    327  | 
|    473 lemma in_no_others_subtree: |    328 lemma in_no_others_subtree: | 
|    474   assumes "th' \<noteq> th" |    329   assumes "th' \<noteq> th" | 
|    475   shows "Th th \<notin> subtree (RAG s') (Th th')" |    330   shows "Th th \<notin> subtree (RAG s) (Th th')" | 
|    476 proof |    331 proof | 
|    477   assume "Th th \<in> subtree (RAG s') (Th th')" |    332   assume "Th th \<in> subtree (RAG s) (Th th')" | 
|    478   thus False |    333   thus False | 
|    479   proof(cases rule:subtreeE) |    334   proof(cases rule:subtreeE) | 
|    480     case 1 |    335     case 1 | 
|    481     with assms show ?thesis by auto |    336     with assms show ?thesis by auto | 
|    482   next |    337   next | 
|    483     case 2 |    338     case 2 | 
|    484     with root_th show ?thesis by (auto simp:root_def) |    339     with root_th show ?thesis by (auto simp:root_def) | 
|    485   qed |    340   qed | 
|    486 qed |    341 qed | 
|    487  |    342  | 
|    488 lemma preced_kept: "the_preced s = the_preced s'" |    343 lemma preced_kept: "the_preced (e#s) = the_preced s" | 
|    489   by (auto simp: s_def the_preced_def preced_def) |    344 proof | 
|         |    345   fix th' | 
|         |    346   show "the_preced (e # s) th' = the_preced s th'" | 
|         |    347     by (unfold the_preced_def is_p preced_def, simp) | 
|         |    348 qed | 
|    490  |    349  | 
|    491 end |    350 end | 
|    492  |    351  | 
|    493 locale step_P_cps_ne =step_P_cps + |    352  | 
|    494   fixes th' |    353 context valid_trace_p_h | 
|    495   assumes ne: "wq s' cs \<noteq> []" |         | 
|    496   defines th'_def: "th' \<equiv> hd (wq s' cs)" |         | 
|    497  |         | 
|    498 locale step_P_cps_e =step_P_cps + |         | 
|    499   assumes ee: "wq s' cs = []" |         | 
|    500  |         | 
|    501 context step_P_cps_e |         | 
|    502 begin |    354 begin | 
|    503  |         | 
|    504 lemma RAG_s: "RAG s = RAG s' \<union> {(Cs cs, Th th)}" |         | 
|    505 proof - |         | 
|    506   from ee and  step_RAG_p[OF vt_s[unfolded s_def], folded s_def] |         | 
|    507   show ?thesis by auto |         | 
|    508 qed |         | 
|    509  |    355  | 
|    510 lemma subtree_kept: |    356 lemma subtree_kept: | 
|    511   assumes "th' \<noteq> th" |    357   assumes "th' \<noteq> th" | 
|    512   shows "subtree (RAG s) (Th th') = subtree (RAG s') (Th th')" |    358   shows "subtree (RAG (e#s)) (Th th') = subtree (RAG s) (Th th')" | 
|    513 proof(unfold RAG_s, rule subtree_insert_next) |    359 proof(unfold RAG_es, rule subtree_insert_next) | 
|    514   from in_no_others_subtree[OF assms]  |    360   from in_no_others_subtree[OF assms]  | 
|    515   show "Th th \<notin> subtree (RAG s') (Th th')" . |    361   show "Th th \<notin> subtree (RAG s) (Th th')" . | 
|    516 qed |    362 qed | 
|    517  |    363  | 
|    518 lemma cp_kept:  |    364 lemma cp_kept:  | 
|    519   assumes "th' \<noteq> th" |    365   assumes "th' \<noteq> th" | 
|    520   shows "cp s th' = cp s' th'" |    366   shows "cp (e#s) th' = cp s th'" | 
|    521 proof - |    367 proof - | 
|    522   have "(the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) = |    368   have "(the_preced (e#s) ` {th'a. Th th'a \<in> subtree (RAG (e#s)) (Th th')}) = | 
|    523         (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})" |    369         (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')})" | 
|    524         by (unfold preced_kept subtree_kept[OF assms], simp) |    370         by (unfold preced_kept subtree_kept[OF assms], simp) | 
|    525   thus ?thesis by (unfold cp_alt_def, simp) |    371   thus ?thesis by (unfold cp_alt_def, simp) | 
|    526 qed |    372 qed | 
|    527  |    373  | 
|    528 end |    374 end | 
|    529  |    375  | 
|    530 context step_P_cps_ne  |    376 context valid_trace_p_w | 
|    531 begin |    377 begin | 
|    532  |    378  | 
|    533 lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}" |    379 interpretation vat_e: valid_trace "e#s" | 
|    534 proof - |    380   by (unfold_locales, insert vt_e, simp) | 
|    535   from step_RAG_p[OF vt_s[unfolded s_def]] and ne |    381  | 
|    536   show ?thesis by (simp add:s_def) |    382 lemma cs_held: "(Cs cs, Th holder) \<in> RAG s" | 
|    537 qed |    383   using holding_s_holder | 
|    538  |    384   by (unfold s_RAG_def, fold holding_eq, auto) | 
|    539 lemma cs_held: "(Cs cs, Th th') \<in> RAG s'" |    385  | 
|    540 proof - |    386 lemma tRAG_s:  | 
|    541   have "(Cs cs, Th th') \<in> hRAG s'" |    387   "tRAG (e#s) = tRAG s \<union> {(Th th, Th holder)}" | 
|         |    388   using local.RAG_tRAG_transfer[OF RAG_es cs_held] . | 
|         |    389  | 
|         |    390 lemma cp_kept: | 
|         |    391   assumes "Th th'' \<notin> ancestors (tRAG (e#s)) (Th th)" | 
|         |    392   shows "cp (e#s) th'' = cp s th''" | 
|         |    393 proof - | 
|         |    394   have h: "subtree (tRAG (e#s)) (Th th'') = subtree (tRAG s) (Th th'')" | 
|    542   proof - |    395   proof - | 
|    543     from ne |    396     have "Th holder \<notin> subtree (tRAG s) (Th th'')" | 
|    544     have " holding s' th' cs" |         | 
|    545       by (unfold th'_def holding_eq cs_holding_def, auto) |         | 
|    546     thus ?thesis  |         | 
|    547       by (unfold hRAG_def, auto) |         | 
|    548   qed |         | 
|    549   thus ?thesis by (unfold RAG_split, auto) |         | 
|    550 qed |         | 
|    551  |         | 
|    552 lemma tRAG_s:  |         | 
|    553   "tRAG s = tRAG s' \<union> {(Th th, Th th')}" |         | 
|    554   using RAG_tRAG_transfer[OF RAG_s cs_held] . |         | 
|    555  |         | 
|    556 lemma cp_kept: |         | 
|    557   assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)" |         | 
|    558   shows "cp s th'' = cp s' th''" |         | 
|    559 proof - |         | 
|    560   have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')" |         | 
|    561   proof - |         | 
|    562     have "Th th' \<notin> subtree (tRAG s') (Th th'')" |         | 
|    563     proof |    397     proof | 
|    564       assume "Th th' \<in> subtree (tRAG s') (Th th'')" |    398       assume "Th holder \<in> subtree (tRAG s) (Th th'')" | 
|    565       thus False |    399       thus False | 
|    566       proof(rule subtreeE) |    400       proof(rule subtreeE) | 
|    567          assume "Th th' = Th th''" |    401          assume "Th holder = Th th''" | 
|    568          from assms[unfolded tRAG_s ancestors_def, folded this] |    402          from assms[unfolded tRAG_s ancestors_def, folded this] | 
|    569          show ?thesis by auto |    403          show ?thesis by auto | 
|    570       next |    404       next | 
|    571          assume "Th th'' \<in> ancestors (tRAG s') (Th th')" |    405          assume "Th th'' \<in> ancestors (tRAG s) (Th holder)" | 
|    572          moreover have "... \<subseteq> ancestors (tRAG s) (Th th')" |    406          moreover have "... \<subseteq> ancestors (tRAG (e#s)) (Th holder)" | 
|    573          proof(rule ancestors_mono) |    407          proof(rule ancestors_mono) | 
|    574             show "tRAG s' \<subseteq> tRAG s" by (unfold tRAG_s, auto) |    408             show "tRAG s \<subseteq> tRAG (e#s)" by (unfold tRAG_s, auto) | 
|    575          qed  |    409          qed  | 
|    576          ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th')" by auto |    410          ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th holder)" by auto | 
|    577          moreover have "Th th' \<in> ancestors (tRAG s) (Th th)" |    411          moreover have "Th holder \<in> ancestors (tRAG (e#s)) (Th th)" | 
|    578            by (unfold tRAG_s, auto simp:ancestors_def) |    412            by (unfold tRAG_s, auto simp:ancestors_def) | 
|    579          ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th)" |    413          ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th th)" | 
|    580                        by (auto simp:ancestors_def) |    414                        by (auto simp:ancestors_def) | 
|    581          with assms show ?thesis by auto |    415          with assms show ?thesis by auto | 
|    582       qed |    416       qed | 
|    583     qed |    417     qed | 
|    584     from subtree_insert_next[OF this] |    418     from subtree_insert_next[OF this] | 
|    585     have "subtree (tRAG s' \<union> {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" . |    419     have "subtree (tRAG s \<union> {(Th th, Th holder)}) (Th th'') = subtree (tRAG s) (Th th'')" . | 
|    586     from this[folded tRAG_s] show ?thesis . |    420     from this[folded tRAG_s] show ?thesis . | 
|    587   qed |    421   qed | 
|    588   show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) |    422   show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) | 
|    589 qed |    423 qed | 
|    590  |    424  | 
|    591 lemma cp_gen_update_stop: (* ddd *) |    425 lemma cp_gen_update_stop: (* ddd *) | 
|    592   assumes "u \<in> ancestors (tRAG s) (Th th)" |    426   assumes "u \<in> ancestors (tRAG (e#s)) (Th th)" | 
|    593   and "cp_gen s u = cp_gen s' u" |    427   and "cp_gen (e#s) u = cp_gen s u" | 
|    594   and "y \<in> ancestors (tRAG s) u" |    428   and "y \<in> ancestors (tRAG (e#s)) u" | 
|    595   shows "cp_gen s y = cp_gen s' y" |    429   shows "cp_gen (e#s) y = cp_gen s y" | 
|    596   using assms(3) |    430   using assms(3) | 
|    597 proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf]) |    431 proof(induct rule:wf_induct[OF vat_e.fsbttRAGs.wf]) | 
|    598   case (1 x) |    432   case (1 x) | 
|    599   show ?case (is "?L = ?R") |    433   show ?case (is "?L = ?R") | 
|    600   proof - |    434   proof - | 
|    601     from tRAG_ancestorsE[OF 1(2)] |    435     from tRAG_ancestorsE[OF 1(2)] | 
|    602     obtain th2 where eq_x: "x = Th th2" by blast |    436     obtain th2 where eq_x: "x = Th th2" by blast | 
|    603     from vat_s.cp_gen_rec[OF this] |    437     from vat_e.cp_gen_rec[OF this] | 
|    604     have "?L =  |    438     have "?L =  | 
|    605           Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" . |    439           Max ({the_preced (e#s) th2} \<union> cp_gen (e#s) ` RTree.children (tRAG (e#s)) x)" . | 
|    606     also have "... =  |    440     also have "... =  | 
|    607           Max ({the_preced s' th2} \<union> cp_gen s' ` RTree.children (tRAG s') x)" |    441           Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" | 
|    608    |         | 
|    609     proof - |    442     proof - | 
|    610       from preced_kept have "the_preced s th2 = the_preced s' th2" by simp |    443       from preced_kept have "the_preced (e#s) th2 = the_preced s th2" by simp | 
|    611       moreover have "cp_gen s ` RTree.children (tRAG s) x = |    444       moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x = | 
|    612                      cp_gen s' ` RTree.children (tRAG s') x" |    445                      cp_gen s ` RTree.children (tRAG s) x" | 
|    613       proof - |    446       proof - | 
|    614         have "RTree.children (tRAG s) x =  RTree.children (tRAG s') x" |    447         have "RTree.children (tRAG (e#s)) x =  RTree.children (tRAG s) x" | 
|    615         proof(unfold tRAG_s, rule children_union_kept) |    448         proof(unfold tRAG_s, rule children_union_kept) | 
|    616           have start: "(Th th, Th th') \<in> tRAG s" |    449           have start: "(Th th, Th holder) \<in> tRAG (e#s)" | 
|    617             by (unfold tRAG_s, auto) |    450             by (unfold tRAG_s, auto) | 
|    618           note x_u = 1(2) |    451           note x_u = 1(2) | 
|    619           show "x \<notin> Range {(Th th, Th th')}" |    452           show "x \<notin> Range {(Th th, Th holder)}" | 
|    620           proof |    453           proof | 
|    621             assume "x \<in> Range {(Th th, Th th')}" |    454             assume "x \<in> Range {(Th th, Th holder)}" | 
|    622             hence eq_x: "x = Th th'" using RangeE by auto |    455             hence eq_x: "x = Th holder" using RangeE by auto | 
|    623             show False |    456             show False | 
|    624             proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start]) |    457             proof(cases rule:vat_e.ancestors_headE[OF assms(1) start]) | 
|    625               case 1 |    458               case 1 | 
|    626               from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG |    459               from x_u[folded this, unfolded eq_x] vat_e.acyclic_tRAG | 
|    627               show ?thesis by (auto simp:ancestors_def acyclic_def) |    460               show ?thesis by (auto simp:ancestors_def acyclic_def) | 
|    628             next |    461             next | 
|    629               case 2 |    462               case 2 | 
|    630               with x_u[unfolded eq_x] |    463               with x_u[unfolded eq_x] | 
|    631               have "(Th th', Th th') \<in> (tRAG s)^+" by (auto simp:ancestors_def) |    464               have "(Th holder, Th holder) \<in> (tRAG (e#s))^+" by (auto simp:ancestors_def) | 
|    632               with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) |    465               with vat_e.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) | 
|    633             qed |    466             qed | 
|    634           qed |    467           qed | 
|    635         qed |    468         qed | 
|    636         moreover have "cp_gen s ` RTree.children (tRAG s) x = |    469         moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x = | 
|    637                        cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A") |    470                        cp_gen s ` RTree.children (tRAG (e#s)) x" (is "?f ` ?A = ?g ` ?A") | 
|    638         proof(rule f_image_eq) |    471         proof(rule f_image_eq) | 
|    639           fix a |    472           fix a | 
|    640           assume a_in: "a \<in> ?A" |    473           assume a_in: "a \<in> ?A" | 
|    641           from 1(2) |    474           from 1(2) | 
|    642           show "?f a = ?g a" |    475           show "?f a = ?g a" | 
|    643           proof(cases rule:vat_s.rtree_s.ancestors_childrenE[case_names in_ch out_ch]) |    476           proof(cases rule:vat_e.rtree_s.ancestors_childrenE[case_names in_ch out_ch]) | 
|    644              case in_ch |    477              case in_ch | 
|    645              show ?thesis |    478              show ?thesis | 
|    646              proof(cases "a = u") |    479              proof(cases "a = u") | 
|    647                 case True |    480                 case True | 
|    648                 from assms(2)[folded this] show ?thesis . |    481                 from assms(2)[folded this] show ?thesis . | 
|    649              next |    482              next | 
|    650                 case False |    483                 case False | 
|    651                 have a_not_in: "a \<notin> ancestors (tRAG s) (Th th)" |    484                 have a_not_in: "a \<notin> ancestors (tRAG (e#s)) (Th th)" | 
|    652                 proof |    485                 proof | 
|    653                   assume a_in': "a \<in> ancestors (tRAG s) (Th th)" |    486                   assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)" | 
|    654                   have "a = u" |    487                   have "a = u" | 
|    655                   proof(rule vat_s.rtree_s.ancestors_children_unique) |    488                   proof(rule vat_e.rtree_s.ancestors_children_unique) | 
|    656                     from a_in' a_in show "a \<in> ancestors (tRAG s) (Th th) \<inter>  |    489                     from a_in' a_in show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter>  | 
|    657                                           RTree.children (tRAG s) x" by auto |    490                                           RTree.children (tRAG (e#s)) x" by auto | 
|    658                   next  |    491                   next  | 
|    659                     from assms(1) in_ch show "u \<in> ancestors (tRAG s) (Th th) \<inter>  |    492                     from assms(1) in_ch show "u \<in> ancestors (tRAG (e#s)) (Th th) \<inter>  | 
|    660                                       RTree.children (tRAG s) x" by auto |    493                                       RTree.children (tRAG (e#s)) x" by auto | 
|    661                   qed |    494                   qed | 
|    662                   with False show False by simp |    495                   with False show False by simp | 
|    663                 qed |    496                 qed | 
|    664                 from a_in obtain th_a where eq_a: "a = Th th_a"  |    497                 from a_in obtain th_a where eq_a: "a = Th th_a"  | 
|    665                     by (unfold RTree.children_def tRAG_alt_def, auto) |    498                     by (unfold RTree.children_def tRAG_alt_def, auto) | 
|    666                 from cp_kept[OF a_not_in[unfolded eq_a]] |    499                 from cp_kept[OF a_not_in[unfolded eq_a]] | 
|    667                 have "cp s th_a = cp s' th_a" . |    500                 have "cp (e#s) th_a = cp s th_a" . | 
|    668                 from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a] |    501                 from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a] | 
|    669                 show ?thesis . |    502                 show ?thesis . | 
|    670              qed |    503              qed | 
|    671           next |    504           next | 
|    672             case (out_ch z) |    505             case (out_ch z) | 
|    673             hence h: "z \<in> ancestors (tRAG s) u" "z \<in> RTree.children (tRAG s) x" by auto |    506             hence h: "z \<in> ancestors (tRAG (e#s)) u" "z \<in> RTree.children (tRAG (e#s)) x" by auto | 
|    674             show ?thesis |    507             show ?thesis | 
|    675             proof(cases "a = z") |    508             proof(cases "a = z") | 
|    676               case True |    509               case True | 
|    677               from h(2) have zx_in: "(z, x) \<in> (tRAG s)" by (auto simp:RTree.children_def) |    510               from h(2) have zx_in: "(z, x) \<in> (tRAG (e#s))" by (auto simp:RTree.children_def) | 
|    678               from 1(1)[rule_format, OF this h(1)] |    511               from 1(1)[rule_format, OF this h(1)] | 
|    679               have eq_cp_gen: "cp_gen s z = cp_gen s' z" . |    512               have eq_cp_gen: "cp_gen (e#s) z = cp_gen s z" . | 
|    680               with True show ?thesis by metis |    513               with True show ?thesis by metis | 
|    681             next |    514             next | 
|    682               case False |    515               case False | 
|    683               from a_in obtain th_a where eq_a: "a = Th th_a" |    516               from a_in obtain th_a where eq_a: "a = Th th_a" | 
|    684                 by (auto simp:RTree.children_def tRAG_alt_def) |    517                 by (auto simp:RTree.children_def tRAG_alt_def) | 
|    685               have "a \<notin> ancestors (tRAG s) (Th th)" |    518               have "a \<notin> ancestors (tRAG (e#s)) (Th th)" | 
|    686               proof |    519               proof | 
|    687                 assume a_in': "a \<in> ancestors (tRAG s) (Th th)" |    520                 assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)" | 
|    688                 have "a = z" |    521                 have "a = z" | 
|    689                 proof(rule vat_s.rtree_s.ancestors_children_unique) |    522                 proof(rule vat_e.rtree_s.ancestors_children_unique) | 
|    690                   from assms(1) h(1) have "z \<in> ancestors (tRAG s) (Th th)" |    523                   from assms(1) h(1) have "z \<in> ancestors (tRAG (e#s)) (Th th)" | 
|    691                       by (auto simp:ancestors_def) |    524                       by (auto simp:ancestors_def) | 
|    692                   with h(2) show " z \<in> ancestors (tRAG s) (Th th) \<inter>  |    525                   with h(2) show " z \<in> ancestors (tRAG (e#s)) (Th th) \<inter>  | 
|    693                                        RTree.children (tRAG s) x" by auto |    526                                        RTree.children (tRAG (e#s)) x" by auto | 
|    694                 next |    527                 next | 
|    695                   from a_in a_in' |    528                   from a_in a_in' | 
|    696                   show "a \<in> ancestors (tRAG s) (Th th) \<inter> RTree.children (tRAG s) x" |    529                   show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter> RTree.children (tRAG (e#s)) x" | 
|    697                     by auto |    530                     by auto | 
|    698                 qed |    531                 qed | 
|    699                 with False show False by auto |    532                 with False show False by auto | 
|    700               qed |    533               qed | 
|    701               from cp_kept[OF this[unfolded eq_a]] |    534               from cp_kept[OF this[unfolded eq_a]] | 
|    702               have "cp s th_a = cp s' th_a" . |    535               have "cp (e#s) th_a = cp s th_a" . | 
|    703               from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a] |    536               from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a] | 
|    704               show ?thesis . |    537               show ?thesis . | 
|    705             qed |    538             qed | 
|    706           qed |    539           qed | 
|    707         qed |    540         qed | 
|    708         ultimately show ?thesis by metis |    541         ultimately show ?thesis by metis | 
|    709       qed |    542       qed | 
|    710       ultimately show ?thesis by simp |    543       ultimately show ?thesis by simp | 
|    711     qed |    544     qed | 
|    712     also have "... = ?R" |    545     also have "... = ?R" | 
|    713       by (fold vat_s'.cp_gen_rec[OF eq_x], simp) |    546       by (fold cp_gen_rec[OF eq_x], simp) | 
|    714     finally show ?thesis . |    547     finally show ?thesis . | 
|    715   qed |    548   qed | 
|    716 qed |    549 qed | 
|    717  |    550  | 
|    718 lemma cp_up: |    551 lemma cp_up: | 
|    719   assumes "(Th th') \<in> ancestors (tRAG s) (Th th)" |    552   assumes "(Th th') \<in> ancestors (tRAG (e#s)) (Th th)" | 
|    720   and "cp s th' = cp s' th'" |    553   and "cp (e#s) th' = cp s th'" | 
|    721   and "(Th th'') \<in> ancestors (tRAG s) (Th th')" |    554   and "(Th th'') \<in> ancestors (tRAG (e#s)) (Th th')" | 
|    722   shows "cp s th'' = cp s' th''" |    555   shows "cp (e#s) th'' = cp s th''" | 
|    723 proof - |    556 proof - | 
|    724   have "cp_gen s (Th th'') = cp_gen s' (Th th'')" |    557   have "cp_gen (e#s) (Th th'') = cp_gen s (Th th'')" | 
|    725   proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)]) |    558   proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)]) | 
|    726     from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]] |    559     from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]] | 
|    727     show "cp_gen s (Th th') = cp_gen s' (Th th')" by metis |    560     show "cp_gen (e#s) (Th th') = cp_gen s (Th th')" by metis | 
|    728   qed |    561   qed | 
|    729   with cp_gen_def_cond[OF refl[of "Th th''"]] |    562   with cp_gen_def_cond[OF refl[of "Th th''"]] | 
|    730   show ?thesis by metis |    563   show ?thesis by metis | 
|    731 qed |    564 qed | 
|    732  |    565  | 
|    733 end |    566 end | 
|    734  |    567  | 
|    735 section {* The @{term Create} operation *} |    568 section {* The @{term Create} operation *} | 
|    736  |    569  | 
|    737 locale step_create_cps = |    570 context valid_trace_create | 
|    738   fixes s' th prio s  |    571 begin  | 
|    739   defines s_def : "s \<equiv> (Create th prio#s')" |    572  | 
|    740   assumes vt_s: "vt s" |    573 interpretation vat_e: valid_trace "e#s" | 
|    741  |    574   by (unfold_locales, insert vt_e, simp) | 
|    742 sublocale step_create_cps < vat_s: valid_trace "s" |    575  | 
|    743   by (unfold_locales, insert vt_s, simp) |    576 lemma tRAG_kept: "tRAG (e#s) = tRAG s" | 
|    744  |    577   by (unfold tRAG_alt_def RAG_unchanged, auto) | 
|    745 sublocale step_create_cps < vat_s': valid_trace "s'" |         | 
|    746   by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp) |         | 
|    747  |         | 
|    748 context step_create_cps |         | 
|    749 begin |         | 
|    750  |         | 
|    751 lemma RAG_kept: "RAG s = RAG s'" |         | 
|    752   by (unfold s_def RAG_create_unchanged, auto) |         | 
|    753  |         | 
|    754 lemma tRAG_kept: "tRAG s = tRAG s'" |         | 
|    755   by (unfold tRAG_alt_def RAG_kept, auto) |         | 
|    756  |    578  | 
|    757 lemma preced_kept: |    579 lemma preced_kept: | 
|    758   assumes "th' \<noteq> th" |    580   assumes "th' \<noteq> th" | 
|    759   shows "the_preced s th' = the_preced s' th'" |    581   shows "the_preced (e#s) th' = the_preced s th'" | 
|    760   by (unfold s_def the_preced_def preced_def, insert assms, auto) |    582   by (unfold the_preced_def preced_def is_create, insert assms, auto) | 
|    761  |    583  | 
|    762 lemma th_not_in: "Th th \<notin> Field (tRAG s')" |    584 lemma th_not_in: "Th th \<notin> Field (tRAG s)" | 
|    763 proof - |    585   by (meson not_in_thread_isolated subsetCE tRAG_Field th_not_live_s) | 
|    764   from vt_s[unfolded s_def] |         | 
|    765   have "PIP s' (Create th prio)" by (cases, simp) |         | 
|    766   hence "th \<notin> threads s'" by(cases, simp) |         | 
|    767   from vat_s'.not_in_thread_isolated[OF this] |         | 
|    768   have "Th th \<notin> Field (RAG s')" . |         | 
|    769   with tRAG_Field show ?thesis by auto |         | 
|    770 qed |         | 
|    771  |    586  | 
|    772 lemma eq_cp: |    587 lemma eq_cp: | 
|    773   assumes neq_th: "th' \<noteq> th" |    588   assumes neq_th: "th' \<noteq> th" | 
|    774   shows "cp s th' = cp s' th'" |    589   shows "cp (e#s) th' = cp s th'" | 
|    775 proof - |    590 proof - | 
|    776   have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') = |    591   have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') = | 
|    777         (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')" |    592         (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')" | 
|    778   proof(unfold tRAG_kept, rule f_image_eq) |    593   proof(unfold tRAG_kept, rule f_image_eq) | 
|    779     fix a |    594     fix a | 
|    780     assume a_in: "a \<in> subtree (tRAG s') (Th th')" |    595     assume a_in: "a \<in> subtree (tRAG s) (Th th')" | 
|    781     then obtain th_a where eq_a: "a = Th th_a"  |    596     then obtain th_a where eq_a: "a = Th th_a"  | 
|    782     proof(cases rule:subtreeE) |    597     proof(cases rule:subtreeE) | 
|    783       case 2 |    598       case 2 | 
|    784       from ancestors_Field[OF 2(2)] |    599       from ancestors_Field[OF 2(2)] | 
|    785       and that show ?thesis by (unfold tRAG_alt_def, auto) |    600       and that show ?thesis by (unfold tRAG_alt_def, auto) | 
|    786     qed auto |    601     qed auto | 
|    787     have neq_th_a: "th_a \<noteq> th" |    602     have neq_th_a: "th_a \<noteq> th" | 
|    788     proof - |    603     proof - | 
|    789       have "(Th th) \<notin> subtree (tRAG s') (Th th')" |    604       have "(Th th) \<notin> subtree (tRAG s) (Th th')" | 
|    790       proof |    605       proof | 
|    791         assume "Th th \<in> subtree (tRAG s') (Th th')" |    606         assume "Th th \<in> subtree (tRAG s) (Th th')" | 
|    792         thus False |    607         thus False | 
|    793         proof(cases rule:subtreeE) |    608         proof(cases rule:subtreeE) | 
|    794           case 2 |    609           case 2 | 
|    795           from ancestors_Field[OF this(2)] |    610           from ancestors_Field[OF this(2)] | 
|    796           and th_not_in[unfolded Field_def] |    611           and th_not_in[unfolded Field_def] | 
|    798         qed (insert assms, auto) |    613         qed (insert assms, auto) | 
|    799       qed |    614       qed | 
|    800       with a_in[unfolded eq_a] show ?thesis by auto |    615       with a_in[unfolded eq_a] show ?thesis by auto | 
|    801     qed |    616     qed | 
|    802     from preced_kept[OF this] |    617     from preced_kept[OF this] | 
|    803     show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a" |    618     show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a" | 
|    804       by (unfold eq_a, simp) |    619       by (unfold eq_a, simp) | 
|    805   qed |    620   qed | 
|    806   thus ?thesis by (unfold cp_alt_def1, simp) |    621   thus ?thesis by (unfold cp_alt_def1, simp) | 
|    807 qed |    622 qed | 
|    808  |    623  | 
|    809 lemma children_of_th: "RTree.children (tRAG s) (Th th) = {}" |    624 lemma children_of_th: "RTree.children (tRAG (e#s)) (Th th) = {}" | 
|    810 proof - |    625 proof - | 
|    811   { fix a |    626   { fix a | 
|    812     assume "a \<in> RTree.children (tRAG s) (Th th)" |    627     assume "a \<in> RTree.children (tRAG (e#s)) (Th th)" | 
|    813     hence "(a, Th th) \<in> tRAG s" by (auto simp:RTree.children_def) |    628     hence "(a, Th th) \<in> tRAG (e#s)" by (auto simp:RTree.children_def) | 
|    814     with th_not_in have False  |    629     with th_not_in have False  | 
|    815      by (unfold Field_def tRAG_kept, auto) |    630      by (unfold Field_def tRAG_kept, auto) | 
|    816   } thus ?thesis by auto |    631   } thus ?thesis by auto | 
|    817 qed |    632 qed | 
|    818  |    633  | 
|    819 lemma eq_cp_th: "cp s th = preced th s" |    634 lemma eq_cp_th: "cp (e#s) th = preced th (e#s)" | 
|    820  by (unfold vat_s.cp_rec children_of_th, simp add:the_preced_def) |    635  by (unfold vat_e.cp_rec children_of_th, simp add:the_preced_def) | 
|    821  |    636  | 
|    822 end |    637 end | 
|    823  |    638  | 
|    824 locale step_exit_cps = |    639  | 
|    825   fixes s' th prio s  |    640 context valid_trace_exit | 
|    826   defines s_def : "s \<equiv> Exit th # s'" |         | 
|    827   assumes vt_s: "vt s" |         | 
|    828  |         | 
|    829 sublocale step_exit_cps < vat_s: valid_trace "s" |         | 
|    830   by (unfold_locales, insert vt_s, simp) |         | 
|    831  |         | 
|    832 sublocale step_exit_cps < vat_s': valid_trace "s'" |         | 
|    833   by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp) |         | 
|    834  |         | 
|    835 context step_exit_cps |         | 
|    836 begin |    641 begin | 
|    837  |    642  | 
|    838 lemma preced_kept: |    643 lemma preced_kept: | 
|    839   assumes "th' \<noteq> th" |    644   assumes "th' \<noteq> th" | 
|    840   shows "the_preced s th' = the_preced s' th'" |    645   shows "the_preced (e#s) th' = the_preced s th'" | 
|    841   by (unfold s_def the_preced_def preced_def, insert assms, auto) |    646   using assms | 
|    842  |    647   by (unfold the_preced_def is_exit preced_def, simp) | 
|    843 lemma RAG_kept: "RAG s = RAG s'" |    648  | 
|    844   by (unfold s_def RAG_exit_unchanged, auto) |    649 lemma tRAG_kept: "tRAG (e#s) = tRAG s" | 
|    845  |    650   by (unfold tRAG_alt_def RAG_unchanged, auto) | 
|    846 lemma tRAG_kept: "tRAG s = tRAG s'" |    651  | 
|    847   by (unfold tRAG_alt_def RAG_kept, auto) |    652 lemma th_RAG: "Th th \<notin> Field (RAG s)" | 
|    848  |    653 proof - | 
|    849 lemma th_ready: "th \<in> readys s'" |    654   have "Th th \<notin> Range (RAG s)" | 
|    850 proof - |         | 
|    851   from vt_s[unfolded s_def] |         | 
|    852   have "PIP s' (Exit th)" by (cases, simp) |         | 
|    853   hence h: "th \<in> runing s' \<and> holdents s' th = {}" by (cases, metis) |         | 
|    854   thus ?thesis by (unfold runing_def, auto) |         | 
|    855 qed |         | 
|    856  |         | 
|    857 lemma th_holdents: "holdents s' th = {}" |         | 
|    858 proof - |         | 
|    859  from vt_s[unfolded s_def] |         | 
|    860   have "PIP s' (Exit th)" by (cases, simp) |         | 
|    861   thus ?thesis by (cases, metis) |         | 
|    862 qed |         | 
|    863  |         | 
|    864 lemma th_RAG: "Th th \<notin> Field (RAG s')" |         | 
|    865 proof - |         | 
|    866   have "Th th \<notin> Range (RAG s')" |         | 
|    867   proof |    655   proof | 
|    868     assume "Th th \<in> Range (RAG s')" |    656     assume "Th th \<in> Range (RAG s)" | 
|    869     then obtain cs where "holding (wq s') th cs" |    657     then obtain cs where "holding (wq s) th cs" | 
|    870       by (unfold Range_iff s_RAG_def, auto) |    658       by (unfold Range_iff s_RAG_def, auto) | 
|    871     with th_holdents[unfolded holdents_def] |    659     with holdents_th_s[unfolded holdents_def] | 
|    872     show False by (unfold eq_holding, auto) |    660     show False by (unfold holding_eq, auto) | 
|    873   qed |    661   qed | 
|    874   moreover have "Th th \<notin> Domain (RAG s')" |    662   moreover have "Th th \<notin> Domain (RAG s)" | 
|    875   proof |    663   proof | 
|    876     assume "Th th \<in> Domain (RAG s')" |    664     assume "Th th \<in> Domain (RAG s)" | 
|    877     then obtain cs where "waiting (wq s') th cs" |    665     then obtain cs where "waiting (wq s) th cs" | 
|    878       by (unfold Domain_iff s_RAG_def, auto) |    666       by (unfold Domain_iff s_RAG_def, auto) | 
|    879     with th_ready show False by (unfold readys_def eq_waiting, auto) |    667     with th_ready_s show False by (unfold readys_def waiting_eq, auto) | 
|    880   qed |    668   qed | 
|    881   ultimately show ?thesis by (auto simp:Field_def) |    669   ultimately show ?thesis by (auto simp:Field_def) | 
|    882 qed |    670 qed | 
|    883  |    671  | 
|    884 lemma th_tRAG: "(Th th) \<notin> Field (tRAG s')" |    672 lemma th_tRAG: "(Th th) \<notin> Field (tRAG s)" | 
|    885   using th_RAG tRAG_Field[of s'] by auto |    673   using th_RAG tRAG_Field by auto | 
|    886  |    674  | 
|    887 lemma eq_cp: |    675 lemma eq_cp: | 
|    888   assumes neq_th: "th' \<noteq> th" |    676   assumes neq_th: "th' \<noteq> th" | 
|    889   shows "cp s th' = cp s' th'" |    677   shows "cp (e#s) th' = cp s th'" | 
|    890 proof - |    678 proof - | 
|    891   have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') = |    679   have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') = | 
|    892         (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')" |    680         (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')" | 
|    893   proof(unfold tRAG_kept, rule f_image_eq) |    681   proof(unfold tRAG_kept, rule f_image_eq) | 
|    894     fix a |    682     fix a | 
|    895     assume a_in: "a \<in> subtree (tRAG s') (Th th')" |    683     assume a_in: "a \<in> subtree (tRAG s) (Th th')" | 
|    896     then obtain th_a where eq_a: "a = Th th_a"  |    684     then obtain th_a where eq_a: "a = Th th_a"  | 
|    897     proof(cases rule:subtreeE) |    685     proof(cases rule:subtreeE) | 
|    898       case 2 |    686       case 2 | 
|    899       from ancestors_Field[OF 2(2)] |    687       from ancestors_Field[OF 2(2)] | 
|    900       and that show ?thesis by (unfold tRAG_alt_def, auto) |    688       and that show ?thesis by (unfold tRAG_alt_def, auto) | 
|    901     qed auto |    689     qed auto | 
|    902     have neq_th_a: "th_a \<noteq> th" |    690     have neq_th_a: "th_a \<noteq> th" | 
|    903     proof - |    691     proof - | 
|    904       from vat_s'.readys_in_no_subtree[OF th_ready assms] |    692       from readys_in_no_subtree[OF th_ready_s assms] | 
|    905       have "(Th th) \<notin> subtree (RAG s') (Th th')" . |    693       have "(Th th) \<notin> subtree (RAG s) (Th th')" . | 
|    906       with tRAG_subtree_RAG[of s' "Th th'"] |    694       with tRAG_subtree_RAG[of s "Th th'"] | 
|    907       have "(Th th) \<notin> subtree (tRAG s') (Th th')" by auto |    695       have "(Th th) \<notin> subtree (tRAG s) (Th th')" by auto | 
|    908       with a_in[unfolded eq_a] show ?thesis by auto |    696       with a_in[unfolded eq_a] show ?thesis by auto | 
|    909     qed |    697     qed | 
|    910     from preced_kept[OF this] |    698     from preced_kept[OF this] | 
|    911     show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a" |    699     show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a" | 
|    912       by (unfold eq_a, simp) |    700       by (unfold eq_a, simp) | 
|    913   qed |    701   qed | 
|    914   thus ?thesis by (unfold cp_alt_def1, simp) |    702   thus ?thesis by (unfold cp_alt_def1, simp) | 
|    915 qed |    703 qed | 
|    916  |    704  | 
|    917 end |    705 end | 
|    918  |    706  | 
|    919 end |    707 end | 
|    920  |    708  | 
|         |    709 ======= | 
|         |    710 theory ExtGG | 
|         |    711 imports PrioG CpsG | 
|         |    712 begin | 
|         |    713  | 
|         |    714 text {*  | 
|         |    715   The following two auxiliary lemmas are used to reason about @{term Max}. | 
|         |    716 *} | 
|         |    717 lemma image_Max_eqI:  | 
|         |    718   assumes "finite B" | 
|         |    719   and "b \<in> B" | 
|         |    720   and "\<forall> x \<in> B. f x \<le> f b" | 
|         |    721   shows "Max (f ` B) = f b" | 
|         |    722   using assms | 
|         |    723   using Max_eqI by blast  | 
|         |    724  | 
|         |    725 lemma image_Max_subset: | 
|         |    726   assumes "finite A" | 
|         |    727   and "B \<subseteq> A" | 
|         |    728   and "a \<in> B" | 
|         |    729   and "Max (f ` A) = f a" | 
|         |    730   shows "Max (f ` B) = f a" | 
|         |    731 proof(rule image_Max_eqI) | 
|         |    732   show "finite B" | 
|         |    733     using assms(1) assms(2) finite_subset by auto  | 
|         |    734 next | 
|         |    735   show "a \<in> B" using assms by simp | 
|         |    736 next | 
|         |    737   show "\<forall>x\<in>B. f x \<le> f a" | 
|         |    738     by (metis Max_ge assms(1) assms(2) assms(4)  | 
|         |    739             finite_imageI image_eqI subsetCE)  | 
|         |    740 qed | 
|         |    741  | 
|         |    742 text {* | 
|         |    743   The following locale @{text "highest_gen"} sets the basic context for our | 
|         |    744   investigation: supposing thread @{text th} holds the highest @{term cp}-value | 
|         |    745   in state @{text s}, which means the task for @{text th} is the  | 
|         |    746   most urgent. We want to show that   | 
|         |    747   @{text th} is treated correctly by PIP, which means | 
|         |    748   @{text th} will not be blocked unreasonably by other less urgent | 
|         |    749   threads.  | 
|         |    750 *} | 
|         |    751 locale highest_gen = | 
|         |    752   fixes s th prio tm | 
|         |    753   assumes vt_s: "vt s" | 
|         |    754   and threads_s: "th \<in> threads s" | 
|         |    755   and highest: "preced th s = Max ((cp s)`threads s)" | 
|         |    756   -- {* The internal structure of @{term th}'s precedence is exposed:*} | 
|         |    757   and preced_th: "preced th s = Prc prio tm"  | 
|         |    758  | 
|         |    759 -- {* @{term s} is a valid trace, so it will inherit all results derived for | 
|         |    760       a valid trace: *} | 
|         |    761 sublocale highest_gen < vat_s: valid_trace "s" | 
|         |    762   by (unfold_locales, insert vt_s, simp) | 
|         |    763  | 
|         |    764 context highest_gen | 
|         |    765 begin | 
|         |    766  | 
|         |    767 text {* | 
|         |    768   @{term tm} is the time when the precedence of @{term th} is set, so  | 
|         |    769   @{term tm} must be a valid moment index into @{term s}. | 
|         |    770 *} | 
|         |    771 lemma lt_tm: "tm < length s" | 
|         |    772   by (insert preced_tm_lt[OF threads_s preced_th], simp) | 
|         |    773  | 
|         |    774 text {* | 
|         |    775   Since @{term th} holds the highest precedence and @{text "cp"} | 
|         |    776   is the highest precedence of all threads in the sub-tree of  | 
|         |    777   @{text "th"} and @{text th} is among these threads,  | 
|         |    778   its @{term cp} must equal to its precedence: | 
|         |    779 *} | 
|         |    780 lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R") | 
|         |    781 proof - | 
|         |    782   have "?L \<le> ?R" | 
|         |    783   by (unfold highest, rule Max_ge,  | 
|         |    784         auto simp:threads_s finite_threads) | 
|         |    785   moreover have "?R \<le> ?L" | 
|         |    786     by (unfold vat_s.cp_rec, rule Max_ge,  | 
|         |    787         auto simp:the_preced_def vat_s.fsbttRAGs.finite_children) | 
|         |    788   ultimately show ?thesis by auto | 
|         |    789 qed | 
|         |    790  | 
|         |    791 (* ccc *) | 
|         |    792 lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)" | 
|         |    793   by (fold max_cp_eq, unfold eq_cp_s_th, insert highest, simp) | 
|         |    794  | 
|         |    795 lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)" | 
|         |    796   by (fold eq_cp_s_th, unfold highest_cp_preced, simp) | 
|         |    797  | 
|         |    798 lemma highest': "cp s th = Max (cp s ` threads s)" | 
|         |    799 proof - | 
|         |    800   from highest_cp_preced max_cp_eq[symmetric] | 
|         |    801   show ?thesis by simp | 
|         |    802 qed | 
|         |    803  | 
|         |    804 end | 
|         |    805  | 
|         |    806 locale extend_highest_gen = highest_gen +  | 
|         |    807   fixes t  | 
|         |    808   assumes vt_t: "vt (t@s)" | 
|         |    809   and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio" | 
|         |    810   and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio" | 
|         |    811   and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th" | 
|         |    812  | 
|         |    813 sublocale extend_highest_gen < vat_t: valid_trace "t@s" | 
|         |    814   by (unfold_locales, insert vt_t, simp) | 
|         |    815  | 
|         |    816 lemma step_back_vt_app:  | 
|         |    817   assumes vt_ts: "vt (t@s)"  | 
|         |    818   shows "vt s" | 
|         |    819 proof - | 
|         |    820   from vt_ts show ?thesis | 
|         |    821   proof(induct t) | 
|         |    822     case Nil | 
|         |    823     from Nil show ?case by auto | 
|         |    824   next | 
|         |    825     case (Cons e t) | 
|         |    826     assume ih: " vt (t @ s) \<Longrightarrow> vt s" | 
|         |    827       and vt_et: "vt ((e # t) @ s)" | 
|         |    828     show ?case | 
|         |    829     proof(rule ih) | 
|         |    830       show "vt (t @ s)" | 
|         |    831       proof(rule step_back_vt) | 
|         |    832         from vt_et show "vt (e # t @ s)" by simp | 
|         |    833       qed | 
|         |    834     qed | 
|         |    835   qed | 
|         |    836 qed | 
|         |    837  | 
|         |    838  | 
|         |    839 locale red_extend_highest_gen = extend_highest_gen + | 
|         |    840    fixes i::nat | 
|         |    841  | 
|         |    842 sublocale red_extend_highest_gen <   red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)" | 
|         |    843   apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric]) | 
|         |    844   apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp) | 
|         |    845   by (unfold highest_gen_def, auto dest:step_back_vt_app) | 
|         |    846  | 
|         |    847  | 
|         |    848 context extend_highest_gen | 
|         |    849 begin | 
|         |    850  | 
|         |    851  lemma ind [consumes 0, case_names Nil Cons, induct type]: | 
|         |    852   assumes  | 
|         |    853     h0: "R []" | 
|         |    854   and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e;  | 
|         |    855                     extend_highest_gen s th prio tm t;  | 
|         |    856                     extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)" | 
|         |    857   shows "R t" | 
|         |    858 proof - | 
|         |    859   from vt_t extend_highest_gen_axioms show ?thesis | 
|         |    860   proof(induct t) | 
|         |    861     from h0 show "R []" . | 
|         |    862   next | 
|         |    863     case (Cons e t') | 
|         |    864     assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'" | 
|         |    865       and vt_e: "vt ((e # t') @ s)" | 
|         |    866       and et: "extend_highest_gen s th prio tm (e # t')" | 
|         |    867     from vt_e and step_back_step have stp: "step (t'@s) e" by auto | 
|         |    868     from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto | 
|         |    869     show ?case | 
|         |    870     proof(rule h2 [OF vt_ts stp _ _ _ ]) | 
|         |    871       show "R t'" | 
|         |    872       proof(rule ih) | 
|         |    873         from et show ext': "extend_highest_gen s th prio tm t'" | 
|         |    874           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt) | 
|         |    875       next | 
|         |    876         from vt_ts show "vt (t' @ s)" . | 
|         |    877       qed | 
|         |    878     next | 
|         |    879       from et show "extend_highest_gen s th prio tm (e # t')" . | 
|         |    880     next | 
|         |    881       from et show ext': "extend_highest_gen s th prio tm t'" | 
|         |    882           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt) | 
|         |    883     qed | 
|         |    884   qed | 
|         |    885 qed | 
|         |    886  | 
|         |    887  | 
|         |    888 lemma th_kept: "th \<in> threads (t @ s) \<and>  | 
|         |    889                  preced th (t@s) = preced th s" (is "?Q t")  | 
|         |    890 proof - | 
|         |    891   show ?thesis | 
|         |    892   proof(induct rule:ind) | 
|         |    893     case Nil | 
|         |    894     from threads_s | 
|         |    895     show ?case | 
|         |    896       by auto | 
|         |    897   next | 
|         |    898     case (Cons e t) | 
|         |    899     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto | 
|         |    900     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto | 
|         |    901     show ?case | 
|         |    902     proof(cases e) | 
|         |    903       case (Create thread prio) | 
|         |    904       show ?thesis | 
|         |    905       proof - | 
|         |    906         from Cons and Create have "step (t@s) (Create thread prio)" by auto | 
|         |    907         hence "th \<noteq> thread" | 
|         |    908         proof(cases) | 
|         |    909           case thread_create | 
|         |    910           with Cons show ?thesis by auto | 
|         |    911         qed | 
|         |    912         hence "preced th ((e # t) @ s)  = preced th (t @ s)" | 
|         |    913           by (unfold Create, auto simp:preced_def) | 
|         |    914         moreover note Cons | 
|         |    915         ultimately show ?thesis | 
|         |    916           by (auto simp:Create) | 
|         |    917       qed | 
|         |    918     next | 
|         |    919       case (Exit thread) | 
|         |    920       from h_e.exit_diff and Exit | 
|         |    921       have neq_th: "thread \<noteq> th" by auto | 
|         |    922       with Cons | 
|         |    923       show ?thesis | 
|         |    924         by (unfold Exit, auto simp:preced_def) | 
|         |    925     next | 
|         |    926       case (P thread cs) | 
|         |    927       with Cons | 
|         |    928       show ?thesis  | 
|         |    929         by (auto simp:P preced_def) | 
|         |    930     next | 
|         |    931       case (V thread cs) | 
|         |    932       with Cons | 
|         |    933       show ?thesis  | 
|         |    934         by (auto simp:V preced_def) | 
|         |    935     next | 
|         |    936       case (Set thread prio') | 
|         |    937       show ?thesis | 
|         |    938       proof - | 
|         |    939         from h_e.set_diff_low and Set | 
|         |    940         have "th \<noteq> thread" by auto | 
|         |    941         hence "preced th ((e # t) @ s)  = preced th (t @ s)" | 
|         |    942           by (unfold Set, auto simp:preced_def) | 
|         |    943         moreover note Cons | 
|         |    944         ultimately show ?thesis | 
|         |    945           by (auto simp:Set) | 
|         |    946       qed | 
|         |    947     qed | 
|         |    948   qed | 
|         |    949 qed | 
|         |    950  | 
|         |    951 text {* | 
|         |    952   According to @{thm th_kept}, thread @{text "th"} has its living status | 
|         |    953   and precedence kept along the way of @{text "t"}. The following lemma | 
|         |    954   shows that this preserved precedence of @{text "th"} remains as the highest | 
|         |    955   along the way of @{text "t"}. | 
|         |    956  | 
|         |    957   The proof goes by induction over @{text "t"} using the specialized | 
|         |    958   induction rule @{thm ind}, followed by case analysis of each possible  | 
|         |    959   operations of PIP. All cases follow the same pattern rendered by the  | 
|         |    960   generalized introduction rule @{thm "image_Max_eqI"}.  | 
|         |    961  | 
|         |    962   The very essence is to show that precedences, no matter whether they are newly introduced  | 
|         |    963   or modified, are always lower than the one held by @{term "th"}, | 
|         |    964   which by @{thm th_kept} is preserved along the way. | 
|         |    965 *} | 
|         |    966 lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s" | 
|         |    967 proof(induct rule:ind) | 
|         |    968   case Nil | 
|         |    969   from highest_preced_thread | 
|         |    970   show ?case | 
|         |    971     by (unfold the_preced_def, simp) | 
|         |    972 next | 
|         |    973   case (Cons e t) | 
|         |    974     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto | 
|         |    975     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto | 
|         |    976   show ?case | 
|         |    977   proof(cases e) | 
|         |    978     case (Create thread prio') | 
|         |    979     show ?thesis (is "Max (?f ` ?A) = ?t") | 
|         |    980     proof - | 
|         |    981       -- {* The following is the common pattern of each branch of the case analysis. *} | 
|         |    982       -- {* The major part is to show that @{text "th"} holds the highest precedence: *} | 
|         |    983       have "Max (?f ` ?A) = ?f th" | 
|         |    984       proof(rule image_Max_eqI) | 
|         |    985         show "finite ?A" using h_e.finite_threads by auto  | 
|         |    986       next | 
|         |    987         show "th \<in> ?A" using h_e.th_kept by auto  | 
|         |    988       next | 
|         |    989         show "\<forall>x\<in>?A. ?f x \<le> ?f th" | 
|         |    990         proof  | 
|         |    991           fix x | 
|         |    992           assume "x \<in> ?A" | 
|         |    993           hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create) | 
|         |    994           thus "?f x \<le> ?f th" | 
|         |    995           proof | 
|         |    996             assume "x = thread" | 
|         |    997             thus ?thesis  | 
|         |    998               apply (simp add:Create the_preced_def preced_def, fold preced_def) | 
|         |    999               using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 preced_th by force | 
|         |   1000           next | 
|         |   1001             assume h: "x \<in> threads (t @ s)" | 
|         |   1002             from Cons(2)[unfolded Create]  | 
|         |   1003             have "x \<noteq> thread" using h by (cases, auto) | 
|         |   1004             hence "?f x = the_preced (t@s) x"  | 
|         |   1005               by (simp add:Create the_preced_def preced_def) | 
|         |   1006             hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))" | 
|         |   1007               by (simp add: h_t.finite_threads h) | 
|         |   1008             also have "... = ?f th" | 
|         |   1009               by (metis Cons.hyps(5) h_e.th_kept the_preced_def)  | 
|         |   1010             finally show ?thesis . | 
|         |   1011           qed | 
|         |   1012         qed | 
|         |   1013       qed | 
|         |   1014      -- {* The minor part is to show that the precedence of @{text "th"}  | 
|         |   1015            equals to preserved one, given by the foregoing lemma @{thm th_kept} *} | 
|         |   1016       also have "... = ?t" using h_e.th_kept the_preced_def by auto | 
|         |   1017       -- {* Then it follows trivially that the precedence preserved | 
|         |   1018             for @{term "th"} remains the maximum of all living threads along the way. *} | 
|         |   1019       finally show ?thesis . | 
|         |   1020     qed  | 
|         |   1021   next  | 
|         |   1022     case (Exit thread) | 
|         |   1023     show ?thesis (is "Max (?f ` ?A) = ?t") | 
|         |   1024     proof - | 
|         |   1025       have "Max (?f ` ?A) = ?f th" | 
|         |   1026       proof(rule image_Max_eqI) | 
|         |   1027         show "finite ?A" using h_e.finite_threads by auto  | 
|         |   1028       next | 
|         |   1029         show "th \<in> ?A" using h_e.th_kept by auto  | 
|         |   1030       next | 
|         |   1031         show "\<forall>x\<in>?A. ?f x \<le> ?f th" | 
|         |   1032         proof  | 
|         |   1033           fix x | 
|         |   1034           assume "x \<in> ?A" | 
|         |   1035           hence "x \<in> threads (t@s)" by (simp add: Exit)  | 
|         |   1036           hence "?f x \<le> Max (?f ` threads (t@s))"  | 
|         |   1037             by (simp add: h_t.finite_threads)  | 
|         |   1038           also have "... \<le> ?f th"  | 
|         |   1039             apply (simp add:Exit the_preced_def preced_def, fold preced_def) | 
|         |   1040             using Cons.hyps(5) h_t.th_kept the_preced_def by auto | 
|         |   1041           finally show "?f x \<le> ?f th" . | 
|         |   1042         qed | 
|         |   1043       qed | 
|         |   1044       also have "... = ?t" using h_e.th_kept the_preced_def by auto | 
|         |   1045       finally show ?thesis . | 
|         |   1046     qed  | 
|         |   1047   next | 
|         |   1048     case (P thread cs) | 
|         |   1049     with Cons | 
|         |   1050     show ?thesis by (auto simp:preced_def the_preced_def) | 
|         |   1051   next | 
|         |   1052     case (V thread cs) | 
|         |   1053     with Cons | 
|         |   1054     show ?thesis by (auto simp:preced_def the_preced_def) | 
|         |   1055   next  | 
|         |   1056     case (Set thread prio') | 
|         |   1057     show ?thesis (is "Max (?f ` ?A) = ?t") | 
|         |   1058     proof - | 
|         |   1059       have "Max (?f ` ?A) = ?f th" | 
|         |   1060       proof(rule image_Max_eqI) | 
|         |   1061         show "finite ?A" using h_e.finite_threads by auto  | 
|         |   1062       next | 
|         |   1063         show "th \<in> ?A" using h_e.th_kept by auto  | 
|         |   1064       next | 
|         |   1065         show "\<forall>x\<in>?A. ?f x \<le> ?f th" | 
|         |   1066         proof  | 
|         |   1067           fix x | 
|         |   1068           assume h: "x \<in> ?A" | 
|         |   1069           show "?f x \<le> ?f th" | 
|         |   1070           proof(cases "x = thread") | 
|         |   1071             case True | 
|         |   1072             moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th" | 
|         |   1073             proof - | 
|         |   1074               have "the_preced (t @ s) th = Prc prio tm"   | 
|         |   1075                 using h_t.th_kept preced_th by (simp add:the_preced_def) | 
|         |   1076               moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto | 
|         |   1077               ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def) | 
|         |   1078             qed | 
|         |   1079             ultimately show ?thesis | 
|         |   1080               by (unfold Set, simp add:the_preced_def preced_def) | 
|         |   1081           next | 
|         |   1082             case False | 
|         |   1083             then have "?f x  = the_preced (t@s) x" | 
|         |   1084               by (simp add:the_preced_def preced_def Set) | 
|         |   1085             also have "... \<le> Max (the_preced (t@s) ` threads (t@s))" | 
|         |   1086               using Set h h_t.finite_threads by auto  | 
|         |   1087             also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def)  | 
|         |   1088             finally show ?thesis . | 
|         |   1089           qed | 
|         |   1090         qed | 
|         |   1091       qed | 
|         |   1092       also have "... = ?t" using h_e.th_kept the_preced_def by auto | 
|         |   1093       finally show ?thesis . | 
|         |   1094     qed  | 
|         |   1095   qed | 
|         |   1096 qed | 
|         |   1097  | 
|         |   1098 lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))" | 
|         |   1099   by (insert th_kept max_kept, auto) | 
|         |   1100  | 
|         |   1101 text {* | 
|         |   1102   The reason behind the following lemma is that: | 
|         |   1103   Since @{term "cp"} is defined as the maximum precedence  | 
|         |   1104   of those threads contained in the sub-tree of node @{term "Th th"}  | 
|         |   1105   in @{term "RAG (t@s)"}, and all these threads are living threads, and  | 
|         |   1106   @{term "th"} is also among them, the maximum precedence of  | 
|         |   1107   them all must be the one for @{text "th"}. | 
|         |   1108 *} | 
|         |   1109 lemma th_cp_max_preced:  | 
|         |   1110   "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R")  | 
|         |   1111 proof - | 
|         |   1112   let ?f = "the_preced (t@s)" | 
|         |   1113   have "?L = ?f th" | 
|         |   1114   proof(unfold cp_alt_def, rule image_Max_eqI) | 
|         |   1115     show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}" | 
|         |   1116     proof - | 
|         |   1117       have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} =  | 
|         |   1118             the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and> | 
|         |   1119                             (\<exists> th'. n = Th th')}" | 
|         |   1120       by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps) | 
|         |   1121       moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree)  | 
|         |   1122       ultimately show ?thesis by simp | 
|         |   1123     qed | 
|         |   1124   next | 
|         |   1125     show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}" | 
|         |   1126       by (auto simp:subtree_def) | 
|         |   1127   next | 
|         |   1128     show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}. | 
|         |   1129                the_preced (t @ s) x \<le> the_preced (t @ s) th" | 
|         |   1130     proof | 
|         |   1131       fix th' | 
|         |   1132       assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}" | 
|         |   1133       hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto | 
|         |   1134       moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}" | 
|         |   1135         by (meson subtree_Field) | 
|         |   1136       ultimately have "Th th' \<in> ..." by auto | 
|         |   1137       hence "th' \<in> threads (t@s)"  | 
|         |   1138       proof | 
|         |   1139         assume "Th th' \<in> {Th th}" | 
|         |   1140         thus ?thesis using th_kept by auto  | 
|         |   1141       next | 
|         |   1142         assume "Th th' \<in> Field (RAG (t @ s))" | 
|         |   1143         thus ?thesis using vat_t.not_in_thread_isolated by blast  | 
|         |   1144       qed | 
|         |   1145       thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th" | 
|         |   1146         by (metis Max_ge finite_imageI finite_threads image_eqI  | 
|         |   1147                max_kept th_kept the_preced_def) | 
|         |   1148     qed | 
|         |   1149   qed | 
|         |   1150   also have "... = ?R" by (simp add: max_preced the_preced_def)  | 
|         |   1151   finally show ?thesis . | 
|         |   1152 qed | 
|         |   1153  | 
|         |   1154 lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))" | 
|         |   1155   using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger | 
|         |   1156  | 
|         |   1157 lemma th_cp_preced: "cp (t@s) th = preced th s" | 
|         |   1158   by (fold max_kept, unfold th_cp_max_preced, simp) | 
|         |   1159  | 
|         |   1160 lemma preced_less: | 
|         |   1161   assumes th'_in: "th' \<in> threads s" | 
|         |   1162   and neq_th': "th' \<noteq> th" | 
|         |   1163   shows "preced th' s < preced th s" | 
|         |   1164   using assms | 
|         |   1165 by (metis Max.coboundedI finite_imageI highest not_le order.trans  | 
|         |   1166     preced_linorder rev_image_eqI threads_s vat_s.finite_threads  | 
|         |   1167     vat_s.le_cp) | 
|         |   1168  | 
|         |   1169 text {* | 
|         |   1170   Counting of the number of @{term "P"} and @{term "V"} operations  | 
|         |   1171   is the cornerstone of a large number of the following proofs.  | 
|         |   1172   The reason is that this counting is quite easy to calculate and  | 
|         |   1173   convenient to use in the reasoning.  | 
|         |   1174  | 
|         |   1175   The following lemma shows that the counting controls whether  | 
|         |   1176   a thread is running or not. | 
|         |   1177 *} | 
|         |   1178  | 
|         |   1179 lemma pv_blocked_pre: | 
|         |   1180   assumes th'_in: "th' \<in> threads (t@s)" | 
|         |   1181   and neq_th': "th' \<noteq> th" | 
|         |   1182   and eq_pv: "cntP (t@s) th' = cntV (t@s) th'" | 
|         |   1183   shows "th' \<notin> runing (t@s)" | 
|         |   1184 proof | 
|         |   1185   assume otherwise: "th' \<in> runing (t@s)" | 
|         |   1186   show False | 
|         |   1187   proof - | 
|         |   1188     have "th' = th" | 
|         |   1189     proof(rule preced_unique) | 
|         |   1190       show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R") | 
|         |   1191       proof - | 
|         |   1192         have "?L = cp (t@s) th'" | 
|         |   1193           by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp) | 
|         |   1194         also have "... = cp (t @ s) th" using otherwise  | 
|         |   1195           by (metis (mono_tags, lifting) mem_Collect_eq  | 
|         |   1196                     runing_def th_cp_max vat_t.max_cp_readys_threads) | 
|         |   1197         also have "... = ?R" by (metis th_cp_preced th_kept)  | 
|         |   1198         finally show ?thesis . | 
|         |   1199       qed | 
|         |   1200     qed (auto simp: th'_in th_kept) | 
|         |   1201     moreover have "th' \<noteq> th" using neq_th' . | 
|         |   1202     ultimately show ?thesis by simp | 
|         |   1203  qed | 
|         |   1204 qed | 
|         |   1205  | 
|         |   1206 lemmas pv_blocked = pv_blocked_pre[folded detached_eq] | 
|         |   1207  | 
|         |   1208 lemma runing_precond_pre: | 
|         |   1209   fixes th' | 
|         |   1210   assumes th'_in: "th' \<in> threads s" | 
|         |   1211   and eq_pv: "cntP s th' = cntV s th'" | 
|         |   1212   and neq_th': "th' \<noteq> th" | 
|         |   1213   shows "th' \<in> threads (t@s) \<and> | 
|         |   1214          cntP (t@s) th' = cntV (t@s) th'" | 
|         |   1215 proof(induct rule:ind) | 
|         |   1216   case (Cons e t) | 
|         |   1217     interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp | 
|         |   1218     interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp | 
|         |   1219     show ?case | 
|         |   1220     proof(cases e) | 
|         |   1221       case (P thread cs) | 
|         |   1222       show ?thesis | 
|         |   1223       proof - | 
|         |   1224         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" | 
|         |   1225         proof - | 
|         |   1226           have "thread \<noteq> th'" | 
|         |   1227           proof - | 
|         |   1228             have "step (t@s) (P thread cs)" using Cons P by auto | 
|         |   1229             thus ?thesis | 
|         |   1230             proof(cases) | 
|         |   1231               assume "thread \<in> runing (t@s)" | 
|         |   1232               moreover have "th' \<notin> runing (t@s)" using Cons(5) | 
|         |   1233                 by (metis neq_th' vat_t.pv_blocked_pre)  | 
|         |   1234               ultimately show ?thesis by auto | 
|         |   1235             qed | 
|         |   1236           qed with Cons show ?thesis | 
|         |   1237             by (unfold P, simp add:cntP_def cntV_def count_def) | 
|         |   1238         qed | 
|         |   1239         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold P, simp) | 
|         |   1240         ultimately show ?thesis by auto | 
|         |   1241       qed | 
|         |   1242     next | 
|         |   1243       case (V thread cs) | 
|         |   1244       show ?thesis | 
|         |   1245       proof - | 
|         |   1246         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" | 
|         |   1247         proof - | 
|         |   1248           have "thread \<noteq> th'" | 
|         |   1249           proof - | 
|         |   1250             have "step (t@s) (V thread cs)" using Cons V by auto | 
|         |   1251             thus ?thesis | 
|         |   1252             proof(cases) | 
|         |   1253               assume "thread \<in> runing (t@s)" | 
|         |   1254               moreover have "th' \<notin> runing (t@s)" using Cons(5) | 
|         |   1255                 by (metis neq_th' vat_t.pv_blocked_pre)  | 
|         |   1256               ultimately show ?thesis by auto | 
|         |   1257             qed | 
|         |   1258           qed with Cons show ?thesis | 
|         |   1259             by (unfold V, simp add:cntP_def cntV_def count_def) | 
|         |   1260         qed | 
|         |   1261         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold V, simp) | 
|         |   1262         ultimately show ?thesis by auto | 
|         |   1263       qed | 
|         |   1264     next | 
|         |   1265       case (Create thread prio') | 
|         |   1266       show ?thesis | 
|         |   1267       proof - | 
|         |   1268         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" | 
|         |   1269         proof - | 
|         |   1270           have "thread \<noteq> th'" | 
|         |   1271           proof - | 
|         |   1272             have "step (t@s) (Create thread prio')" using Cons Create by auto | 
|         |   1273             thus ?thesis using Cons(5) by (cases, auto) | 
|         |   1274           qed with Cons show ?thesis | 
|         |   1275             by (unfold Create, simp add:cntP_def cntV_def count_def) | 
|         |   1276         qed | 
|         |   1277         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold Create, simp) | 
|         |   1278         ultimately show ?thesis by auto | 
|         |   1279       qed | 
|         |   1280     next | 
|         |   1281       case (Exit thread) | 
|         |   1282       show ?thesis | 
|         |   1283       proof - | 
|         |   1284         have neq_thread: "thread \<noteq> th'" | 
|         |   1285         proof - | 
|         |   1286           have "step (t@s) (Exit thread)" using Cons Exit by auto | 
|         |   1287           thus ?thesis apply (cases) using Cons(5) | 
|         |   1288                 by (metis neq_th' vat_t.pv_blocked_pre)  | 
|         |   1289         qed  | 
|         |   1290         hence "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" using Cons | 
|         |   1291             by (unfold Exit, simp add:cntP_def cntV_def count_def) | 
|         |   1292         moreover have "th' \<in> threads ((e # t) @ s)" using Cons neq_thread  | 
|         |   1293           by (unfold Exit, simp)  | 
|         |   1294         ultimately show ?thesis by auto | 
|         |   1295       qed | 
|         |   1296     next | 
|         |   1297       case (Set thread prio') | 
|         |   1298       with Cons | 
|         |   1299       show ?thesis  | 
|         |   1300         by (auto simp:cntP_def cntV_def count_def) | 
|         |   1301     qed | 
|         |   1302 next | 
|         |   1303   case Nil | 
|         |   1304   with assms | 
|         |   1305   show ?case by auto | 
|         |   1306 qed | 
|         |   1307  | 
|         |   1308 text {* Changing counting balance to detachedness *} | 
|         |   1309 lemmas runing_precond_pre_dtc = runing_precond_pre | 
|         |   1310          [folded vat_t.detached_eq vat_s.detached_eq] | 
|         |   1311  | 
|         |   1312 lemma runing_precond: | 
|         |   1313   fixes th' | 
|         |   1314   assumes th'_in: "th' \<in> threads s" | 
|         |   1315   and neq_th': "th' \<noteq> th" | 
|         |   1316   and is_runing: "th' \<in> runing (t@s)" | 
|         |   1317   shows "cntP s th' > cntV s th'" | 
|         |   1318   using assms | 
|         |   1319 proof - | 
|         |   1320   have "cntP s th' \<noteq> cntV s th'" | 
|         |   1321     by (metis is_runing neq_th' pv_blocked_pre runing_precond_pre th'_in) | 
|         |   1322   moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto | 
|         |   1323   ultimately show ?thesis by auto | 
|         |   1324 qed | 
|         |   1325  | 
|         |   1326 lemma moment_blocked_pre: | 
|         |   1327   assumes neq_th': "th' \<noteq> th" | 
|         |   1328   and th'_in: "th' \<in> threads ((moment i t)@s)" | 
|         |   1329   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'" | 
|         |   1330   shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and> | 
|         |   1331          th' \<in> threads ((moment (i+j) t)@s)" | 
|         |   1332 proof - | 
|         |   1333   interpret h_i: red_extend_highest_gen _ _ _ _ _ i | 
|         |   1334       by (unfold_locales) | 
|         |   1335   interpret h_j: red_extend_highest_gen _ _ _ _ _ "i+j" | 
|         |   1336       by (unfold_locales) | 
|         |   1337   interpret h:  extend_highest_gen "((moment i t)@s)" th prio tm "moment j (restm i t)" | 
|         |   1338   proof(unfold_locales) | 
|         |   1339     show "vt (moment i t @ s)" by (metis h_i.vt_t)  | 
|         |   1340   next | 
|         |   1341     show "th \<in> threads (moment i t @ s)" by (metis h_i.th_kept) | 
|         |   1342   next | 
|         |   1343     show "preced th (moment i t @ s) =  | 
|         |   1344             Max (cp (moment i t @ s) ` threads (moment i t @ s))" | 
|         |   1345               by (metis h_i.th_cp_max h_i.th_cp_preced h_i.th_kept) | 
|         |   1346   next | 
|         |   1347     show "preced th (moment i t @ s) = Prc prio tm" by (metis h_i.th_kept preced_th)  | 
|         |   1348   next | 
|         |   1349     show "vt (moment j (restm i t) @ moment i t @ s)" | 
|         |   1350       using moment_plus_split by (metis add.commute append_assoc h_j.vt_t) | 
|         |   1351   next | 
|         |   1352     fix th' prio' | 
|         |   1353     assume "Create th' prio' \<in> set (moment j (restm i t))" | 
|         |   1354     thus "prio' \<le> prio" using assms | 
|         |   1355        by (metis Un_iff add.commute h_j.create_low moment_plus_split set_append) | 
|         |   1356   next | 
|         |   1357     fix th' prio' | 
|         |   1358     assume "Set th' prio' \<in> set (moment j (restm i t))" | 
|         |   1359     thus "th' \<noteq> th \<and> prio' \<le> prio" | 
|         |   1360     by (metis Un_iff add.commute h_j.set_diff_low moment_plus_split set_append) | 
|         |   1361   next | 
|         |   1362     fix th' | 
|         |   1363     assume "Exit th' \<in> set (moment j (restm i t))" | 
|         |   1364     thus "th' \<noteq> th" | 
|         |   1365       by (metis Un_iff add.commute h_j.exit_diff moment_plus_split set_append) | 
|         |   1366   qed | 
|         |   1367   show ?thesis  | 
|         |   1368     by (metis add.commute append_assoc eq_pv h.runing_precond_pre | 
|         |   1369           moment_plus_split neq_th' th'_in) | 
|         |   1370 qed | 
|         |   1371  | 
|         |   1372 lemma moment_blocked_eqpv: | 
|         |   1373   assumes neq_th': "th' \<noteq> th" | 
|         |   1374   and th'_in: "th' \<in> threads ((moment i t)@s)" | 
|         |   1375   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'" | 
|         |   1376   and le_ij: "i \<le> j" | 
|         |   1377   shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and> | 
|         |   1378          th' \<in> threads ((moment j t)@s) \<and> | 
|         |   1379          th' \<notin> runing ((moment j t)@s)" | 
|         |   1380 proof - | 
|         |   1381   from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij | 
|         |   1382   have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'" | 
|         |   1383    and h2: "th' \<in> threads ((moment j t)@s)" by auto | 
|         |   1384   moreover have "th' \<notin> runing ((moment j t)@s)" | 
|         |   1385   proof - | 
|         |   1386     interpret h: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales) | 
|         |   1387     show ?thesis | 
|         |   1388       using h.pv_blocked_pre h1 h2 neq_th' by auto  | 
|         |   1389   qed | 
|         |   1390   ultimately show ?thesis by auto | 
|         |   1391 qed | 
|         |   1392  | 
|         |   1393 (* The foregoing two lemmas are preparation for this one, but | 
|         |   1394    in long run can be combined. Maybe I am wrong. | 
|         |   1395 *) | 
|         |   1396 lemma moment_blocked: | 
|         |   1397   assumes neq_th': "th' \<noteq> th" | 
|         |   1398   and th'_in: "th' \<in> threads ((moment i t)@s)" | 
|         |   1399   and dtc: "detached (moment i t @ s) th'" | 
|         |   1400   and le_ij: "i \<le> j" | 
|         |   1401   shows "detached (moment j t @ s) th' \<and> | 
|         |   1402          th' \<in> threads ((moment j t)@s) \<and> | 
|         |   1403          th' \<notin> runing ((moment j t)@s)" | 
|         |   1404 proof - | 
|         |   1405   interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales) | 
|         |   1406   interpret h_j: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)  | 
|         |   1407   have cnt_i: "cntP (moment i t @ s) th' = cntV (moment i t @ s) th'" | 
|         |   1408                 by (metis dtc h_i.detached_elim) | 
|         |   1409   from moment_blocked_eqpv[OF neq_th' th'_in cnt_i le_ij] | 
|         |   1410   show ?thesis by (metis h_j.detached_intro)  | 
|         |   1411 qed | 
|         |   1412  | 
|         |   1413 lemma runing_preced_inversion: | 
|         |   1414   assumes runing': "th' \<in> runing (t@s)" | 
|         |   1415   shows "cp (t@s) th' = preced th s" (is "?L = ?R") | 
|         |   1416 proof - | 
|         |   1417   have "?L = Max (cp (t @ s) ` readys (t @ s))" using assms | 
|         |   1418       by (unfold runing_def, auto) | 
|         |   1419   also have "\<dots> = ?R" | 
|         |   1420       by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads)  | 
|         |   1421   finally show ?thesis . | 
|         |   1422 qed | 
|         |   1423  | 
|         |   1424 text {* | 
|         |   1425   The situation when @{term "th"} is blocked is analyzed by the following lemmas. | 
|         |   1426 *} | 
|         |   1427  | 
|         |   1428 text {* | 
|         |   1429   The following lemmas shows the running thread @{text "th'"}, if it is different from | 
|         |   1430   @{term th}, must be live at the very beginning. By the term {\em the very beginning}, | 
|         |   1431   we mean the moment where the formal investigation starts, i.e. the moment (or state) | 
|         |   1432   @{term s}.  | 
|         |   1433 *} | 
|         |   1434  | 
|         |   1435 lemma runing_inversion_0: | 
|         |   1436   assumes neq_th': "th' \<noteq> th" | 
|         |   1437   and runing': "th' \<in> runing (t@s)" | 
|         |   1438   shows "th' \<in> threads s" | 
|         |   1439 proof - | 
|         |   1440     -- {* The proof is by contradiction: *} | 
|         |   1441     { assume otherwise: "\<not> ?thesis" | 
|         |   1442       have "th' \<notin> runing (t @ s)" | 
|         |   1443       proof - | 
|         |   1444         -- {* Since @{term "th'"} is running at time @{term "t@s"}, so it exists that time. *} | 
|         |   1445         have th'_in: "th' \<in> threads (t@s)" using runing' by (simp add:runing_def readys_def) | 
|         |   1446         -- {* However, @{text "th'"} does not exist at very beginning. *} | 
|         |   1447         have th'_notin: "th' \<notin> threads (moment 0 t @ s)" using otherwise | 
|         |   1448           by (metis append.simps(1) moment_zero) | 
|         |   1449         -- {* Therefore, there must be a moment during @{text "t"}, when  | 
|         |   1450               @{text "th'"} came into being. *} | 
|         |   1451         -- {* Let us suppose the moment being @{text "i"}: *} | 
|         |   1452         from p_split_gen[OF th'_in th'_notin] | 
|         |   1453         obtain i where lt_its: "i < length t" | 
|         |   1454                  and le_i: "0 \<le> i" | 
|         |   1455                  and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre") | 
|         |   1456                  and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by (auto) | 
|         |   1457         interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales) | 
|         |   1458         interpret h_i': red_extend_highest_gen _ _ _ _ _ "(Suc i)" by (unfold_locales) | 
|         |   1459         from lt_its have "Suc i \<le> length t" by auto | 
|         |   1460         -- {* Let us also suppose the event which makes this change is @{text e}: *} | 
|         |   1461         from moment_head[OF this] obtain e where  | 
|         |   1462           eq_me: "moment (Suc i) t = e # moment i t" by blast | 
|         |   1463         hence "vt (e # (moment i t @ s))" by (metis append_Cons h_i'.vt_t)  | 
|         |   1464         hence "PIP (moment i t @ s) e" by (cases, simp) | 
|         |   1465         -- {* It can be derived that this event @{text "e"}, which  | 
|         |   1466               gives birth to @{term "th'"} must be a @{term "Create"}: *} | 
|         |   1467         from create_pre[OF this, of th'] | 
|         |   1468         obtain prio where eq_e: "e = Create th' prio" | 
|         |   1469             by (metis append_Cons eq_me lessI post pre)  | 
|         |   1470         have h1: "th' \<in> threads (moment (Suc i) t @ s)" using post by auto  | 
|         |   1471         have h2: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'" | 
|         |   1472         proof - | 
|         |   1473           have "cntP (moment i t@s) th' = cntV (moment i t@s) th'" | 
|         |   1474             by (metis h_i.cnp_cnv_eq pre) | 
|         |   1475           thus ?thesis by (simp add:eq_me eq_e cntP_def cntV_def count_def) | 
|         |   1476         qed | 
|         |   1477         show ?thesis  | 
|         |   1478           using moment_blocked_eqpv [OF neq_th' h1 h2, of "length t"] lt_its moment_ge | 
|         |   1479             by auto | 
|         |   1480       qed | 
|         |   1481       with `th' \<in> runing (t@s)` | 
|         |   1482       have False by simp | 
|         |   1483     } thus ?thesis by auto | 
|         |   1484 qed | 
|         |   1485  | 
|         |   1486 text {*  | 
|         |   1487   The second lemma says, if the running thread @{text th'} is different from  | 
|         |   1488   @{term th}, then this @{text th'} must in the possession of some resources | 
|         |   1489   at the very beginning.  | 
|         |   1490  | 
|         |   1491   To ease the reasoning of resource possession of one particular thread,  | 
|         |   1492   we used two auxiliary functions @{term cntV} and @{term cntP},  | 
|         |   1493   which are the counters of @{term P}-operations and  | 
|         |   1494   @{term V}-operations respectively.  | 
|         |   1495   If the number of @{term V}-operation is less than the number of  | 
|         |   1496   @{term "P"}-operations, the thread must have some unreleased resource.  | 
|         |   1497 *} | 
|         |   1498  | 
|         |   1499 lemma runing_inversion_1: (* ddd *) | 
|         |   1500   assumes neq_th': "th' \<noteq> th" | 
|         |   1501   and runing': "th' \<in> runing (t@s)" | 
|         |   1502   -- {* thread @{term "th'"} is a live on in state @{term "s"} and  | 
|         |   1503         it has some unreleased resource. *} | 
|         |   1504   shows "th' \<in> threads s \<and> cntV s th' < cntP s th'" | 
|         |   1505 proof - | 
|         |   1506   -- {* The proof is a simple composition of @{thm runing_inversion_0} and  | 
|         |   1507         @{thm runing_precond}: *} | 
|         |   1508   -- {* By applying @{thm runing_inversion_0} to assumptions, | 
|         |   1509         it can be shown that @{term th'} is live in state @{term s}: *} | 
|         |   1510   have "th' \<in> threads s"  using runing_inversion_0[OF assms(1,2)] . | 
|         |   1511   -- {* Then the thesis is derived easily by applying @{thm runing_precond}: *} | 
|         |   1512   with runing_precond [OF this neq_th' runing'] show ?thesis by simp | 
|         |   1513 qed | 
|         |   1514  | 
|         |   1515 text {*  | 
|         |   1516   The following lemma is just a rephrasing of @{thm runing_inversion_1}: | 
|         |   1517 *} | 
|         |   1518 lemma runing_inversion_2: | 
|         |   1519   assumes runing': "th' \<in> runing (t@s)" | 
|         |   1520   shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')" | 
|         |   1521 proof - | 
|         |   1522   from runing_inversion_1[OF _ runing'] | 
|         |   1523   show ?thesis by auto | 
|         |   1524 qed | 
|         |   1525  | 
|         |   1526 lemma runing_inversion_3: | 
|         |   1527   assumes runing': "th' \<in> runing (t@s)" | 
|         |   1528   and neq_th: "th' \<noteq> th" | 
|         |   1529   shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)" | 
|         |   1530   by (metis neq_th runing' runing_inversion_2 runing_preced_inversion) | 
|         |   1531  | 
|         |   1532 lemma runing_inversion_4: | 
|         |   1533   assumes runing': "th' \<in> runing (t@s)" | 
|         |   1534   and neq_th: "th' \<noteq> th" | 
|         |   1535   shows "th' \<in> threads s" | 
|         |   1536   and    "\<not>detached s th'" | 
|         |   1537   and    "cp (t@s) th' = preced th s" | 
|         |   1538   apply (metis neq_th runing' runing_inversion_2) | 
|         |   1539   apply (metis neq_th pv_blocked runing' runing_inversion_2 runing_precond_pre_dtc) | 
|         |   1540   by (metis neq_th runing' runing_inversion_3) | 
|         |   1541  | 
|         |   1542  | 
|         |   1543 text {*  | 
|         |   1544   Suppose @{term th} is not running, it is first shown that | 
|         |   1545   there is a path in RAG leading from node @{term th} to another thread @{text "th'"}  | 
|         |   1546   in the @{term readys}-set (So @{text "th'"} is an ancestor of @{term th}}). | 
|         |   1547  | 
|         |   1548   Now, since @{term readys}-set is non-empty, there must be | 
|         |   1549   one in it which holds the highest @{term cp}-value, which, by definition,  | 
|         |   1550   is the @{term runing}-thread. However, we are going to show more: this running thread | 
|         |   1551   is exactly @{term "th'"}. | 
|         |   1552      *} | 
|         |   1553 lemma th_blockedE: (* ddd *) | 
|         |   1554   assumes "th \<notin> runing (t@s)" | 
|         |   1555   obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)" | 
|         |   1556                     "th' \<in> runing (t@s)" | 
|         |   1557 proof - | 
|         |   1558   -- {* According to @{thm vat_t.th_chain_to_ready}, either  | 
|         |   1559         @{term "th"} is in @{term "readys"} or there is path leading from it to  | 
|         |   1560         one thread in @{term "readys"}. *} | 
|         |   1561   have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)"  | 
|         |   1562     using th_kept vat_t.th_chain_to_ready by auto | 
|         |   1563   -- {* However, @{term th} can not be in @{term readys}, because otherwise, since  | 
|         |   1564        @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *} | 
|         |   1565   moreover have "th \<notin> readys (t@s)"  | 
|         |   1566     using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto  | 
|         |   1567   -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in  | 
|         |   1568         term @{term readys}: *} | 
|         |   1569   ultimately obtain th' where th'_in: "th' \<in> readys (t@s)" | 
|         |   1570                           and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto | 
|         |   1571   -- {* We are going to show that this @{term th'} is running. *} | 
|         |   1572   have "th' \<in> runing (t@s)" | 
|         |   1573   proof - | 
|         |   1574     -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *} | 
|         |   1575     have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R") | 
|         |   1576     proof - | 
|         |   1577       have "?L =  Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))" | 
|         |   1578         by (unfold cp_alt_def1, simp) | 
|         |   1579       also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)" | 
|         |   1580       proof(rule image_Max_subset) | 
|         |   1581         show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads) | 
|         |   1582       next | 
|         |   1583         show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)" | 
|         |   1584           by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread)  | 
|         |   1585       next | 
|         |   1586         show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp | 
|         |   1587                     by (unfold tRAG_subtree_eq, auto simp:subtree_def) | 
|         |   1588       next | 
|         |   1589         show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) = | 
|         |   1590                       (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _") | 
|         |   1591         proof - | 
|         |   1592           have "?L = the_preced (t @ s) `  threads (t @ s)"  | 
|         |   1593                      by (unfold image_comp, rule image_cong, auto) | 
|         |   1594           thus ?thesis using max_preced the_preced_def by auto | 
|         |   1595         qed | 
|         |   1596       qed | 
|         |   1597       also have "... = ?R" | 
|         |   1598         using th_cp_max th_cp_preced th_kept  | 
|         |   1599               the_preced_def vat_t.max_cp_readys_threads by auto | 
|         |   1600       finally show ?thesis . | 
|         |   1601     qed  | 
|         |   1602     -- {* Now, since @{term th'} holds the highest @{term cp}  | 
|         |   1603           and we have already show it is in @{term readys}, | 
|         |   1604           it is @{term runing} by definition. *} | 
|         |   1605     with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def)  | 
|         |   1606   qed | 
|         |   1607   -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *} | 
|         |   1608   moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"  | 
|         |   1609     using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def) | 
|         |   1610   ultimately show ?thesis using that by metis | 
|         |   1611 qed | 
|         |   1612  | 
|         |   1613 text {* | 
|         |   1614   Now it is easy to see there is always a thread to run by case analysis | 
|         |   1615   on whether thread @{term th} is running: if the answer is Yes, the  | 
|         |   1616   the running thread is obviously @{term th} itself; otherwise, the running | 
|         |   1617   thread is the @{text th'} given by lemma @{thm th_blockedE}. | 
|         |   1618 *} | 
|         |   1619 lemma live: "runing (t@s) \<noteq> {}" | 
|         |   1620 proof(cases "th \<in> runing (t@s)")  | 
|         |   1621   case True thus ?thesis by auto | 
|         |   1622 next | 
|         |   1623   case False | 
|         |   1624   thus ?thesis using th_blockedE by auto | 
|         |   1625 qed | 
|         |   1626  | 
|         |   1627 end | 
|         |   1628 end | 
|         |   1629  |