ExtGG.thy
changeset 0 110247f9d47e
child 32 e861aff29655
equal deleted inserted replaced
-1:000000000000 0:110247f9d47e
       
     1 theory ExtGG
       
     2 imports PrioG
       
     3 begin
       
     4 
       
     5 lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
       
     6   apply (induct s, simp)
       
     7 proof -
       
     8   fix a s
       
     9   assume ih: "s \<noteq> [] \<Longrightarrow> birthtime th s < length s"
       
    10     and eq_as: "a # s \<noteq> []"
       
    11   show "birthtime th (a # s) < length (a # s)"
       
    12   proof(cases "s \<noteq> []")
       
    13     case False
       
    14     from False show ?thesis
       
    15       by (cases a, auto simp:birthtime.simps)
       
    16   next
       
    17     case True
       
    18     from ih [OF True] show ?thesis
       
    19       by (cases a, auto simp:birthtime.simps)
       
    20   qed
       
    21 qed
       
    22 
       
    23 lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
       
    24   by (induct s, auto simp:threads.simps)
       
    25 
       
    26 lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
       
    27   apply (drule_tac th_in_ne)
       
    28   by (unfold preced_def, auto intro: birth_time_lt)
       
    29 
       
    30 locale highest_gen =
       
    31   fixes s th prio tm
       
    32   assumes vt_s: "vt s"
       
    33   and threads_s: "th \<in> threads s"
       
    34   and highest: "preced th s = Max ((cp s)`threads s)"
       
    35   and preced_th: "preced th s = Prc prio tm"
       
    36 
       
    37 context highest_gen
       
    38 begin
       
    39 
       
    40 
       
    41 
       
    42 lemma lt_tm: "tm < length s"
       
    43   by (insert preced_tm_lt[OF threads_s preced_th], simp)
       
    44 
       
    45 lemma eq_cp_s_th: "cp s th = preced th s"
       
    46 proof -
       
    47   from highest and max_cp_eq[OF vt_s]
       
    48   have is_max: "preced th s = Max ((\<lambda>th. preced th s) ` threads s)" by simp
       
    49   have sbs: "({th} \<union> dependents (wq s) th) \<subseteq> threads s"
       
    50   proof -
       
    51     from threads_s and dependents_threads[OF vt_s, of th]
       
    52     show ?thesis by auto
       
    53   qed
       
    54   show ?thesis
       
    55   proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
       
    56     show "preced th s \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)" by simp
       
    57   next
       
    58     fix y 
       
    59     assume "y \<in> (\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th)"
       
    60     then obtain th1 where th1_in: "th1 \<in> ({th} \<union> dependents (wq s) th)"
       
    61       and eq_y: "y = preced th1 s" by auto
       
    62     show "y \<le> preced th s"
       
    63     proof(unfold is_max, rule Max_ge)
       
    64       from finite_threads[OF vt_s] 
       
    65       show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
    66     next
       
    67       from sbs th1_in and eq_y 
       
    68       show "y \<in> (\<lambda>th. preced th s) ` threads s" by auto
       
    69     qed
       
    70   next
       
    71     from sbs and finite_threads[OF vt_s]
       
    72     show "finite ((\<lambda>th. preced th s) ` ({th} \<union> dependents (wq s) th))"
       
    73       by (auto intro:finite_subset)
       
    74   qed
       
    75 qed
       
    76 
       
    77 lemma highest_cp_preced: "cp s th = Max ((\<lambda> th'. preced th' s) ` threads s)"
       
    78   by (fold max_cp_eq[OF vt_s], unfold eq_cp_s_th, insert highest, simp)
       
    79 
       
    80 lemma highest_preced_thread: "preced th s = Max ((\<lambda> th'. preced th' s) ` threads s)"
       
    81   by (fold eq_cp_s_th, unfold highest_cp_preced, simp)
       
    82 
       
    83 lemma highest': "cp s th = Max (cp s ` threads s)"
       
    84 proof -
       
    85   from highest_cp_preced max_cp_eq[OF vt_s, symmetric]
       
    86   show ?thesis by simp
       
    87 qed
       
    88 
       
    89 end
       
    90 
       
    91 locale extend_highest_gen = highest_gen + 
       
    92   fixes t 
       
    93   assumes vt_t: "vt (t@s)"
       
    94   and create_low: "Create th' prio' \<in> set t \<Longrightarrow> prio' \<le> prio"
       
    95   and set_diff_low: "Set th' prio' \<in> set t \<Longrightarrow> th' \<noteq> th \<and> prio' \<le> prio"
       
    96   and exit_diff: "Exit th' \<in> set t \<Longrightarrow> th' \<noteq> th"
       
    97 
       
    98 lemma step_back_vt_app: 
       
    99   assumes vt_ts: "vt (t@s)" 
       
   100   shows "vt s"
       
   101 proof -
       
   102   from vt_ts show ?thesis
       
   103   proof(induct t)
       
   104     case Nil
       
   105     from Nil show ?case by auto
       
   106   next
       
   107     case (Cons e t)
       
   108     assume ih: " vt (t @ s) \<Longrightarrow> vt s"
       
   109       and vt_et: "vt ((e # t) @ s)"
       
   110     show ?case
       
   111     proof(rule ih)
       
   112       show "vt (t @ s)"
       
   113       proof(rule step_back_vt)
       
   114         from vt_et show "vt (e # t @ s)" by simp
       
   115       qed
       
   116     qed
       
   117   qed
       
   118 qed
       
   119 
       
   120 context extend_highest_gen
       
   121 begin
       
   122 
       
   123 thm extend_highest_gen_axioms_def
       
   124 
       
   125 lemma red_moment:
       
   126   "extend_highest_gen s th prio tm (moment i t)"
       
   127   apply (insert extend_highest_gen_axioms, subst (asm) (1) moment_restm_s [of i t, symmetric])
       
   128   apply (unfold extend_highest_gen_def extend_highest_gen_axioms_def, clarsimp)
       
   129   by (unfold highest_gen_def, auto dest:step_back_vt_app)
       
   130 
       
   131 lemma ind [consumes 0, case_names Nil Cons, induct type]:
       
   132   assumes 
       
   133     h0: "R []"
       
   134   and h2: "\<And> e t. \<lbrakk>vt (t@s); step (t@s) e; 
       
   135                     extend_highest_gen s th prio tm t; 
       
   136                     extend_highest_gen s th prio tm (e#t); R t\<rbrakk> \<Longrightarrow> R (e#t)"
       
   137   shows "R t"
       
   138 proof -
       
   139   from vt_t extend_highest_gen_axioms show ?thesis
       
   140   proof(induct t)
       
   141     from h0 show "R []" .
       
   142   next
       
   143     case (Cons e t')
       
   144     assume ih: "\<lbrakk>vt (t' @ s); extend_highest_gen s th prio tm t'\<rbrakk> \<Longrightarrow> R t'"
       
   145       and vt_e: "vt ((e # t') @ s)"
       
   146       and et: "extend_highest_gen s th prio tm (e # t')"
       
   147     from vt_e and step_back_step have stp: "step (t'@s) e" by auto
       
   148     from vt_e and step_back_vt have vt_ts: "vt (t'@s)" by auto
       
   149     show ?case
       
   150     proof(rule h2 [OF vt_ts stp _ _ _ ])
       
   151       show "R t'"
       
   152       proof(rule ih)
       
   153         from et show ext': "extend_highest_gen s th prio tm t'"
       
   154           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
       
   155       next
       
   156         from vt_ts show "vt (t' @ s)" .
       
   157       qed
       
   158     next
       
   159       from et show "extend_highest_gen s th prio tm (e # t')" .
       
   160     next
       
   161       from et show ext': "extend_highest_gen s th prio tm t'"
       
   162           by (unfold extend_highest_gen_def extend_highest_gen_axioms_def, auto dest:step_back_vt)
       
   163     qed
       
   164   qed
       
   165 qed
       
   166 
       
   167 lemma th_kept: "th \<in> threads (t @ s) \<and> 
       
   168                  preced th (t@s) = preced th s" (is "?Q t")
       
   169 proof -
       
   170   show ?thesis
       
   171   proof(induct rule:ind)
       
   172     case Nil
       
   173     from threads_s
       
   174     show "th \<in> threads ([] @ s) \<and> preced th ([] @ s) = preced th s"
       
   175       by auto
       
   176   next
       
   177     case (Cons e t)
       
   178     show ?case
       
   179     proof(cases e)
       
   180       case (Create thread prio)
       
   181       assume eq_e: " e = Create thread prio"
       
   182       show ?thesis
       
   183       proof -
       
   184         from Cons and eq_e have "step (t@s) (Create thread prio)" by auto
       
   185         hence "th \<noteq> thread"
       
   186         proof(cases)
       
   187           assume "thread \<notin> threads (t @ s)"
       
   188           with Cons show ?thesis by auto
       
   189         qed
       
   190         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
       
   191           by (unfold eq_e, auto simp:preced_def)
       
   192         moreover note Cons
       
   193         ultimately show ?thesis
       
   194           by (auto simp:eq_e)
       
   195       qed
       
   196     next
       
   197       case (Exit thread)
       
   198       assume eq_e: "e = Exit thread"
       
   199       from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
       
   200       from extend_highest_gen.exit_diff [OF this] and eq_e
       
   201       have neq_th: "thread \<noteq> th" by auto
       
   202       with Cons
       
   203       show ?thesis
       
   204         by (unfold eq_e, auto simp:preced_def)
       
   205     next
       
   206       case (P thread cs)
       
   207       assume eq_e: "e = P thread cs"
       
   208       with Cons
       
   209       show ?thesis 
       
   210         by (auto simp:eq_e preced_def)
       
   211     next
       
   212       case (V thread cs)
       
   213       assume eq_e: "e = V thread cs"
       
   214       with Cons
       
   215       show ?thesis 
       
   216         by (auto simp:eq_e preced_def)
       
   217     next
       
   218       case (Set thread prio')
       
   219       assume eq_e: " e = Set thread prio'"
       
   220       show ?thesis
       
   221       proof -
       
   222         from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
       
   223         from extend_highest_gen.set_diff_low[OF this] and eq_e
       
   224         have "th \<noteq> thread" by auto
       
   225         hence "preced th ((e # t) @ s)  = preced th (t @ s)"
       
   226           by (unfold eq_e, auto simp:preced_def)
       
   227         moreover note Cons
       
   228         ultimately show ?thesis
       
   229           by (auto simp:eq_e)
       
   230       qed
       
   231     qed
       
   232   qed
       
   233 qed
       
   234 
       
   235 lemma max_kept: "Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s))) = preced th s"
       
   236 proof(induct rule:ind)
       
   237   case Nil
       
   238   from highest_preced_thread
       
   239   show "Max ((\<lambda>th'. preced th' ([] @ s)) ` threads ([] @ s)) = preced th s"
       
   240     by simp
       
   241 next
       
   242   case (Cons e t)
       
   243   show ?case
       
   244   proof(cases e)
       
   245     case (Create thread prio')
       
   246     assume eq_e: " e = Create thread prio'"
       
   247     from Cons and eq_e have stp: "step (t@s) (Create thread prio')" by auto
       
   248     hence neq_thread: "thread \<noteq> th"
       
   249     proof(cases)
       
   250       assume "thread \<notin> threads (t @ s)"
       
   251       moreover have "th \<in> threads (t@s)"
       
   252       proof -
       
   253         from Cons have "extend_highest_gen s th prio tm t" by auto
       
   254         from extend_highest_gen.th_kept[OF this] show ?thesis by (simp)
       
   255       qed
       
   256       ultimately show ?thesis by auto
       
   257     qed
       
   258     from Cons have "extend_highest_gen s th prio tm t" by auto
       
   259     from extend_highest_gen.th_kept[OF this]
       
   260     have h': " th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" 
       
   261       by (auto)
       
   262     from stp
       
   263     have thread_ts: "thread \<notin> threads (t @ s)"
       
   264       by (cases, auto)
       
   265     show ?thesis (is "Max (?f ` ?A) = ?t")
       
   266     proof -
       
   267       have "Max (?f ` ?A) = Max(insert (?f thread) (?f ` (threads (t@s))))"
       
   268         by (unfold eq_e, simp)
       
   269       moreover have "\<dots> = max (?f thread) (Max (?f ` (threads (t@s))))"
       
   270       proof(rule Max_insert)
       
   271         from Cons have "vt (t @ s)" by auto
       
   272         from finite_threads[OF this]
       
   273         show "finite (?f ` (threads (t@s)))" by simp
       
   274       next
       
   275         from h' show "(?f ` (threads (t@s))) \<noteq> {}" by auto
       
   276       qed
       
   277       moreover have "(Max (?f ` (threads (t@s)))) = ?t"
       
   278       proof -
       
   279         have "(\<lambda>th'. preced th' ((e # t) @ s)) ` threads (t @ s) = 
       
   280           (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" (is "?f1 ` ?B = ?f2 ` ?B")
       
   281         proof -
       
   282           { fix th' 
       
   283             assume "th' \<in> ?B"
       
   284             with thread_ts eq_e
       
   285             have "?f1 th' = ?f2 th'" by (auto simp:preced_def)
       
   286           } thus ?thesis 
       
   287             apply (auto simp:Image_def)
       
   288           proof -
       
   289             fix th'
       
   290             assume h: "\<And>th'. th' \<in> threads (t @ s) \<Longrightarrow> 
       
   291               preced th' (e # t @ s) = preced th' (t @ s)"
       
   292               and h1: "th' \<in> threads (t @ s)"
       
   293             show "preced th' (t @ s) \<in> (\<lambda>th'. preced th' (e # t @ s)) ` threads (t @ s)"
       
   294             proof -
       
   295               from h1 have "?f1 th' \<in> ?f1 ` ?B" by auto
       
   296               moreover from h[OF h1] have "?f1 th' = ?f2 th'" by simp
       
   297               ultimately show ?thesis by simp
       
   298             qed
       
   299           qed
       
   300         qed
       
   301         with Cons show ?thesis by auto
       
   302       qed
       
   303       moreover have "?f thread < ?t"
       
   304       proof -
       
   305         from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
       
   306         from extend_highest_gen.create_low[OF this] and eq_e
       
   307         have "prio' \<le> prio" by auto
       
   308         thus ?thesis
       
   309         by (unfold preced_th, unfold eq_e, insert lt_tm, 
       
   310           auto simp:preced_def precedence_less_def preced_th)
       
   311     qed
       
   312     ultimately show ?thesis by (auto simp:max_def)
       
   313   qed
       
   314 next
       
   315     case (Exit thread)
       
   316     assume eq_e: "e = Exit thread"
       
   317     from Cons have vt_e: "vt (e#(t @ s))" by auto
       
   318     from Cons and eq_e have stp: "step (t@s) (Exit thread)" by auto
       
   319     from stp have thread_ts: "thread \<in> threads (t @ s)"
       
   320       by(cases, unfold runing_def readys_def, auto)
       
   321     from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
       
   322     from extend_highest_gen.exit_diff[OF this] and eq_e
       
   323     have neq_thread: "thread \<noteq> th" by auto
       
   324     from Cons have "extend_highest_gen s th prio tm t" by auto
       
   325     from extend_highest_gen.th_kept[OF this]
       
   326     have h': "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
       
   327     show ?thesis (is "Max (?f ` ?A) = ?t")
       
   328     proof -
       
   329       have "threads (t@s) = insert thread ?A"
       
   330         by (insert stp thread_ts, unfold eq_e, auto)
       
   331       hence "Max (?f ` (threads (t@s))) = Max (?f ` \<dots>)" by simp
       
   332       also from this have "\<dots> = Max (insert (?f thread) (?f ` ?A))" by simp
       
   333       also have "\<dots> = max (?f thread) (Max (?f ` ?A))"
       
   334       proof(rule Max_insert)
       
   335         from finite_threads [OF vt_e]
       
   336         show "finite (?f ` ?A)" by simp
       
   337       next
       
   338         from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
       
   339         from extend_highest_gen.th_kept[OF this]
       
   340         show "?f ` ?A \<noteq> {}" by  auto
       
   341       qed
       
   342       finally have "Max (?f ` (threads (t@s))) = max (?f thread) (Max (?f ` ?A))" .
       
   343       moreover have "Max (?f ` (threads (t@s))) = ?t"
       
   344       proof -
       
   345         from Cons show ?thesis
       
   346           by (unfold eq_e, auto simp:preced_def)
       
   347       qed
       
   348       ultimately have "max (?f thread) (Max (?f ` ?A)) = ?t" by simp
       
   349       moreover have "?f thread < ?t"
       
   350       proof(unfold eq_e, simp add:preced_def, fold preced_def)
       
   351         show "preced thread (t @ s) < ?t"
       
   352         proof -
       
   353           have "preced thread (t @ s) \<le> ?t" 
       
   354           proof -
       
   355             from Cons
       
   356             have "?t = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" 
       
   357               (is "?t = Max (?g ` ?B)") by simp
       
   358             moreover have "?g thread \<le> \<dots>"
       
   359             proof(rule Max_ge)
       
   360               have "vt (t@s)" by fact
       
   361               from finite_threads [OF this]
       
   362               show "finite (?g ` ?B)" by simp
       
   363             next
       
   364               from thread_ts
       
   365               show "?g thread \<in> (?g ` ?B)" by auto
       
   366             qed
       
   367             ultimately show ?thesis by auto
       
   368           qed
       
   369           moreover have "preced thread (t @ s) \<noteq> ?t"
       
   370           proof
       
   371             assume "preced thread (t @ s) = preced th s"
       
   372             with h' have "preced thread (t @ s) = preced th (t@s)" by simp
       
   373             from preced_unique [OF this] have "thread = th"
       
   374             proof
       
   375               from h' show "th \<in> threads (t @ s)" by simp
       
   376             next
       
   377               from thread_ts show "thread \<in> threads (t @ s)" .
       
   378             qed(simp)
       
   379             with neq_thread show "False" by simp
       
   380           qed
       
   381           ultimately show ?thesis by auto
       
   382         qed
       
   383       qed
       
   384       ultimately show ?thesis 
       
   385         by (auto simp:max_def split:if_splits)
       
   386     qed
       
   387   next
       
   388     case (P thread cs)
       
   389     with Cons
       
   390     show ?thesis by (auto simp:preced_def)
       
   391   next
       
   392     case (V thread cs)
       
   393     with Cons
       
   394     show ?thesis by (auto simp:preced_def)
       
   395   next
       
   396     case (Set thread prio')
       
   397     show ?thesis (is "Max (?f ` ?A) = ?t")
       
   398     proof -
       
   399       let ?B = "threads (t@s)"
       
   400       from Cons have "extend_highest_gen s th prio tm (e # t)" by auto
       
   401       from extend_highest_gen.set_diff_low[OF this] and Set
       
   402       have neq_thread: "thread \<noteq> th" and le_p: "prio' \<le> prio" by auto
       
   403       from Set have "Max (?f ` ?A) = Max (?f ` ?B)" by simp
       
   404       also have "\<dots> = ?t"
       
   405       proof(rule Max_eqI)
       
   406         fix y
       
   407         assume y_in: "y \<in> ?f ` ?B"
       
   408         then obtain th1 where 
       
   409           th1_in: "th1 \<in> ?B" and eq_y: "y = ?f th1" by auto
       
   410         show "y \<le> ?t"
       
   411         proof(cases "th1 = thread")
       
   412           case True
       
   413           with neq_thread le_p eq_y Set
       
   414           show ?thesis
       
   415             apply (subst preced_th, insert lt_tm)
       
   416             by (auto simp:preced_def precedence_le_def)
       
   417         next
       
   418           case False
       
   419           with Set eq_y
       
   420           have "y  = preced th1 (t@s)"
       
   421             by (simp add:preced_def)
       
   422           moreover have "\<dots> \<le> ?t"
       
   423           proof -
       
   424             from Cons
       
   425             have "?t = Max ((\<lambda> th'. preced th' (t@s)) ` (threads (t@s)))"
       
   426               by auto
       
   427             moreover have "preced th1 (t@s) \<le> \<dots>"
       
   428             proof(rule Max_ge)
       
   429               from th1_in 
       
   430               show "preced th1 (t @ s) \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)"
       
   431                 by simp
       
   432             next
       
   433               show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
       
   434               proof -
       
   435                 from Cons have "vt (t @ s)" by auto
       
   436                 from finite_threads[OF this] show ?thesis by auto
       
   437               qed
       
   438             qed
       
   439             ultimately show ?thesis by auto
       
   440           qed
       
   441           ultimately show ?thesis by auto
       
   442         qed
       
   443       next
       
   444         from Cons and finite_threads
       
   445         show "finite (?f ` ?B)" by auto
       
   446       next
       
   447         from Cons have "extend_highest_gen s th prio tm t" by auto
       
   448         from extend_highest_gen.th_kept [OF this]
       
   449         have h: "th \<in> threads (t @ s) \<and> preced th (t @ s) = preced th s" .
       
   450         show "?t \<in> (?f ` ?B)" 
       
   451         proof -
       
   452           from neq_thread Set h
       
   453           have "?t = ?f th" by (auto simp:preced_def)
       
   454           with h show ?thesis by auto
       
   455         qed
       
   456       qed
       
   457       finally show ?thesis .
       
   458     qed
       
   459   qed
       
   460 qed
       
   461 
       
   462 lemma max_preced: "preced th (t@s) = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))"
       
   463   by (insert th_kept max_kept, auto)
       
   464 
       
   465 lemma th_cp_max_preced: "cp (t@s) th = Max ((\<lambda> th'. preced th' (t @ s)) ` (threads (t@s)))" 
       
   466   (is "?L = ?R")
       
   467 proof -
       
   468   have "?L = cpreced (wq (t@s)) (t@s) th" 
       
   469     by (unfold cp_eq_cpreced, simp)
       
   470   also have "\<dots> = ?R"
       
   471   proof(unfold cpreced_def)
       
   472     show "Max ((\<lambda>th. preced th (t @ s)) ` ({th} \<union> dependents (wq (t @ s)) th)) =
       
   473           Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))"
       
   474       (is "Max (?f ` ({th} \<union> ?A)) = Max (?f ` ?B)")
       
   475     proof(cases "?A = {}")
       
   476       case False
       
   477       have "Max (?f ` ({th} \<union> ?A)) = Max (insert (?f th) (?f ` ?A))" by simp
       
   478       moreover have "\<dots> = max (?f th) (Max (?f ` ?A))"
       
   479       proof(rule Max_insert)
       
   480         show "finite (?f ` ?A)"
       
   481         proof -
       
   482           from dependents_threads[OF vt_t]
       
   483           have "?A \<subseteq> threads (t@s)" .
       
   484           moreover from finite_threads[OF vt_t] have "finite \<dots>" .
       
   485           ultimately show ?thesis 
       
   486             by (auto simp:finite_subset)
       
   487         qed
       
   488       next
       
   489         from False show "(?f ` ?A) \<noteq> {}" by simp
       
   490       qed
       
   491       moreover have "\<dots> = Max (?f ` ?B)"
       
   492       proof -
       
   493         from max_preced have "?f th = Max (?f ` ?B)" .
       
   494         moreover have "Max (?f ` ?A) \<le> \<dots>" 
       
   495         proof(rule Max_mono)
       
   496           from False show "(?f ` ?A) \<noteq> {}" by simp
       
   497         next
       
   498           show "?f ` ?A \<subseteq> ?f ` ?B" 
       
   499           proof -
       
   500             have "?A \<subseteq> ?B" by (rule dependents_threads[OF vt_t])
       
   501             thus ?thesis by auto
       
   502           qed
       
   503         next
       
   504           from finite_threads[OF vt_t] 
       
   505           show "finite (?f ` ?B)" by simp
       
   506         qed
       
   507         ultimately show ?thesis
       
   508           by (auto simp:max_def)
       
   509       qed
       
   510       ultimately show ?thesis by auto
       
   511     next
       
   512       case True
       
   513       with max_preced show ?thesis by auto
       
   514     qed
       
   515   qed
       
   516   finally show ?thesis .
       
   517 qed
       
   518 
       
   519 lemma th_cp_max: "cp (t@s) th = Max (cp (t@s) ` threads (t@s))"
       
   520   by (unfold max_cp_eq[OF vt_t] th_cp_max_preced, simp)
       
   521 
       
   522 lemma th_cp_preced: "cp (t@s) th = preced th s"
       
   523   by (fold max_kept, unfold th_cp_max_preced, simp)
       
   524 
       
   525 lemma preced_less:
       
   526   fixes th'
       
   527   assumes th'_in: "th' \<in> threads s"
       
   528   and neq_th': "th' \<noteq> th"
       
   529   shows "preced th' s < preced th s"
       
   530 proof -
       
   531   have "preced th' s \<le> Max ((\<lambda>th'. preced th' s) ` threads s)"
       
   532   proof(rule Max_ge)
       
   533     from finite_threads [OF vt_s]
       
   534     show "finite ((\<lambda>th'. preced th' s) ` threads s)" by simp
       
   535   next
       
   536     from th'_in show "preced th' s \<in> (\<lambda>th'. preced th' s) ` threads s"
       
   537       by simp
       
   538   qed
       
   539   moreover have "preced th' s \<noteq> preced th s"
       
   540   proof
       
   541     assume "preced th' s = preced th s"
       
   542     from preced_unique[OF this th'_in] neq_th' threads_s
       
   543     show "False" by  (auto simp:readys_def)
       
   544   qed
       
   545   ultimately show ?thesis using highest_preced_thread
       
   546     by auto
       
   547 qed
       
   548 
       
   549 lemma pv_blocked_pre:
       
   550   fixes th'
       
   551   assumes th'_in: "th' \<in> threads (t@s)"
       
   552   and neq_th': "th' \<noteq> th"
       
   553   and eq_pv: "cntP (t@s) th' = cntV (t@s) th'"
       
   554   shows "th' \<notin> runing (t@s)"
       
   555 proof
       
   556   assume "th' \<in> runing (t@s)"
       
   557   hence "cp (t@s) th' = Max (cp (t@s) ` readys (t@s))" 
       
   558     by (auto simp:runing_def)
       
   559   with max_cp_readys_threads [OF vt_t]
       
   560   have "cp (t @ s) th' = Max (cp (t@s) ` threads (t@s))"
       
   561     by auto
       
   562   moreover from th_cp_max have "cp (t @ s) th = \<dots>" by simp
       
   563   ultimately have "cp (t @ s) th' = cp (t @ s) th" by simp
       
   564   moreover from th_cp_preced and th_kept have "\<dots> = preced th (t @ s)"
       
   565     by simp
       
   566   finally have h: "cp (t @ s) th' = preced th (t @ s)" .
       
   567   show False
       
   568   proof -
       
   569     have "dependents (wq (t @ s)) th' = {}" 
       
   570       by (rule count_eq_dependents [OF vt_t eq_pv])
       
   571     moreover have "preced th' (t @ s) \<noteq> preced th (t @ s)"
       
   572     proof
       
   573       assume "preced th' (t @ s) = preced th (t @ s)"
       
   574       hence "th' = th"
       
   575       proof(rule preced_unique)
       
   576         from th_kept show "th \<in> threads (t @ s)" by simp
       
   577       next
       
   578         from th'_in show "th' \<in> threads (t @ s)" by simp
       
   579       qed
       
   580       with assms show False by simp
       
   581     qed
       
   582     ultimately show ?thesis
       
   583       by (insert h, unfold cp_eq_cpreced cpreced_def, simp)
       
   584   qed
       
   585 qed
       
   586 
       
   587 lemmas pv_blocked = pv_blocked_pre[folded detached_eq [OF vt_t]]
       
   588 
       
   589 lemma runing_precond_pre:
       
   590   fixes th'
       
   591   assumes th'_in: "th' \<in> threads s"
       
   592   and eq_pv: "cntP s th' = cntV s th'"
       
   593   and neq_th': "th' \<noteq> th"
       
   594   shows "th' \<in> threads (t@s) \<and>
       
   595          cntP (t@s) th' = cntV (t@s) th'"
       
   596 proof -
       
   597   show ?thesis
       
   598   proof(induct rule:ind)
       
   599     case (Cons e t)
       
   600     from Cons
       
   601     have in_thread: "th' \<in> threads (t @ s)"
       
   602       and not_holding: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
       
   603     from Cons have "extend_highest_gen s th prio tm t" by auto
       
   604     then have not_runing: "th' \<notin> runing (t @ s)" 
       
   605       apply(rule extend_highest_gen.pv_blocked) 
       
   606       using Cons(1) in_thread neq_th' not_holding
       
   607       apply(simp_all add: detached_eq)
       
   608       done
       
   609     show ?case
       
   610     proof(cases e)
       
   611       case (V thread cs)
       
   612       from Cons and V have vt_v: "vt (V thread cs#(t@s))" by auto
       
   613 
       
   614       show ?thesis
       
   615       proof -
       
   616         from Cons and V have "step (t@s) (V thread cs)" by auto
       
   617         hence neq_th': "thread \<noteq> th'"
       
   618         proof(cases)
       
   619           assume "thread \<in> runing (t@s)"
       
   620           moreover have "th' \<notin> runing (t@s)" by fact
       
   621           ultimately show ?thesis by auto
       
   622         qed
       
   623         with not_holding have cnt_eq: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'" 
       
   624           by (unfold V, simp add:cntP_def cntV_def count_def)
       
   625         moreover from in_thread
       
   626         have in_thread': "th' \<in> threads ((e # t) @ s)" by (unfold V, simp)
       
   627         ultimately show ?thesis by auto
       
   628       qed
       
   629     next
       
   630       case (P thread cs)
       
   631       from Cons and P have "step (t@s) (P thread cs)" by auto
       
   632       hence neq_th': "thread \<noteq> th'"
       
   633       proof(cases)
       
   634         assume "thread \<in> runing (t@s)"
       
   635         moreover note not_runing
       
   636         ultimately show ?thesis by auto
       
   637       qed
       
   638       with Cons and P have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   639         by (auto simp:cntP_def cntV_def count_def)
       
   640       moreover from Cons and P have in_thread': "th' \<in> threads ((e # t) @ s)"
       
   641         by auto
       
   642       ultimately show ?thesis by auto
       
   643     next
       
   644       case (Create thread prio')
       
   645       from Cons and Create have "step (t@s) (Create thread prio')" by auto
       
   646       hence neq_th': "thread \<noteq> th'"
       
   647       proof(cases)
       
   648         assume "thread \<notin> threads (t @ s)"
       
   649         moreover have "th' \<in> threads (t@s)" by fact
       
   650         ultimately show ?thesis by auto
       
   651       qed
       
   652       with Cons and Create 
       
   653       have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   654         by (auto simp:cntP_def cntV_def count_def)
       
   655       moreover from Cons and Create 
       
   656       have in_thread': "th' \<in> threads ((e # t) @ s)" by auto
       
   657       ultimately show ?thesis by auto
       
   658     next
       
   659       case (Exit thread)
       
   660       from Cons and Exit have "step (t@s) (Exit thread)" by auto
       
   661       hence neq_th': "thread \<noteq> th'"
       
   662       proof(cases)
       
   663         assume "thread \<in> runing (t @ s)"
       
   664         moreover note not_runing
       
   665         ultimately show ?thesis by auto
       
   666       qed
       
   667       with Cons and Exit 
       
   668       have eq_cnt: "cntP ((e # t) @ s) th' = cntV ((e # t) @ s) th'"
       
   669         by (auto simp:cntP_def cntV_def count_def)
       
   670       moreover from Cons and Exit and neq_th' 
       
   671       have in_thread': "th' \<in> threads ((e # t) @ s)"
       
   672         by auto
       
   673       ultimately show ?thesis by auto
       
   674     next
       
   675       case (Set thread prio')
       
   676       with Cons
       
   677       show ?thesis 
       
   678         by (auto simp:cntP_def cntV_def count_def)
       
   679     qed
       
   680   next
       
   681     case Nil
       
   682     with assms
       
   683     show ?case by auto
       
   684   qed
       
   685 qed
       
   686 
       
   687 (*
       
   688 lemma runing_precond:
       
   689   fixes th'
       
   690   assumes th'_in: "th' \<in> threads s"
       
   691   and eq_pv: "cntP s th' = cntV s th'"
       
   692   and neq_th': "th' \<noteq> th"
       
   693   shows "th' \<notin> runing (t@s)"
       
   694 proof -
       
   695   from runing_precond_pre[OF th'_in eq_pv neq_th']
       
   696   have h1: "th' \<in> threads (t @ s)"  and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
       
   697   from pv_blocked[OF h1 neq_th' h2] 
       
   698   show ?thesis .
       
   699 qed
       
   700 *)
       
   701 
       
   702 lemmas runing_precond_pre_dtc = runing_precond_pre[folded detached_eq[OF vt_t] detached_eq[OF vt_s]]
       
   703 
       
   704 lemma runing_precond:
       
   705   fixes th'
       
   706   assumes th'_in: "th' \<in> threads s"
       
   707   and neq_th': "th' \<noteq> th"
       
   708   and is_runing: "th' \<in> runing (t@s)"
       
   709   shows "cntP s th' > cntV s th'"
       
   710 proof -
       
   711   have "cntP s th' \<noteq> cntV s th'"
       
   712   proof
       
   713     assume eq_pv: "cntP s th' = cntV s th'"
       
   714     from runing_precond_pre[OF th'_in eq_pv neq_th']
       
   715     have h1: "th' \<in> threads (t @ s)"  
       
   716       and h2: "cntP (t @ s) th' = cntV (t @ s) th'" by auto
       
   717     from pv_blocked_pre[OF h1 neq_th' h2] have " th' \<notin> runing (t @ s)" .
       
   718     with is_runing show "False" by simp
       
   719   qed
       
   720   moreover from cnp_cnv_cncs[OF vt_s, of th'] 
       
   721   have "cntV s th' \<le> cntP s th'" by auto
       
   722   ultimately show ?thesis by auto
       
   723 qed
       
   724 
       
   725 lemma moment_blocked_pre:
       
   726   assumes neq_th': "th' \<noteq> th"
       
   727   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
   728   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
       
   729   shows "cntP ((moment (i+j) t)@s) th' = cntV ((moment (i+j) t)@s) th' \<and>
       
   730          th' \<in> threads ((moment (i+j) t)@s)"
       
   731 proof(induct j)
       
   732   case (Suc k)
       
   733   show ?case
       
   734   proof -
       
   735     { assume True: "Suc (i+k) \<le> length t"
       
   736       from moment_head [OF this] 
       
   737       obtain e where
       
   738         eq_me: "moment (Suc(i+k)) t = e#(moment (i+k) t)"
       
   739         by blast
       
   740       from red_moment[of "Suc(i+k)"]
       
   741       and eq_me have "extend_highest_gen s th prio tm (e # moment (i + k) t)" by simp
       
   742       hence vt_e: "vt (e#(moment (i + k) t)@s)"
       
   743         by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
       
   744                           highest_gen_def, auto)
       
   745       have not_runing': "th' \<notin>  runing (moment (i + k) t @ s)"
       
   746       proof -
       
   747         show "th' \<notin> runing (moment (i + k) t @ s)"
       
   748         proof(rule extend_highest_gen.pv_blocked)
       
   749           from Suc show "th' \<in> threads (moment (i + k) t @ s)"
       
   750             by simp
       
   751         next
       
   752           from neq_th' show "th' \<noteq> th" .
       
   753         next
       
   754           from red_moment show "extend_highest_gen s th prio tm (moment (i + k) t)" .
       
   755         next
       
   756           from Suc vt_e show "detached (moment (i + k) t @ s) th'"
       
   757             apply(subst detached_eq)
       
   758             apply(auto intro: vt_e evt_cons)
       
   759             done
       
   760         qed
       
   761       qed
       
   762       from step_back_step[OF vt_e]
       
   763       have "step ((moment (i + k) t)@s) e" .
       
   764       hence "cntP (e#(moment (i + k) t)@s) th' = cntV (e#(moment (i + k) t)@s) th' \<and>
       
   765         th' \<in> threads (e#(moment (i + k) t)@s)"
       
   766       proof(cases)
       
   767         case (thread_create thread prio)
       
   768         with Suc show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   769       next
       
   770         case (thread_exit thread)
       
   771         moreover have "thread \<noteq> th'"
       
   772         proof -
       
   773           have "thread \<in> runing (moment (i + k) t @ s)" by fact
       
   774           moreover note not_runing'
       
   775           ultimately show ?thesis by auto
       
   776         qed
       
   777         moreover note Suc 
       
   778         ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   779       next
       
   780         case (thread_P thread cs)
       
   781         moreover have "thread \<noteq> th'"
       
   782         proof -
       
   783           have "thread \<in> runing (moment (i + k) t @ s)" by fact
       
   784           moreover note not_runing'
       
   785           ultimately show ?thesis by auto
       
   786         qed
       
   787         moreover note Suc 
       
   788         ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   789       next
       
   790         case (thread_V thread cs)
       
   791         moreover have "thread \<noteq> th'"
       
   792         proof -
       
   793           have "thread \<in> runing (moment (i + k) t @ s)" by fact
       
   794           moreover note not_runing'
       
   795           ultimately show ?thesis by auto
       
   796         qed
       
   797         moreover note Suc 
       
   798         ultimately show ?thesis by (auto simp:cntP_def cntV_def count_def)
       
   799       next
       
   800         case (thread_set thread prio')
       
   801         with Suc show ?thesis
       
   802           by (auto simp:cntP_def cntV_def count_def)
       
   803       qed
       
   804       with eq_me have ?thesis using eq_me by auto 
       
   805     } note h = this
       
   806     show ?thesis
       
   807     proof(cases "Suc (i+k) \<le> length t")
       
   808       case True
       
   809       from h [OF this] show ?thesis .
       
   810     next
       
   811       case False
       
   812       with moment_ge
       
   813       have eq_m: "moment (i + Suc k) t = moment (i+k) t" by auto
       
   814       with Suc show ?thesis by auto
       
   815     qed
       
   816   qed
       
   817 next
       
   818   case 0
       
   819   from assms show ?case by auto
       
   820 qed
       
   821 
       
   822 lemma moment_blocked_eqpv:
       
   823   assumes neq_th': "th' \<noteq> th"
       
   824   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
   825   and eq_pv: "cntP ((moment i t)@s) th' = cntV ((moment i t)@s) th'"
       
   826   and le_ij: "i \<le> j"
       
   827   shows "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th' \<and>
       
   828          th' \<in> threads ((moment j t)@s) \<and>
       
   829          th' \<notin> runing ((moment j t)@s)"
       
   830 proof -
       
   831   from moment_blocked_pre [OF neq_th' th'_in eq_pv, of "j-i"] and le_ij
       
   832   have h1: "cntP ((moment j t)@s) th' = cntV ((moment j t)@s) th'"
       
   833     and h2: "th' \<in> threads ((moment j t)@s)" by auto
       
   834   with extend_highest_gen.pv_blocked 
       
   835   show ?thesis 
       
   836     using  red_moment [of j] h2 neq_th' h1
       
   837     apply(auto)
       
   838     by (metis extend_highest_gen.pv_blocked_pre)
       
   839 qed
       
   840 
       
   841 lemma moment_blocked:
       
   842   assumes neq_th': "th' \<noteq> th"
       
   843   and th'_in: "th' \<in> threads ((moment i t)@s)"
       
   844   and dtc: "detached (moment i t @ s) th'"
       
   845   and le_ij: "i \<le> j"
       
   846   shows "detached (moment j t @ s) th' \<and>
       
   847          th' \<in> threads ((moment j t)@s) \<and>
       
   848          th' \<notin> runing ((moment j t)@s)"
       
   849 proof -
       
   850   from vt_moment[OF vt_t, of "i+length s"] moment_prefix[of i t s]
       
   851   have vt_i: "vt (moment i t @ s)" by auto
       
   852   from vt_moment[OF vt_t, of "j+length s"] moment_prefix[of j t s]
       
   853   have vt_j: "vt  (moment j t @ s)" by auto
       
   854   from moment_blocked_eqpv [OF neq_th' th'_in detached_elim [OF vt_i dtc] le_ij, 
       
   855   folded detached_eq[OF vt_j]]
       
   856   show ?thesis .
       
   857 qed
       
   858 
       
   859 lemma runing_inversion_1:
       
   860   assumes neq_th': "th' \<noteq> th"
       
   861   and runing': "th' \<in> runing (t@s)"
       
   862   shows "th' \<in> threads s \<and> cntV s th' < cntP s th'"
       
   863 proof(cases "th' \<in> threads s")
       
   864   case True
       
   865   with runing_precond [OF this neq_th' runing'] show ?thesis by simp
       
   866 next
       
   867   case False
       
   868   let ?Q = "\<lambda> t. th' \<in> threads (t@s)"
       
   869   let ?q = "moment 0 t"
       
   870   from moment_eq and False have not_thread: "\<not> ?Q ?q" by simp
       
   871   from runing' have "th' \<in> threads (t@s)" by (simp add:runing_def readys_def)
       
   872   from p_split_gen [of ?Q, OF this not_thread]
       
   873   obtain i where lt_its: "i < length t"
       
   874     and le_i: "0 \<le> i"
       
   875     and pre: " th' \<notin> threads (moment i t @ s)" (is "th' \<notin> threads ?pre")
       
   876     and post: "(\<forall>i'>i. th' \<in> threads (moment i' t @ s))" by auto
       
   877   from lt_its have "Suc i \<le> length t" by auto
       
   878   from moment_head[OF this] obtain e where 
       
   879    eq_me: "moment (Suc i) t = e # moment i t" by blast
       
   880   from red_moment[of "Suc i"] and eq_me
       
   881   have "extend_highest_gen s th prio tm (e # moment i t)" by simp
       
   882   hence vt_e: "vt (e#(moment i t)@s)"
       
   883     by (unfold extend_highest_gen_def extend_highest_gen_axioms_def 
       
   884       highest_gen_def, auto)
       
   885   from step_back_step[OF this] have stp_i: "step (moment i t @ s) e" .
       
   886   from post[rule_format, of "Suc i"] and eq_me 
       
   887   have not_in': "th' \<in> threads (e # moment i t@s)" by auto
       
   888   from create_pre[OF stp_i pre this] 
       
   889   obtain prio where eq_e: "e = Create th' prio" .
       
   890   have "cntP (moment i t@s) th' = cntV (moment i t@s) th'"
       
   891   proof(rule cnp_cnv_eq)
       
   892     from step_back_vt [OF vt_e] 
       
   893     show "vt (moment i t @ s)" .
       
   894   next
       
   895     from eq_e and stp_i 
       
   896     have "step (moment i t @ s) (Create th' prio)" by simp
       
   897     thus "th' \<notin> threads (moment i t @ s)" by (cases, simp)
       
   898   qed
       
   899   with eq_e
       
   900   have "cntP ((e#moment i t)@s) th' = cntV ((e#moment i t)@s) th'" 
       
   901     by (simp add:cntP_def cntV_def count_def)
       
   902   with eq_me[symmetric]
       
   903   have h1: "cntP (moment (Suc i) t @ s) th' = cntV (moment (Suc i) t@ s) th'"
       
   904     by simp
       
   905   from eq_e have "th' \<in> threads ((e#moment i t)@s)" by simp
       
   906   with eq_me [symmetric]
       
   907   have h2: "th' \<in> threads (moment (Suc i) t @ s)" by simp
       
   908   from moment_blocked_eqpv [OF neq_th' h2 h1, of "length t"] and lt_its
       
   909   and moment_ge
       
   910   have "th' \<notin> runing (t @ s)" by auto
       
   911   with runing'
       
   912   show ?thesis by auto
       
   913 qed
       
   914 
       
   915 lemma runing_inversion_2:
       
   916   assumes runing': "th' \<in> runing (t@s)"
       
   917   shows "th' = th \<or> (th' \<noteq> th \<and> th' \<in> threads s \<and> cntV s th' < cntP s th')"
       
   918 proof -
       
   919   from runing_inversion_1[OF _ runing']
       
   920   show ?thesis by auto
       
   921 qed
       
   922 
       
   923 lemma runing_preced_inversion:
       
   924   assumes runing': "th' \<in> runing (t@s)"
       
   925   shows "cp (t@s) th' = preced th s"
       
   926 proof -
       
   927   from runing' have "cp (t@s) th' = Max (cp (t @ s) ` readys (t @ s))"
       
   928     by (unfold runing_def, auto)
       
   929   also have "\<dots> = preced th s"
       
   930   proof -
       
   931     from max_cp_readys_threads[OF vt_t]
       
   932     have "\<dots> =  Max (cp (t @ s) ` threads (t @ s))" .
       
   933     also have "\<dots> = preced th s"
       
   934     proof -
       
   935       from max_kept
       
   936       and max_cp_eq [OF vt_t]
       
   937       show ?thesis by auto
       
   938     qed
       
   939     finally show ?thesis .
       
   940   qed
       
   941   finally show ?thesis .
       
   942 qed
       
   943 
       
   944 lemma runing_inversion_3:
       
   945   assumes runing': "th' \<in> runing (t@s)"
       
   946   and neq_th: "th' \<noteq> th"
       
   947   shows "th' \<in> threads s \<and> (cntV s th' < cntP s th' \<and> cp (t@s) th' = preced th s)"
       
   948 proof -
       
   949   from runing_inversion_2 [OF runing'] 
       
   950     and neq_th 
       
   951     and runing_preced_inversion[OF runing']
       
   952   show ?thesis by auto
       
   953 qed
       
   954 
       
   955 lemma runing_inversion_4:
       
   956   assumes runing': "th' \<in> runing (t@s)"
       
   957   and neq_th: "th' \<noteq> th"
       
   958   shows "th' \<in> threads s"
       
   959   and    "\<not>detached s th'"
       
   960   and    "cp (t@s) th' = preced th s"
       
   961 using runing_inversion_3 [OF runing'] 
       
   962   and neq_th 
       
   963   and runing_preced_inversion[OF runing']
       
   964 apply(auto simp add: detached_eq[OF vt_s])
       
   965 done
       
   966 
       
   967 
       
   968 
       
   969 lemma live: "runing (t@s) \<noteq> {}"
       
   970 proof(cases "th \<in> runing (t@s)")
       
   971   case True thus ?thesis by auto
       
   972 next
       
   973   case False
       
   974   then have not_ready: "th \<notin> readys (t@s)"
       
   975     apply (unfold runing_def, 
       
   976             insert th_cp_max max_cp_readys_threads[OF vt_t, symmetric])
       
   977     by auto
       
   978   from th_kept have "th \<in> threads (t@s)" by auto
       
   979   from th_chain_to_ready[OF vt_t this] and not_ready
       
   980   obtain th' where th'_in: "th' \<in> readys (t@s)"
       
   981     and dp: "(Th th, Th th') \<in> (depend (t @ s))\<^sup>+" by auto
       
   982   have "th' \<in> runing (t@s)"
       
   983   proof -
       
   984     have "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"
       
   985     proof -
       
   986       have " Max ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')) = 
       
   987                preced th (t@s)"
       
   988       proof(rule Max_eqI)
       
   989         fix y
       
   990         assume "y \<in> (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
       
   991         then obtain th1 where
       
   992           h1: "th1 = th' \<or> th1 \<in>  dependents (wq (t @ s)) th'"
       
   993           and eq_y: "y = preced th1 (t@s)" by auto
       
   994         show "y \<le> preced th (t @ s)"
       
   995         proof -
       
   996           from max_preced
       
   997           have "preced th (t @ s) = Max ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" .
       
   998           moreover have "y \<le> \<dots>"
       
   999           proof(rule Max_ge)
       
  1000             from h1
       
  1001             have "th1 \<in> threads (t@s)"
       
  1002             proof
       
  1003               assume "th1 = th'"
       
  1004               with th'_in show ?thesis by (simp add:readys_def)
       
  1005             next
       
  1006               assume "th1 \<in> dependents (wq (t @ s)) th'"
       
  1007               with dependents_threads [OF vt_t]
       
  1008               show "th1 \<in> threads (t @ s)" by auto
       
  1009             qed
       
  1010             with eq_y show " y \<in> (\<lambda>th'. preced th' (t @ s)) ` threads (t @ s)" by simp
       
  1011           next
       
  1012             from finite_threads[OF vt_t]
       
  1013             show "finite ((\<lambda>th'. preced th' (t @ s)) ` threads (t @ s))" by simp
       
  1014           qed
       
  1015           ultimately show ?thesis by auto
       
  1016         qed
       
  1017       next
       
  1018         from finite_threads[OF vt_t] dependents_threads [OF vt_t, of th']
       
  1019         show "finite ((\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th'))"
       
  1020           by (auto intro:finite_subset)
       
  1021       next
       
  1022         from dp
       
  1023         have "th \<in> dependents (wq (t @ s)) th'" 
       
  1024           by (unfold cs_dependents_def, auto simp:eq_depend)
       
  1025         thus "preced th (t @ s) \<in> 
       
  1026                 (\<lambda>th. preced th (t @ s)) ` ({th'} \<union> dependents (wq (t @ s)) th')"
       
  1027           by auto
       
  1028       qed
       
  1029       moreover have "\<dots> = Max (cp (t @ s) ` readys (t @ s))"
       
  1030       proof -
       
  1031         from max_preced and max_cp_eq[OF vt_t, symmetric]
       
  1032         have "preced th (t @ s) = Max (cp (t @ s) ` threads (t @ s))" by simp
       
  1033         with max_cp_readys_threads[OF vt_t] show ?thesis by simp
       
  1034       qed
       
  1035       ultimately show ?thesis by (unfold cp_eq_cpreced cpreced_def, simp)
       
  1036     qed
       
  1037     with th'_in show ?thesis by (auto simp:runing_def)
       
  1038   qed
       
  1039   thus ?thesis by auto
       
  1040 qed
       
  1041 
       
  1042 end
       
  1043 end
       
  1044 
       
  1045 
       
  1046