ExtGG.thy
changeset 63 b620a2a0806a
parent 62 031d2ae9c9b8
equal deleted inserted replaced
62:031d2ae9c9b8 63:b620a2a0806a
     1 theory ExtGG
     1 theory ExtGG
     2 imports PrioG CpsG
     2 imports PrioG CpsG
     3 begin
     3 begin
     4 
     4 
     5 lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
     5 text {* 
     6   apply (induct s, simp)
     6   The following two auxiliary lemmas are used to reason about @{term Max}.
     7 proof -
     7 *}
     8   fix a s
     8 lemma image_Max_eqI: 
     9   assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
     9   assumes "finite B"
    10     and eq_as: "a # s \<noteq> []"
    10   and "b \<in> B"
    11   show "last_set th (a # s) < length (a # s)"
    11   and "\<forall> x \<in> B. f x \<le> f b"
    12   proof(cases "s \<noteq> []")
    12   shows "Max (f ` B) = f b"
    13     case False
    13   using assms
    14     from False show ?thesis
    14   using Max_eqI by blast 
    15       by (cases a, auto simp:last_set.simps)
    15 
    16   next
    16 lemma image_Max_subset:
    17     case True
    17   assumes "finite A"
    18     from ih [OF True] show ?thesis
    18   and "B \<subseteq> A"
    19       by (cases a, auto simp:last_set.simps)
    19   and "a \<in> B"
    20   qed
    20   and "Max (f ` A) = f a"
    21 qed
    21   shows "Max (f ` B) = f a"
    22 
    22 proof(rule image_Max_eqI)
    23 lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
    23   show "finite B"
    24   by (induct s, auto simp:threads.simps)
    24     using assms(1) assms(2) finite_subset by auto 
    25 
    25 next
    26 lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
    26   show "a \<in> B" using assms by simp
    27   apply (drule_tac th_in_ne)
    27 next
    28   by (unfold preced_def, auto intro: birth_time_lt)
    28   show "\<forall>x\<in>B. f x \<le> f a"
    29 
    29     by (metis Max_ge assms(1) assms(2) assms(4) 
       
    30             finite_imageI image_eqI subsetCE) 
       
    31 qed
       
    32 
       
    33 text {*
       
    34   The following locale @{text "highest_gen"} sets the basic context for our
       
    35   investigation: supposing thread @{text th} holds the highest @{term cp}-value
       
    36   in state @{text s}, which means the task for @{text th} is the 
       
    37   most urgent. We want to show that  
       
    38   @{text th} is treated correctly by PIP, which means
       
    39   @{text th} will not be blocked unreasonably by other less urgent
       
    40   threads. 
       
    41 *}
    30 locale highest_gen =
    42 locale highest_gen =
    31   fixes s th prio tm
    43   fixes s th prio tm
    32   assumes vt_s: "vt s"
    44   assumes vt_s: "vt s"
    33   and threads_s: "th \<in> threads s"
    45   and threads_s: "th \<in> threads s"
    34   and highest: "preced th s = Max ((cp s)`threads s)"
    46   and highest: "preced th s = Max ((cp s)`threads s)"
    35   and preced_th: "preced th s = Prc prio tm"
    47   -- {* The internal structure of @{term th}'s precedence is exposed:*}
    36 
    48   and preced_th: "preced th s = Prc prio tm" 
       
    49 
       
    50 -- {* @{term s} is a valid trace, so it will inherit all results derived for
       
    51       a valid trace: *}
    37 sublocale highest_gen < vat_s: valid_trace "s"
    52 sublocale highest_gen < vat_s: valid_trace "s"
    38   by (unfold_locales, insert vt_s, simp)
    53   by (unfold_locales, insert vt_s, simp)
    39 
    54 
    40 context highest_gen
    55 context highest_gen
    41 begin
    56 begin
    42 
    57 
       
    58 text {*
       
    59   @{term tm} is the time when the precedence of @{term th} is set, so 
       
    60   @{term tm} must be a valid moment index into @{term s}.
       
    61 *}
    43 lemma lt_tm: "tm < length s"
    62 lemma lt_tm: "tm < length s"
    44   by (insert preced_tm_lt[OF threads_s preced_th], simp)
    63   by (insert preced_tm_lt[OF threads_s preced_th], simp)
    45 
    64 
       
    65 text {*
       
    66   Since @{term th} holds the highest precedence and @{text "cp"}
       
    67   is the highest precedence of all threads in the sub-tree of 
       
    68   @{text "th"} and @{text th} is among these threads, 
       
    69   its @{term cp} must equal to its precedence:
       
    70 *}
    46 lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
    71 lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
    47 proof -
    72 proof -
    48   have "?L \<le> ?R"
    73   have "?L \<le> ?R"
    49   by (unfold highest, rule Max_ge, 
    74   by (unfold highest, rule Max_ge, 
    50         auto simp:threads_s finite_threads[OF vt_s])
    75         auto simp:threads_s finite_threads)
    51   moreover have "?R \<le> ?L"
    76   moreover have "?R \<le> ?L"
    52     by (unfold vat_s.cp_rec, rule Max_ge, 
    77     by (unfold vat_s.cp_rec, rule Max_ge, 
    53         auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
    78         auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
    54   ultimately show ?thesis by auto
    79   ultimately show ?thesis by auto
    55 qed
    80 qed
    56 
    81 
       
    82 (* ccc *)
    57 lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
    83 lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
    58   by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp)
    84   by (fold max_cp_eq, unfold eq_cp_s_th, insert highest, simp)
    59 
    85 
    60 lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
    86 lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
    61   by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
    87   by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
    62 
    88 
    63 lemma highest': "cp s th = Max (cp s ` threads s)"
    89 lemma highest': "cp s th = Max (cp s ` threads s)"
    64 proof -
    90 proof -
    65   from highest_cp_preced max_cp_eq[OF vt_s, symmetric]
    91   from highest_cp_preced max_cp_eq[symmetric]
    66   show ?thesis by simp
    92   show ?thesis by simp
    67 qed
    93 qed
    68 
    94 
    69 end
    95 end
    70 
    96 
    72   fixes t 
    98   fixes t 
    73   assumes vt_t: "vt (t@s)"
    99   assumes vt_t: "vt (t@s)"
    74   and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
   100   and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
    75   and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
   101   and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
    76   and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
   102   and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
       
   103 
       
   104 sublocale extend_highest_gen < vat_t: valid_trace "t@s"
       
   105   by (unfold_locales, insert vt_t, simp)
    77 
   106 
    78 lemma step_back_vt_app: 
   107 lemma step_back_vt_app: 
    79   assumes vt_ts: "vt (t@s)" 
   108   assumes vt_ts: "vt (t@s)" 
    80   shows "vt s"
   109   shows "vt s"
    81 proof -
   110 proof -
   107   by (unfold highest_gen_def, auto dest:step_back_vt_app)
   136   by (unfold highest_gen_def, auto dest:step_back_vt_app)
   108 
   137 
   109 
   138 
   110 context extend_highest_gen
   139 context extend_highest_gen
   111 begin
   140 begin
   112 
       
   113 (*
       
   114  lemma red_moment:
       
   115   "extend_highest_gen s th prio tm (moment i t)"
       
   116   apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
       
   117   apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
       
   118   by (unfold highest_gen_def, auto dest:step_back_vt_app) 
       
   119 *)
       
   120 
   141 
   121  lemma ind [consumes 0, case_names Nil Cons, induct type]:
   142  lemma ind [consumes 0, case_names Nil Cons, induct type]:
   122   assumes 
   143   assumes 
   123     h0: "R []"
   144     h0: "R []"
   124   and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
   145   and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
   216       qed
   237       qed
   217     qed
   238     qed
   218   qed
   239   qed
   219 qed
   240 qed
   220 
   241 
   221 lemma Max_remove_less:
   242 text {*
   222   assumes "finite A"
   243   According to @{thm th_kept}, thread @{text "th"} has its living status
   223   and "a \<in> A"
   244   and precedence kept along the way of @{text "t"}. The following lemma
   224   and "b \<in> A"
   245   shows that this preserved precedence of @{text "th"} remains as the highest
   225   and "a \<noteq> b"
   246   along the way of @{text "t"}.
   226   and "inj_on f A"
   247 
   227   and "f a = Max (f ` A)" 
   248   The proof goes by induction over @{text "t"} using the specialized
   228   shows "Max (f ` (A - {b})) = (Max (f ` A))"
   249   induction rule @{thm ind}, followed by case analysis of each possible 
   229 proof -
   250   operations of PIP. All cases follow the same pattern rendered by the 
   230   have "Max (f ` (A - {b})) = Max (f`A - {f b})"
   251   generalized introduction rule @{thm "image_Max_eqI"}. 
   231   proof -
   252 
   232     have "f ` (A - {b}) = f ` A - f ` {b}"
   253   The very essence is to show that precedences, no matter whether they are newly introduced 
   233     by (rule inj_on_image_set_diff[OF assms(5)], insert assms(3), auto)
   254   or modified, are always lower than the one held by @{term "th"},
   234     thus ?thesis by simp
   255   which by @{thm th_kept} is preserved along the way.
   235   qed
   256 *}
   236   also have "... =  
       
   237        (if f ` A - {f b} - {f a} = {} then f a else max (f a) (Max (f ` A - {f b} - {f a})))" 
       
   238   proof(subst Max.remove)
       
   239     from assms show "f a \<in> f ` A - {f b}"
       
   240       by (meson DiffI empty_iff imageI inj_on_eq_iff insert_iff) 
       
   241   next
       
   242     from assms(1) show "finite (f ` A - {f b})" by auto
       
   243   qed auto
       
   244   also have "... = Max (f ` A)" (is "?L = ?R")
       
   245   proof(cases "f ` A - {f b} - {f a} = {}")
       
   246     case True
       
   247     with assms show ?thesis by auto
       
   248   next
       
   249     case False
       
   250     hence "?L =  max (f a) (Max (f ` A - {f b} - {f a}))" 
       
   251       by simp
       
   252     also have "... = ?R" 
       
   253     proof -
       
   254       from assms False
       
   255       have "(Max (f ` A - {f b} - {f a})) \<le> f a" by auto
       
   256       thus ?thesis by (simp add: assms(6) max_def) 
       
   257     qed
       
   258     finally show ?thesis .
       
   259   qed
       
   260   finally show ?thesis .
       
   261 qed
       
   262 
       
   263 lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
   257 lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
   264 proof(induct rule:ind)
   258 proof(induct rule:ind)
   265   case Nil
   259   case Nil
   266   from highest_preced_thread
   260   from highest_preced_thread
   267   show ?case
   261   show ?case
   271     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
   265     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
   272     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
   266     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
   273   show ?case
   267   show ?case
   274   proof(cases e)
   268   proof(cases e)
   275     case (Create thread prio')
   269     case (Create thread prio')
   276     from Cons(2)[unfolded this] 
       
   277     have thread_not_in: "thread \<notin> threads (t@s)" by (cases, simp)
       
   278     show ?thesis (is "Max (?f ` ?A) = ?t")
   270     show ?thesis (is "Max (?f ` ?A) = ?t")
   279     proof -
   271     proof -
   280       have "Max (?f ` ?A) = Max (insert (?f thread) (?f ` (threads (t@s))))"
   272       -- {* The following is the common pattern of each branch of the case analysis. *}
   281         by (unfold Create, simp)
   273       -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
   282       moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
   274       have "Max (?f ` ?A) = ?f th"
   283       proof(rule Max.insert)
   275       proof(rule image_Max_eqI)
   284         from finite_threads[OF Cons(1)]
   276         show "finite ?A" using h_e.finite_threads by auto 
   285         show "finite (?f ` (threads (t@s)))" by simp
   277       next
   286       qed (insert h_t.th_kept, auto)
   278         show "th \<in> ?A" using h_e.th_kept by auto 
   287       moreover have "(Max (?f ` (threads (t@s)))) = ?t" 
   279       next
   288       proof -
   280         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
   289         have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
   281         proof 
   290                 (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" 
   282           fix x
   291         by (intro f_image_eq, insert thread_not_in, auto simp:Create preced_def)
   283           assume "x \<in> ?A"
   292         with Cons show ?thesis by (auto simp:the_preced_def)
   284           hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
   293       qed
   285           thus "?f x \<le> ?f th"
   294       moreover have "?f thread < ?t"
   286           proof
   295       proof -
   287             assume "x = thread"
   296         from h_e.create_low and Create
   288             thus ?thesis 
   297         have "prio' \<le> prio" by auto
   289               apply (simp add:Create the_preced_def preced_def, fold preced_def)
   298         thus ?thesis
   290               using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 preced_th by force
   299         by (unfold preced_th, unfold Create, insert lt_tm, 
   291           next
   300           auto simp:preced_def precedence_less_def preced_th the_preced_def)
   292             assume h: "x \<in> threads (t @ s)"
   301      qed
   293             from Cons(2)[unfolded Create] 
   302      ultimately show ?thesis by (auto simp:max_def)
   294             have "x \<noteq> thread" using h by (cases, auto)
   303     qed
   295             hence "?f x = the_preced (t@s) x" 
       
   296               by (simp add:Create the_preced_def preced_def)
       
   297             hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
       
   298               by (simp add: h_t.finite_threads h)
       
   299             also have "... = ?f th"
       
   300               by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
       
   301             finally show ?thesis .
       
   302           qed
       
   303         qed
       
   304       qed
       
   305      -- {* The minor part is to show that the precedence of @{text "th"} 
       
   306            equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
       
   307       also have "... = ?t" using h_e.th_kept the_preced_def by auto
       
   308       -- {* Then it follows trivially that the precedence preserved
       
   309             for @{term "th"} remains the maximum of all living threads along the way. *}
       
   310       finally show ?thesis .
       
   311     qed 
   304   next 
   312   next 
   305     case (Exit thread)
   313     case (Exit thread)
   306     show ?thesis
   314     show ?thesis (is "Max (?f ` ?A) = ?t")
   307     proof -
   315     proof -
   308       have "Max (the_preced (t @ s) ` (threads (t @ s) - {thread})) = 
   316       have "Max (?f ` ?A) = ?f th"
   309             Max (the_preced (t @ s) ` (threads (t @ s)))"
   317       proof(rule image_Max_eqI)
   310       proof(rule Max_remove_less)
   318         show "finite ?A" using h_e.finite_threads by auto 
   311         show "th \<noteq> thread" using Exit h_e.exit_diff by auto 
   319       next
   312       next
   320         show "th \<in> ?A" using h_e.th_kept by auto 
   313         from Cons(2)[unfolded Exit]
   321       next
   314         show "thread \<in> threads (t @ s)" 
   322         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
   315           by (cases, simp add: readys_def runing_def)
   323         proof 
   316       next
   324           fix x
   317         show "finite (threads (t @ s))" by (simp add: finite_threads h_t.vt_t) 
   325           assume "x \<in> ?A"
   318       next
   326           hence "x \<in> threads (t@s)" by (simp add: Exit) 
   319         show "th \<in> threads (t @ s)" by (simp add: h_t.th_kept) 
   327           hence "?f x \<le> Max (?f ` threads (t@s))" 
   320       next
   328             by (simp add: h_t.finite_threads) 
   321         show "inj_on (the_preced (t @ s)) (threads (t @ s))" by (simp add: inj_the_preced) 
   329           also have "... \<le> ?f th" 
   322       next
   330             apply (simp add:Exit the_preced_def preced_def, fold preced_def)
   323         show "the_preced (t @ s) th = Max (the_preced (t @ s) ` threads (t @ s))"
   331             using Cons.hyps(5) h_t.th_kept the_preced_def by auto
   324             by (simp add: Cons.hyps(5) h_t.th_kept the_preced_def)
   332           finally show "?f x \<le> ?f th" .
   325       qed
   333         qed
   326       from this[unfolded Cons(5)]
   334       qed
   327       have "Max (the_preced (t @ s) ` (threads (t @ s) - {thread})) = preced th s" .
   335       also have "... = ?t" using h_e.th_kept the_preced_def by auto
   328       moreover have "the_preced ((e # t) @ s) = the_preced (t@s)"
   336       finally show ?thesis .
   329                              by (auto simp:Exit the_preced_def preced_def)
   337     qed 
   330       ultimately show ?thesis by (simp add:Exit)
       
   331     qed
       
   332   next
   338   next
   333     case (P thread cs)
   339     case (P thread cs)
   334     with Cons
   340     with Cons
   335     show ?thesis by (auto simp:preced_def the_preced_def)
   341     show ?thesis by (auto simp:preced_def the_preced_def)
   336   next
   342   next
   337     case (V thread cs)
   343     case (V thread cs)
   338     with Cons
   344     with Cons
   339     show ?thesis by (auto simp:preced_def the_preced_def)
   345     show ?thesis by (auto simp:preced_def the_preced_def)
   340   next (* ccc *)
   346   next 
   341     case (Set thread prio')
   347     case (Set thread prio')
   342     show ?thesis
   348     show ?thesis (is "Max (?f ` ?A) = ?t")
   343     apply (unfold Set, simp, insert Cons(5)) (* ccc *)
       
   344     find_theorems priority Set
       
   345     find_theorems preced Set
       
   346     proof -
   349     proof -
   347       let ?B = "threads (t@s)"
   350       have "Max (?f ` ?A) = ?f th"
   348       from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
   351       proof(rule image_Max_eqI)
   349       from extend_highest_gen.set_diff_low[OF this] and Set
   352         show "finite ?A" using h_e.finite_threads by auto 
   350       have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
   353       next
   351       from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
   354         show "th \<in> ?A" using h_e.th_kept by auto 
   352       also have "\<dots> = ?t"
   355       next
   353       proof(rule Max_eqI)
   356         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
   354         fix y
   357         proof 
   355         assume y_in: "y \<in> ?f ` ?B"
   358           fix x
   356         then obtain th1 where 
   359           assume h: "x \<in> ?A"
   357           th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
   360           show "?f x \<le> ?f th"
   358         show "y \<le> ?t"
   361           proof(cases "x = thread")
   359         proof(cases "th1 = thread")
   362             case True
   360           case True
   363             moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
   361           with neq_thread le_p eq_y Set
   364             proof -
   362           show ?thesis
   365               have "the_preced (t @ s) th = Prc prio tm"  
   363             apply (subst preced_th, insert lt_tm)
   366                 using h_t.th_kept preced_th by (simp add:the_preced_def)
   364             by (auto simp:preced_def precedence_le_def)
   367               moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
   365         next
   368               ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
   366           case False
       
   367           with Set eq_y
       
   368           have "y  = preced th1 (t@s)"
       
   369             by (simp add:preced_def)
       
   370           moreover have "\<dots> \<le> ?t"
       
   371           proof -
       
   372             from Cons
       
   373             have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
       
   374               by auto
       
   375             moreover have "preced th1 (t@s) \<le> \<dots>"
       
   376             proof(rule Max_ge)
       
   377               from th1_in 
       
   378               show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
       
   379                 by simp
       
   380             next
       
   381               show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
       
   382               proof -
       
   383                 from Cons have "vt (t @ s)" by auto
       
   384                 from finite_threads[OF this] show ?thesis by auto
       
   385               qed
       
   386             qed
   369             qed
   387             ultimately show ?thesis by auto
   370             ultimately show ?thesis
       
   371               by (unfold Set, simp add:the_preced_def preced_def)
       
   372           next
       
   373             case False
       
   374             then have "?f x  = the_preced (t@s) x"
       
   375               by (simp add:the_preced_def preced_def Set)
       
   376             also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
       
   377               using Set h h_t.finite_threads by auto 
       
   378             also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
       
   379             finally show ?thesis .
   388           qed
   380           qed
   389           ultimately show ?thesis by auto
       
   390         qed
   381         qed
   391       next
   382       qed
   392         from Cons and finite_threads
   383       also have "... = ?t" using h_e.th_kept the_preced_def by auto
   393         show "finite (?f ` ?B)" by auto
       
   394       next
       
   395         from Cons have "extend_highest_gen s th prio tm t" by auto
       
   396         from extend_highest_gen.th_kept [OF this]
       
   397         have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
       
   398         show "?t \<in> (?f ` ?B)" 
       
   399         proof -
       
   400           from neq_thread Set h
       
   401           have "?t = ?f th" by (auto simp:preced_def)
       
   402           with h show ?thesis by auto
       
   403         qed
       
   404       qed
       
   405       finally show ?thesis .
   384       finally show ?thesis .
       
   385     qed 
       
   386   qed
       
   387 qed
       
   388 
       
   389 lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
       
   390   by (insert th_kept max_kept, auto)
       
   391 
       
   392 text {*
       
   393   The reason behind the following lemma is that:
       
   394   Since @{term "cp"} is defined as the maximum precedence 
       
   395   of those threads contained in the sub-tree of node @{term "Th th"} 
       
   396   in @{term "RAG (t@s)"}, and all these threads are living threads, and 
       
   397   @{term "th"} is also among them, the maximum precedence of 
       
   398   them all must be the one for @{text "th"}.
       
   399 *}
       
   400 lemma th_cp_max_preced: 
       
   401   "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R") 
       
   402 proof -
       
   403   let ?f = "the_preced (t@s)"
       
   404   have "?L = ?f th"
       
   405   proof(unfold cp_alt_def, rule image_Max_eqI)
       
   406     show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
       
   407     proof -
       
   408       have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} = 
       
   409             the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
       
   410                             (\<exists> th'. n = Th th')}"
       
   411       by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
       
   412       moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree) 
       
   413       ultimately show ?thesis by simp
       
   414     qed
       
   415   next
       
   416     show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
       
   417       by (auto simp:subtree_def)
       
   418   next
       
   419     show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
       
   420                the_preced (t @ s) x \<le> the_preced (t @ s) th"
       
   421     proof
       
   422       fix th'
       
   423       assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
       
   424       hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
       
   425       moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
       
   426         by (meson subtree_Field)
       
   427       ultimately have "Th th' \<in> ..." by auto
       
   428       hence "th' \<in> threads (t@s)" 
       
   429       proof
       
   430         assume "Th th' \<in> {Th th}"
       
   431         thus ?thesis using th_kept by auto 
       
   432       next
       
   433         assume "Th th' \<in> Field (RAG (t @ s))"
       
   434         thus ?thesis using vat_t.not_in_thread_isolated by blast 
       
   435       qed
       
   436       thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
       
   437         by (metis Max_ge finite_imageI finite_threads image_eqI 
       
   438                max_kept th_kept the_preced_def)
   406     qed
   439     qed
   407   qed
   440   qed
   408 qed
   441   also have "... = ?R" by (simp add: max_preced the_preced_def) 
   409 
       
   410 lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
       
   411   by (insert th_kept max_kept, auto)
       
   412 
       
   413 lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
       
   414   (is "?L = ?R")
       
   415 proof -
       
   416   have "?L = cpreced (wq (t@s)) (t@s) th" 
       
   417     by (unfold cp_eq_cpreced, simp)
       
   418   also have "\<dots> = ?R"
       
   419   proof(unfold cpreced_def)
       
   420     show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependants (wq (t @ s)) th)) =
       
   421           Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
       
   422       (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
       
   423     proof(cases "?A = {}")
       
   424       case False
       
   425       have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
       
   426       moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
       
   427       proof(rule Max_insert)
       
   428         show "finite (?f ` ?A)"
       
   429         proof -
       
   430           from dependants_threads[OF vt_t]
       
   431           have "?A \<subseteq> threads (t@s)" .
       
   432           moreover from finite_threads[OF vt_t] have "finite \<dots>" .
       
   433           ultimately show ?thesis 
       
   434             by (auto simp:finite_subset)
       
   435         qed
       
   436       next
       
   437         from False show "(?f ` ?A) \<noteq> {}" by simp
       
   438       qed
       
   439       moreover have "\<dots> = Max (?f ` ?B)"
       
   440       proof -
       
   441         from max_preced have "?f th = Max (?f ` ?B)" .
       
   442         moreover have "Max (?f ` ?A) \<le> \<dots>" 
       
   443         proof(rule Max_mono)
       
   444           from False show "(?f ` ?A) \<noteq> {}" by simp
       
   445         next
       
   446           show "?f ` ?A \<subseteq> ?f ` ?B" 
       
   447           proof -
       
   448             have "?A \<subseteq> ?B" by (rule dependants_threads[OF vt_t])
       
   449             thus ?thesis by auto
       
   450           qed
       
   451         next
       
   452           from finite_threads[OF vt_t] 
       
   453           show "finite (?f ` ?B)" by simp
       
   454         qed
       
   455         ultimately show ?thesis
       
   456           by (auto simp:max_def)
       
   457       qed
       
   458       ultimately show ?thesis by auto
       
   459     next
       
   460       case True
       
   461       with max_preced show ?thesis by auto
       
   462     qed
       
   463   qed
       
   464   finally show ?thesis .
   442   finally show ?thesis .
   465 qed
   443 qed
   466 
   444 
   467 lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
   445 lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
   468   by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)
   446   using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
   469 
   447 
   470 lemma th_cp_preced: "cp (t@s) th = preced th s"
   448 lemma th_cp_preced: "cp (t@s) th = preced th s"
   471   by (fold max_kept, unfold th_cp_max_preced, simp)
   449   by (fold max_kept, unfold th_cp_max_preced, simp)
   472 
   450 
   473 lemma preced_less:
   451 lemma preced_less:
   474   fixes th'
       
   475   assumes th'_in: "th' \<in> threads s"
   452   assumes th'_in: "th' \<in> threads s"
   476   and neq_th': "th' \<noteq> th"
   453   and neq_th': "th' \<noteq> th"
   477   shows "preced th' s < preced th s"
   454   shows "preced th' s < preced th s"
   478 proof -
   455   using assms
   479   have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
   456 by (metis Max.coboundedI finite_imageI highest not_le order.trans 
   480   proof(rule Max_ge)
   457     preced_linorder rev_image_eqI threads_s vat_s.finite_threads 
   481     from finite_threads [OF vt_s]
   458     vat_s.le_cp)
   482     show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
   459 
   483   next
   460 text {*
   484     from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
   461   Counting of the number of @{term "P"} and @{term "V"} operations 
   485       by simp
   462   is the cornerstone of a large number of the following proofs. 
   486   qed
   463   The reason is that this counting is quite easy to calculate and 
   487   moreover have "preced th' s \<noteq> preced th s"
   464   convenient to use in the reasoning. 
   488   proof
   465 
   489     assume "preced th' s = preced th s"
   466   The following lemma shows that the counting controls whether 
   490     from preced_unique[OF this th'_in] neq_th' threads_s
   467   a thread is running or not.
   491     show "False" by  (auto simp:readys_def)
   468 *}
   492   qed
       
   493   ultimately show ?thesis using highest_preced_thread
       
   494     by auto
       
   495 qed
       
   496 
   469 
   497 lemma pv_blocked_pre:
   470 lemma pv_blocked_pre:
   498   fixes th'
       
   499   assumes th'_in: "th' \<in> threads (t@s)"
   471   assumes th'_in: "th' \<in> threads (t@s)"
   500   and neq_th': "th' \<noteq> th"
   472   and neq_th': "th' \<noteq> th"
   501   and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
   473   and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
   502   shows "th' \<notin> runing (t@s)"
   474   shows "th' \<notin> runing (t@s)"
   503 proof
   475 proof
   504   assume "th' \<in> runing (t@s)"
   476   assume otherwise: "th' \<in> runing (t@s)"
   505   hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
       
   506     by (auto simp:runing_def)
       
   507   with max_cp_readys_threads [OF vt_t]
       
   508   have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
       
   509     by auto
       
   510   moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
       
   511   ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
       
   512   moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
       
   513     by simp
       
   514   finally have h: "cp (t @ s) th' = preced th (t @ s)" .
       
   515   show False
   477   show False
   516   proof -
   478   proof -
   517     have "dependants (wq (t @ s)) th' = {}" 
   479     have "th' = th"
   518       by (rule count_eq_dependants [OF vt_t eq_pv])
   480     proof(rule preced_unique)
   519     moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
   481       show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
   520     proof
   482       proof -
   521       assume "preced th' (t @ s) = preced th (t @ s)"
   483         have "?L = cp (t@s) th'"
   522       hence "th' = th"
   484           by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
   523       proof(rule preced_unique)
   485         also have "... = cp (t @ s) th" using otherwise 
   524         from th_kept show "th \<in> threads (t @ s)" by simp
   486           by (metis (mono_tags, lifting) mem_Collect_eq 
   525       next
   487                     runing_def th_cp_max vat_t.max_cp_readys_threads)
   526         from th'_in show "th' \<in> threads (t @ s)" by simp
   488         also have "... = ?R" by (metis th_cp_preced th_kept) 
   527       qed
   489         finally show ?thesis .
   528       with assms show False by simp
   490       qed
   529     qed
   491     qed (auto simp: th'_in th_kept)
   530     ultimately show ?thesis
   492     moreover have "th' \<noteq> th" using neq_th' .
   531       by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
   493     ultimately show ?thesis by simp
   532   qed
   494  qed
   533 qed
   495 qed
   534 
   496 
   535 lemmas pv_blocked = pv_blocked_pre[folded detached_eq [OF vt_t]]
   497 lemmas pv_blocked = pv_blocked_pre[folded detached_eq]
   536 
   498 
   537 lemma runing_precond_pre:
   499 lemma runing_precond_pre:
   538   fixes th'
   500   fixes th'
   539   assumes th'_in: "th' \<in> threads s"
   501   assumes th'_in: "th' \<in> threads s"
   540   and eq_pv: "cntP s th' = cntV s th'"
   502   and eq_pv: "cntP s th' = cntV s th'"
   541   and neq_th': "th' \<noteq> th"
   503   and neq_th': "th' \<noteq> th"
   542   shows "th' \<in> threads (t@s) \<and>
   504   shows "th' \<in> threads (t@s) \<and>
   543          cntP (t@s) th' = cntV (t@s) th'"
   505          cntP (t@s) th' = cntV (t@s) th'"
   544 proof -
   506 proof(induct rule:ind)
   545   show ?thesis
   507   case (Cons e t)
   546   proof(induct rule:ind)
   508     interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
   547     case (Cons e t)
   509     interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
   548     from Cons
       
   549     have in_thread: "th' \<in> threads (t @ s)"
       
   550       and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
       
   551     from Cons have "extend_highest_gen s th prio tm t" by auto
       
   552     then have not_runing: "th' \<notin> runing (t @ s)" 
       
   553       apply(rule extend_highest_gen.pv_blocked) 
       
   554       using Cons(1) in_thread neq_th' not_holding
       
   555       apply(simp_all add: detached_eq)
       
   556       done
       
   557     show ?case
   510     show ?case
   558     proof(cases e)
   511     proof(cases e)
   559       case (V thread cs)
   512       case (P thread cs)
   560       from Cons and V have vt_v: "vt (V thread cs#(t@s))" by auto
       
   561 
       
   562       show ?thesis
   513       show ?thesis
   563       proof -
   514       proof -
   564         from Cons and V have "step (t@s) (V thread cs)" by auto
   515         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
   565         hence neq_th': "thread \<noteq> th'"
   516         proof -
   566         proof(cases)
   517           have "thread \<noteq> th'"
   567           assume "thread \<in> runing (t@s)"
   518           proof -
   568           moreover have "th' \<notin> runing (t@s)" by fact
   519             have "step (t@s) (P thread cs)" using Cons P by auto
   569           ultimately show ?thesis by auto
   520             thus ?thesis
       
   521             proof(cases)
       
   522               assume "thread \<in> runing (t@s)"
       
   523               moreover have "th' \<notin> runing (t@s)" using Cons(5)
       
   524                 by (metis neq_th' vat_t.pv_blocked_pre) 
       
   525               ultimately show ?thesis by auto
       
   526             qed
       
   527           qed with Cons show ?thesis
       
   528             by (unfold P, simp add:cntP_def cntV_def count_def)
   570         qed
   529         qed
   571         with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
   530         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold P, simp)
   572           by (unfold V, simp add:cntP_def cntV_def count_def)
       
   573         moreover from in_thread
       
   574         have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
       
   575         ultimately show ?thesis by auto
   531         ultimately show ?thesis by auto
   576       qed
   532       qed
   577     next
   533     next
   578       case (P thread cs)
   534       case (V thread cs)
   579       from Cons and P have "step (t@s) (P thread cs)" by auto
   535       show ?thesis
   580       hence neq_th': "thread \<noteq> th'"
   536       proof -
   581       proof(cases)
   537         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
   582         assume "thread \<in> runing (t@s)"
   538         proof -
   583         moreover note not_runing
   539           have "thread \<noteq> th'"
       
   540           proof -
       
   541             have "step (t@s) (V thread cs)" using Cons V by auto
       
   542             thus ?thesis
       
   543             proof(cases)
       
   544               assume "thread \<in> runing (t@s)"
       
   545               moreover have "th' \<notin> runing (t@s)" using Cons(5)
       
   546                 by (metis neq_th' vat_t.pv_blocked_pre) 
       
   547               ultimately show ?thesis by auto
       
   548             qed
       
   549           qed with Cons show ?thesis
       
   550             by (unfold V, simp add:cntP_def cntV_def count_def)
       
   551         qed
       
   552         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold V, simp)
   584         ultimately show ?thesis by auto
   553         ultimately show ?thesis by auto
   585       qed
   554       qed
   586       with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   587         by (auto simp:cntP_def cntV_def count_def)
       
   588       moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
       
   589         by auto
       
   590       ultimately show ?thesis by auto
       
   591     next
   555     next
   592       case (Create thread prio')
   556       case (Create thread prio')
   593       from Cons and Create have "step (t@s) (Create thread prio')" by auto
   557       show ?thesis
   594       hence neq_th': "thread \<noteq> th'"
   558       proof -
   595       proof(cases)
   559         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
   596         assume "thread \<notin> threads (t @ s)"
   560         proof -
   597         moreover have "th' \<in> threads (t@s)" by fact
   561           have "thread \<noteq> th'"
       
   562           proof -
       
   563             have "step (t@s) (Create thread prio')" using Cons Create by auto
       
   564             thus ?thesis using Cons(5) by (cases, auto)
       
   565           qed with Cons show ?thesis
       
   566             by (unfold Create, simp add:cntP_def cntV_def count_def)
       
   567         qed
       
   568         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold Create, simp)
   598         ultimately show ?thesis by auto
   569         ultimately show ?thesis by auto
   599       qed
   570       qed
   600       with Cons and Create 
       
   601       have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   602         by (auto simp:cntP_def cntV_def count_def)
       
   603       moreover from Cons and Create 
       
   604       have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
       
   605       ultimately show ?thesis by auto
       
   606     next
   571     next
   607       case (Exit thread)
   572       case (Exit thread)
   608       from Cons and Exit have "step (t@s) (Exit thread)" by auto
   573       show ?thesis
   609       hence neq_th': "thread \<noteq> th'"
   574       proof -
   610       proof(cases)
   575         have neq_thread: "thread \<noteq> th'"
   611         assume "thread \<in> runing (t @ s)"
   576         proof -
   612         moreover note not_runing
   577           have "step (t@s) (Exit thread)" using Cons Exit by auto
       
   578           thus ?thesis apply (cases) using Cons(5)
       
   579                 by (metis neq_th' vat_t.pv_blocked_pre) 
       
   580         qed 
       
   581         hence "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" using Cons
       
   582             by (unfold Exit, simp add:cntP_def cntV_def count_def)
       
   583         moreover have "th' \<in> threads ((e # t) @ s)" using Cons neq_thread 
       
   584           by (unfold Exit, simp) 
   613         ultimately show ?thesis by auto
   585         ultimately show ?thesis by auto
   614       qed
   586       qed
   615       with Cons and Exit 
       
   616       have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   617         by (auto simp:cntP_def cntV_def count_def)
       
   618       moreover from Cons and Exit and neq_th' 
       
   619       have in_thread': "th' \<in> threads ((e # t) @ s)"
       
   620         by auto
       
   621       ultimately show ?thesis by auto
       
   622     next
   587     next
   623       case (Set thread prio')
   588       case (Set thread prio')
   624       with Cons
   589       with Cons
   625       show ?thesis 
   590       show ?thesis 
   626         by (auto simp:cntP_def cntV_def count_def)
   591         by (auto simp:cntP_def cntV_def count_def)
   627     qed
   592     qed
   628   next
   593 next
   629     case Nil
   594   case Nil
   630     with assms
   595   with assms
   631     show ?case by auto
   596   show ?case by auto
   632   qed
   597 qed
   633 qed
   598 
   634 
   599 text {* Changing counting balance to detachedness *}
   635 (*
   600 lemmas runing_precond_pre_dtc = runing_precond_pre
   636 lemma runing_precond:
   601          [folded vat_t.detached_eq vat_s.detached_eq]
   637   fixes th'
       
   638   assumes th'_in: "th' \<in> threads s"
       
   639   and eq_pv: "cntP s th' = cntV s th'"
       
   640   and neq_th': "th' \<noteq> th"
       
   641   shows "th' \<notin> runing (t@s)"
       
   642 proof -
       
   643   from runing_precond_pre[OF th'_in eq_pv neq_th']
       
   644   have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
       
   645   from pv_blocked[OF h1 neq_th' h2] 
       
   646   show ?thesis .
       
   647 qed
       
   648 *)
       
   649 
       
   650 lemmas runing_precond_pre_dtc = runing_precond_pre[folded detached_eq[OF vt_t] detached_eq[OF vt_s]]
       
   651 
   602 
   652 lemma runing_precond:
   603 lemma runing_precond:
   653   fixes th'
   604   fixes th'
   654   assumes th'_in: "th' \<in> threads s"
   605   assumes th'_in: "th' \<in> threads s"
   655   and neq_th': "th' \<noteq> th"
   606   and neq_th': "th' \<noteq> th"
   656   and is_runing: "th' \<in> runing (t@s)"
   607   and is_runing: "th' \<in> runing (t@s)"
   657   shows "cntP s th' > cntV s th'"
   608   shows "cntP s th' > cntV s th'"
       
   609   using assms
   658 proof -
   610 proof -
   659   have "cntP s th' \<noteq> cntV s th'"
   611   have "cntP s th' \<noteq> cntV s th'"
   660   proof
   612     by (metis is_runing neq_th' pv_blocked_pre runing_precond_pre th'_in)
   661     assume eq_pv: "cntP s th' = cntV s th'"
   613   moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
   662     from runing_precond_pre[OF th'_in eq_pv neq_th']
       
   663     have h1: "th' \<in> threads (t @ s)"  
       
   664       and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
       
   665     from pv_blocked_pre[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
       
   666     with is_runing show "False" by simp
       
   667   qed
       
   668   moreover from cnp_cnv_cncs[OF vt_s, of th'] 
       
   669   have "cntV s th' \<le> cntP s th'" by auto
       
   670   ultimately show ?thesis by auto
   614   ultimately show ?thesis by auto
   671 qed
   615 qed
   672 
   616 
   673 lemma moment_blocked_pre:
   617 lemma moment_blocked_pre:
   674   assumes neq_th': "th' \<noteq> th"
   618   assumes neq_th': "th' \<noteq> th"
   675   and th'_in: "th' \<in> threads ((moment i t)@s)"
   619   and th'_in: "th' \<in> threads ((moment i t)@s)"
   676   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
   620   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
   677   shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
   621   shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
   678          th' \<in> threads ((moment (i+j) t)@s)"
   622          th' \<in> threads ((moment (i+j) t)@s)"
   679 proof(induct j)
   623 proof -
   680   case (Suc k)
   624   interpret h_i: red_extend_highest_gen _ _ _ _ _ i
   681   show ?case
   625       by (unfold_locales)
   682   proof -
   626   interpret h_j: red_extend_highest_gen _ _ _ _ _ "i+j"
   683     { assume True: "Suc (i+k) \<le> length t"
   627       by (unfold_locales)
   684       from moment_head [OF this] 
   628   interpret h:  extend_highest_gen "((moment i t)@s)" th prio tm "moment j (restm i t)"
   685       obtain e where
   629   proof(unfold_locales)
   686         eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
   630     show "vt (moment i t @ s)" by (metis h_i.vt_t) 
   687         by blast
   631   next
   688       from red_moment[of "Suc(i+k)"]
   632     show "th \<in> threads (moment i t @ s)" by (metis h_i.th_kept)
   689       and eq_me have "extend_highest_gen s th prio tm (e # moment (i + k) t)" by simp
   633   next
   690       hence vt_e: "vt (e#(moment (i + k) t)@s)"
   634     show "preced th (moment i t @ s) = 
   691         by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
   635             Max (cp (moment i t @ s) ` threads (moment i t @ s))"
   692                           highest_gen_def, auto)
   636               by (metis h_i.th_cp_max h_i.th_cp_preced h_i.th_kept)
   693       have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
   637   next
   694       proof -
   638     show "preced th (moment i t @ s) = Prc prio tm" by (metis h_i.th_kept preced_th) 
   695         show "th' \<notin> runing (moment (i + k) t @ s)"
   639   next
   696         proof(rule extend_highest_gen.pv_blocked)
   640     show "vt (moment j (restm i t) @ moment i t @ s)"
   697           from Suc show "th' \<in> threads (moment (i + k) t @ s)"
   641       using moment_plus_split by (metis add.commute append_assoc h_j.vt_t)
   698             by simp
   642   next
   699         next
   643     fix th' prio'
   700           from neq_th' show "th' \<noteq> th" .
   644     assume "Create th' prio' \<in> set (moment j (restm i t))"
   701         next
   645     thus "prio' \<le> prio" using assms
   702           from red_moment show "extend_highest_gen s th prio tm (moment (i + k) t)" .
   646        by (metis Un_iff add.commute h_j.create_low moment_plus_split set_append)
   703         next
   647   next
   704           from Suc vt_e show "detached (moment (i + k) t @ s) th'"
   648     fix th' prio'
   705             apply(subst detached_eq)
   649     assume "Set th' prio' \<in> set (moment j (restm i t))"
   706             apply(auto intro: vt_e evt_cons)
   650     thus "th' \<noteq> th \<and> prio' \<le> prio"
   707             done
   651     by (metis Un_iff add.commute h_j.set_diff_low moment_plus_split set_append)
   708         qed
   652   next
   709       qed
   653     fix th'
   710       from step_back_step[OF vt_e]
   654     assume "Exit th' \<in> set (moment j (restm i t))"
   711       have "step ((moment (i + k) t)@s) e" .
   655     thus "th' \<noteq> th"
   712       hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
   656       by (metis Un_iff add.commute h_j.exit_diff moment_plus_split set_append)
   713         th' \<in> threads (e#(moment (i + k) t)@s)"
       
   714       proof(cases)
       
   715         case (thread_create thread prio)
       
   716         with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   717       next
       
   718         case (thread_exit thread)
       
   719         moreover have "thread \<noteq> th'"
       
   720         proof -
       
   721           have "thread \<in> runing (moment (i + k) t @ s)" by fact
       
   722           moreover note not_runing'
       
   723           ultimately show ?thesis by auto
       
   724         qed
       
   725         moreover note Suc 
       
   726         ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   727       next
       
   728         case (thread_P thread cs)
       
   729         moreover have "thread \<noteq> th'"
       
   730         proof -
       
   731           have "thread \<in> runing (moment (i + k) t @ s)" by fact
       
   732           moreover note not_runing'
       
   733           ultimately show ?thesis by auto
       
   734         qed
       
   735         moreover note Suc 
       
   736         ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   737       next
       
   738         case (thread_V thread cs)
       
   739         moreover have "thread \<noteq> th'"
       
   740         proof -
       
   741           have "thread \<in> runing (moment (i + k) t @ s)" by fact
       
   742           moreover note not_runing'
       
   743           ultimately show ?thesis by auto
       
   744         qed
       
   745         moreover note Suc 
       
   746         ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   747       next
       
   748         case (thread_set thread prio')
       
   749         with Suc show ?thesis
       
   750           by (auto simp:cntP_def cntV_def count_def)
       
   751       qed
       
   752       with eq_me have ?thesis using eq_me by auto 
       
   753     } note h = this
       
   754     show ?thesis
       
   755     proof(cases "Suc (i+k) \<le> length t")
       
   756       case True
       
   757       from h [OF this] show ?thesis .
       
   758     next
       
   759       case False
       
   760       with moment_ge
       
   761       have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
       
   762       with Suc show ?thesis by auto
       
   763     qed
       
   764   qed
   657   qed
   765 next
   658   show ?thesis 
   766   case 0
   659     by (metis add.commute append_assoc eq_pv h.runing_precond_pre
   767   from assms show ?case by auto
   660           moment_plus_split neq_th' th'_in)
   768 qed
   661 qed
   769 
   662 
   770 lemma moment_blocked_eqpv:
   663 lemma moment_blocked_eqpv:
   771   assumes neq_th': "th' \<noteq> th"
   664   assumes neq_th': "th' \<noteq> th"
   772   and th'_in: "th' \<in> threads ((moment i t)@s)"
   665   and th'_in: "th' \<in> threads ((moment i t)@s)"
   776          th' \<in> threads ((moment j t)@s) \<and>
   669          th' \<in> threads ((moment j t)@s) \<and>
   777          th' \<notin> runing ((moment j t)@s)"
   670          th' \<notin> runing ((moment j t)@s)"
   778 proof -
   671 proof -
   779   from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
   672   from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
   780   have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
   673   have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
   781     and h2: "th' \<in> threads ((moment j t)@s)" by auto
   674    and h2: "th' \<in> threads ((moment j t)@s)" by auto
   782   with extend_highest_gen.pv_blocked 
   675   moreover have "th' \<notin> runing ((moment j t)@s)"
   783   show ?thesis 
   676   proof -
   784     using  red_moment [of j] h2 neq_th' h1
   677     interpret h: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)
   785     apply(auto)
   678     show ?thesis
   786     by (metis extend_highest_gen.pv_blocked_pre)
   679       using h.pv_blocked_pre h1 h2 neq_th' by auto 
   787 qed
   680   qed
   788 
   681   ultimately show ?thesis by auto
       
   682 qed
       
   683 
       
   684 (* The foregoing two lemmas are preparation for this one, but
       
   685    in long run can be combined. Maybe I am wrong.
       
   686 *)
   789 lemma moment_blocked:
   687 lemma moment_blocked:
   790   assumes neq_th': "th' \<noteq> th"
   688   assumes neq_th': "th' \<noteq> th"
   791   and th'_in: "th' \<in> threads ((moment i t)@s)"
   689   and th'_in: "th' \<in> threads ((moment i t)@s)"
   792   and dtc: "detached (moment i t @ s) th'"
   690   and dtc: "detached (moment i t @ s) th'"
   793   and le_ij: "i \<le> j"
   691   and le_ij: "i \<le> j"
   794   shows "detached (moment j t @ s) th' \<and>
   692   shows "detached (moment j t @ s) th' \<and>
   795          th' \<in> threads ((moment j t)@s) \<and>
   693          th' \<in> threads ((moment j t)@s) \<and>
   796          th' \<notin> runing ((moment j t)@s)"
   694          th' \<notin> runing ((moment j t)@s)"
   797 proof -
   695 proof -
   798   from vt_moment[OF vt_t, of "i+length s"] moment_prefix[of i t s]
   696   interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
   799   have vt_i: "vt (moment i t @ s)" by auto
   697   interpret h_j: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales) 
   800   from vt_moment[OF vt_t, of "j+length s"] moment_prefix[of j t s]
   698   have cnt_i: "cntP (moment i t @ s) th' = cntV (moment i t @ s) th'"
   801   have vt_j: "vt  (moment j t @ s)" by auto
   699                 by (metis dtc h_i.detached_elim)
   802   from moment_blocked_eqpv [OF neq_th' th'_in detached_elim [OF vt_i dtc] le_ij, 
   700   from moment_blocked_eqpv[OF neq_th' th'_in cnt_i le_ij]
   803   folded detached_eq[OF vt_j]]
   701   show ?thesis by (metis h_j.detached_intro) 
   804   show ?thesis .
   702 qed
   805 qed
   703 
   806 
   704 lemma runing_preced_inversion:
   807 lemma runing_inversion_1:
   705   assumes runing': "th' \<in> runing (t@s)"
       
   706   shows "cp (t@s) th' = preced th s" (is "?L = ?R")
       
   707 proof -
       
   708   have "?L = Max (cp (t @ s) ` readys (t @ s))" using assms
       
   709       by (unfold runing_def, auto)
       
   710   also have "\<dots> = ?R"
       
   711       by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads) 
       
   712   finally show ?thesis .
       
   713 qed
       
   714 
       
   715 text {*
       
   716   The situation when @{term "th"} is blocked is analyzed by the following lemmas.
       
   717 *}
       
   718 
       
   719 text {*
       
   720   The following lemmas shows the running thread @{text "th'"}, if it is different from
       
   721   @{term th}, must be live at the very beginning. By the term {\em the very beginning},
       
   722   we mean the moment where the formal investigation starts, i.e. the moment (or state)
       
   723   @{term s}. 
       
   724 *}
       
   725 
       
   726 lemma runing_inversion_0:
   808   assumes neq_th': "th' \<noteq> th"
   727   assumes neq_th': "th' \<noteq> th"
   809   and runing': "th' \<in> runing (t@s)"
   728   and runing': "th' \<in> runing (t@s)"
       
   729   shows "th' \<in> threads s"
       
   730 proof -
       
   731     -- {* The proof is by contradiction: *}
       
   732     { assume otherwise: "\<not> ?thesis"
       
   733       have "th' \<notin> runing (t @ s)"
       
   734       proof -
       
   735         -- {* Since @{term "th'"} is running at time @{term "t@s"}, so it exists that time. *}
       
   736         have th'_in: "th' \<in> threads (t@s)" using runing' by (simp add:runing_def readys_def)
       
   737         -- {* However, @{text "th'"} does not exist at very beginning. *}
       
   738         have th'_notin: "th' \<notin> threads (moment 0 t @ s)" using otherwise
       
   739           by (metis append.simps(1) moment_zero)
       
   740         -- {* Therefore, there must be a moment during @{text "t"}, when 
       
   741               @{text "th'"} came into being. *}
       
   742         -- {* Let us suppose the moment being @{text "i"}: *}
       
   743         from p_split_gen[OF th'_in th'_notin]
       
   744         obtain i where lt_its: "i < length t"
       
   745                  and le_i: "0 \<le> i"
       
   746                  and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
       
   747                  and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by (auto)
       
   748         interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
       
   749         interpret h_i': red_extend_highest_gen _ _ _ _ _ "(Suc i)" by (unfold_locales)
       
   750         from lt_its have "Suc i \<le> length t" by auto
       
   751         -- {* Let us also suppose the event which makes this change is @{text e}: *}
       
   752         from moment_head[OF this] obtain e where 
       
   753           eq_me: "moment (Suc i) t = e # moment i t" by blast
       
   754         hence "vt (e # (moment i t @ s))" by (metis append_Cons h_i'.vt_t) 
       
   755         hence "PIP (moment i t @ s) e" by (cases, simp)
       
   756         -- {* It can be derived that this event @{text "e"}, which 
       
   757               gives birth to @{term "th'"} must be a @{term "Create"}: *}
       
   758         from create_pre[OF this, of th']
       
   759         obtain prio where eq_e: "e = Create th' prio"
       
   760             by (metis append_Cons eq_me lessI post pre) 
       
   761         have h1: "th' \<in> threads (moment (Suc i) t @ s)" using post by auto 
       
   762         have h2: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
       
   763         proof -
       
   764           have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
       
   765             by (metis h_i.cnp_cnv_eq pre)
       
   766           thus ?thesis by (simp add:eq_me eq_e cntP_def cntV_def count_def)
       
   767         qed
       
   768         show ?thesis 
       
   769           using moment_blocked_eqpv [OF neq_th' h1 h2, of "length t"] lt_its moment_ge
       
   770             by auto
       
   771       qed
       
   772       with `th' \<in> runing (t@s)`
       
   773       have False by simp
       
   774     } thus ?thesis by auto
       
   775 qed
       
   776 
       
   777 text {* 
       
   778   The second lemma says, if the running thread @{text th'} is different from 
       
   779   @{term th}, then this @{text th'} must in the possession of some resources
       
   780   at the very beginning. 
       
   781 
       
   782   To ease the reasoning of resource possession of one particular thread, 
       
   783   we used two auxiliary functions @{term cntV} and @{term cntP}, 
       
   784   which are the counters of @{term P}-operations and 
       
   785   @{term V}-operations respectively. 
       
   786   If the number of @{term V}-operation is less than the number of 
       
   787   @{term "P"}-operations, the thread must have some unreleased resource. 
       
   788 *}
       
   789 
       
   790 lemma runing_inversion_1: (* ddd *)
       
   791   assumes neq_th': "th' \<noteq> th"
       
   792   and runing': "th' \<in> runing (t@s)"
       
   793   -- {* thread @{term "th'"} is a live on in state @{term "s"} and 
       
   794         it has some unreleased resource. *}
   810   shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
   795   shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
   811 proof(cases "th' \<in> threads s")
   796 proof -
   812   case True
   797   -- {* The proof is a simple composition of @{thm runing_inversion_0} and 
       
   798         @{thm runing_precond}: *}
       
   799   -- {* By applying @{thm runing_inversion_0} to assumptions,
       
   800         it can be shown that @{term th'} is live in state @{term s}: *}
       
   801   have "th' \<in> threads s"  using runing_inversion_0[OF assms(1,2)] .
       
   802   -- {* Then the thesis is derived easily by applying @{thm runing_precond}: *}
   813   with runing_precond [OF this neq_th' runing'] show ?thesis by simp
   803   with runing_precond [OF this neq_th' runing'] show ?thesis by simp
   814 next
   804 qed
   815   case False
   805 
   816   let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
   806 text {* 
   817   let ?q = "moment 0 t"
   807   The following lemma is just a rephrasing of @{thm runing_inversion_1}:
   818   from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
   808 *}
   819   from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
       
   820   from p_split_gen [of ?Q, OF this not_thread]
       
   821   obtain i where lt_its: "i < length t"
       
   822     and le_i: "0 \<le> i"
       
   823     and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
       
   824     and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
       
   825   from lt_its have "Suc i \<le> length t" by auto
       
   826   from moment_head[OF this] obtain e where 
       
   827    eq_me: "moment (Suc i) t = e # moment i t" by blast
       
   828   from red_moment[of "Suc i"] and eq_me
       
   829   have "extend_highest_gen s th prio tm (e # moment i t)" by simp
       
   830   hence vt_e: "vt (e#(moment i t)@s)"
       
   831     by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
       
   832       highest_gen_def, auto)
       
   833   from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
       
   834   from post[rule_format, of "Suc i"] and eq_me 
       
   835   have not_in': "th' \<in> threads (e # moment i t@s)" by auto
       
   836   from create_pre[OF stp_i pre this] 
       
   837   obtain prio where eq_e: "e = Create th' prio" .
       
   838   have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
       
   839   proof(rule cnp_cnv_eq)
       
   840     from step_back_vt [OF vt_e] 
       
   841     show "vt (moment i t @ s)" .
       
   842   next
       
   843     from eq_e and stp_i 
       
   844     have "step (moment i t @ s) (Create th' prio)" by simp
       
   845     thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
       
   846   qed
       
   847   with eq_e
       
   848   have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
       
   849     by (simp add:cntP_def cntV_def count_def)
       
   850   with eq_me[symmetric]
       
   851   have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
       
   852     by simp
       
   853   from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
       
   854   with eq_me [symmetric]
       
   855   have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
       
   856   from moment_blocked_eqpv [OF neq_th' h2 h1, of "length t"] and lt_its
       
   857   and moment_ge
       
   858   have "th' \<notin> runing (t @ s)" by auto
       
   859   with runing'
       
   860   show ?thesis by auto
       
   861 qed
       
   862 
       
   863 lemma runing_inversion_2:
   809 lemma runing_inversion_2:
   864   assumes runing': "th' \<in> runing (t@s)"
   810   assumes runing': "th' \<in> runing (t@s)"
   865   shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
   811   shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
   866 proof -
   812 proof -
   867   from runing_inversion_1[OF _ runing']
   813   from runing_inversion_1[OF _ runing']
   868   show ?thesis by auto
   814   show ?thesis by auto
   869 qed
       
   870 
       
   871 lemma runing_preced_inversion:
       
   872   assumes runing': "th' \<in> runing (t@s)"
       
   873   shows "cp (t@s) th' = preced th s"
       
   874 proof -
       
   875   from runing' have "cp (t@s) th' = Max (cp (t @ s) ` readys (t @ s))"
       
   876     by (unfold runing_def, auto)
       
   877   also have "\<dots> = preced th s"
       
   878   proof -
       
   879     from max_cp_readys_threads[OF vt_t]
       
   880     have "\<dots> =  Max (cp (t @ s) ` threads (t @ s))" .
       
   881     also have "\<dots> = preced th s"
       
   882     proof -
       
   883       from max_kept
       
   884       and max_cp_eq [OF vt_t]
       
   885       show ?thesis by auto
       
   886     qed
       
   887     finally show ?thesis .
       
   888   qed
       
   889   finally show ?thesis .
       
   890 qed
   815 qed
   891 
   816 
   892 lemma runing_inversion_3:
   817 lemma runing_inversion_3:
   893   assumes runing': "th' \<in> runing (t@s)"
   818   assumes runing': "th' \<in> runing (t@s)"
   894   and neq_th: "th' \<noteq> th"
   819   and neq_th: "th' \<noteq> th"
   895   shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
   820   shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
   896 proof -
   821   by (metis neq_th runing' runing_inversion_2 runing_preced_inversion)
   897   from runing_inversion_2 [OF runing'] 
       
   898     and neq_th 
       
   899     and runing_preced_inversion[OF runing']
       
   900   show ?thesis by auto
       
   901 qed
       
   902 
   822 
   903 lemma runing_inversion_4:
   823 lemma runing_inversion_4:
   904   assumes runing': "th' \<in> runing (t@s)"
   824   assumes runing': "th' \<in> runing (t@s)"
   905   and neq_th: "th' \<noteq> th"
   825   and neq_th: "th' \<noteq> th"
   906   shows "th' \<in> threads s"
   826   shows "th' \<in> threads s"
   907   and    "\<not>detached s th'"
   827   and    "\<not>detached s th'"
   908   and    "cp (t@s) th' = preced th s"
   828   and    "cp (t@s) th' = preced th s"
   909 using runing_inversion_3 [OF runing'] 
   829   apply (metis neq_th runing' runing_inversion_2)
   910   and neq_th 
   830   apply (metis neq_th pv_blocked runing' runing_inversion_2 runing_precond_pre_dtc)
   911   and runing_preced_inversion[OF runing']
   831   by (metis neq_th runing' runing_inversion_3)
   912 apply(auto simp add: detached_eq[OF vt_s])
   832 
   913 done
   833 
   914 
   834 text {* 
       
   835   Suppose @{term th} is not running, it is first shown that
       
   836   there is a path in RAG leading from node @{term th} to another thread @{text "th'"} 
       
   837   in the @{term readys}-set (So @{text "th'"} is an ancestor of @{term th}}).
       
   838 
       
   839   Now, since @{term readys}-set is non-empty, there must be
       
   840   one in it which holds the highest @{term cp}-value, which, by definition, 
       
   841   is the @{term runing}-thread. However, we are going to show more: this running thread
       
   842   is exactly @{term "th'"}.
       
   843      *}
       
   844 lemma th_blockedE: (* ddd *)
       
   845   assumes "th \<notin> runing (t@s)"
       
   846   obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
       
   847                     "th' \<in> runing (t@s)"
       
   848 proof -
       
   849   -- {* According to @{thm vat_t.th_chain_to_ready}, either 
       
   850         @{term "th"} is in @{term "readys"} or there is path leading from it to 
       
   851         one thread in @{term "readys"}. *}
       
   852   have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)" 
       
   853     using th_kept vat_t.th_chain_to_ready by auto
       
   854   -- {* However, @{term th} can not be in @{term readys}, because otherwise, since 
       
   855        @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
       
   856   moreover have "th \<notin> readys (t@s)" 
       
   857     using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto 
       
   858   -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in 
       
   859         term @{term readys}: *}
       
   860   ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
       
   861                           and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
       
   862   -- {* We are going to show that this @{term th'} is running. *}
       
   863   have "th' \<in> runing (t@s)"
       
   864   proof -
       
   865     -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
       
   866     have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
       
   867     proof -
       
   868       have "?L =  Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
       
   869         by (unfold cp_alt_def1, simp)
       
   870       also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
       
   871       proof(rule image_Max_subset)
       
   872         show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
       
   873       next
       
   874         show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
       
   875           by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread) 
       
   876       next
       
   877         show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
       
   878                     by (unfold tRAG_subtree_eq, auto simp:subtree_def)
       
   879       next
       
   880         show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
       
   881                       (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
       
   882         proof -
       
   883           have "?L = the_preced (t @ s) `  threads (t @ s)" 
       
   884                      by (unfold image_comp, rule image_cong, auto)
       
   885           thus ?thesis using max_preced the_preced_def by auto
       
   886         qed
       
   887       qed
       
   888       also have "... = ?R"
       
   889         using th_cp_max th_cp_preced th_kept 
       
   890               the_preced_def vat_t.max_cp_readys_threads by auto
       
   891       finally show ?thesis .
       
   892     qed 
       
   893     -- {* Now, since @{term th'} holds the highest @{term cp} 
       
   894           and we have already show it is in @{term readys},
       
   895           it is @{term runing} by definition. *}
       
   896     with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def) 
       
   897   qed
       
   898   -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
       
   899   moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)" 
       
   900     using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
       
   901   ultimately show ?thesis using that by metis
       
   902 qed
       
   903 
       
   904 text {*
       
   905   Now it is easy to see there is always a thread to run by case analysis
       
   906   on whether thread @{term th} is running: if the answer is Yes, the 
       
   907   the running thread is obviously @{term th} itself; otherwise, the running
       
   908   thread is the @{text th'} given by lemma @{thm th_blockedE}.
       
   909 *}
   915 lemma live: "runing (t@s) \<noteq> {}"
   910 lemma live: "runing (t@s) \<noteq> {}"
   916 proof(cases "th \<in> runing (t@s)")
   911 proof(cases "th \<in> runing (t@s)") 
   917   case True thus ?thesis by auto
   912   case True thus ?thesis by auto
   918 next
   913 next
   919   case False
   914   case False
   920   then have not_ready: "th \<notin> readys (t@s)"
   915   thus ?thesis using th_blockedE by auto
   921     apply (unfold runing_def, 
       
   922             insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
       
   923     by auto
       
   924   from th_kept have "th \<in> threads (t@s)" by auto
       
   925   from th_chain_to_ready[OF vt_t this] and not_ready
       
   926   obtain th' where th'_in: "th' \<in> readys (t@s)"
       
   927     and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
       
   928   have "th' \<in> runing (t@s)"
       
   929   proof -
       
   930     have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
       
   931     proof -
       
   932       have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependants (wq (t @ s)) th')) = 
       
   933                preced th (t@s)"
       
   934       proof(rule Max_eqI)
       
   935         fix y
       
   936         assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependants (wq (t @ s)) th')"
       
   937         then obtain th1 where
       
   938           h1: "th1 = th' \<or> th1 \<in>  dependants (wq (t @ s)) th'"
       
   939           and eq_y: "y = preced th1 (t@s)" by auto
       
   940         show "y \<le> preced th (t @ s)"
       
   941         proof -
       
   942           from max_preced
       
   943           have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
       
   944           moreover have "y \<le> \<dots>"
       
   945           proof(rule Max_ge)
       
   946             from h1
       
   947             have "th1 \<in> threads (t@s)"
       
   948             proof
       
   949               assume "th1 = th'"
       
   950               with th'_in show ?thesis by (simp add:readys_def)
       
   951             next
       
   952               assume "th1 \<in> dependants (wq (t @ s)) th'"
       
   953               with dependants_threads [OF vt_t]
       
   954               show "th1 \<in> threads (t @ s)" by auto
       
   955             qed
       
   956             with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
       
   957           next
       
   958             from finite_threads[OF vt_t]
       
   959             show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
       
   960           qed
       
   961           ultimately show ?thesis by auto
       
   962         qed
       
   963       next
       
   964         from finite_threads[OF vt_t] dependants_threads [OF vt_t, of th']
       
   965         show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependants (wq (t @ s)) th'))"
       
   966           by (auto intro:finite_subset)
       
   967       next
       
   968         from dp
       
   969         have "th \<in> dependants (wq (t @ s)) th'" 
       
   970           by (unfold cs_dependants_def, auto simp:eq_RAG)
       
   971         thus "preced th (t @ s) \<in> 
       
   972                 (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependants (wq (t @ s)) th')"
       
   973           by auto
       
   974       qed
       
   975       moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
       
   976       proof -
       
   977         from max_preced and max_cp_eq[OF vt_t, symmetric]
       
   978         have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
       
   979         with max_cp_readys_threads[OF vt_t] show ?thesis by simp
       
   980       qed
       
   981       ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
       
   982     qed
       
   983     with th'_in show ?thesis by (auto simp:runing_def)
       
   984   qed
       
   985   thus ?thesis by auto
       
   986 qed
   916 qed
   987 
   917 
   988 end
   918 end
   989 end
   919 end
   990 
   920