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