| 86 |      1 | theory Correctness
 | 
|  |      2 | imports PIPBasics
 | 
|  |      3 | begin
 | 
|  |      4 | 
 | 
|  |      5 | 
 | 
|  |      6 | text {* 
 | 
|  |      7 |   The following two auxiliary lemmas are used to reason about @{term Max}.
 | 
|  |      8 | *}
 | 
|  |      9 | lemma image_Max_eqI: 
 | 
|  |     10 |   assumes "finite B"
 | 
|  |     11 |   and "b \<in> B"
 | 
|  |     12 |   and "\<forall> x \<in> B. f x \<le> f b"
 | 
|  |     13 |   shows "Max (f ` B) = f b"
 | 
|  |     14 |   using assms
 | 
|  |     15 |   using Max_eqI by blast 
 | 
|  |     16 | 
 | 
|  |     17 | lemma image_Max_subset:
 | 
|  |     18 |   assumes "finite A"
 | 
|  |     19 |   and "B \<subseteq> A"
 | 
|  |     20 |   and "a \<in> B"
 | 
|  |     21 |   and "Max (f ` A) = f a"
 | 
|  |     22 |   shows "Max (f ` B) = f a"
 | 
|  |     23 | proof(rule image_Max_eqI)
 | 
|  |     24 |   show "finite B"
 | 
|  |     25 |     using assms(1) assms(2) finite_subset by auto 
 | 
|  |     26 | next
 | 
|  |     27 |   show "a \<in> B" using assms by simp
 | 
|  |     28 | next
 | 
|  |     29 |   show "\<forall>x\<in>B. f x \<le> f a"
 | 
|  |     30 |     by (metis Max_ge assms(1) assms(2) assms(4) 
 | 
|  |     31 |             finite_imageI image_eqI subsetCE) 
 | 
|  |     32 | qed
 | 
|  |     33 | 
 | 
|  |     34 | text {*
 | 
|  |     35 |   The following locale @{text "highest_gen"} sets the basic context for our
 | 
|  |     36 |   investigation: supposing thread @{text th} holds the highest @{term cp}-value
 | 
|  |     37 |   in state @{text s}, which means the task for @{text th} is the 
 | 
|  |     38 |   most urgent. We want to show that  
 | 
|  |     39 |   @{text th} is treated correctly by PIP, which means
 | 
|  |     40 |   @{text th} will not be blocked unreasonably by other less urgent
 | 
|  |     41 |   threads. 
 | 
|  |     42 | *}
 | 
|  |     43 | locale highest_gen =
 | 
|  |     44 |   fixes s th prio tm
 | 
|  |     45 |   assumes vt_s: "vt s"
 | 
|  |     46 |   and threads_s: "th \<in> threads s"
 | 
|  |     47 |   and highest: "preced th s = Max ((cp s)`threads s)"
 | 
