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