Implementation.thy
changeset 93 524bd3caa6b6
parent 92 4763aa246dbd
child 97 c7ba70dc49bd
child 102 3a801bbd2687
equal deleted inserted replaced
92:4763aa246dbd 93:524bd3caa6b6
     1 section {*
     1 section {*
     2   This file contains lemmas used to guide the recalculation of current precedence 
     2   This file contains lemmas used to guide the recalculation of current precedence 
     3   after every system call (or system operation)
     3   after every system call (or system operation)
     4 *}
     4 *}
     5 theory ExtGG
     5 theory Implementation
     6 imports CpsG
     6 imports PIPBasics
     7 begin
     7 begin
     8 
     8 
     9 text {* (* ddd *)
     9 text {* (* ddd *)
    10   One beauty of our modelling is that we follow the definitional extension tradition of HOL.
    10   One beauty of our modelling is that we follow the definitional extension tradition of HOL.
    11   The benefit of such a concise and miniature model is that  large number of intuitively 
    11   The benefit of such a concise and miniature model is that  large number of intuitively 
   704 
   704 
   705 end
   705 end
   706 
   706 
   707 end
   707 end
   708 
   708 
   709 =======
   709 
   710 theory ExtGG
       
   711 imports PrioG CpsG
       
   712 begin
       
   713 
       
   714 text {* 
       
   715   The following two auxiliary lemmas are used to reason about @{term Max}.
       
   716 *}
       
   717 lemma image_Max_eqI: 
       
   718   assumes "finite B"
       
   719   and "b \<in> B"
       
   720   and "\<forall> x \<in> B. f x \<le> f b"
       
   721   shows "Max (f ` B) = f b"
       
   722   using assms
       
   723   using Max_eqI by blast 
       
   724 
       
   725 lemma image_Max_subset:
       
   726   assumes "finite A"
       
   727   and "B \<subseteq> A"
       
   728   and "a \<in> B"
       
   729   and "Max (f ` A) = f a"
       
   730   shows "Max (f ` B) = f a"
       
   731 proof(rule image_Max_eqI)
       
   732   show "finite B"
       
   733     using assms(1) assms(2) finite_subset by auto 
       
   734 next
       
   735   show "a \<in> B" using assms by simp
       
   736 next
       
   737   show "\<forall>x\<in>B. f x \<le> f a"
       
   738     by (metis Max_ge assms(1) assms(2) assms(4) 
       
   739             finite_imageI image_eqI subsetCE) 
       
   740 qed
       
   741 
       
   742 text {*
       
   743   The following locale @{text "highest_gen"} sets the basic context for our
       
   744   investigation: supposing thread @{text th} holds the highest @{term cp}-value
       
   745   in state @{text s}, which means the task for @{text th} is the 
       
   746   most urgent. We want to show that  
       
   747   @{text th} is treated correctly by PIP, which means
       
   748   @{text th} will not be blocked unreasonably by other less urgent
       
   749   threads. 
       
   750 *}
       
   751 locale highest_gen =
       
   752   fixes s th prio tm
       
   753   assumes vt_s: "vt s"
       
   754   and threads_s: "th \<in> threads s"
       
   755   and highest: "preced th s = Max ((cp s)`threads s)"
       
   756   -- {* The internal structure of @{term th}'s precedence is exposed:*}
       
   757   and preced_th: "preced th s = Prc prio tm" 
       
   758 
       
   759 -- {* @{term s} is a valid trace, so it will inherit all results derived for
       
   760       a valid trace: *}
       
   761 sublocale highest_gen < vat_s: valid_trace "s"
       
   762   by (unfold_locales, insert vt_s, simp)
       
   763 
       
   764 context highest_gen
       
   765 begin
       
   766 
       
   767 text {*
       
   768   @{term tm} is the time when the precedence of @{term th} is set, so 
       
   769   @{term tm} must be a valid moment index into @{term s}.
       
   770 *}
       
   771 lemma lt_tm: "tm < length s"
       
   772   by (insert preced_tm_lt[OF threads_s preced_th], simp)
       
   773 
       
   774 text {*
       
   775   Since @{term th} holds the highest precedence and @{text "cp"}
       
   776   is the highest precedence of all threads in the sub-tree of 
       
   777   @{text "th"} and @{text th} is among these threads, 
       
   778   its @{term cp} must equal to its precedence:
       
   779 *}
       
   780 lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
       
   781 proof -
       
   782   have "?L \<le> ?R"
       
   783   by (unfold highest, rule Max_ge, 
       
   784         auto simp:threads_s finite_threads)
       
   785   moreover have "?R \<le> ?L"
       
   786     by (unfold vat_s.cp_rec, rule Max_ge, 
       
   787         auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
       
   788   ultimately show ?thesis by auto
       
   789 qed
       
   790 
       
   791 (* ccc *)
       
   792 lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
       
   793   by (fold max_cp_eq, unfold eq_cp_s_th, insert highest, simp)
       
   794 
       
   795 lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
       
   796   by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
       
   797 
       
   798 lemma highest': "cp s th = Max (cp s ` threads s)"
       
   799 proof -
       
   800   from highest_cp_preced max_cp_eq[symmetric]
       
   801   show ?thesis by simp
       
   802 qed
       
   803 
       
   804 end
       
   805 
       
   806 locale extend_highest_gen = highest_gen + 
       
   807   fixes t 
       
   808   assumes vt_t: "vt (t@s)"
       
   809   and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
       
   810   and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
       
   811   and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
       
   812 
       
   813 sublocale extend_highest_gen < vat_t: valid_trace "t@s"
       
   814   by (unfold_locales, insert vt_t, simp)
       
   815 
       
   816 lemma step_back_vt_app: 
       
   817   assumes vt_ts: "vt (t@s)" 
       
   818   shows "vt s"
       
   819 proof -
       
   820   from vt_ts show ?thesis
       
   821   proof(induct t)
       
   822     case Nil
       
   823     from Nil show ?case by auto
       
   824   next
       
   825     case (Cons e t)
       
   826     assume ih: " vt (t @ s) \<Longrightarrow> vt s"
       
   827       and vt_et: "vt ((e # t) @ s)"
       
   828     show ?case
       
   829     proof(rule ih)
       
   830       show "vt (t @ s)"
       
   831       proof(rule step_back_vt)
       
   832         from vt_et show "vt (e # t @ s)" by simp
       
   833       qed
       
   834     qed
       
   835   qed
       
   836 qed
       
   837 
       
   838 
       
   839 locale red_extend_highest_gen = extend_highest_gen +
       
   840    fixes i::nat
       
   841 
       
   842 sublocale red_extend_highest_gen <   red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
       
   843   apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
       
   844   apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
       
   845   by (unfold highest_gen_def, auto dest:step_back_vt_app)
       
   846 
       
   847 
       
   848 context extend_highest_gen
       
   849 begin
       
   850 
       
   851  lemma ind [consumes 0, case_names Nil Cons, induct type]:
       
   852   assumes 
       
   853     h0: "R []"
       
   854   and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
       
   855                     extend_highest_gen s th prio tm t; 
       
   856                     extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
       
   857   shows "R t"
       
   858 proof -
       
   859   from vt_t extend_highest_gen_axioms show ?thesis
       
   860   proof(induct t)
       
   861     from h0 show "R []" .
       
   862   next
       
   863     case (Cons e t')
       
   864     assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
       
   865       and vt_e: "vt ((e # t') @ s)"
       
   866       and et: "extend_highest_gen s th prio tm (e # t')"
       
   867     from vt_e and step_back_step have stp: "step (t'@s) e" by auto
       
   868     from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
       
   869     show ?case
       
   870     proof(rule h2 [OF vt_ts stp _ _ _ ])
       
   871       show "R t'"
       
   872       proof(rule ih)
       
   873         from et show ext': "extend_highest_gen s th prio tm t'"
       
   874           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
       
   875       next
       
   876         from vt_ts show "vt (t' @ s)" .
       
   877       qed
       
   878     next
       
   879       from et show "extend_highest_gen s th prio tm (e # t')" .
       
   880     next
       
   881       from et show ext': "extend_highest_gen s th prio tm t'"
       
   882           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
       
   883     qed
       
   884   qed
       
   885 qed
       
   886 
       
   887 
       
   888 lemma th_kept: "th \<in> threads (t @ s) \<and> 
       
   889                  preced th (t@s) = preced th s" (is "?Q t") 
       
   890 proof -
       
   891   show ?thesis
       
   892   proof(induct rule:ind)
       
   893     case Nil
       
   894     from threads_s
       
   895     show ?case
       
   896       by auto
       
   897   next
       
   898     case (Cons e t)
       
   899     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
       
   900     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
       
   901     show ?case
       
   902     proof(cases e)
       
   903       case (Create thread prio)
       
   904       show ?thesis
       
   905       proof -
       
   906         from Cons and Create have "step (t@s) (Create thread prio)" by auto
       
   907         hence "th \<noteq> thread"
       
   908         proof(cases)
       
   909           case thread_create
       
   910           with Cons show ?thesis by auto
       
   911         qed
       
   912         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
       
   913           by (unfold Create, auto simp:preced_def)
       
   914         moreover note Cons
       
   915         ultimately show ?thesis
       
   916           by (auto simp:Create)
       
   917       qed
       
   918     next
       
   919       case (Exit thread)
       
   920       from h_e.exit_diff and Exit
       
   921       have neq_th: "thread \<noteq> th" by auto
       
   922       with Cons
       
   923       show ?thesis
       
   924         by (unfold Exit, auto simp:preced_def)
       
   925     next
       
   926       case (P thread cs)
       
   927       with Cons
       
   928       show ?thesis 
       
   929         by (auto simp:P preced_def)
       
   930     next
       
   931       case (V thread cs)
       
   932       with Cons
       
   933       show ?thesis 
       
   934         by (auto simp:V preced_def)
       
   935     next
       
   936       case (Set thread prio')
       
   937       show ?thesis
       
   938       proof -
       
   939         from h_e.set_diff_low and Set
       
   940         have "th \<noteq> thread" by auto
       
   941         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
       
   942           by (unfold Set, auto simp:preced_def)
       
   943         moreover note Cons
       
   944         ultimately show ?thesis
       
   945           by (auto simp:Set)
       
   946       qed
       
   947     qed
       
   948   qed
       
   949 qed
       
   950 
       
   951 text {*
       
   952   According to @{thm th_kept}, thread @{text "th"} has its living status
       
   953   and precedence kept along the way of @{text "t"}. The following lemma
       
   954   shows that this preserved precedence of @{text "th"} remains as the highest
       
   955   along the way of @{text "t"}.
       
   956 
       
   957   The proof goes by induction over @{text "t"} using the specialized
       
   958   induction rule @{thm ind}, followed by case analysis of each possible 
       
   959   operations of PIP. All cases follow the same pattern rendered by the 
       
   960   generalized introduction rule @{thm "image_Max_eqI"}. 
       
   961 
       
   962   The very essence is to show that precedences, no matter whether they are newly introduced 
       
   963   or modified, are always lower than the one held by @{term "th"},
       
   964   which by @{thm th_kept} is preserved along the way.
       
   965 *}
       
   966 lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
       
   967 proof(induct rule:ind)
       
   968   case Nil
       
   969   from highest_preced_thread
       
   970   show ?case
       
   971     by (unfold the_preced_def, simp)
       
   972 next
       
   973   case (Cons e t)
       
   974     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
       
   975     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
       
   976   show ?case
       
   977   proof(cases e)
       
   978     case (Create thread prio')
       
   979     show ?thesis (is "Max (?f ` ?A) = ?t")
       
   980     proof -
       
   981       -- {* The following is the common pattern of each branch of the case analysis. *}
       
   982       -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
       
   983       have "Max (?f ` ?A) = ?f th"
       
   984       proof(rule image_Max_eqI)
       
   985         show "finite ?A" using h_e.finite_threads by auto 
       
   986       next
       
   987         show "th \<in> ?A" using h_e.th_kept by auto 
       
   988       next
       
   989         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
       
   990         proof 
       
   991           fix x
       
   992           assume "x \<in> ?A"
       
   993           hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
       
   994           thus "?f x \<le> ?f th"
       
   995           proof
       
   996             assume "x = thread"
       
   997             thus ?thesis 
       
   998               apply (simp add:Create the_preced_def preced_def, fold preced_def)
       
   999               using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 preced_th by force
       
  1000           next
       
  1001             assume h: "x \<in> threads (t @ s)"
       
  1002             from Cons(2)[unfolded Create] 
       
  1003             have "x \<noteq> thread" using h by (cases, auto)
       
  1004             hence "?f x = the_preced (t@s) x" 
       
  1005               by (simp add:Create the_preced_def preced_def)
       
  1006             hence "?f x \<le> Max (the_preced (t@s) ` threads (t@s))"
       
  1007               by (simp add: h_t.finite_threads h)
       
  1008             also have "... = ?f th"
       
  1009               by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
       
  1010             finally show ?thesis .
       
  1011           qed
       
  1012         qed
       
  1013       qed
       
  1014      -- {* The minor part is to show that the precedence of @{text "th"} 
       
  1015            equals to preserved one, given by the foregoing lemma @{thm th_kept} *}
       
  1016       also have "... = ?t" using h_e.th_kept the_preced_def by auto
       
  1017       -- {* Then it follows trivially that the precedence preserved
       
  1018             for @{term "th"} remains the maximum of all living threads along the way. *}
       
  1019       finally show ?thesis .
       
  1020     qed 
       
  1021   next 
       
  1022     case (Exit thread)
       
  1023     show ?thesis (is "Max (?f ` ?A) = ?t")
       
  1024     proof -
       
  1025       have "Max (?f ` ?A) = ?f th"
       
  1026       proof(rule image_Max_eqI)
       
  1027         show "finite ?A" using h_e.finite_threads by auto 
       
  1028       next
       
  1029         show "th \<in> ?A" using h_e.th_kept by auto 
       
  1030       next
       
  1031         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
       
  1032         proof 
       
  1033           fix x
       
  1034           assume "x \<in> ?A"
       
  1035           hence "x \<in> threads (t@s)" by (simp add: Exit) 
       
  1036           hence "?f x \<le> Max (?f ` threads (t@s))" 
       
  1037             by (simp add: h_t.finite_threads) 
       
  1038           also have "... \<le> ?f th" 
       
  1039             apply (simp add:Exit the_preced_def preced_def, fold preced_def)
       
  1040             using Cons.hyps(5) h_t.th_kept the_preced_def by auto
       
  1041           finally show "?f x \<le> ?f th" .
       
  1042         qed
       
  1043       qed
       
  1044       also have "... = ?t" using h_e.th_kept the_preced_def by auto
       
  1045       finally show ?thesis .
       
  1046     qed 
       
  1047   next
       
  1048     case (P thread cs)
       
  1049     with Cons
       
  1050     show ?thesis by (auto simp:preced_def the_preced_def)
       
  1051   next
       
  1052     case (V thread cs)
       
  1053     with Cons
       
  1054     show ?thesis by (auto simp:preced_def the_preced_def)
       
  1055   next 
       
  1056     case (Set thread prio')
       
  1057     show ?thesis (is "Max (?f ` ?A) = ?t")
       
  1058     proof -
       
  1059       have "Max (?f ` ?A) = ?f th"
       
  1060       proof(rule image_Max_eqI)
       
  1061         show "finite ?A" using h_e.finite_threads by auto 
       
  1062       next
       
  1063         show "th \<in> ?A" using h_e.th_kept by auto 
       
  1064       next
       
  1065         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
       
  1066         proof 
       
  1067           fix x
       
  1068           assume h: "x \<in> ?A"
       
  1069           show "?f x \<le> ?f th"
       
  1070           proof(cases "x = thread")
       
  1071             case True
       
  1072             moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
       
  1073             proof -
       
  1074               have "the_preced (t @ s) th = Prc prio tm"  
       
  1075                 using h_t.th_kept preced_th by (simp add:the_preced_def)
       
  1076               moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
       
  1077               ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
       
  1078             qed
       
  1079             ultimately show ?thesis
       
  1080               by (unfold Set, simp add:the_preced_def preced_def)
       
  1081           next
       
  1082             case False
       
  1083             then have "?f x  = the_preced (t@s) x"
       
  1084               by (simp add:the_preced_def preced_def Set)
       
  1085             also have "... \<le> Max (the_preced (t@s) ` threads (t@s))"
       
  1086               using Set h h_t.finite_threads by auto 
       
  1087             also have "... = ?f th" by (metis Cons.hyps(5) h_e.th_kept the_preced_def) 
       
  1088             finally show ?thesis .
       
  1089           qed
       
  1090         qed
       
  1091       qed
       
  1092       also have "... = ?t" using h_e.th_kept the_preced_def by auto
       
  1093       finally show ?thesis .
       
  1094     qed 
       
  1095   qed
       
  1096 qed
       
  1097 
       
  1098 lemma max_preced: "preced th (t@s) = Max (the_preced (t@s) ` (threads (t@s)))"
       
  1099   by (insert th_kept max_kept, auto)
       
  1100 
       
  1101 text {*
       
  1102   The reason behind the following lemma is that:
       
  1103   Since @{term "cp"} is defined as the maximum precedence 
       
  1104   of those threads contained in the sub-tree of node @{term "Th th"} 
       
  1105   in @{term "RAG (t@s)"}, and all these threads are living threads, and 
       
  1106   @{term "th"} is also among them, the maximum precedence of 
       
  1107   them all must be the one for @{text "th"}.
       
  1108 *}
       
  1109 lemma th_cp_max_preced: 
       
  1110   "cp (t@s) th = Max (the_preced (t@s) ` (threads (t@s)))" (is "?L = ?R") 
       
  1111 proof -
       
  1112   let ?f = "the_preced (t@s)"
       
  1113   have "?L = ?f th"
       
  1114   proof(unfold cp_alt_def, rule image_Max_eqI)
       
  1115     show "finite {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
       
  1116     proof -
       
  1117       have "{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)} = 
       
  1118             the_thread ` {n . n \<in> subtree (RAG (t @ s)) (Th th) \<and>
       
  1119                             (\<exists> th'. n = Th th')}"
       
  1120       by (smt Collect_cong Setcompr_eq_image mem_Collect_eq the_thread.simps)
       
  1121       moreover have "finite ..." by (simp add: vat_t.fsbtRAGs.finite_subtree) 
       
  1122       ultimately show ?thesis by simp
       
  1123     qed
       
  1124   next
       
  1125     show "th \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
       
  1126       by (auto simp:subtree_def)
       
  1127   next
       
  1128     show "\<forall>x\<in>{th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}.
       
  1129                the_preced (t @ s) x \<le> the_preced (t @ s) th"
       
  1130     proof
       
  1131       fix th'
       
  1132       assume "th' \<in> {th'. Th th' \<in> subtree (RAG (t @ s)) (Th th)}"
       
  1133       hence "Th th' \<in> subtree (RAG (t @ s)) (Th th)" by auto
       
  1134       moreover have "... \<subseteq> Field (RAG (t @ s)) \<union> {Th th}"
       
  1135         by (meson subtree_Field)
       
  1136       ultimately have "Th th' \<in> ..." by auto
       
  1137       hence "th' \<in> threads (t@s)" 
       
  1138       proof
       
  1139         assume "Th th' \<in> {Th th}"
       
  1140         thus ?thesis using th_kept by auto 
       
  1141       next
       
  1142         assume "Th th' \<in> Field (RAG (t @ s))"
       
  1143         thus ?thesis using vat_t.not_in_thread_isolated by blast 
       
  1144       qed
       
  1145       thus "the_preced (t @ s) th' \<le> the_preced (t @ s) th"
       
  1146         by (metis Max_ge finite_imageI finite_threads image_eqI 
       
  1147                max_kept th_kept the_preced_def)
       
  1148     qed
       
  1149   qed
       
  1150   also have "... = ?R" by (simp add: max_preced the_preced_def) 
       
  1151   finally show ?thesis .
       
  1152 qed
       
  1153 
       
  1154 lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
       
  1155   using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
       
  1156 
       
  1157 lemma th_cp_preced: "cp (t@s) th = preced th s"
       
  1158   by (fold max_kept, unfold th_cp_max_preced, simp)
       
  1159 
       
  1160 lemma preced_less:
       
  1161   assumes th'_in: "th' \<in> threads s"
       
  1162   and neq_th': "th' \<noteq> th"
       
  1163   shows "preced th' s < preced th s"
       
  1164   using assms
       
  1165 by (metis Max.coboundedI finite_imageI highest not_le order.trans 
       
  1166     preced_linorder rev_image_eqI threads_s vat_s.finite_threads 
       
  1167     vat_s.le_cp)
       
  1168 
       
  1169 text {*
       
  1170   Counting of the number of @{term "P"} and @{term "V"} operations 
       
  1171   is the cornerstone of a large number of the following proofs. 
       
  1172   The reason is that this counting is quite easy to calculate and 
       
  1173   convenient to use in the reasoning. 
       
  1174 
       
  1175   The following lemma shows that the counting controls whether 
       
  1176   a thread is running or not.
       
  1177 *}
       
  1178 
       
  1179 lemma pv_blocked_pre:
       
  1180   assumes th'_in: "th' \<in> threads (t@s)"
       
  1181   and neq_th': "th' \<noteq> th"
       
  1182   and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
       
  1183   shows "th' \<notin> runing (t@s)"
       
  1184 proof
       
  1185   assume otherwise: "th' \<in> runing (t@s)"
       
  1186   show False
       
  1187   proof -
       
  1188     have "th' = th"
       
  1189     proof(rule preced_unique)
       
  1190       show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
       
  1191       proof -
       
  1192         have "?L = cp (t@s) th'"
       
  1193           by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
       
  1194         also have "... = cp (t @ s) th" using otherwise 
       
  1195           by (metis (mono_tags, lifting) mem_Collect_eq 
       
  1196                     runing_def th_cp_max vat_t.max_cp_readys_threads)
       
  1197         also have "... = ?R" by (metis th_cp_preced th_kept) 
       
  1198         finally show ?thesis .
       
  1199       qed
       
  1200     qed (auto simp: th'_in th_kept)
       
  1201     moreover have "th' \<noteq> th" using neq_th' .
       
  1202     ultimately show ?thesis by simp
       
  1203  qed
       
  1204 qed
       
  1205 
       
  1206 lemmas pv_blocked = pv_blocked_pre[folded detached_eq]
       
  1207 
       
  1208 lemma runing_precond_pre:
       
  1209   fixes th'
       
  1210   assumes th'_in: "th' \<in> threads s"
       
  1211   and eq_pv: "cntP s th' = cntV s th'"
       
  1212   and neq_th': "th' \<noteq> th"
       
  1213   shows "th' \<in> threads (t@s) \<and>
       
  1214          cntP (t@s) th' = cntV (t@s) th'"
       
  1215 proof(induct rule:ind)
       
  1216   case (Cons e t)
       
  1217     interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
       
  1218     interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
       
  1219     show ?case
       
  1220     proof(cases e)
       
  1221       case (P thread cs)
       
  1222       show ?thesis
       
  1223       proof -
       
  1224         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
  1225         proof -
       
  1226           have "thread \<noteq> th'"
       
  1227           proof -
       
  1228             have "step (t@s) (P thread cs)" using Cons P by auto
       
  1229             thus ?thesis
       
  1230             proof(cases)
       
  1231               assume "thread \<in> runing (t@s)"
       
  1232               moreover have "th' \<notin> runing (t@s)" using Cons(5)
       
  1233                 by (metis neq_th' vat_t.pv_blocked_pre) 
       
  1234               ultimately show ?thesis by auto
       
  1235             qed
       
  1236           qed with Cons show ?thesis
       
  1237             by (unfold P, simp add:cntP_def cntV_def count_def)
       
  1238         qed
       
  1239         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold P, simp)
       
  1240         ultimately show ?thesis by auto
       
  1241       qed
       
  1242     next
       
  1243       case (V thread cs)
       
  1244       show ?thesis
       
  1245       proof -
       
  1246         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
  1247         proof -
       
  1248           have "thread \<noteq> th'"
       
  1249           proof -
       
  1250             have "step (t@s) (V thread cs)" using Cons V by auto
       
  1251             thus ?thesis
       
  1252             proof(cases)
       
  1253               assume "thread \<in> runing (t@s)"
       
  1254               moreover have "th' \<notin> runing (t@s)" using Cons(5)
       
  1255                 by (metis neq_th' vat_t.pv_blocked_pre) 
       
  1256               ultimately show ?thesis by auto
       
  1257             qed
       
  1258           qed with Cons show ?thesis
       
  1259             by (unfold V, simp add:cntP_def cntV_def count_def)
       
  1260         qed
       
  1261         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold V, simp)
       
  1262         ultimately show ?thesis by auto
       
  1263       qed
       
  1264     next
       
  1265       case (Create thread prio')
       
  1266       show ?thesis
       
  1267       proof -
       
  1268         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
  1269         proof -
       
  1270           have "thread \<noteq> th'"
       
  1271           proof -
       
  1272             have "step (t@s) (Create thread prio')" using Cons Create by auto
       
  1273             thus ?thesis using Cons(5) by (cases, auto)
       
  1274           qed with Cons show ?thesis
       
  1275             by (unfold Create, simp add:cntP_def cntV_def count_def)
       
  1276         qed
       
  1277         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold Create, simp)
       
  1278         ultimately show ?thesis by auto
       
  1279       qed
       
  1280     next
       
  1281       case (Exit thread)
       
  1282       show ?thesis
       
  1283       proof -
       
  1284         have neq_thread: "thread \<noteq> th'"
       
  1285         proof -
       
  1286           have "step (t@s) (Exit thread)" using Cons Exit by auto
       
  1287           thus ?thesis apply (cases) using Cons(5)
       
  1288                 by (metis neq_th' vat_t.pv_blocked_pre) 
       
  1289         qed 
       
  1290         hence "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" using Cons
       
  1291             by (unfold Exit, simp add:cntP_def cntV_def count_def)
       
  1292         moreover have "th' \<in> threads ((e # t) @ s)" using Cons neq_thread 
       
  1293           by (unfold Exit, simp) 
       
  1294         ultimately show ?thesis by auto
       
  1295       qed
       
  1296     next
       
  1297       case (Set thread prio')
       
  1298       with Cons
       
  1299       show ?thesis 
       
  1300         by (auto simp:cntP_def cntV_def count_def)
       
  1301     qed
       
  1302 next
       
  1303   case Nil
       
  1304   with assms
       
  1305   show ?case by auto
       
  1306 qed
       
  1307 
       
  1308 text {* Changing counting balance to detachedness *}
       
  1309 lemmas runing_precond_pre_dtc = runing_precond_pre
       
  1310          [folded vat_t.detached_eq vat_s.detached_eq]
       
  1311 
       
  1312 lemma runing_precond:
       
  1313   fixes th'
       
  1314   assumes th'_in: "th' \<in> threads s"
       
  1315   and neq_th': "th' \<noteq> th"
       
  1316   and is_runing: "th' \<in> runing (t@s)"
       
  1317   shows "cntP s th' > cntV s th'"
       
  1318   using assms
       
  1319 proof -
       
  1320   have "cntP s th' \<noteq> cntV s th'"
       
  1321     by (metis is_runing neq_th' pv_blocked_pre runing_precond_pre th'_in)
       
  1322   moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
       
  1323   ultimately show ?thesis by auto
       
  1324 qed
       
  1325 
       
  1326 lemma moment_blocked_pre:
       
  1327   assumes neq_th': "th' \<noteq> th"
       
  1328   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
  1329   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
       
  1330   shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
       
  1331          th' \<in> threads ((moment (i+j) t)@s)"
       
  1332 proof -
       
  1333   interpret h_i: red_extend_highest_gen _ _ _ _ _ i
       
  1334       by (unfold_locales)
       
  1335   interpret h_j: red_extend_highest_gen _ _ _ _ _ "i+j"
       
  1336       by (unfold_locales)
       
  1337   interpret h:  extend_highest_gen "((moment i t)@s)" th prio tm "moment j (restm i t)"
       
  1338   proof(unfold_locales)
       
  1339     show "vt (moment i t @ s)" by (metis h_i.vt_t) 
       
  1340   next
       
  1341     show "th \<in> threads (moment i t @ s)" by (metis h_i.th_kept)
       
  1342   next
       
  1343     show "preced th (moment i t @ s) = 
       
  1344             Max (cp (moment i t @ s) ` threads (moment i t @ s))"
       
  1345               by (metis h_i.th_cp_max h_i.th_cp_preced h_i.th_kept)
       
  1346   next
       
  1347     show "preced th (moment i t @ s) = Prc prio tm" by (metis h_i.th_kept preced_th) 
       
  1348   next
       
  1349     show "vt (moment j (restm i t) @ moment i t @ s)"
       
  1350       using moment_plus_split by (metis add.commute append_assoc h_j.vt_t)
       
  1351   next
       
  1352     fix th' prio'
       
  1353     assume "Create th' prio' \<in> set (moment j (restm i t))"
       
  1354     thus "prio' \<le> prio" using assms
       
  1355        by (metis Un_iff add.commute h_j.create_low moment_plus_split set_append)
       
  1356   next
       
  1357     fix th' prio'
       
  1358     assume "Set th' prio' \<in> set (moment j (restm i t))"
       
  1359     thus "th' \<noteq> th \<and> prio' \<le> prio"
       
  1360     by (metis Un_iff add.commute h_j.set_diff_low moment_plus_split set_append)
       
  1361   next
       
  1362     fix th'
       
  1363     assume "Exit th' \<in> set (moment j (restm i t))"
       
  1364     thus "th' \<noteq> th"
       
  1365       by (metis Un_iff add.commute h_j.exit_diff moment_plus_split set_append)
       
  1366   qed
       
  1367   show ?thesis 
       
  1368     by (metis add.commute append_assoc eq_pv h.runing_precond_pre
       
  1369           moment_plus_split neq_th' th'_in)
       
  1370 qed
       
  1371 
       
  1372 lemma moment_blocked_eqpv:
       
  1373   assumes neq_th': "th' \<noteq> th"
       
  1374   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
  1375   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
       
  1376   and le_ij: "i \<le> j"
       
  1377   shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
       
  1378          th' \<in> threads ((moment j t)@s) \<and>
       
  1379          th' \<notin> runing ((moment j t)@s)"
       
  1380 proof -
       
  1381   from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
       
  1382   have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
       
  1383    and h2: "th' \<in> threads ((moment j t)@s)" by auto
       
  1384   moreover have "th' \<notin> runing ((moment j t)@s)"
       
  1385   proof -
       
  1386     interpret h: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)
       
  1387     show ?thesis
       
  1388       using h.pv_blocked_pre h1 h2 neq_th' by auto 
       
  1389   qed
       
  1390   ultimately show ?thesis by auto
       
  1391 qed
       
  1392 
       
  1393 (* The foregoing two lemmas are preparation for this one, but
       
  1394    in long run can be combined. Maybe I am wrong.
       
  1395 *)
       
  1396 lemma moment_blocked:
       
  1397   assumes neq_th': "th' \<noteq> th"
       
  1398   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
  1399   and dtc: "detached (moment i t @ s) th'"
       
  1400   and le_ij: "i \<le> j"
       
  1401   shows "detached (moment j t @ s) th' \<and>
       
  1402          th' \<in> threads ((moment j t)@s) \<and>
       
  1403          th' \<notin> runing ((moment j t)@s)"
       
  1404 proof -
       
  1405   interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
       
  1406   interpret h_j: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales) 
       
  1407   have cnt_i: "cntP (moment i t @ s) th' = cntV (moment i t @ s) th'"
       
  1408                 by (metis dtc h_i.detached_elim)
       
  1409   from moment_blocked_eqpv[OF neq_th' th'_in cnt_i le_ij]
       
  1410   show ?thesis by (metis h_j.detached_intro) 
       
  1411 qed
       
  1412 
       
  1413 lemma runing_preced_inversion:
       
  1414   assumes runing': "th' \<in> runing (t@s)"
       
  1415   shows "cp (t@s) th' = preced th s" (is "?L = ?R")
       
  1416 proof -
       
  1417   have "?L = Max (cp (t @ s) ` readys (t @ s))" using assms
       
  1418       by (unfold runing_def, auto)
       
  1419   also have "\<dots> = ?R"
       
  1420       by (metis th_cp_max th_cp_preced vat_t.max_cp_readys_threads) 
       
  1421   finally show ?thesis .
       
  1422 qed
       
  1423 
       
  1424 text {*
       
  1425   The situation when @{term "th"} is blocked is analyzed by the following lemmas.
       
  1426 *}
       
  1427 
       
  1428 text {*
       
  1429   The following lemmas shows the running thread @{text "th'"}, if it is different from
       
  1430   @{term th}, must be live at the very beginning. By the term {\em the very beginning},
       
  1431   we mean the moment where the formal investigation starts, i.e. the moment (or state)
       
  1432   @{term s}. 
       
  1433 *}
       
  1434 
       
  1435 lemma runing_inversion_0:
       
  1436   assumes neq_th': "th' \<noteq> th"
       
  1437   and runing': "th' \<in> runing (t@s)"
       
  1438   shows "th' \<in> threads s"
       
  1439 proof -
       
  1440     -- {* The proof is by contradiction: *}
       
  1441     { assume otherwise: "\<not> ?thesis"
       
  1442       have "th' \<notin> runing (t @ s)"
       
  1443       proof -
       
  1444         -- {* Since @{term "th'"} is running at time @{term "t@s"}, so it exists that time. *}
       
  1445         have th'_in: "th' \<in> threads (t@s)" using runing' by (simp add:runing_def readys_def)
       
  1446         -- {* However, @{text "th'"} does not exist at very beginning. *}
       
  1447         have th'_notin: "th' \<notin> threads (moment 0 t @ s)" using otherwise
       
  1448           by (metis append.simps(1) moment_zero)
       
  1449         -- {* Therefore, there must be a moment during @{text "t"}, when 
       
  1450               @{text "th'"} came into being. *}
       
  1451         -- {* Let us suppose the moment being @{text "i"}: *}
       
  1452         from p_split_gen[OF th'_in th'_notin]
       
  1453         obtain i where lt_its: "i < length t"
       
  1454                  and le_i: "0 \<le> i"
       
  1455                  and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
       
  1456                  and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by (auto)
       
  1457         interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
       
  1458         interpret h_i': red_extend_highest_gen _ _ _ _ _ "(Suc i)" by (unfold_locales)
       
  1459         from lt_its have "Suc i \<le> length t" by auto
       
  1460         -- {* Let us also suppose the event which makes this change is @{text e}: *}
       
  1461         from moment_head[OF this] obtain e where 
       
  1462           eq_me: "moment (Suc i) t = e # moment i t" by blast
       
  1463         hence "vt (e # (moment i t @ s))" by (metis append_Cons h_i'.vt_t) 
       
  1464         hence "PIP (moment i t @ s) e" by (cases, simp)
       
  1465         -- {* It can be derived that this event @{text "e"}, which 
       
  1466               gives birth to @{term "th'"} must be a @{term "Create"}: *}
       
  1467         from create_pre[OF this, of th']
       
  1468         obtain prio where eq_e: "e = Create th' prio"
       
  1469             by (metis append_Cons eq_me lessI post pre) 
       
  1470         have h1: "th' \<in> threads (moment (Suc i) t @ s)" using post by auto 
       
  1471         have h2: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
       
  1472         proof -
       
  1473           have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
       
  1474             by (metis h_i.cnp_cnv_eq pre)
       
  1475           thus ?thesis by (simp add:eq_me eq_e cntP_def cntV_def count_def)
       
  1476         qed
       
  1477         show ?thesis 
       
  1478           using moment_blocked_eqpv [OF neq_th' h1 h2, of "length t"] lt_its moment_ge
       
  1479             by auto
       
  1480       qed
       
  1481       with `th' \<in> runing (t@s)`
       
  1482       have False by simp
       
  1483     } thus ?thesis by auto
       
  1484 qed
       
  1485 
       
  1486 text {* 
       
  1487   The second lemma says, if the running thread @{text th'} is different from 
       
  1488   @{term th}, then this @{text th'} must in the possession of some resources
       
  1489   at the very beginning. 
       
  1490 
       
  1491   To ease the reasoning of resource possession of one particular thread, 
       
  1492   we used two auxiliary functions @{term cntV} and @{term cntP}, 
       
  1493   which are the counters of @{term P}-operations and 
       
  1494   @{term V}-operations respectively. 
       
  1495   If the number of @{term V}-operation is less than the number of 
       
  1496   @{term "P"}-operations, the thread must have some unreleased resource. 
       
  1497 *}
       
  1498 
       
  1499 lemma runing_inversion_1: (* ddd *)
       
  1500   assumes neq_th': "th' \<noteq> th"
       
  1501   and runing': "th' \<in> runing (t@s)"
       
  1502   -- {* thread @{term "th'"} is a live on in state @{term "s"} and 
       
  1503         it has some unreleased resource. *}
       
  1504   shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
       
  1505 proof -
       
  1506   -- {* The proof is a simple composition of @{thm runing_inversion_0} and 
       
  1507         @{thm runing_precond}: *}
       
  1508   -- {* By applying @{thm runing_inversion_0} to assumptions,
       
  1509         it can be shown that @{term th'} is live in state @{term s}: *}
       
  1510   have "th' \<in> threads s"  using runing_inversion_0[OF assms(1,2)] .
       
  1511   -- {* Then the thesis is derived easily by applying @{thm runing_precond}: *}
       
  1512   with runing_precond [OF this neq_th' runing'] show ?thesis by simp
       
  1513 qed
       
  1514 
       
  1515 text {* 
       
  1516   The following lemma is just a rephrasing of @{thm runing_inversion_1}:
       
  1517 *}
       
  1518 lemma runing_inversion_2:
       
  1519   assumes runing': "th' \<in> runing (t@s)"
       
  1520   shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
       
  1521 proof -
       
  1522   from runing_inversion_1[OF _ runing']
       
  1523   show ?thesis by auto
       
  1524 qed
       
  1525 
       
  1526 lemma runing_inversion_3:
       
  1527   assumes runing': "th' \<in> runing (t@s)"
       
  1528   and neq_th: "th' \<noteq> th"
       
  1529   shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
       
  1530   by (metis neq_th runing' runing_inversion_2 runing_preced_inversion)
       
  1531 
       
  1532 lemma runing_inversion_4:
       
  1533   assumes runing': "th' \<in> runing (t@s)"
       
  1534   and neq_th: "th' \<noteq> th"
       
  1535   shows "th' \<in> threads s"
       
  1536   and    "\<not>detached s th'"
       
  1537   and    "cp (t@s) th' = preced th s"
       
  1538   apply (metis neq_th runing' runing_inversion_2)
       
  1539   apply (metis neq_th pv_blocked runing' runing_inversion_2 runing_precond_pre_dtc)
       
  1540   by (metis neq_th runing' runing_inversion_3)
       
  1541 
       
  1542 
       
  1543 text {* 
       
  1544   Suppose @{term th} is not running, it is first shown that
       
  1545   there is a path in RAG leading from node @{term th} to another thread @{text "th'"} 
       
  1546   in the @{term readys}-set (So @{text "th'"} is an ancestor of @{term th}}).
       
  1547 
       
  1548   Now, since @{term readys}-set is non-empty, there must be
       
  1549   one in it which holds the highest @{term cp}-value, which, by definition, 
       
  1550   is the @{term runing}-thread. However, we are going to show more: this running thread
       
  1551   is exactly @{term "th'"}.
       
  1552      *}
       
  1553 lemma th_blockedE: (* ddd *)
       
  1554   assumes "th \<notin> runing (t@s)"
       
  1555   obtains th' where "Th th' \<in> ancestors (RAG (t @ s)) (Th th)"
       
  1556                     "th' \<in> runing (t@s)"
       
  1557 proof -
       
  1558   -- {* According to @{thm vat_t.th_chain_to_ready}, either 
       
  1559         @{term "th"} is in @{term "readys"} or there is path leading from it to 
       
  1560         one thread in @{term "readys"}. *}
       
  1561   have "th \<in> readys (t @ s) \<or> (\<exists>th'. th' \<in> readys (t @ s) \<and> (Th th, Th th') \<in> (RAG (t @ s))\<^sup>+)" 
       
  1562     using th_kept vat_t.th_chain_to_ready by auto
       
  1563   -- {* However, @{term th} can not be in @{term readys}, because otherwise, since 
       
  1564        @{term th} holds the highest @{term cp}-value, it must be @{term "runing"}. *}
       
  1565   moreover have "th \<notin> readys (t@s)" 
       
  1566     using assms runing_def th_cp_max vat_t.max_cp_readys_threads by auto 
       
  1567   -- {* So, there must be a path from @{term th} to another thread @{text "th'"} in 
       
  1568         term @{term readys}: *}
       
  1569   ultimately obtain th' where th'_in: "th' \<in> readys (t@s)"
       
  1570                           and dp: "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+" by auto
       
  1571   -- {* We are going to show that this @{term th'} is running. *}
       
  1572   have "th' \<in> runing (t@s)"
       
  1573   proof -
       
  1574     -- {* We only need to show that this @{term th'} holds the highest @{term cp}-value: *}
       
  1575     have "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" (is "?L = ?R")
       
  1576     proof -
       
  1577       have "?L =  Max ((the_preced (t @ s) \<circ> the_thread) ` subtree (tRAG (t @ s)) (Th th'))"
       
  1578         by (unfold cp_alt_def1, simp)
       
  1579       also have "... = (the_preced (t @ s) \<circ> the_thread) (Th th)"
       
  1580       proof(rule image_Max_subset)
       
  1581         show "finite (Th ` (threads (t@s)))" by (simp add: vat_t.finite_threads)
       
  1582       next
       
  1583         show "subtree (tRAG (t @ s)) (Th th') \<subseteq> Th ` threads (t @ s)"
       
  1584           by (metis Range.intros dp trancl_range vat_t.range_in vat_t.subtree_tRAG_thread) 
       
  1585       next
       
  1586         show "Th th \<in> subtree (tRAG (t @ s)) (Th th')" using dp
       
  1587                     by (unfold tRAG_subtree_eq, auto simp:subtree_def)
       
  1588       next
       
  1589         show "Max ((the_preced (t @ s) \<circ> the_thread) ` Th ` threads (t @ s)) =
       
  1590                       (the_preced (t @ s) \<circ> the_thread) (Th th)" (is "Max ?L = _")
       
  1591         proof -
       
  1592           have "?L = the_preced (t @ s) `  threads (t @ s)" 
       
  1593                      by (unfold image_comp, rule image_cong, auto)
       
  1594           thus ?thesis using max_preced the_preced_def by auto
       
  1595         qed
       
  1596       qed
       
  1597       also have "... = ?R"
       
  1598         using th_cp_max th_cp_preced th_kept 
       
  1599               the_preced_def vat_t.max_cp_readys_threads by auto
       
  1600       finally show ?thesis .
       
  1601     qed 
       
  1602     -- {* Now, since @{term th'} holds the highest @{term cp} 
       
  1603           and we have already show it is in @{term readys},
       
  1604           it is @{term runing} by definition. *}
       
  1605     with `th' \<in> readys (t@s)` show ?thesis by (simp add: runing_def) 
       
  1606   qed
       
  1607   -- {* It is easy to show @{term th'} is an ancestor of @{term th}: *}
       
  1608   moreover have "Th th' \<in> ancestors (RAG (t @ s)) (Th th)" 
       
  1609     using `(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+` by (auto simp:ancestors_def)
       
  1610   ultimately show ?thesis using that by metis
       
  1611 qed
       
  1612 
       
  1613 text {*
       
  1614   Now it is easy to see there is always a thread to run by case analysis
       
  1615   on whether thread @{term th} is running: if the answer is Yes, the 
       
  1616   the running thread is obviously @{term th} itself; otherwise, the running
       
  1617   thread is the @{text th'} given by lemma @{thm th_blockedE}.
       
  1618 *}
       
  1619 lemma live: "runing (t@s) \<noteq> {}"
       
  1620 proof(cases "th \<in> runing (t@s)") 
       
  1621   case True thus ?thesis by auto
       
  1622 next
       
  1623   case False
       
  1624   thus ?thesis using th_blockedE by auto
       
  1625 qed
       
  1626 
       
  1627 end
       
  1628 end
       
  1629