|  |     48 |   -- {* The internal structure of @{term th}'s precedence is exposed:*}
 | 
|  |     49 |   and preced_th: "preced th s = Prc prio tm" 
 | 
|  |     50 | 
 | 
|  |     51 | -- {* @{term s} is a valid trace, so it will inherit all results derived for
 | 
|  |     52 |       a valid trace: *}
 | 
|  |     53 | sublocale highest_gen < vat_s: valid_trace "s"
 | 
|  |     54 |   by (unfold_locales, insert vt_s, simp)
 | 
|  |     55 | 
 | 
|  |     56 | context highest_gen
 | 
|  |     57 | begin
 | 
|  |     58 | 
 | 
|  |     59 | text {*
 | 
|  |     60 |   @{term tm} is the time when the precedence of @{term th} is set, so 
 | 
|  |     61 |   @{term tm} must be a valid moment index into @{term s}.
 | 
|  |     62 | *}
 | 
|  |     63 | lemma lt_tm: "tm < length s"
 | 
|  |     64 |   by (insert preced_tm_lt[OF threads_s preced_th], simp)
 | 
|  |     65 | 
 | 
|  |     66 | text {*
 | 
|  |     67 |   Since @{term th} holds the highest precedence and @{text "cp"}
 | 
|  |     68 |   is the highest precedence of all threads in the sub-tree of 
 | 
|  |     69 |   @{text "th"} and @{text th} is among these threads, 
 | 
|  |     70 |   its @{term cp} must equal to its precedence:
 | 
|  |     71 | *}
 | 
|  |     72 | lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
 | 
|  |     73 | proof -
 | 
|  |     74 |   have "?L \<le> ?R"
 | 
|  |     75 |   by (unfold highest, rule Max_ge, 
 | 
|  |     76 |         auto simp:threads_s finite_threads)
 | 
|  |     77 |   moreover have "?R \<le> ?L"
 | 
|  |     78 |     by (unfold vat_s.cp_rec, rule Max_ge, 
 | 
|  |     79 |         auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
 | 
|  |     80 |   ultimately show ?thesis by auto
 | 
|  |     81 | qed
 | 
|  |     82 | 
 | 
|  |     83 | lemma highest_cp_preced: "cp s th = Max (the_preced s ` threads s)"
 | 
|  |     84 |   using eq_cp_s_th highest max_cp_eq the_preced_def by presburger
 | 
|  |     85 |   
 | 
|  |     86 | 
 | 
|  |     87 | lemma highest_preced_thread: "preced th s = Max (the_preced s ` threads s)"
 | 
|  |     88 |   by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
 | 
|  |     89 | 
 | 
|  |     90 | lemma highest': "cp s th = Max (cp s ` threads s)"
 | 
|  |     91 |   by (simp add: eq_cp_s_th highest)
 | 
|  |     92 | 
 | 
|  |     93 | end
 | 
|  |     94 | 
 | 
|  |     95 | locale extend_highest_gen = highest_gen + 
 | 
|  |     96 |   fixes t 
 | 
|  |     97 |   assumes vt_t: "vt (t@s)"
 | 
|  |     98 |   and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
 | 
|  |     99 |   and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
 | 
|  |    100 |   and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
 | 
|  |    101 | 
 | 
|  |    102 | sublocale extend_highest_gen < vat_t: valid_trace "t@s"
 | 
|  |    103 |   by (unfold_locales, insert vt_t, simp)
 | 
|  |    104 | 
 | 
|  |    105 | lemma step_back_vt_app: 
 | 
|  |    106 |   assumes vt_ts: "vt (t@s)" 
 | 
|  |    107 |   shows "vt s"
 | 
|  |    108 | proof -
 | 
|  |    109 |   from vt_ts show ?thesis
 | 
|  |    110 |   proof(induct t)
 | 
|  |    111 |     case Nil
 | 
|  |    112 |     from Nil show ?case by auto
 | 
|  |    113 |   next
 | 
|  |    114 |     case (Cons e t)
 | 
|  |    115 |     assume ih: " vt (t @ s) \<Longrightarrow> vt s"
 | 
|  |    116 |       and vt_et: "vt ((e # t) @ s)"
 | 
|  |    117 |     show ?case
 | 
|  |    118 |     proof(rule ih)
 | 
|  |    119 |       show "vt (t @ s)"
 | 
|  |    120 |       proof(rule step_back_vt)
 | 
|  |    121 |         from vt_et show "vt (e # t @ s)" by simp
 | 
|  |    122 |       qed
 | 
|  |    123 |     qed
 | 
|  |    124 |   qed
 | 
|  |    125 | qed
 | 
|  |    126 | 
 | 
|  |    127 | (* locale red_extend_highest_gen = extend_highest_gen +
 | 
|  |    128 |    fixes i::nat
 | 
|  |    129 | *)
 | 
|  |    130 | 
 | 
|  |    131 | (*
 | 
|  |    132 | sublocale red_extend_highest_gen <   red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
 | 
|  |    133 |   apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
 | 
|  |    134 |   apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
 | 
|  |    135 |   by (unfold highest_gen_def, auto dest:step_back_vt_app)
 | 
|  |    136 | *)
 | 
|  |    137 | 
 | 
|  |    138 | context extend_highest_gen
 | 
|  |    139 | begin
 | 
|  |    140 | 
 | 
|  |    141 |  lemma ind [consumes 0, case_names Nil Cons, induct type]:
 | 
|  |    142 |   assumes 
 | 
|  |    143 |     h0: "R []"
 | 
|  |    144 |   and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
 | 
|  |    145 |                     extend_highest_gen s th prio tm t; 
 | 
|  |    146 |                     extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
 | 
|  |    147 |   shows "R t"
 | 
|  |    148 | proof -
 | 
|  |    149 |   from vt_t extend_highest_gen_axioms show ?thesis
 | 
|  |    150 |   proof(induct t)
 | 
|  |    151 |     from h0 show "R []" .
 | 
|  |    152 |   next
 | 
|  |    153 |     case (Cons e t')
 | 
|  |    154 |     assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
 | 
|  |    155 |       and vt_e: "vt ((e # t') @ s)"
 | 
|  |    156 |       and et: "extend_highest_gen s th prio tm (e # t')"
 | 
|  |    157 |     from vt_e and step_back_step have stp: "step (t'@s) e" by auto
 | 
|  |    158 |     from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
 | 
|  |    159 |     show ?case
 | 
|  |    160 |     proof(rule h2 [OF vt_ts stp _ _ _ ])
 | 
|  |    161 |       show "R t'"
 | 
|  |    162 |       proof(rule ih)
 | 
|  |    163 |         from et show ext': "extend_highest_gen s th prio tm t'"
 | 
|  |    164 |           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
 | 
|  |    165 |       next
 | 
|  |    166 |         from vt_ts show "vt (t' @ s)" .
 | 
|  |    167 |       qed
 | 
|  |    168 |     next
 | 
|  |    169 |       from et show "extend_highest_gen s th prio tm (e # t')" .
 | 
|  |    170 |     next
 | 
|  |    171 |       from et show ext': "extend_highest_gen s th prio tm t'"
 | 
|  |    172 |           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
 | 
|  |    173 |     qed
 | 
|  |    174 |   qed
 | 
|  |    175 | qed
 | 
|  |    176 | 
 | 
|  |    177 | 
 | 
|  |    178 | lemma th_kept: "th \<in> threads (t @ s) \<and> 
 | 
|  |    179 |                  preced th (t@s) = preced th s" (is "?Q t") 
 | 
|  |    180 | proof -
 | 
|  |    181 |   show ?thesis
 | 
|  |    182 |   proof(induct rule:ind)
 | 
|  |    183 |     case Nil
 | 
|  |    184 |     from threads_s
 | 
|  |    185 |     show ?case
 | 
|  |    186 |       by auto
 | 
|  |    187 |   next
 | 
|  |    188 |     case (Cons e t)
 | 
|  |    189 |     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
 | 
|  |    190 |     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
 | 
|  |    191 |     show ?case
 | 
|  |    192 |     proof(cases e)
 | 
|  |    193 |       case (Create thread prio)
 | 
|  |    194 |       show ?thesis
 | 
|  |    195 |       proof -
 | 
|  |    196 |         from Cons and Create have "step (t@s) (Create thread prio)" by auto
 | 
|  |    197 |         hence "th \<noteq> thread"
 | 
|  |    198 |         proof(cases)
 | 
|  |    199 |           case thread_create
 | 
|  |    200 |           with Cons show ?thesis by auto
 | 
|  |    201 |         qed
 | 
|  |    202 |         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
 | 
|  |    203 |           by (unfold Create, auto simp:preced_def)
 | 
|  |    204 |         moreover note Cons
 | 
|  |    205 |         ultimately show ?thesis
 | 
|  |    206 |           by (auto simp:Create)
 | 
|  |    207 |       qed
 | 
|  |    208 |     next
 | 
|  |    209 |       case (Exit thread)
 | 
|  |    210 |       from h_e.exit_diff and Exit
 | 
|  |    211 |       have neq_th: "thread \<noteq> th" by auto
 | 
|  |    212 |       with Cons
 | 
|  |    213 |       show ?thesis
 | 
|  |    214 |         by (unfold Exit, auto simp:preced_def)
 | 
|  |    215 |     next
 | 
|  |    216 |       case (P thread cs)
 | 
|  |    217 |       with Cons
 | 
|  |    218 |       show ?thesis 
 | 
|  |    219 |         by (auto simp:P preced_def)
 | 
|  |    220 |     next
 | 
|  |    221 |       case (V thread cs)
 | 
|  |    222 |       with Cons
 | 
|  |    223 |       show ?thesis 
 | 
|  |    224 |         by (auto simp:V preced_def)
 | 
|  |    225 |     next
 | 
|  |    226 |       case (Set thread prio')
 | 
|  |    227 |       show ?thesis
 | 
|  |    228 |       proof -
 | 
|  |    229 |         from h_e.set_diff_low and Set
 | 
|  |    230 |         have "th \<noteq> thread" by auto
 | 
|  |    231 |         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
 | 
|  |    232 |           by (unfold Set, auto simp:preced_def)
 | 
|  |    233 |         moreover note Cons
 | 
|  |    234 |         ultimately show ?thesis
 | 
|  |    235 |           by (auto simp:Set)
 | 
|  |    236 |       qed
 | 
|  |    237 |     qed
 | 
|  |    238 |   qed
 | 
|  |    239 | qed
 | 
|  |    240 | 
 | 
|  |    241 | text {*
 | 
|  |    242 |   According to @{thm th_kept}, thread @{text "th"} has its living status
 | 
|  |    243 |   and precedence kept along the way of @{text "t"}. The following lemma
 | 
|  |    244 |   shows that this preserved precedence of @{text "th"} remains as the highest
 | 
|  |    245 |   along the way of @{text "t"}.
 | 
|  |    246 | 
 | 
|  |    247 |   The proof goes by induction over @{text "t"} using the specialized
 | 
|  |    248 |   induction rule @{thm ind}, followed by case analysis of each possible 
 | 
|  |    249 |   operations of PIP. All cases follow the same pattern rendered by the 
 | 
|  |    250 |   generalized introduction rule @{thm "image_Max_eqI"}. 
 | 
|  |    251 | 
 | 
|  |    252 |   The very essence is to show that precedences, no matter whether they 
 | 
|  |    253 |   are newly introduced or modified, are always lower than the one held 
 | 
|  |    254 |   by @{term "th"}, which by @{thm th_kept} is preserved along the way.
 | 
|  |    255 | *}
 | 
|  |    256 | lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
 | 
|  |    257 | proof(induct rule:ind)
 | 
|  |    258 |   case Nil
 | 
|  |    259 |   from highest_preced_thread
 | 
|  |    260 |   show ?case by simp
 | 
|  |    261 | next
 | 
|  |    262 |   case (Cons e t)
 | 
|  |    263 |     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
 | 
|  |    264 |     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
 | 
|  |    265 |   show ?case
 | 
|  |    266 |   proof(cases e)
 | 
|  |    267 |     case (Create thread prio')
 | 
|  |    268 |     show ?thesis (is "Max (?f ` ?A) = ?t")
 | 
|  |    269 |     proof -
 | 
|  |    270 |       -- {* The following is the common pattern of each branch of the case analysis. *}
 | 
|  |    271 |       -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
 | 
|  |    272 |       have "Max (?f ` ?A) = ?f th"
 | 
|  |    273 |       proof(rule image_Max_eqI)
 | 
|  |    274 |         show "finite ?A" using h_e.finite_threads by auto 
 | 
|  |    275 |       next
 | 
|  |    276 |         show "th \<in> ?A" using h_e.th_kept by auto 
 | 
|  |    277 |       next
 | 
|  |    278 |         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
 | 
|  |    279 |         proof 
 | 
|  |    280 |           fix x
 | 
|  |    281 |           assume "x \<in> ?A"
 | 
|  |    282 |           hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
 | 
|  |    283 |           thus "?f x \<le> ?f th"
 | 
|  |    284 |           proof
 | 
|  |    285 |             assume "x = thread"
 | 
|  |    286 |             thus ?thesis 
 | 
|  |    287 |               apply (simp add:Create the_preced_def preced_def, fold preced_def)
 | 
|  |    288 |               using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 
 | 
|  |    289 |               preced_th by force
 | 
|  |    290 |           next
 | 
|  |    291 |             assume h: "x \<in> threads (t @ s)"
 | 
|  |    292 |             from Cons(2)[unfolded Create] 
 | 
|  |    293 |             have "x \<noteq> thread" using h by (cases, auto)
 | 
|  |    294 |             hence "?f x = the_preced (t@s) x" 
 | 
|  |    295 |               by (simp add:Create the_preced_def preced_def)
 | 
|  |    296 |             hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
 | 
|  |    297 |               by (simp add: h_t.finite_threads h)
 | 
|  |    298 |             also have "... = ?f th"
 | 
|  |    299 |               by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
 | 
|  |    300 |             finally show ?thesis .
 | 
|  |    301 |           qed
 | 
|  |    302 |         qed
 | 
|  |    303 |       qed
 | 
|  |    304 |      -- {* The minor part is to show that the precedence of @{text "th"} 
 | 
|  |    305 |            equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
 | 
|  |    306 |       also have "... = ?t" using h_e.th_kept the_preced_def by auto
 | 
|  |    307 |       -- {* Then it follows trivially that the precedence preserved
 | 
|  |    308 |             for @{term "th"} remains the maximum of all living threads along the way. *}
 | 
|  |    309 |       finally show ?thesis .
 | 
|  |    310 |     qed 
 | 
|  |    311 |   next 
 | 
|  |    312 |     case (Exit thread)
 | 
|  |    313 |     show ?thesis (is "Max (?f ` ?A) = ?t")
 | 
|  |    314 |     proof -
 | 
|  |    315 |       have "Max (?f ` ?A) = ?f th"
 | 
|  |    316 |       proof(rule image_Max_eqI)
 | 
|  |    317 |         show "finite ?A" using h_e.finite_threads by auto 
 | 
|  |    318 |       next
 | 
|  |    319 |         show "th \<in> ?A" using h_e.th_kept by auto 
 | 
|  |    320 |       next
 | 
|  |    321 |         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
 | 
|  |    322 |         proof 
 | 
|  |    323 |           fix x
 | 
|  |    324 |           assume "x \<in> ?A"
 | 
|  |    325 |           hence "x \<in> threads (t@s)" by (simp add: Exit) 
 | 
|  |    326 |           hence "?f x \<le> Max (?f ` threads (t@s))" 
 | 
|  |    327 |             by (simp add: h_t.finite_threads) 
 | 
|  |    328 |           also have "... \<le> ?f th" 
 | 
|  |    329 |             apply (simp add:Exit the_preced_def preced_def, fold preced_def)
 | 
|  |    330 |             using Cons.hyps(5) h_t.th_kept the_preced_def by auto
 | 
|  |    331 |           finally show "?f x \<le> ?f th" .
 | 
|  |    332 |         qed
 | 
|  |    333 |       qed
 | 
|  |    334 |       also have "... = ?t" using h_e.th_kept the_preced_def by auto
 | 
|  |    335 |       finally show ?thesis .
 | 
|  |    336 |     qed 
 | 
|  |    337 |   next
 | 
|  |    338 |     case (P thread cs)
 | 
|  |    339 |     with Cons
 | 
|  |    340 |     show ?thesis by (auto simp:preced_def the_preced_def)
 | 
|  |    341 |   next
 | 
|  |    342 |     case (V thread cs)
 | 
|  |    343 |     with Cons
 | 
|  |    344 |     show ?thesis by (auto simp:preced_def the_preced_def)
 | 
|  |    345 |   next 
 | 
|  |    346 |     case (Set thread prio')
 | 
|  |    347 |     show ?thesis (is "Max (?f ` ?A) = ?t")
 | 
|  |    348 |     proof -
 | 
|  |    349 |       have "Max (?f ` ?A) = ?f th"
 | 
|  |    350 |       proof(rule image_Max_eqI)
 | 
|  |    351 |         show "finite ?A" using h_e.finite_threads by auto 
 | 
|  |    352 |       next
 | 
|  |    353 |         show "th \<in> ?A" using h_e.th_kept by auto 
 | 
|  |    354 |       next
 | 
|  |    355 |         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
 | 
|  |    356 |         proof 
 | 
|  |    357 |           fix x
 | 
|  |    358 |           assume h: "x \<in> ?A"
 | 
|  |    359 |           show "?f x \<le> ?f th"
 | 
|  |    360 |           proof(cases "x = thread")
 | 
|  |    361 |             case True
 | 
|  |    362 |             moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
 | 
|  |    363 |             proof -
 | 
|  |    364 |               have "the_preced (t @ s) th = Prc prio tm"  
 | 
|  |    365 |                 using h_t.th_kept preced_th by (simp add:the_preced_def)
 | 
|  |    366 |               moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
 | 
|  |    367 |               ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
 | 
|  |    368 |             qed
 | 
|  |    369 |             ultimately show ?thesis
 | 
|  |    370 |               by (unfold Set, simp add:the_preced_def preced_def)
 | 
|  |    371 |           next
 | 
|  |    372 |             case False
 | 
|  |    373 |             then have "?f x  = the_preced (t@s) x"
 | 
|  |    374 |               by (simp add:the_preced_def preced_def Set)
 | 
|  |    375 |             also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
 | 
|  |    376 |               using Set h h_t.finite_threads by auto 
 | 
|  |    377 |             also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
 | 
|  |    378 |             finally show ?thesis .
 | 
|  |    379 |           qed
 | 
|  |    380 |         qed
 | 
|  |    381 |       qed
 | 
|  |    382 |       also have "... = ?t" using h_e.th_kept the_preced_def by auto
 | 
|  |    383 |       finally show ?thesis .
 | 
|  |    384 |     qed 
 | 
|  |    385 |   qed
 | 
|  |    386 | qed
 | 
|  |    387 | 
 | 
|  |    388 | lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
 | 
|  |    389 |   by (insert th_kept max_kept, auto)
 | 
|  |    390 | 
 | 
|  |    391 | text {*
 | 
|  |    392 |   The reason behind the following lemma is that:
 | 
|  |    393 |   Since @{term "cp"} is defined as the maximum precedence 
 | 
|  |    394 |   of those threads contained in the sub-tree of node @{term "Th th"} 
 | 
|  |    395 |   in @{term "RAG (t@s)"}, and all these threads are living threads, and 
 | 
|  |    396 |   @{term "th"} is also among them, the maximum precedence of 
 | 
|  |    397 |   them all must be the one for @{text "th"}.
 | 
|  |    398 | *}
 | 
|  |    399 | lemma th_cp_max_preced: 
 | 
|  |    400 |   "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R") 
 | 
|  |    401 | proof -
 | 
|  |    402 |   let ?f = "the_preced (t@s)"
 | 
|  |    403 |   have "?L = ?f th"
 | 
|  |    404 |   proof(unfold cp_alt_def, rule image_Max_eqI)
 | 
|  |    405 |     show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
 | 
|  |    406 |     proof -
 | 
|  |    407 |       have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} = 
 | 
|  |    408 |             the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
 | 
|  |    409 |                             (\<exists> th'. n = Th th')}"
 | 
|  |    410 |       by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
 | 
|  |    411 |       moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree) 
 | 
|  |    412 |       ultimately show ?thesis by simp
 | 
|  |    413 |     qed
 | 
|  |    414 |   next
 | 
|  |    415 |     show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
 | 
|  |    416 |       by (auto simp:subtree_def)
 | 
|  |    417 |   next
 | 
|  |    418 |     show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
 | 
|  |    419 |                the_preced (t @ s) x \<le> the_preced (t @ s) th"
 | 
|  |    420 |     proof
 | 
|  |    421 |       fix th'
 | 
|  |    422 |       assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
 | 
|  |    423 |       hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
 | 
|  |    424 |       moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
 | 
|  |    425 |         by (meson subtree_Field)
 | 
|  |    426 |       ultimately have "Th th' \<in> ..." by auto
 | 
|  |    427 |       hence "th' \<in> threads (t@s)" 
 | 
|  |    428 |       proof
 | 
|  |    429 |         assume "Th th' \<in> {Th th}"
 | 
|  |    430 |         thus ?thesis using th_kept by auto 
 | 
|  |    431 |       next
 | 
|  |    432 |         assume "Th th' \<in> Field (RAG (t @ s))"
 | 
|  |    433 |         thus ?thesis using vat_t.not_in_thread_isolated by blast 
 | 
|  |    434 |       qed
 | 
|  |    435 |       thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
 | 
|  |    436 |         by (metis Max_ge finite_imageI finite_threads image_eqI 
 | 
|  |    437 |                max_kept th_kept the_preced_def)
 | 
|  |    438 |     qed
 | 
|  |    439 |   qed
 | 
|  |    440 |   also have "... = ?R" by (simp add: max_preced the_preced_def) 
 | 
|  |    441 |   finally show ?thesis .
 | 
|  |    442 | qed
 | 
|  |    443 | 
 | 
|  |    444 | lemma th_cp_max[simp]: "Max (cp (t@s) ` threads (t@s)) = cp (t@s) th"
 | 
|  |    445 |   using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
 | 
|  |    446 | 
 | 
|  |    447 | lemma [simp]: "Max (cp (t@s) ` threads (t@s)) = Max (the_preced (t@s) ` threads (t@s))"
 | 
|  |    448 |   by (simp add: th_cp_max_preced)
 | 
|  |    449 |   
 | 
|  |    450 | lemma [simp]: "Max (the_preced (t@s) ` threads (t@s)) = the_preced (t@s) th"
 | 
|  |    451 |   using max_kept th_kept the_preced_def by auto
 | 
|  |    452 | 
 | 
|  |    453 | lemma [simp]: "the_preced (t@s) th = preced th (t@s)"
 | 
|  |    454 |   using the_preced_def by auto
 | 
|  |    455 | 
 | 
|  |    456 | lemma [simp]: "preced th (t@s) = preced th s"
 | 
|  |    457 |   by (simp add: th_kept)
 | 
|  |    458 | 
 | 
|  |    459 | lemma [simp]: "cp s th = preced th s"
 | 
|  |    460 |   by (simp add: eq_cp_s_th)
 | 
|  |    461 | 
 | 
|  |    462 | lemma th_cp_preced [simp]: "cp (t@s) th = preced th s"
 | 
|  |    463 |   by (fold max_kept, unfold th_cp_max_preced, simp)
 | 
|  |    464 | 
 | 
|  |    465 | lemma preced_less:
 | 
|  |    466 |   assumes th'_in: "th' \<in> threads s"
 | 
|  |    467 |   and neq_th': "th' \<noteq> th"
 | 
|  |    468 |   shows "preced th' s < preced th s"
 | 
|  |    469 |   using assms
 | 
|  |    470 | by (metis Max.coboundedI finite_imageI highest not_le order.trans 
 | 
|  |    471 |     preced_linorder rev_image_eqI threads_s vat_s.finite_threads 
 | 
|  |    472 |     vat_s.le_cp)
 | 
|  |    473 | 
 | 
|  |    474 | section {* The `blocking thread` *}
 | 
|  |    475 | 
 | 
|  |    476 | text {* 
 | 
|  |    477 | 
 | 
|  |    478 |   The purpose of PIP is to ensure that the most urgent thread @{term
 | 
|  |    479 |   th} is not blocked unreasonably. Therefore, below, we will derive
 | 
|  |    480 |   properties of the blocking thread. By blocking thread, we mean a
 | 
|  |    481 |   thread in running state t @ s, but is different from thread @{term
 | 
|  |    482 |   th}.
 | 
|  |    483 | 
 | 
|  |    484 |   The first lemmas shows that the @{term cp}-value of the blocking
 | 
|  |    485 |   thread @{text th'} equals to the highest precedence in the whole
 | 
|  |    486 |   system.
 | 
|  |    487 | 
 | 
|  |    488 | *}
 | 
|  |    489 | 
 | 
|  |    490 | lemma runing_preced_inversion:
 | 
|  |    491 |   assumes runing': "th' \<in> runing (t @ s)"
 | 
|  |    492 |   shows "cp (t @ s) th' = preced th s" 
 | 
|  |    493 | proof -
 | 
|  |    494 |   have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))" 
 | 
|  |    495 |     using assms by (unfold runing_def, auto)
 | 
|  |    496 |   also have "\<dots> = preced th s"
 | 
|  |    497 |     by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads) 
 | 
|  |    498 |   finally show ?thesis .
 | 
|  |    499 | qed
 | 
|  |    500 | 
 | 
|  |    501 | text {*
 | 
|  |    502 | 
 | 
|  |    503 |   The next lemma shows how the counters for @{term "P"} and @{term
 | 
|  |    504 |   "V"} operations relate to the running threads in the states @{term
 | 
|  |    505 |   s} and @{term "t @ s"}: if a thread's @{term "P"}-count equals its
 | 
|  |    506 |   @{term "V"}-count (which means it no longer has any resource in its
 | 
|  |    507 |   possession), it cannot be a running thread.
 | 
|  |    508 | 
 | 
|  |    509 |   The proof is by contraction with the assumption @{text "th' \<noteq> th"}.
 | 
|  |    510 |   The key is the use of @{thm count_eq_dependants} to derive the
 | 
|  |    511 |   emptiness of @{text th'}s @{term dependants}-set from the balance of
 | 
|  |    512 |   its @{term P} and @{term V} counts.  From this, it can be shown
 | 
|  |    513 |   @{text th'}s @{term cp}-value equals to its own precedence.
 | 
|  |    514 | 
 | 
|  |    515 |   On the other hand, since @{text th'} is running, by @{thm
 | 
|  |    516 |   runing_preced_inversion}, its @{term cp}-value equals to the
 | 
|  |    517 |   precedence of @{term th}.
 | 
|  |    518 | 
 | 
|  |    519 |   Combining the above two results we have that @{text th'} and @{term
 | 
|  |    520 |   th} have the same precedence. By uniqueness of precedences, we have
 | 
|  |    521 |   @{text "th' = th"}, which is in contradiction with the assumption
 | 
|  |    522 |   @{text "th' \<noteq> th"}.
 | 
|  |    523 | 
 | 
|  |    524 | *} 
 | 
|  |    525 |                       
 | 
|  |    526 | lemma eq_pv_blocked: (* ddd *)
 | 
|  |    527 |   assumes neq_th': "th' \<noteq> th"
 | 
|  |    528 |   and eq_pv: "cntP (t @ s) th' = cntV (t @ s) th'"
 | 
|  |    529 |   shows "th' \<notin> runing (t @ s)"
 | 
|  |    530 | proof
 | 
|  |    531 |   assume otherwise: "th' \<in> runing (t @ s)"
 | 
|  |    532 |   show False
 | 
|  |    533 |   proof -
 | 
|  |    534 |     have th'_in: "th' \<in> threads (t @ s)"
 | 
|  |    535 |         using otherwise readys_threads runing_def by auto 
 | 
|  |    536 |     have "th' = th"
 | 
|  |    537 |     proof(rule preced_unique)
 | 
|  |    538 |       -- {* The proof goes like this: 
 | 
|  |    539 |             it is first shown that the @{term preced}-value of @{term th'} 
 | 
|  |    540 |             equals to that of @{term th}, then by uniqueness 
 | 
|  |    541 |             of @{term preced}-values (given by lemma @{thm preced_unique}), 
 | 
|  |    542 |             @{term th'} equals to @{term th}: *}
 | 
|  |    543 |       show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
 | 
|  |    544 |       proof -
 | 
|  |    545 |         -- {* Since the counts of @{term th'} are balanced, the subtree
 | 
|  |    546 |               of it contains only itself, so, its @{term cp}-value
 | 
|  |    547 |               equals its @{term preced}-value: *}
 | 
|  |    548 |         have "?L = cp (t @ s) th'"
 | 
|  |    549 |           by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
 | 
|  |    550 |         -- {* Since @{term "th'"} is running, by @{thm runing_preced_inversion},
 | 
|  |    551 |               its @{term cp}-value equals @{term "preced th s"}, 
 | 
|  |    552 |               which equals to @{term "?R"} by simplification: *}
 | 
|  |    553 |         also have "... = ?R" 
 | 
|  |    554 |             using runing_preced_inversion[OF otherwise] by simp
 | 
|  |    555 |         finally show ?thesis .
 | 
|  |    556 |       qed
 | 
|  |    557 |     qed (auto simp: th'_in th_kept)
 | 
|  |    558 |     with `th' \<noteq> th` show ?thesis by simp
 | 
|  |    559 |  qed
 | 
|  |    560 | qed
 | 
|  |    561 | 
 | 
|  |    562 | text {*
 | 
|  |    563 |   The following lemma is the extrapolation of @{thm eq_pv_blocked}.
 | 
|  |    564 |   It says if a thread, different from @{term th}, 
 | 
|  |    565 |   does not hold any resource at the very beginning,
 | 
|  |    566 |   it will keep hand-emptied in the future @{term "t@s"}.
 | 
|  |    567 | *}
 | 
|  |    568 | lemma eq_pv_persist: (* ddd *)
 | 
|  |    569 |   assumes neq_th': "th' \<noteq> th"
 | 
|  |    570 |   and eq_pv: "cntP s th' = cntV s th'"
 | 
|  |    571 |   shows "cntP (t @ s) th' = cntV (t @ s) th'"
 | 
|  |    572 | proof(induction rule: ind) 
 | 
|  |    573 |   -- {* The nontrivial case is for the @{term Cons}: *}
 | 
|  |    574 |   case (Cons e t)
 | 
|  |    575 |   -- {* All results derived so far hold for both @{term s} and @{term "t@s"}: *}
 | 
|  |    576 |   interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
 | 
|  |    577 |   interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
 | 
|  |    578 |   show ?case
 | 
|  |    579 |   proof -
 | 
|  |    580 |     -- {* It can be proved that @{term cntP}-value of @{term th'} does not change
 | 
|  |    581 |           by the happening of event @{term e}: *}
 | 
|  |    582 |     have "cntP ((e#t)@s) th' = cntP (t@s) th'"
 | 
|  |    583 |     proof(rule ccontr) -- {* Proof by contradiction. *}
 | 
|  |    584 |       -- {* Suppose @{term cntP}-value of @{term th'} is changed by @{term e}: *}
 | 
|  |    585 |       assume otherwise: "cntP ((e # t) @ s) th' \<noteq> cntP (t @ s) th'"
 | 
|  |    586 |       -- {* Then the actor of @{term e} must be @{term th'} and @{term e}
 | 
|  |    587 |             must be a @{term P}-event: *}
 | 
|  |    588 |       hence "isP e" "actor e = th'" by (auto simp:cntP_diff_inv) 
 | 
|  |    589 |       with vat_t.actor_inv[OF Cons(2)]
 | 
|  |    590 |       -- {* According to @{thm actor_inv}, @{term th'} must be running at 
 | 
|  |    591 |             the moment @{term "t@s"}: *}
 | 
|  |    592 |       have "th' \<in> runing (t@s)" by (cases e, auto)
 | 
|  |    593 |       -- {* However, an application of @{thm eq_pv_blocked} to induction hypothesis
 | 
|  |    594 |             shows @{term th'} can not be running at moment  @{term "t@s"}: *}
 | 
|  |    595 |       moreover have "th' \<notin> runing (t@s)" 
 | 
|  |    596 |                using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
 | 
|  |    597 |       -- {* Contradiction is finally derived: *}
 | 
|  |    598 |       ultimately show False by simp
 | 
|  |    599 |     qed
 | 
|  |    600 |     -- {* It can also be proved that @{term cntV}-value of @{term th'} does not change
 | 
|  |    601 |           by the happening of event @{term e}: *}
 | 
|  |    602 |     -- {* The proof follows exactly the same pattern as the case for @{term cntP}-value: *}
 | 
|  |    603 |     moreover have "cntV ((e#t)@s) th' = cntV (t@s) th'"
 | 
|  |    604 |     proof(rule ccontr) -- {* Proof by contradiction. *}
 | 
|  |    605 |       assume otherwise: "cntV ((e # t) @ s) th' \<noteq> cntV (t @ s) th'"
 | 
|  |    606 |       hence "isV e" "actor e = th'" by (auto simp:cntV_diff_inv) 
 | 
|  |    607 |       with vat_t.actor_inv[OF Cons(2)]
 | 
|  |    608 |       have "th' \<in> runing (t@s)" by (cases e, auto)
 | 
|  |    609 |       moreover have "th' \<notin> runing (t@s)"
 | 
|  |    610 |           using vat_t.eq_pv_blocked[OF neq_th' Cons(5)] .
 | 
|  |    611 |       ultimately show False by simp
 | 
|  |    612 |     qed
 | 
|  |    613 |     -- {* Finally, it can be shown that the @{term cntP} and @{term cntV} 
 | 
|  |    614 |           value for @{term th'} are still in balance, so @{term th'} 
 | 
|  |    615 |           is still hand-emptied after the execution of event @{term e}: *}
 | 
|  |    616 |     ultimately show ?thesis using Cons(5) by metis
 | 
|  |    617 |   qed
 | 
|  |    618 | qed (auto simp:eq_pv)
 | 
|  |    619 | 
 | 
|  |    620 | text {*
 | 
|  |    621 | 
 | 
|  |    622 |   By combining @{thm eq_pv_blocked} and @{thm eq_pv_persist}, it can
 | 
|  |    623 |   be derived easily that @{term th'} can not be running in the future:
 | 
|  |    624 | 
 | 
|  |    625 | *}
 | 
|  |    626 | 
 | 
|  |    627 | lemma eq_pv_blocked_persist:
 | 
|  |    628 |   assumes neq_th': "th' \<noteq> th"
 | 
|  |    629 |   and eq_pv: "cntP s th' = cntV s th'"
 | 
|  |    630 |   shows "th' \<notin> runing (t @ s)"
 | 
|  |    631 |   using assms
 | 
|  |    632 |   by (simp add: eq_pv_blocked eq_pv_persist) 
 | 
|  |    633 | 
 | 
|  |    634 | text {*
 | 
|  |    635 | 
 | 
|  |    636 |   The following lemma shows the blocking thread @{term th'} must hold
 | 
|  |    637 |   some resource in the very beginning.
 | 
|  |    638 | 
 | 
|  |    639 | *}
 | 
|  |    640 | 
 | 
|  |    641 | lemma runing_cntP_cntV_inv: (* ddd *)
 | 
|  |    642 |   assumes is_runing: "th' \<in> runing (t @ s)"
 | 
|  |    643 |   and neq_th': "th' \<noteq> th"
 | 
|  |    644 |   shows "cntP s th' > cntV s th'"
 | 
|  |    645 |   using assms
 | 
|  |    646 | proof -
 | 
|  |    647 |   -- {* First, it can be shown that the number of @{term P} and
 | 
|  |    648 |         @{term V} operations can not be equal for thred @{term th'} *}
 | 
|  |    649 |   have "cntP s th' \<noteq> cntV s th'"
 | 
|  |    650 |   proof
 | 
|  |    651 |      -- {* The proof goes by contradiction, suppose otherwise: *}
 | 
|  |    652 |     assume otherwise: "cntP s th' = cntV s th'"
 | 
|  |    653 |     -- {* By applying @{thm  eq_pv_blocked_persist} to this: *}
 | 
|  |    654 |     from eq_pv_blocked_persist[OF neq_th' otherwise] 
 | 
|  |    655 |     -- {* we have that @{term th'} can not be running at moment @{term "t@s"}: *}
 | 
|  |    656 |     have "th' \<notin> runing (t@s)" .
 | 
|  |    657 |     -- {* This is obvious in contradiction with assumption @{thm is_runing}  *}
 | 
|  |    658 |     thus False using is_runing by simp
 | 
|  |    659 |   qed
 | 
|  |    660 |   -- {* However, the number of @{term V} is always less or equal to @{term P}: *}
 | 
|  |    661 |   moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
 | 
|  |    662 |   -- {* Thesis is finally derived by combining the these two results: *}
 | 
|  |    663 |   ultimately show ?thesis by auto
 | 
|  |    664 | qed
 | 
|  |    665 | 
 | 
|  |    666 | 
 | 
|  |    667 | text {*
 | 
|  |    668 | 
 | 
|  |    669 |   The following lemmas shows the blocking thread @{text th'} must be
 | 
|  |    670 |   live at the very beginning, i.e. the moment (or state) @{term s}.
 | 
|  |    671 |   The proof is a  simple combination of the results above:
 | 
|  |    672 | 
 | 
|  |    673 | *}
 | 
|  |    674 | 
 | 
|  |    675 | lemma runing_threads_inv: 
 | 
|  |    676 |   assumes runing': "th' \<in> runing (t@s)"
 | 
|  |    677 |   and neq_th': "th' \<noteq> th"
 | 
|  |    678 |   shows "th' \<in> threads s"
 | 
|  |    679 | proof(rule ccontr) -- {* Proof by contradiction: *}
 | 
|  |    680 |   assume otherwise: "th' \<notin> threads s" 
 | 
|  |    681 |   have "th' \<notin> runing (t @ s)"
 | 
|  |    682 |   proof -
 | 
|  |    683 |     from vat_s.cnp_cnv_eq[OF otherwise]
 | 
|  |    684 |     have "cntP s th' = cntV s th'" .
 | 
|  |    685 |     from eq_pv_blocked_persist[OF neq_th' this]
 | 
|  |    686 |     show ?thesis .
 | 
|  |    687 |   qed
 | 
|  |    688 |   with runing' show False by simp
 | 
|  |    689 | qed
 | 
|  |    690 | 
 | 
|  |    691 | text {*
 | 
|  |    692 | 
 | 
|  |    693 |   The following lemma summarises the above lemmas to give an overall
 | 
|  |    694 |   characterisationof the blocking thread @{text "th'"}:
 | 
|  |    695 | 
 | 
|  |    696 | *}
 | 
|  |    697 | 
 | 
|  |    698 | lemma runing_inversion: (* ddd, one of the main lemmas to present *)
 | 
|  |    699 |   assumes runing': "th' \<in> runing (t@s)"
 | 
|  |    700 |   and neq_th: "th' \<noteq> th"
 | 
|  |    701 |   shows "th' \<in> threads s"
 | 
|  |    702 |   and    "\<not>detached s th'"
 | 
|  |    703 |   and    "cp (t@s) th' = preced th s"
 | 
|  |    704 | proof -
 | 
|  |    705 |   from runing_threads_inv[OF assms]
 | 
|  |    706 |   show "th' \<in> threads s" .
 | 
|  |    707 | next
 | 
|  |    708 |   from runing_cntP_cntV_inv[OF runing' neq_th]
 | 
|  |    709 |   show "\<not>detached s th'" using vat_s.detached_eq by simp
 | 
|  |    710 | next
 | 
|  |    711 |   from runing_preced_inversion[OF runing']
 | 
|  |    712 |   show "cp (t@s) th' = preced th s" .
 | 
|  |    713 | qed
 | 
|  |    714 | 
 | 
|  |    715 | 
 | 
|  |    716 | section {* The existence of `blocking thread` *}
 | 
|  |    717 | 
 | 
|  |    718 | text {* 
 | 
|  |    719 | 
 | 
|  |    720 |   Suppose @{term th} is not running, it is first shown that there is a
 | 
|  |    721 |   path in RAG leading from node @{term th} to another thread @{text
 | 
|  |    722 |   "th'"} in the @{term readys}-set (So @{text "th'"} is an ancestor of
 | 
|  |    723 |   @{term th}}).
 | 
|  |    724 | 
 | 
|  |    725 |   Now, since @{term readys}-set is non-empty, there must be one in it
 | 
|  |    726 |   which holds the highest @{term cp}-value, which, by definition, is
 | 
|  |    727 |   the @{term runing}-thread. However, we are going to show more: this
 | 
|  |    728 |   running thread is exactly @{term "th'"}.
 | 
|  |    729 | 
 | 
|  |    730 | *}
 | 
|  |    731 | 
 | 
|  |    732 | lemma th_blockedE: (* ddd, the other main lemma to be presented: *)
 | 
|  |    733 |   assumes "th \<notin> runing (t@s)"
 | 
|  |    734 |   obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
 | 
|  |    735 |                     "th' \<in> runing (t@s)"
 | 
|  |    736 | proof -
 | 
|  |    737 |   -- {* According to @{thm vat_t.th_chain_to_ready}, either 
 | 
|  |    738 |         @{term "th"} is in @{term "readys"} or there is path leading from it to 
 | 
|  |    739 |         one thread in @{term "readys"}. *}
 | 
|  |    740 |   have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)" 
 | 
|  |    741 |     using th_kept vat_t.th_chain_to_ready by auto
 | 
|  |    742 |   -- {* However, @{term th} can not be in @{term readys}, because otherwise, since 
 | 
|  |    743 |        @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
 | 
|  |    744 |   moreover have "th \<notin> readys (t@s)" 
 | 
|  |    745 |     using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto 
 | 
|  |    746 |   -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in 
 | 
|  |    747 |         term @{term readys}: *}
 | 
|  |    748 |   ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
 | 
|  |    749 |                           and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
 | 
|  |    750 |   -- {* We are going to show that this @{term th'} is running. *}
 | 
|  |    751 |   have "th' \<in> runing (t@s)"
 | 
|  |    752 |   proof -
 | 
|  |    753 |     -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
 | 
|  |    754 |     have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
 | 
|  |    755 |     proof -
 | 
|  |    756 |       have "?L =  Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
 | 
|  |    757 |         by (unfold cp_alt_def1, simp)
 | 
|  |    758 |       also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
 | 
|  |    759 |       proof(rule image_Max_subset)
 | 
|  |    760 |         show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
 | 
|  |    761 |       next
 | 
|  |    762 |         show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
 | 
|  |    763 |           by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread) 
 | 
|  |    764 |       next
 | 
|  |    765 |         show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
 | 
|  |    766 |                     by (unfold tRAG_subtree_eq, auto simp:subtree_def)
 | 
|  |    767 |       next
 | 
|  |    768 |         show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
 | 
|  |    769 |                       (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
 | 
|  |    770 |         proof -
 | 
|  |    771 |           have "?L = the_preced (t @ s) `  threads (t @ s)" 
 | 
|  |    772 |                      by (unfold image_comp, rule image_cong, auto)
 | 
|  |    773 |           thus ?thesis using max_preced the_preced_def by auto
 | 
|  |    774 |         qed
 | 
|  |    775 |       qed
 | 
|  |    776 |       also have "... = ?R"
 | 
|  |    777 |         using th_cp_max th_cp_preced th_kept 
 | 
|  |    778 |               the_preced_def vat_t.max_cp_readys_threads by auto
 | 
|  |    779 |       finally show ?thesis .
 | 
|  |    780 |     qed 
 | 
|  |    781 |     -- {* Now, since @{term th'} holds the highest @{term cp} 
 | 
|  |    782 |           and we have already show it is in @{term readys},
 | 
|  |    783 |           it is @{term runing} by definition. *}
 | 
|  |    784 |     with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def) 
 | 
|  |    785 |   qed
 | 
|  |    786 |   -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
 | 
|  |    787 |   moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)" 
 | 
|  |    788 |     using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
 | 
|  |    789 |   ultimately show ?thesis using that by metis
 | 
|  |    790 | qed
 | 
|  |    791 | 
 | 
|  |    792 | text {*
 | 
|  |    793 | 
 | 
|  |    794 |   Now it is easy to see there is always a thread to run by case
 | 
|  |    795 |   analysis on whether thread @{term th} is running: if the answer is
 | 
|  |    796 |   yes, the the running thread is obviously @{term th} itself;
 | 
|  |    797 |   otherwise, the running thread is the @{text th'} given by lemma
 | 
|  |    798 |   @{thm th_blockedE}.
 | 
|  |    799 | 
 | 
|  |    800 | *}
 | 
|  |    801 | 
 | 
|  |    802 | lemma live: "runing (t@s) \<noteq> {}"
 | 
|  |    803 | proof(cases "th \<in> runing (t@s)") 
 | 
|  |    804 |   case True thus ?thesis by auto
 | 
|  |    805 | next
 | 
|  |    806 |   case False
 | 
|  |    807 |   thus ?thesis using th_blockedE by auto
 | 
|  |    808 | qed
 | 
|  |    809 | 
 | 
|  |    810 | 
 | 
|  |    811 | end
 | 
|  |    812 | end
 |