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