ExtGG.thy~
changeset 97 c7ba70dc49bd
parent 96 4805c6333fef
parent 93 524bd3caa6b6
child 98 382293d415f3
equal deleted inserted replaced
96:4805c6333fef 97:c7ba70dc49bd
     1 theory ExtGG
       
     2 imports PrioG CpsG
       
     3 begin
       
     4 
       
     5 text {* 
       
     6   The following two auxiliary lemmas are used to reason about @{term Max}.
       
     7 *}
       
     8 lemma image_Max_eqI: 
       
     9   assumes "finite B"
       
    10   and "b \<in> B"
       
    11   and "\<forall> x \<in> B. f x \<le> f b"
       
    12   shows "Max (f ` B) = f b"
       
    13   using assms
       
    14   using Max_eqI by blast 
       
    15 
       
    16 lemma image_Max_subset:
       
    17   assumes "finite A"
       
    18   and "B \<subseteq> A"
       
    19   and "a \<in> B"
       
    20   and "Max (f ` A) = f a"
       
    21   shows "Max (f ` B) = f a"
       
    22 proof(rule image_Max_eqI)
       
    23   show "finite B"
       
    24     using assms(1) assms(2) finite_subset by auto 
       
    25 next
       
    26   show "a \<in> B" using assms by simp
       
    27 next
       
    28   show "\<forall>x\<in>B. f x \<le> f a"
       
    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 *}
       
    42 locale highest_gen =
       
    43   fixes s th prio tm
       
    44   assumes vt_s: "vt s"
       
    45   and threads_s: "th \<in> threads s"
       
    46   and highest: "preced th s = Max ((cp s)`threads s)"
       
    47   -- {* The internal structure of @{term th}'s precedence is exposed:*}
       
    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: *}
       
    52 sublocale highest_gen < vat_s: valid_trace "s"
       
    53   by (unfold_locales, insert vt_s, simp)
       
    54 
       
    55 context highest_gen
       
    56 begin
       
    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 *}
       
    62 lemma lt_tm: "tm < length s"
       
    63   by (insert preced_tm_lt[OF threads_s preced_th], simp)
       
    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 *}
       
    71 lemma eq_cp_s_th: "cp s th = preced th s" (is "?L = ?R")
       
    72 proof -
       
    73   have "?L \<le> ?R"
       
    74   by (unfold highest, rule Max_ge, 
       
    75         auto simp:threads_s finite_threads)
       
    76   moreover have "?R \<le> ?L"
       
    77     by (unfold vat_s.cp_rec, rule Max_ge, 
       
    78         auto simp:the_preced_def vat_s.fsbttRAGs.finite_children)
       
    79   ultimately show ?thesis by auto
       
    80 qed
       
    81 
       
    82 (* ccc *)
       
    83 lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
       
    84   by (fold max_cp_eq, unfold eq_cp_s_th, insert highest, simp)
       
    85 
       
    86 lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
       
    87   by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
       
    88 
       
    89 lemma highest': "cp s th = Max (cp s ` threads s)"
       
    90 proof -
       
    91   from highest_cp_preced max_cp_eq[symmetric]
       
    92   show ?thesis by simp
       
    93 qed
       
    94 
       
    95 end
       
    96 
       
    97 locale extend_highest_gen = highest_gen + 
       
    98   fixes t 
       
    99   assumes vt_t: "vt (t@s)"
       
   100   and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
       
   101   and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
       
   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)
       
   106 
       
   107 lemma step_back_vt_app: 
       
   108   assumes vt_ts: "vt (t@s)" 
       
   109   shows "vt s"
       
   110 proof -
       
   111   from vt_ts show ?thesis
       
   112   proof(induct t)
       
   113     case Nil
       
   114     from Nil show ?case by auto
       
   115   next
       
   116     case (Cons e t)
       
   117     assume ih: " vt (t @ s) \<Longrightarrow> vt s"
       
   118       and vt_et: "vt ((e # t) @ s)"
       
   119     show ?case
       
   120     proof(rule ih)
       
   121       show "vt (t @ s)"
       
   122       proof(rule step_back_vt)
       
   123         from vt_et show "vt (e # t @ s)" by simp
       
   124       qed
       
   125     qed
       
   126   qed
       
   127 qed
       
   128 
       
   129 
       
   130 locale red_extend_highest_gen = extend_highest_gen +
       
   131    fixes i::nat
       
   132 
       
   133 sublocale red_extend_highest_gen <   red_moment: extend_highest_gen "s" "th" "prio" "tm" "(moment i t)"
       
   134   apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
       
   135   apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
       
   136   by (unfold highest_gen_def, auto dest:step_back_vt_app)
       
   137 
       
   138 
       
   139 context extend_highest_gen
       
   140 begin
       
   141 
       
   142  lemma ind [consumes 0, case_names Nil Cons, induct type]:
       
   143   assumes 
       
   144     h0: "R []"
       
   145   and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
       
   146                     extend_highest_gen s th prio tm t; 
       
   147                     extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
       
   148   shows "R t"
       
   149 proof -
       
   150   from vt_t extend_highest_gen_axioms show ?thesis
       
   151   proof(induct t)
       
   152     from h0 show "R []" .
       
   153   next
       
   154     case (Cons e t')
       
   155     assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
       
   156       and vt_e: "vt ((e # t') @ s)"
       
   157       and et: "extend_highest_gen s th prio tm (e # t')"
       
   158     from vt_e and step_back_step have stp: "step (t'@s) e" by auto
       
   159     from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
       
   160     show ?case
       
   161     proof(rule h2 [OF vt_ts stp _ _ _ ])
       
   162       show "R t'"
       
   163       proof(rule ih)
       
   164         from et show ext': "extend_highest_gen s th prio tm t'"
       
   165           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
       
   166       next
       
   167         from vt_ts show "vt (t' @ s)" .
       
   168       qed
       
   169     next
       
   170       from et show "extend_highest_gen s th prio tm (e # t')" .
       
   171     next
       
   172       from et show ext': "extend_highest_gen s th prio tm t'"
       
   173           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
       
   174     qed
       
   175   qed
       
   176 qed
       
   177 
       
   178 
       
   179 lemma th_kept: "th \<in> threads (t @ s) \<and> 
       
   180                  preced th (t@s) = preced th s" (is "?Q t") 
       
   181 proof -
       
   182   show ?thesis
       
   183   proof(induct rule:ind)
       
   184     case Nil
       
   185     from threads_s
       
   186     show ?case
       
   187       by auto
       
   188   next
       
   189     case (Cons e t)
       
   190     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
       
   191     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
       
   192     show ?case
       
   193     proof(cases e)
       
   194       case (Create thread prio)
       
   195       show ?thesis
       
   196       proof -
       
   197         from Cons and Create have "step (t@s) (Create thread prio)" by auto
       
   198         hence "th \<noteq> thread"
       
   199         proof(cases)
       
   200           case thread_create
       
   201           with Cons show ?thesis by auto
       
   202         qed
       
   203         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
       
   204           by (unfold Create, auto simp:preced_def)
       
   205         moreover note Cons
       
   206         ultimately show ?thesis
       
   207           by (auto simp:Create)
       
   208       qed
       
   209     next
       
   210       case (Exit thread)
       
   211       from h_e.exit_diff and Exit
       
   212       have neq_th: "thread \<noteq> th" by auto
       
   213       with Cons
       
   214       show ?thesis
       
   215         by (unfold Exit, auto simp:preced_def)
       
   216     next
       
   217       case (P thread cs)
       
   218       with Cons
       
   219       show ?thesis 
       
   220         by (auto simp:P preced_def)
       
   221     next
       
   222       case (V thread cs)
       
   223       with Cons
       
   224       show ?thesis 
       
   225         by (auto simp:V preced_def)
       
   226     next
       
   227       case (Set thread prio')
       
   228       show ?thesis
       
   229       proof -
       
   230         from h_e.set_diff_low and Set
       
   231         have "th \<noteq> thread" by auto
       
   232         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
       
   233           by (unfold Set, auto simp:preced_def)
       
   234         moreover note Cons
       
   235         ultimately show ?thesis
       
   236           by (auto simp:Set)
       
   237       qed
       
   238     qed
       
   239   qed
       
   240 qed
       
   241 
       
   242 text {*
       
   243   According to @{thm th_kept}, thread @{text "th"} has its living status
       
   244   and precedence kept along the way of @{text "t"}. The following lemma
       
   245   shows that this preserved precedence of @{text "th"} remains as the highest
       
   246   along the way of @{text "t"}.
       
   247 
       
   248   The proof goes by induction over @{text "t"} using the specialized
       
   249   induction rule @{thm ind}, followed by case analysis of each possible 
       
   250   operations of PIP. All cases follow the same pattern rendered by the 
       
   251   generalized introduction rule @{thm "image_Max_eqI"}. 
       
   252 
       
   253   The very essence is to show that precedences, no matter whether they are newly introduced 
       
   254   or modified, are always lower than the one held by @{term "th"},
       
   255   which by @{thm th_kept} is preserved along the way.
       
   256 *}
       
   257 lemma max_kept: "Max (the_preced (t @ s) ` (threads (t@s))) = preced th s"
       
   258 proof(induct rule:ind)
       
   259   case Nil
       
   260   from highest_preced_thread
       
   261   show ?case
       
   262     by (unfold the_preced_def, simp)
       
   263 next
       
   264   case (Cons e t)
       
   265     interpret h_e: extend_highest_gen _ _ _ _ "(e # t)" using Cons by auto
       
   266     interpret h_t: extend_highest_gen _ _ _ _ t using Cons by auto
       
   267   show ?case
       
   268   proof(cases e)
       
   269     case (Create thread prio')
       
   270     show ?thesis (is "Max (?f ` ?A) = ?t")
       
   271     proof -
       
   272       -- {* The following is the common pattern of each branch of the case analysis. *}
       
   273       -- {* The major part is to show that @{text "th"} holds the highest precedence: *}
       
   274       have "Max (?f ` ?A) = ?f th"
       
   275       proof(rule image_Max_eqI)
       
   276         show "finite ?A" using h_e.finite_threads by auto 
       
   277       next
       
   278         show "th \<in> ?A" using h_e.th_kept by auto 
       
   279       next
       
   280         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
       
   281         proof 
       
   282           fix x
       
   283           assume "x \<in> ?A"
       
   284           hence "x = thread \<or> x \<in> threads (t@s)" by (auto simp:Create)
       
   285           thus "?f x \<le> ?f th"
       
   286           proof
       
   287             assume "x = thread"
       
   288             thus ?thesis 
       
   289               apply (simp add:Create the_preced_def preced_def, fold preced_def)
       
   290               using Create h_e.create_low h_t.th_kept lt_tm preced_leI2 preced_th by force
       
   291           next
       
   292             assume h: "x \<in> threads (t @ s)"
       
   293             from Cons(2)[unfolded Create] 
       
   294             have "x \<noteq> thread" using h by (cases, auto)
       
   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 
       
   312   next 
       
   313     case (Exit thread)
       
   314     show ?thesis (is "Max (?f ` ?A) = ?t")
       
   315     proof -
       
   316       have "Max (?f ` ?A) = ?f th"
       
   317       proof(rule image_Max_eqI)
       
   318         show "finite ?A" using h_e.finite_threads by auto 
       
   319       next
       
   320         show "th \<in> ?A" using h_e.th_kept by auto 
       
   321       next
       
   322         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
       
   323         proof 
       
   324           fix x
       
   325           assume "x \<in> ?A"
       
   326           hence "x \<in> threads (t@s)" by (simp add: Exit) 
       
   327           hence "?f x \<le> Max (?f ` threads (t@s))" 
       
   328             by (simp add: h_t.finite_threads) 
       
   329           also have "... \<le> ?f th" 
       
   330             apply (simp add:Exit the_preced_def preced_def, fold preced_def)
       
   331             using Cons.hyps(5) h_t.th_kept the_preced_def by auto
       
   332           finally show "?f x \<le> ?f th" .
       
   333         qed
       
   334       qed
       
   335       also have "... = ?t" using h_e.th_kept the_preced_def by auto
       
   336       finally show ?thesis .
       
   337     qed 
       
   338   next
       
   339     case (P thread cs)
       
   340     with Cons
       
   341     show ?thesis by (auto simp:preced_def the_preced_def)
       
   342   next
       
   343     case (V thread cs)
       
   344     with Cons
       
   345     show ?thesis by (auto simp:preced_def the_preced_def)
       
   346   next 
       
   347     case (Set thread prio')
       
   348     show ?thesis (is "Max (?f ` ?A) = ?t")
       
   349     proof -
       
   350       have "Max (?f ` ?A) = ?f th"
       
   351       proof(rule image_Max_eqI)
       
   352         show "finite ?A" using h_e.finite_threads by auto 
       
   353       next
       
   354         show "th \<in> ?A" using h_e.th_kept by auto 
       
   355       next
       
   356         show "\<forall>x\<in>?A. ?f x \<le> ?f th"
       
   357         proof 
       
   358           fix x
       
   359           assume h: "x \<in> ?A"
       
   360           show "?f x \<le> ?f th"
       
   361           proof(cases "x = thread")
       
   362             case True
       
   363             moreover have "the_preced (Set thread prio' # t @ s) thread \<le> the_preced (t @ s) th"
       
   364             proof -
       
   365               have "the_preced (t @ s) th = Prc prio tm"  
       
   366                 using h_t.th_kept preced_th by (simp add:the_preced_def)
       
   367               moreover have "prio' \<le> prio" using Set h_e.set_diff_low by auto
       
   368               ultimately show ?thesis by (insert lt_tm, auto simp:the_preced_def preced_def)
       
   369             qed
       
   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 .
       
   380           qed
       
   381         qed
       
   382       qed
       
   383       also have "... = ?t" using h_e.th_kept the_preced_def by auto
       
   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)
       
   439     qed
       
   440   qed
       
   441   also have "... = ?R" by (simp add: max_preced the_preced_def) 
       
   442   finally show ?thesis .
       
   443 qed
       
   444 
       
   445 lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
       
   446   using max_cp_eq th_cp_max_preced the_preced_def vt_t by presburger
       
   447 
       
   448 lemma th_cp_preced: "cp (t@s) th = preced th s"
       
   449   by (fold max_kept, unfold th_cp_max_preced, simp)
       
   450 
       
   451 lemma preced_less:
       
   452   assumes th'_in: "th' \<in> threads s"
       
   453   and neq_th': "th' \<noteq> th"
       
   454   shows "preced th' s < preced th s"
       
   455   using assms
       
   456 by (metis Max.coboundedI finite_imageI highest not_le order.trans 
       
   457     preced_linorder rev_image_eqI threads_s vat_s.finite_threads 
       
   458     vat_s.le_cp)
       
   459 
       
   460 text {*
       
   461   Counting of the number of @{term "P"} and @{term "V"} operations 
       
   462   is the cornerstone of a large number of the following proofs. 
       
   463   The reason is that this counting is quite easy to calculate and 
       
   464   convenient to use in the reasoning. 
       
   465 
       
   466   The following lemma shows that the counting controls whether 
       
   467   a thread is running or not.
       
   468 *}
       
   469 
       
   470 lemma pv_blocked_pre:
       
   471   assumes th'_in: "th' \<in> threads (t@s)"
       
   472   and neq_th': "th' \<noteq> th"
       
   473   and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
       
   474   shows "th' \<notin> runing (t@s)"
       
   475 proof
       
   476   assume otherwise: "th' \<in> runing (t@s)"
       
   477   show False
       
   478   proof -
       
   479     have "th' = th"
       
   480     proof(rule preced_unique)
       
   481       show "preced th' (t @ s) = preced th (t @ s)" (is "?L = ?R")
       
   482       proof -
       
   483         have "?L = cp (t@s) th'"
       
   484           by (unfold cp_eq_cpreced cpreced_def count_eq_dependants[OF eq_pv], simp)
       
   485         also have "... = cp (t @ s) th" using otherwise 
       
   486           by (metis (mono_tags, lifting) mem_Collect_eq 
       
   487                     runing_def th_cp_max vat_t.max_cp_readys_threads)
       
   488         also have "... = ?R" by (metis th_cp_preced th_kept) 
       
   489         finally show ?thesis .
       
   490       qed
       
   491     qed (auto simp: th'_in th_kept)
       
   492     moreover have "th' \<noteq> th" using neq_th' .
       
   493     ultimately show ?thesis by simp
       
   494  qed
       
   495 qed
       
   496 
       
   497 lemmas pv_blocked = pv_blocked_pre[folded detached_eq]
       
   498 
       
   499 lemma runing_precond_pre:
       
   500   fixes th'
       
   501   assumes th'_in: "th' \<in> threads s"
       
   502   and eq_pv: "cntP s th' = cntV s th'"
       
   503   and neq_th': "th' \<noteq> th"
       
   504   shows "th' \<in> threads (t@s) \<and>
       
   505          cntP (t@s) th' = cntV (t@s) th'"
       
   506 proof(induct rule:ind)
       
   507   case (Cons e t)
       
   508     interpret vat_t: extend_highest_gen s th prio tm t using Cons by simp
       
   509     interpret vat_e: extend_highest_gen s th prio tm "(e # t)" using Cons by simp
       
   510     show ?case
       
   511     proof(cases e)
       
   512       case (P thread cs)
       
   513       show ?thesis
       
   514       proof -
       
   515         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   516         proof -
       
   517           have "thread \<noteq> th'"
       
   518           proof -
       
   519             have "step (t@s) (P thread cs)" using Cons P 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)
       
   529         qed
       
   530         moreover have "th' \<in> threads ((e # t) @ s)" using Cons by (unfold P, simp)
       
   531         ultimately show ?thesis by auto
       
   532       qed
       
   533     next
       
   534       case (V thread cs)
       
   535       show ?thesis
       
   536       proof -
       
   537         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   538         proof -
       
   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)
       
   553         ultimately show ?thesis by auto
       
   554       qed
       
   555     next
       
   556       case (Create thread prio')
       
   557       show ?thesis
       
   558       proof -
       
   559         have "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   560         proof -
       
   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)
       
   569         ultimately show ?thesis by auto
       
   570       qed
       
   571     next
       
   572       case (Exit thread)
       
   573       show ?thesis
       
   574       proof -
       
   575         have neq_thread: "thread \<noteq> th'"
       
   576         proof -
       
   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) 
       
   585         ultimately show ?thesis by auto
       
   586       qed
       
   587     next
       
   588       case (Set thread prio')
       
   589       with Cons
       
   590       show ?thesis 
       
   591         by (auto simp:cntP_def cntV_def count_def)
       
   592     qed
       
   593 next
       
   594   case Nil
       
   595   with assms
       
   596   show ?case by auto
       
   597 qed
       
   598 
       
   599 text {* Changing counting balance to detachedness *}
       
   600 lemmas runing_precond_pre_dtc = runing_precond_pre
       
   601          [folded vat_t.detached_eq vat_s.detached_eq]
       
   602 
       
   603 lemma runing_precond:
       
   604   fixes th'
       
   605   assumes th'_in: "th' \<in> threads s"
       
   606   and neq_th': "th' \<noteq> th"
       
   607   and is_runing: "th' \<in> runing (t@s)"
       
   608   shows "cntP s th' > cntV s th'"
       
   609   using assms
       
   610 proof -
       
   611   have "cntP s th' \<noteq> cntV s th'"
       
   612     by (metis is_runing neq_th' pv_blocked_pre runing_precond_pre th'_in)
       
   613   moreover have "cntV s th' \<le> cntP s th'" using vat_s.cnp_cnv_cncs by auto
       
   614   ultimately show ?thesis by auto
       
   615 qed
       
   616 
       
   617 lemma moment_blocked_pre:
       
   618   assumes neq_th': "th' \<noteq> th"
       
   619   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
   620   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
       
   621   shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
       
   622          th' \<in> threads ((moment (i+j) t)@s)"
       
   623 proof -
       
   624   interpret h_i: red_extend_highest_gen _ _ _ _ _ i
       
   625       by (unfold_locales)
       
   626   interpret h_j: red_extend_highest_gen _ _ _ _ _ "i+j"
       
   627       by (unfold_locales)
       
   628   interpret h:  extend_highest_gen "((moment i t)@s)" th prio tm "moment j (restm i t)"
       
   629   proof(unfold_locales)
       
   630     show "vt (moment i t @ s)" by (metis h_i.vt_t) 
       
   631   next
       
   632     show "th \<in> threads (moment i t @ s)" by (metis h_i.th_kept)
       
   633   next
       
   634     show "preced th (moment i t @ s) = 
       
   635             Max (cp (moment i t @ s) ` threads (moment i t @ s))"
       
   636               by (metis h_i.th_cp_max h_i.th_cp_preced h_i.th_kept)
       
   637   next
       
   638     show "preced th (moment i t @ s) = Prc prio tm" by (metis h_i.th_kept preced_th) 
       
   639   next
       
   640     show "vt (moment j (restm i t) @ moment i t @ s)"
       
   641       using moment_plus_split by (metis add.commute append_assoc h_j.vt_t)
       
   642   next
       
   643     fix th' prio'
       
   644     assume "Create th' prio' \<in> set (moment j (restm i t))"
       
   645     thus "prio' \<le> prio" using assms
       
   646        by (metis Un_iff add.commute h_j.create_low moment_plus_split set_append)
       
   647   next
       
   648     fix th' prio'
       
   649     assume "Set th' prio' \<in> set (moment j (restm i t))"
       
   650     thus "th' \<noteq> th \<and> prio' \<le> prio"
       
   651     by (metis Un_iff add.commute h_j.set_diff_low moment_plus_split set_append)
       
   652   next
       
   653     fix th'
       
   654     assume "Exit th' \<in> set (moment j (restm i t))"
       
   655     thus "th' \<noteq> th"
       
   656       by (metis Un_iff add.commute h_j.exit_diff moment_plus_split set_append)
       
   657   qed
       
   658   show ?thesis 
       
   659     by (metis add.commute append_assoc eq_pv h.runing_precond_pre
       
   660           moment_plus_split neq_th' th'_in)
       
   661 qed
       
   662 
       
   663 lemma moment_blocked_eqpv:
       
   664   assumes neq_th': "th' \<noteq> th"
       
   665   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
   666   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
       
   667   and le_ij: "i \<le> j"
       
   668   shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
       
   669          th' \<in> threads ((moment j t)@s) \<and>
       
   670          th' \<notin> runing ((moment j t)@s)"
       
   671 proof -
       
   672   from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
       
   673   have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
       
   674    and h2: "th' \<in> threads ((moment j t)@s)" by auto
       
   675   moreover have "th' \<notin> runing ((moment j t)@s)"
       
   676   proof -
       
   677     interpret h: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales)
       
   678     show ?thesis
       
   679       using h.pv_blocked_pre h1 h2 neq_th' by auto 
       
   680   qed
       
   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 *)
       
   687 lemma moment_blocked:
       
   688   assumes neq_th': "th' \<noteq> th"
       
   689   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
   690   and dtc: "detached (moment i t @ s) th'"
       
   691   and le_ij: "i \<le> j"
       
   692   shows "detached (moment j t @ s) th' \<and>
       
   693          th' \<in> threads ((moment j t)@s) \<and>
       
   694          th' \<notin> runing ((moment j t)@s)"
       
   695 proof -
       
   696   interpret h_i: red_extend_highest_gen _ _ _ _ _ i by (unfold_locales)
       
   697   interpret h_j: red_extend_highest_gen _ _ _ _ _ j by (unfold_locales) 
       
   698   have cnt_i: "cntP (moment i t @ s) th' = cntV (moment i t @ s) th'"
       
   699                 by (metis dtc h_i.detached_elim)
       
   700   from moment_blocked_eqpv[OF neq_th' th'_in cnt_i le_ij]
       
   701   show ?thesis by (metis h_j.detached_intro) 
       
   702 qed
       
   703 
       
   704 lemma runing_preced_inversion:
       
   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:
       
   727   assumes neq_th': "th' \<noteq> th"
       
   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. *}
       
   795   shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
       
   796 proof -
       
   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}: *}
       
   803   with runing_precond [OF this neq_th' runing'] show ?thesis by simp
       
   804 qed
       
   805 
       
   806 text {* 
       
   807   The following lemma is just a rephrasing of @{thm runing_inversion_1}:
       
   808 *}
       
   809 lemma runing_inversion_2:
       
   810   assumes runing': "th' \<in> runing (t@s)"
       
   811   shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
       
   812 proof -
       
   813   from runing_inversion_1[OF _ runing']
       
   814   show ?thesis by auto
       
   815 qed
       
   816 
       
   817 lemma runing_inversion_3:
       
   818   assumes runing': "th' \<in> runing (t@s)"
       
   819   and neq_th: "th' \<noteq> th"
       
   820   shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
       
   821   by (metis neq_th runing' runing_inversion_2 runing_preced_inversion)
       
   822 
       
   823 lemma runing_inversion_4:
       
   824   assumes runing': "th' \<in> runing (t@s)"
       
   825   and neq_th: "th' \<noteq> th"
       
   826   shows "th' \<in> threads s"
       
   827   and    "\<not>detached s th'"
       
   828   and    "cp (t@s) th' = preced th s"
       
   829   apply (metis neq_th runing' runing_inversion_2)
       
   830   apply (metis neq_th pv_blocked runing' runing_inversion_2 runing_precond_pre_dtc)
       
   831   by (metis neq_th runing' runing_inversion_3)
       
   832 
       
   833 
       
   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 *}
       
   910 lemma live: "runing (t@s) \<noteq> {}"
       
   911 proof(cases "th \<in> runing (t@s)") 
       
   912   case True thus ?thesis by auto
       
   913 next
       
   914   case False
       
   915   thus ?thesis using th_blockedE by auto
       
   916 qed
       
   917 
       
   918 end
       
   919 end
       
   920 
       
   921 
       
   922