prio/PrioG.thy
changeset 262 4190df6f4488
child 264 24199eb2c423
equal deleted inserted replaced
261:12e9aa68d5db 262:4190df6f4488
       
     1 theory PrioG
       
     2 imports PrioGDef
       
     3 begin
       
     4 
       
     5 lemma runing_ready: "runing s \<subseteq> readys s"
       
     6   by (auto simp only:runing_def readys_def)
       
     7 
       
     8 lemma wq_v_neq:
       
     9    "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
       
    10   by (auto simp:wq_def Let_def cp_def split:list.splits)
       
    11 
       
    12 lemma wq_distinct: "vt step s \<Longrightarrow> distinct (wq s cs)"
       
    13 proof(erule_tac vt.induct, simp add:wq_def)
       
    14   fix s e
       
    15   assume h1: "step s e"
       
    16   and h2: "distinct (wq s cs)"
       
    17   thus "distinct (wq (e # s) cs)"
       
    18   proof(induct rule:step.induct, auto simp: wq_def Let_def split:list.splits)
       
    19     fix thread s
       
    20     assume h1: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
       
    21       and h2: "thread \<in> set (waiting_queue (schs s) cs)"
       
    22       and h3: "thread \<in> runing s"
       
    23     show "False" 
       
    24     proof -
       
    25       from h3 have "\<And> cs. thread \<in>  set (waiting_queue (schs s) cs) \<Longrightarrow> 
       
    26                              thread = hd ((waiting_queue (schs s) cs))" 
       
    27         by (simp add:runing_def readys_def s_waiting_def wq_def)
       
    28       from this [OF h2] have "thread = hd (waiting_queue (schs s) cs)" .
       
    29       with h2
       
    30       have "(Cs cs, Th thread) \<in> (depend s)"
       
    31         by (simp add:s_depend_def s_holding_def wq_def cs_holding_def)
       
    32       with h1 show False by auto
       
    33     qed
       
    34   next
       
    35     fix thread s a list
       
    36     assume dst: "distinct list"
       
    37     show "distinct (SOME q. distinct q \<and> set q = set list)"
       
    38     proof(rule someI2)
       
    39       from dst show  "distinct list \<and> set list = set list" by auto
       
    40     next
       
    41       fix q assume "distinct q \<and> set q = set list"
       
    42       thus "distinct q" by auto
       
    43     qed
       
    44   qed
       
    45 qed
       
    46 
       
    47 lemma step_back_vt: "vt ccs (e#s) \<Longrightarrow> vt ccs s"
       
    48   by(ind_cases "vt ccs (e#s)", simp)
       
    49 
       
    50 lemma step_back_step: "vt ccs (e#s) \<Longrightarrow> ccs s e"
       
    51   by(ind_cases "vt ccs (e#s)", simp)
       
    52 
       
    53 lemma block_pre: 
       
    54   fixes thread cs s
       
    55   assumes vt_e: "vt step (e#s)"
       
    56   and s_ni: "thread \<notin>  set (wq s cs)"
       
    57   and s_i: "thread \<in> set (wq (e#s) cs)"
       
    58   shows "e = P thread cs"
       
    59 proof -
       
    60   show ?thesis
       
    61   proof(cases e)
       
    62     case (P th cs)
       
    63     with assms
       
    64     show ?thesis
       
    65       by (auto simp:wq_def Let_def split:if_splits)
       
    66   next
       
    67     case (Create th prio)
       
    68     with assms show ?thesis
       
    69       by (auto simp:wq_def Let_def split:if_splits)
       
    70   next
       
    71     case (Exit th)
       
    72     with assms show ?thesis
       
    73       by (auto simp:wq_def Let_def split:if_splits)
       
    74   next
       
    75     case (Set th prio)
       
    76     with assms show ?thesis
       
    77       by (auto simp:wq_def Let_def split:if_splits)
       
    78   next
       
    79     case (V th cs)
       
    80     with assms show ?thesis
       
    81       apply (auto simp:wq_def Let_def split:if_splits)
       
    82     proof -
       
    83       fix q qs
       
    84       assume h1: "thread \<notin> set (waiting_queue (schs s) cs)"
       
    85         and h2: "q # qs = waiting_queue (schs s) cs"
       
    86         and h3: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
       
    87         and vt: "vt step (V th cs # s)"
       
    88       from h1 and h2[symmetric] have "thread \<notin> set (q # qs)" by simp
       
    89       moreover have "thread \<in> set qs"
       
    90       proof -
       
    91         have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
       
    92         proof(rule someI2)
       
    93           from wq_distinct [OF step_back_vt[OF vt], of cs]
       
    94           and h2[symmetric, folded wq_def]
       
    95           show "distinct qs \<and> set qs = set qs" by auto
       
    96         next
       
    97           fix x assume "distinct x \<and> set x = set qs"
       
    98           thus "set x = set qs" by auto
       
    99         qed
       
   100         with h3 show ?thesis by simp
       
   101       qed
       
   102       ultimately show "False" by auto
       
   103       qed
       
   104   qed
       
   105 qed
       
   106 
       
   107 lemma p_pre: "\<lbrakk>vt step ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
       
   108   thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (depend s)^+"
       
   109 apply (ind_cases "vt step ((P thread cs)#s)")
       
   110 apply (ind_cases "step s (P thread cs)")
       
   111 by auto
       
   112 
       
   113 lemma abs1:
       
   114   fixes e es
       
   115   assumes ein: "e \<in> set es"
       
   116   and neq: "hd es \<noteq> hd (es @ [x])"
       
   117   shows "False"
       
   118 proof -
       
   119   from ein have "es \<noteq> []" by auto
       
   120   then obtain e ess where "es = e # ess" by (cases es, auto)
       
   121   with neq show ?thesis by auto
       
   122 qed
       
   123 
       
   124 lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]"
       
   125   by (cases es, auto)
       
   126 
       
   127 inductive_cases evt_cons: "vt cs (a#s)"
       
   128 
       
   129 lemma abs2:
       
   130   assumes vt: "vt step (e#s)"
       
   131   and inq: "thread \<in> set (wq s cs)"
       
   132   and nh: "thread = hd (wq s cs)"
       
   133   and qt: "thread \<noteq> hd (wq (e#s) cs)"
       
   134   and inq': "thread \<in> set (wq (e#s) cs)"
       
   135   shows "False"
       
   136 proof -
       
   137   from assms show "False"
       
   138     apply (cases e)
       
   139     apply ((simp split:if_splits add:Let_def wq_def)[1])+
       
   140     apply (insert abs1, fast)[1]
       
   141     apply (auto simp:wq_def simp:Let_def split:if_splits list.splits)
       
   142   proof -
       
   143     fix th qs
       
   144     assume vt: "vt step (V th cs # s)"
       
   145       and th_in: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
       
   146       and eq_wq: "waiting_queue (schs s) cs = thread # qs"
       
   147     show "False"
       
   148     proof -
       
   149       from wq_distinct[OF step_back_vt[OF vt], of cs]
       
   150         and eq_wq[folded wq_def] have "distinct (thread#qs)" by simp
       
   151       moreover have "thread \<in> set qs"
       
   152       proof -
       
   153         have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
       
   154         proof(rule someI2)
       
   155           from wq_distinct [OF step_back_vt[OF vt], of cs]
       
   156           and eq_wq [folded wq_def]
       
   157           show "distinct qs \<and> set qs = set qs" by auto
       
   158         next
       
   159           fix x assume "distinct x \<and> set x = set qs"
       
   160           thus "set x = set qs" by auto
       
   161         qed
       
   162         with th_in show ?thesis by auto
       
   163       qed
       
   164       ultimately show ?thesis by auto
       
   165     qed
       
   166   qed
       
   167 qed
       
   168 
       
   169 lemma vt_moment: "\<And> t. \<lbrakk>vt cs s; t \<le> length s\<rbrakk> \<Longrightarrow> vt cs (moment t s)"
       
   170 proof(induct s, simp)
       
   171   fix a s t
       
   172   assume h: "\<And>t.\<lbrakk>vt cs s; t \<le> length s\<rbrakk> \<Longrightarrow> vt cs (moment t s)"
       
   173     and vt_a: "vt cs (a # s)"
       
   174     and le_t: "t \<le> length (a # s)"
       
   175   show "vt cs (moment t (a # s))"
       
   176   proof(cases "t = length (a#s)")
       
   177     case True
       
   178     from True have "moment t (a#s) = a#s" by simp
       
   179     with vt_a show ?thesis by simp
       
   180   next
       
   181     case False
       
   182     with le_t have le_t1: "t \<le> length s" by simp
       
   183     from vt_a have "vt cs s"
       
   184       by (erule_tac evt_cons, simp)
       
   185     from h [OF this le_t1] have "vt cs (moment t s)" .
       
   186     moreover have "moment t (a#s) = moment t s"
       
   187     proof -
       
   188       from moment_app [OF le_t1, of "[a]"] 
       
   189       show ?thesis by simp
       
   190     qed
       
   191     ultimately show ?thesis by auto
       
   192   qed
       
   193 qed
       
   194 
       
   195 (* Wrong:
       
   196     lemma \<lbrakk>thread \<in> set (waiting_queue cs1 s); thread \<in> set (waiting_queue cs2 s)\<rbrakk> \<Longrightarrow> cs1 = cs2"
       
   197 *)
       
   198 
       
   199 lemma waiting_unique_pre:
       
   200   fixes cs1 cs2 s thread
       
   201   assumes vt: "vt step s"
       
   202   and h11: "thread \<in> set (wq s cs1)"
       
   203   and h12: "thread \<noteq> hd (wq s cs1)"
       
   204   assumes h21: "thread \<in> set (wq s cs2)"
       
   205   and h22: "thread \<noteq> hd (wq s cs2)"
       
   206   and neq12: "cs1 \<noteq> cs2"
       
   207   shows "False"
       
   208 proof -
       
   209   let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
       
   210   from h11 and h12 have q1: "?Q cs1 s" by simp
       
   211   from h21 and h22 have q2: "?Q cs2 s" by simp
       
   212   have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
       
   213   have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
       
   214   from p_split [of "?Q cs1", OF q1 nq1]
       
   215   obtain t1 where lt1: "t1 < length s"
       
   216     and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and>
       
   217         thread \<noteq> hd (wq (moment t1 s) cs1))"
       
   218     and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and>
       
   219              thread \<noteq> hd (wq (moment i' s) cs1))" by auto
       
   220   from p_split [of "?Q cs2", OF q2 nq2]
       
   221   obtain t2 where lt2: "t2 < length s"
       
   222     and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and>
       
   223         thread \<noteq> hd (wq (moment t2 s) cs2))"
       
   224     and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and>
       
   225              thread \<noteq> hd (wq (moment i' s) cs2))" by auto
       
   226   show ?thesis
       
   227   proof -
       
   228     { 
       
   229       assume lt12: "t1 < t2"
       
   230       let ?t3 = "Suc t2"
       
   231       from lt2 have le_t3: "?t3 \<le> length s" by auto
       
   232       from moment_plus [OF this] 
       
   233       obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
       
   234       have "t2 < ?t3" by simp
       
   235       from nn2 [rule_format, OF this] and eq_m
       
   236       have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   237         h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   238       have vt_e: "vt step (e#moment t2 s)"
       
   239       proof -
       
   240         from vt_moment [OF vt le_t3]
       
   241         have "vt step (moment ?t3 s)" .
       
   242         with eq_m show ?thesis by simp
       
   243       qed
       
   244       have ?thesis
       
   245       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   246         case True
       
   247         from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
       
   248           by auto
       
   249         from abs2 [OF vt_e True eq_th h2 h1]
       
   250         show ?thesis by auto
       
   251       next
       
   252         case False
       
   253         from block_pre [OF vt_e False h1]
       
   254         have "e = P thread cs2" .
       
   255         with vt_e have "vt step ((P thread cs2)# moment t2 s)" by simp
       
   256         from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp
       
   257         with runing_ready have "thread \<in> readys (moment t2 s)" by auto
       
   258         with nn1 [rule_format, OF lt12]
       
   259         show ?thesis  by (simp add:readys_def s_waiting_def, auto)
       
   260       qed
       
   261     } moreover {
       
   262       assume lt12: "t2 < t1"
       
   263       let ?t3 = "Suc t1"
       
   264       from lt1 have le_t3: "?t3 \<le> length s" by auto
       
   265       from moment_plus [OF this] 
       
   266       obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
       
   267       have lt_t3: "t1 < ?t3" by simp
       
   268       from nn1 [rule_format, OF this] and eq_m
       
   269       have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   270         h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
       
   271       have vt_e: "vt step (e#moment t1 s)"
       
   272       proof -
       
   273         from vt_moment [OF vt le_t3]
       
   274         have "vt step (moment ?t3 s)" .
       
   275         with eq_m show ?thesis by simp
       
   276       qed
       
   277       have ?thesis
       
   278       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
       
   279         case True
       
   280         from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
       
   281           by auto
       
   282         from abs2 [OF vt_e True eq_th h2 h1]
       
   283         show ?thesis by auto
       
   284       next
       
   285         case False
       
   286         from block_pre [OF vt_e False h1]
       
   287         have "e = P thread cs1" .
       
   288         with vt_e have "vt step ((P thread cs1)# moment t1 s)" by simp
       
   289         from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp
       
   290         with runing_ready have "thread \<in> readys (moment t1 s)" by auto
       
   291         with nn2 [rule_format, OF lt12]
       
   292         show ?thesis  by (simp add:readys_def s_waiting_def, auto)
       
   293       qed
       
   294     } moreover {
       
   295       assume eqt12: "t1 = t2"
       
   296       let ?t3 = "Suc t1"
       
   297       from lt1 have le_t3: "?t3 \<le> length s" by auto
       
   298       from moment_plus [OF this] 
       
   299       obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
       
   300       have lt_t3: "t1 < ?t3" by simp
       
   301       from nn1 [rule_format, OF this] and eq_m
       
   302       have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   303         h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
       
   304       have vt_e: "vt step (e#moment t1 s)"
       
   305       proof -
       
   306         from vt_moment [OF vt le_t3]
       
   307         have "vt step (moment ?t3 s)" .
       
   308         with eq_m show ?thesis by simp
       
   309       qed
       
   310       have ?thesis
       
   311       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
       
   312         case True
       
   313         from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
       
   314           by auto
       
   315         from abs2 [OF vt_e True eq_th h2 h1]
       
   316         show ?thesis by auto
       
   317       next
       
   318         case False
       
   319         from block_pre [OF vt_e False h1]
       
   320         have eq_e1: "e = P thread cs1" .
       
   321         have lt_t3: "t1 < ?t3" by simp
       
   322         with eqt12 have "t2 < ?t3" by simp
       
   323         from nn2 [rule_format, OF this] and eq_m and eqt12
       
   324         have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   325           h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   326         show ?thesis
       
   327         proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   328           case True
       
   329           from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
       
   330             by auto
       
   331           from vt_e and eqt12 have "vt step (e#moment t2 s)" by simp 
       
   332           from abs2 [OF this True eq_th h2 h1]
       
   333           show ?thesis .
       
   334         next
       
   335           case False
       
   336           have vt_e: "vt step (e#moment t2 s)"
       
   337           proof -
       
   338             from vt_moment [OF vt le_t3] eqt12
       
   339             have "vt step (moment (Suc t2) s)" by auto
       
   340             with eq_m eqt12 show ?thesis by simp
       
   341           qed
       
   342           from block_pre [OF vt_e False h1]
       
   343           have "e = P thread cs2" .
       
   344           with eq_e1 neq12 show ?thesis by auto
       
   345         qed
       
   346       qed
       
   347     } ultimately show ?thesis by arith
       
   348   qed
       
   349 qed
       
   350 
       
   351 lemma waiting_unique:
       
   352   assumes "vt step s"
       
   353   and "waiting s th cs1"
       
   354   and "waiting s th cs2"
       
   355   shows "cs1 = cs2"
       
   356 proof -
       
   357   from waiting_unique_pre and prems
       
   358   show ?thesis
       
   359     by (auto simp add:s_waiting_def)
       
   360 qed
       
   361 
       
   362 lemma holded_unique:
       
   363   assumes "vt step s"
       
   364   and "holding s th1 cs"
       
   365   and "holding s th2 cs"
       
   366   shows "th1 = th2"
       
   367 proof -
       
   368   from prems show ?thesis
       
   369     unfolding s_holding_def
       
   370     by auto
       
   371 qed
       
   372 
       
   373 lemma birthtime_lt: "th \<in> threads s \<Longrightarrow> birthtime th s < length s"
       
   374   apply (induct s, auto)
       
   375   by (case_tac a, auto split:if_splits)
       
   376 
       
   377 lemma birthtime_unique: 
       
   378   "\<lbrakk>birthtime th1 s = birthtime th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
       
   379           \<Longrightarrow> th1 = th2"
       
   380   apply (induct s, auto)
       
   381   by (case_tac a, auto split:if_splits dest:birthtime_lt)
       
   382 
       
   383 lemma preced_unique : 
       
   384   assumes pcd_eq: "preced th1 s = preced th2 s"
       
   385   and th_in1: "th1 \<in> threads s"
       
   386   and th_in2: " th2 \<in> threads s"
       
   387   shows "th1 = th2"
       
   388 proof -
       
   389   from pcd_eq have "birthtime th1 s = birthtime th2 s" by (simp add:preced_def)
       
   390   from birthtime_unique [OF this th_in1 th_in2]
       
   391   show ?thesis .
       
   392 qed
       
   393 
       
   394 lemma preced_linorder: 
       
   395   assumes neq_12: "th1 \<noteq> th2"
       
   396   and th_in1: "th1 \<in> threads s"
       
   397   and th_in2: " th2 \<in> threads s"
       
   398   shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
       
   399 proof -
       
   400   from preced_unique [OF _ th_in1 th_in2] and neq_12 
       
   401   have "preced th1 s \<noteq> preced th2 s" by auto
       
   402   thus ?thesis by auto
       
   403 qed
       
   404 
       
   405 lemma unique_minus:
       
   406   fixes x y z r
       
   407   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   408   and xy: "(x, y) \<in> r"
       
   409   and xz: "(x, z) \<in> r^+"
       
   410   and neq: "y \<noteq> z"
       
   411   shows "(y, z) \<in> r^+"
       
   412 proof -
       
   413  from xz and neq show ?thesis
       
   414  proof(induct)
       
   415    case (base ya)
       
   416    have "(x, ya) \<in> r" by fact
       
   417    from unique [OF xy this] have "y = ya" .
       
   418    with base show ?case by auto
       
   419  next
       
   420    case (step ya z)
       
   421    show ?case
       
   422    proof(cases "y = ya")
       
   423      case True
       
   424      from step True show ?thesis by simp
       
   425    next
       
   426      case False
       
   427      from step False
       
   428      show ?thesis by auto
       
   429    qed
       
   430  qed
       
   431 qed
       
   432 
       
   433 lemma unique_base:
       
   434   fixes r x y z
       
   435   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   436   and xy: "(x, y) \<in> r"
       
   437   and xz: "(x, z) \<in> r^+"
       
   438   and neq_yz: "y \<noteq> z"
       
   439   shows "(y, z) \<in> r^+"
       
   440 proof -
       
   441   from xz neq_yz show ?thesis
       
   442   proof(induct)
       
   443     case (base ya)
       
   444     from xy unique base show ?case by auto
       
   445   next
       
   446     case (step ya z)
       
   447     show ?case
       
   448     proof(cases "y = ya")
       
   449       case True
       
   450       from True step show ?thesis by auto
       
   451     next
       
   452       case False
       
   453       from False step 
       
   454       have "(y, ya) \<in> r\<^sup>+" by auto
       
   455       with step show ?thesis by auto
       
   456     qed
       
   457   qed
       
   458 qed
       
   459 
       
   460 lemma unique_chain:
       
   461   fixes r x y z
       
   462   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   463   and xy: "(x, y) \<in> r^+"
       
   464   and xz: "(x, z) \<in> r^+"
       
   465   and neq_yz: "y \<noteq> z"
       
   466   shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
       
   467 proof -
       
   468   from xy xz neq_yz show ?thesis
       
   469   proof(induct)
       
   470     case (base y)
       
   471     have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
       
   472     from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
       
   473   next
       
   474     case (step y za)
       
   475     show ?case
       
   476     proof(cases "y = z")
       
   477       case True
       
   478       from True step show ?thesis by auto
       
   479     next
       
   480       case False
       
   481       from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
       
   482       thus ?thesis
       
   483       proof
       
   484         assume "(z, y) \<in> r\<^sup>+"
       
   485         with step have "(z, za) \<in> r\<^sup>+" by auto
       
   486         thus ?thesis by auto
       
   487       next
       
   488         assume h: "(y, z) \<in> r\<^sup>+"
       
   489         from step have yza: "(y, za) \<in> r" by simp
       
   490         from step have "za \<noteq> z" by simp
       
   491         from unique_minus [OF _ yza h this] and unique
       
   492         have "(za, z) \<in> r\<^sup>+" by auto
       
   493         thus ?thesis by auto
       
   494       qed
       
   495     qed
       
   496   qed
       
   497 qed
       
   498 
       
   499 lemma depend_set_unchanged: "(depend (Set th prio # s)) = depend s"
       
   500 apply (unfold s_depend_def s_waiting_def wq_def)
       
   501 by (simp add:Let_def)
       
   502 
       
   503 lemma depend_create_unchanged: "(depend (Create th prio # s)) = depend s"
       
   504 apply (unfold s_depend_def s_waiting_def wq_def)
       
   505 by (simp add:Let_def)
       
   506 
       
   507 lemma depend_exit_unchanged: "(depend (Exit th # s)) = depend s"
       
   508 apply (unfold s_depend_def s_waiting_def wq_def)
       
   509 by (simp add:Let_def)
       
   510 
       
   511 
       
   512 
       
   513 lemma step_v_hold_inv[elim_format]:
       
   514   "\<And>c t. \<lbrakk>vt step (V th cs # s); 
       
   515   \<not> holding (wq s) t c; holding (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> next_th s th cs t \<and> c = cs"
       
   516 proof -
       
   517   fix c t
       
   518   assume vt: "vt step (V th cs # s)"
       
   519     and nhd: "\<not> holding (wq s) t c"
       
   520     and hd: "holding (wq (V th cs # s)) t c"
       
   521   show "next_th s th cs t \<and> c = cs"
       
   522   proof(cases "c = cs")
       
   523     case False
       
   524     with nhd hd show ?thesis
       
   525       by (unfold cs_holding_def wq_def, auto simp:Let_def)
       
   526   next
       
   527     case True
       
   528     with step_back_step [OF vt] 
       
   529     have "step s (V th c)" by simp
       
   530     hence "next_th s th cs t"
       
   531     proof(cases)
       
   532       assume "holding s th c"
       
   533       with nhd hd show ?thesis
       
   534         apply (unfold s_holding_def cs_holding_def wq_def next_th_def,
       
   535                auto simp:Let_def split:list.splits if_splits)
       
   536         proof -
       
   537           assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
       
   538           moreover have "\<dots> = set []"
       
   539           proof(rule someI2)
       
   540             show "distinct [] \<and> [] = []" by auto
       
   541           next
       
   542             fix x assume "distinct x \<and> x = []"
       
   543             thus "set x = set []" by auto
       
   544           qed
       
   545           ultimately show False by auto
       
   546         next
       
   547           assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
       
   548           moreover have "\<dots> = set []"
       
   549           proof(rule someI2)
       
   550             show "distinct [] \<and> [] = []" by auto
       
   551           next
       
   552             fix x assume "distinct x \<and> x = []"
       
   553             thus "set x = set []" by auto
       
   554           qed
       
   555           ultimately show False by auto
       
   556         qed
       
   557     qed
       
   558     with True show ?thesis by auto
       
   559   qed
       
   560 qed
       
   561 
       
   562 lemma step_v_wait_inv[elim_format]:
       
   563     "\<And>t c. \<lbrakk>vt step (V th cs # s); \<not> waiting (wq (V th cs # s)) t c; waiting (wq s) t c
       
   564            \<rbrakk>
       
   565           \<Longrightarrow> (next_th s th cs t \<and> cs = c)"
       
   566 proof -
       
   567   fix t c 
       
   568   assume vt: "vt step (V th cs # s)"
       
   569     and nw: "\<not> waiting (wq (V th cs # s)) t c"
       
   570     and wt: "waiting (wq s) t c"
       
   571   show "next_th s th cs t \<and> cs = c"
       
   572   proof(cases "cs = c")
       
   573     case False
       
   574     with nw wt show ?thesis
       
   575       by (auto simp:cs_waiting_def wq_def Let_def)
       
   576   next
       
   577     case True
       
   578     from nw[folded True] wt[folded True]
       
   579     have "next_th s th cs t"
       
   580       apply (unfold next_th_def, auto simp:cs_waiting_def wq_def Let_def split:list.splits)
       
   581     proof -
       
   582       fix a list
       
   583       assume t_in: "t \<in> set list"
       
   584         and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
       
   585         and eq_wq: "waiting_queue (schs s) cs = a # list"
       
   586       have " set (SOME q. distinct q \<and> set q = set list) = set list"
       
   587       proof(rule someI2)
       
   588         from wq_distinct[OF step_back_vt[OF vt], of cs] and eq_wq[folded wq_def]
       
   589         show "distinct list \<and> set list = set list" by auto
       
   590       next
       
   591         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
       
   592           by auto
       
   593       qed
       
   594       with t_ni and t_in show "a = th" by auto
       
   595     next
       
   596       fix a list
       
   597       assume t_in: "t \<in> set list"
       
   598         and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
       
   599         and eq_wq: "waiting_queue (schs s) cs = a # list"
       
   600       have " set (SOME q. distinct q \<and> set q = set list) = set list"
       
   601       proof(rule someI2)
       
   602         from wq_distinct[OF step_back_vt[OF vt], of cs] and eq_wq[folded wq_def]
       
   603         show "distinct list \<and> set list = set list" by auto
       
   604       next
       
   605         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
       
   606           by auto
       
   607       qed
       
   608       with t_ni and t_in show "t = hd (SOME q. distinct q \<and> set q = set list)" by auto
       
   609     next
       
   610       fix a list
       
   611       assume eq_wq: "waiting_queue (schs s) cs = a # list"
       
   612       from step_back_step[OF vt]
       
   613       show "a = th"
       
   614       proof(cases)
       
   615         assume "holding s th cs"
       
   616         with eq_wq show ?thesis
       
   617           by (unfold s_holding_def wq_def, auto)
       
   618       qed
       
   619     qed
       
   620     with True show ?thesis by simp
       
   621   qed
       
   622 qed
       
   623 
       
   624 lemma step_v_not_wait[consumes 3]:
       
   625   "\<lbrakk>vt step (V th cs # s); next_th s th cs t; waiting (wq (V th cs # s)) t cs\<rbrakk> \<Longrightarrow> False"
       
   626   by (unfold next_th_def cs_waiting_def wq_def, auto simp:Let_def)
       
   627 
       
   628 lemma step_v_release:
       
   629   "\<lbrakk>vt step (V th cs # s); holding (wq (V th cs # s)) th cs\<rbrakk> \<Longrightarrow> False"
       
   630 proof -
       
   631   assume vt: "vt step (V th cs # s)"
       
   632     and hd: "holding (wq (V th cs # s)) th cs"
       
   633   from step_back_step [OF vt] and hd
       
   634   show "False"
       
   635   proof(cases)
       
   636     assume "holding (wq (V th cs # s)) th cs" and "holding s th cs"
       
   637     thus ?thesis
       
   638       apply (unfold s_holding_def wq_def cs_holding_def)
       
   639       apply (auto simp:Let_def split:list.splits)
       
   640     proof -
       
   641       fix list
       
   642       assume eq_wq[folded wq_def]: 
       
   643         "waiting_queue (schs s) cs = hd (SOME q. distinct q \<and> set q = set list) # list"
       
   644       and hd_in: "hd (SOME q. distinct q \<and> set q = set list)
       
   645             \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   646       have "set (SOME q. distinct q \<and> set q = set list) = set list"
       
   647       proof(rule someI2)
       
   648         from wq_distinct[OF step_back_vt[OF vt], of cs] and eq_wq
       
   649         show "distinct list \<and> set list = set list" by auto
       
   650       next
       
   651         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
       
   652           by auto
       
   653       qed
       
   654       moreover have "distinct  (hd (SOME q. distinct q \<and> set q = set list) # list)"
       
   655       proof -
       
   656         from wq_distinct[OF step_back_vt[OF vt], of cs] and eq_wq
       
   657         show ?thesis by auto
       
   658       qed
       
   659       moreover note eq_wq and hd_in
       
   660       ultimately show "False" by auto
       
   661     qed
       
   662   qed
       
   663 qed
       
   664 
       
   665 lemma step_v_get_hold:
       
   666   "\<And>th'. \<lbrakk>vt step (V th cs # s); \<not> holding (wq (V th cs # s)) th' cs; next_th s th cs th'\<rbrakk> \<Longrightarrow> False"
       
   667   apply (unfold cs_holding_def next_th_def wq_def,
       
   668          auto simp:Let_def)
       
   669 proof -
       
   670   fix rest
       
   671   assume vt: "vt step (V th cs # s)"
       
   672     and eq_wq[folded wq_def]: " waiting_queue (schs s) cs = th # rest"
       
   673     and nrest: "rest \<noteq> []"
       
   674     and ni: "hd (SOME q. distinct q \<and> set q = set rest)
       
   675             \<notin> set (SOME q. distinct q \<and> set q = set rest)"
       
   676   have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
   677   proof(rule someI2)
       
   678     from wq_distinct[OF step_back_vt[OF vt], of cs] and eq_wq
       
   679     show "distinct rest \<and> set rest = set rest" by auto
       
   680   next
       
   681     fix x assume "distinct x \<and> set x = set rest"
       
   682     hence "set x = set rest" by auto
       
   683     with nrest
       
   684     show "x \<noteq> []" by (case_tac x, auto)
       
   685   qed
       
   686   with ni show "False" by auto
       
   687 qed
       
   688 
       
   689 lemma step_v_release_inv[elim_format]:
       
   690 "\<And>c t. \<lbrakk>vt step (V th cs # s); \<not> holding (wq (V th cs # s)) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> 
       
   691   c = cs \<and> t = th"
       
   692   apply (unfold cs_holding_def wq_def, auto simp:Let_def split:if_splits list.splits)
       
   693   proof -
       
   694     fix a list
       
   695     assume vt: "vt step (V th cs # s)" and eq_wq: "waiting_queue (schs s) cs = a # list"
       
   696     from step_back_step [OF vt] show "a = th"
       
   697     proof(cases)
       
   698       assume "holding s th cs" with eq_wq
       
   699       show ?thesis
       
   700         by (unfold s_holding_def wq_def, auto)
       
   701     qed
       
   702   next
       
   703     fix a list
       
   704     assume vt: "vt step (V th cs # s)" and eq_wq: "waiting_queue (schs s) cs = a # list"
       
   705     from step_back_step [OF vt] show "a = th"
       
   706     proof(cases)
       
   707       assume "holding s th cs" with eq_wq
       
   708       show ?thesis
       
   709         by (unfold s_holding_def wq_def, auto)
       
   710     qed
       
   711   qed
       
   712 
       
   713 lemma step_v_waiting_mono:
       
   714   "\<And>t c. \<lbrakk>vt step (V th cs # s); waiting (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> waiting (wq s) t c"
       
   715 proof -
       
   716   fix t c
       
   717   let ?s' = "(V th cs # s)"
       
   718   assume vt: "vt step ?s'" 
       
   719     and wt: "waiting (wq ?s') t c"
       
   720   show "waiting (wq s) t c"
       
   721   proof(cases "c = cs")
       
   722     case False
       
   723     assume neq_cs: "c \<noteq> cs"
       
   724     hence "waiting (wq ?s') t c = waiting (wq s) t c"
       
   725       by (unfold cs_waiting_def wq_def, auto simp:Let_def)
       
   726     with wt show ?thesis by simp
       
   727   next
       
   728     case True
       
   729     with wt show ?thesis
       
   730       apply (unfold cs_waiting_def wq_def, auto simp:Let_def split:list.splits)
       
   731     proof -
       
   732       fix a list
       
   733       assume not_in: "t \<notin> set list"
       
   734         and is_in: "t \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   735         and eq_wq: "waiting_queue (schs s) cs = a # list"
       
   736       have "set (SOME q. distinct q \<and> set q = set list) = set list"
       
   737       proof(rule someI2)
       
   738         from wq_distinct [OF step_back_vt[OF vt], of cs]
       
   739         and eq_wq[folded wq_def]
       
   740         show "distinct list \<and> set list = set list" by auto
       
   741       next
       
   742         fix x assume "distinct x \<and> set x = set list"
       
   743         thus "set x = set list" by auto
       
   744       qed
       
   745       with not_in is_in show "t = a" by auto
       
   746     next
       
   747       fix list
       
   748       assume is_waiting: "waiting (wq (V th cs # s)) t cs"
       
   749       and eq_wq: "waiting_queue (schs s) cs = t # list"
       
   750       hence "t \<in> set list"
       
   751         apply (unfold wq_def, auto simp:Let_def cs_waiting_def)
       
   752       proof -
       
   753         assume " t \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   754         moreover have "\<dots> = set list" 
       
   755         proof(rule someI2)
       
   756           from wq_distinct [OF step_back_vt[OF vt], of cs]
       
   757             and eq_wq[folded wq_def]
       
   758           show "distinct list \<and> set list = set list" by auto
       
   759         next
       
   760           fix x assume "distinct x \<and> set x = set list" 
       
   761           thus "set x = set list" by auto
       
   762         qed
       
   763         ultimately show "t \<in> set list" by simp
       
   764       qed
       
   765       with eq_wq and wq_distinct [OF step_back_vt[OF vt], of cs, unfolded wq_def]
       
   766       show False by auto
       
   767     qed
       
   768   qed
       
   769 qed
       
   770 
       
   771 lemma step_depend_v:
       
   772 assumes vt:
       
   773   "vt step (V th cs#s)"
       
   774 shows "
       
   775   depend (V th cs # s) =
       
   776   depend s - {(Cs cs, Th th)} -
       
   777   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
       
   778   {(Cs cs, Th th') |th'.  next_th s th cs th'}"
       
   779   apply (insert vt, unfold s_depend_def) 
       
   780   apply (auto split:if_splits list.splits simp:Let_def)
       
   781   apply (auto elim: step_v_waiting_mono step_v_hold_inv 
       
   782               step_v_release step_v_wait_inv
       
   783               step_v_get_hold step_v_release_inv)
       
   784   apply (erule_tac step_v_not_wait, auto)
       
   785   done
       
   786 
       
   787 lemma step_depend_p:
       
   788   "vt step (P th cs#s) \<Longrightarrow>
       
   789   depend (P th cs # s) =  (if (wq s cs = []) then depend s \<union> {(Cs cs, Th th)}
       
   790                                              else depend s \<union> {(Th th, Cs cs)})"
       
   791   apply(unfold s_depend_def wq_def)
       
   792   apply (auto split:list.splits prod.splits simp:Let_def cs_waiting_def cs_holding_def)
       
   793   apply(case_tac "c = cs", auto)
       
   794   apply(fold wq_def)
       
   795   apply(drule_tac step_back_step)
       
   796   by (ind_cases " step s (P (hd (wq s cs)) cs)", 
       
   797     auto simp:s_depend_def wq_def cs_holding_def)
       
   798 
       
   799 lemma simple_A:
       
   800   fixes A
       
   801   assumes h: "\<And> x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x = y"
       
   802   shows "A = {} \<or> (\<exists> a. A = {a})"
       
   803 proof(cases "A = {}")
       
   804   case True thus ?thesis by simp
       
   805 next
       
   806   case False then obtain a where "a \<in> A" by auto
       
   807   with h have "A = {a}" by auto
       
   808   thus ?thesis by simp
       
   809 qed
       
   810 
       
   811 lemma depend_target_th: "(Th th, x) \<in> depend (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
       
   812   by (unfold s_depend_def, auto)
       
   813 
       
   814 lemma acyclic_depend: 
       
   815   fixes s
       
   816   assumes vt: "vt step s"
       
   817   shows "acyclic (depend s)"
       
   818 proof -
       
   819   from vt show ?thesis
       
   820   proof(induct)
       
   821     case (vt_cons s e)
       
   822     assume ih: "acyclic (depend s)"
       
   823       and stp: "step s e"
       
   824       and vt: "vt step s"
       
   825     show ?case
       
   826     proof(cases e)
       
   827       case (Create th prio)
       
   828       with ih
       
   829       show ?thesis by (simp add:depend_create_unchanged)
       
   830     next
       
   831       case (Exit th)
       
   832       with ih show ?thesis by (simp add:depend_exit_unchanged)
       
   833     next
       
   834       case (V th cs)
       
   835       from V vt stp have vtt: "vt step (V th cs#s)" by auto
       
   836       from step_depend_v [OF this]
       
   837       have eq_de: 
       
   838         "depend (e # s) = 
       
   839             depend s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
       
   840             {(Cs cs, Th th') |th'. next_th s th cs th'}"
       
   841         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
       
   842       from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
       
   843       from step_back_step [OF vtt]
       
   844       have "step s (V th cs)" .
       
   845       thus ?thesis
       
   846       proof(cases)
       
   847         assume "holding s th cs"
       
   848         hence th_in: "th \<in> set (wq s cs)" and
       
   849           eq_hd: "th = hd (wq s cs)" by (unfold s_holding_def, auto)
       
   850         then obtain rest where
       
   851           eq_wq: "wq s cs = th#rest"
       
   852           by (cases "wq s cs", auto)
       
   853         show ?thesis
       
   854         proof(cases "rest = []")
       
   855           case False
       
   856           let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
   857           from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
       
   858             by (unfold next_th_def, auto)
       
   859           let ?E = "(?A - ?B - ?C)"
       
   860           have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*"
       
   861           proof
       
   862             assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
       
   863             hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
       
   864             from tranclD [OF this]
       
   865             obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
       
   866             hence th_d: "(Th ?th', x) \<in> ?A" by simp
       
   867             from depend_target_th [OF this]
       
   868             obtain cs' where eq_x: "x = Cs cs'" by auto
       
   869             with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
       
   870             hence wt_th': "waiting s ?th' cs'"
       
   871               unfolding s_depend_def s_waiting_def cs_waiting_def by simp
       
   872             hence "cs' = cs"
       
   873             proof(rule waiting_unique [OF vt])
       
   874               from eq_wq wq_distinct[OF vt, of cs]
       
   875               show "waiting s ?th' cs" 
       
   876                 apply (unfold s_waiting_def, auto)
       
   877               proof -
       
   878                 assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
   879                 and eq_wq: "wq s cs = th # rest"
       
   880                 have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
   881                 proof(rule someI2)
       
   882                   from wq_distinct[OF vt, of cs] and eq_wq
       
   883                   show "distinct rest \<and> set rest = set rest" by auto
       
   884                 next
       
   885                   fix x assume "distinct x \<and> set x = set rest"
       
   886                   with False show "x \<noteq> []" by auto
       
   887                 qed
       
   888                 hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
       
   889                                   set (SOME q. distinct q \<and> set q = set rest)" by auto
       
   890                 moreover have "\<dots> = set rest" 
       
   891                 proof(rule someI2)
       
   892                   from wq_distinct[OF vt, of cs] and eq_wq
       
   893                   show "distinct rest \<and> set rest = set rest" by auto
       
   894                 next
       
   895                   show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
   896                 qed
       
   897                 moreover note hd_in
       
   898                 ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
       
   899               next
       
   900                 assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
   901                 and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
       
   902                 have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
   903                 proof(rule someI2)
       
   904                   from wq_distinct[OF vt, of cs] and eq_wq
       
   905                   show "distinct rest \<and> set rest = set rest" by auto
       
   906                 next
       
   907                   fix x assume "distinct x \<and> set x = set rest"
       
   908                   with False show "x \<noteq> []" by auto
       
   909                 qed
       
   910                 hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
       
   911                                   set (SOME q. distinct q \<and> set q = set rest)" by auto
       
   912                 moreover have "\<dots> = set rest" 
       
   913                 proof(rule someI2)
       
   914                   from wq_distinct[OF vt, of cs] and eq_wq
       
   915                   show "distinct rest \<and> set rest = set rest" by auto
       
   916                 next
       
   917                   show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
   918                 qed
       
   919                 moreover note hd_in
       
   920                 ultimately show False by auto
       
   921               qed
       
   922             qed
       
   923             with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp
       
   924             with False
       
   925             show "False" by (auto simp: next_th_def eq_wq)
       
   926           qed
       
   927           with acyclic_insert[symmetric] and ac
       
   928             and eq_de eq_D show ?thesis by auto
       
   929         next
       
   930           case True
       
   931           with eq_wq
       
   932           have eq_D: "?D = {}"
       
   933             by (unfold next_th_def, auto)
       
   934           with eq_de ac
       
   935           show ?thesis by auto
       
   936         qed 
       
   937       qed
       
   938   next
       
   939     case (P th cs)
       
   940     from P vt stp have vtt: "vt step (P th cs#s)" by auto
       
   941     from step_depend_p [OF this] P
       
   942     have "depend (e # s) = 
       
   943       (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
       
   944       depend s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
       
   945       by simp
       
   946     moreover have "acyclic ?R"
       
   947     proof(cases "wq s cs = []")
       
   948       case True
       
   949       hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
       
   950       have "(Th th, Cs cs) \<notin> (depend s)\<^sup>*"
       
   951       proof
       
   952         assume "(Th th, Cs cs) \<in> (depend s)\<^sup>*"
       
   953         hence "(Th th, Cs cs) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
       
   954         from tranclD2 [OF this]
       
   955         obtain x where "(x, Cs cs) \<in> depend s" by auto
       
   956         with True show False by (auto simp:s_depend_def cs_waiting_def)
       
   957       qed
       
   958       with acyclic_insert ih eq_r show ?thesis by auto
       
   959     next
       
   960       case False
       
   961       hence eq_r: "?R =  depend s \<union> {(Th th, Cs cs)}" by simp
       
   962       have "(Cs cs, Th th) \<notin> (depend s)\<^sup>*"
       
   963       proof
       
   964         assume "(Cs cs, Th th) \<in> (depend s)\<^sup>*"
       
   965         hence "(Cs cs, Th th) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
       
   966           moreover from step_back_step [OF vtt] have "step s (P th cs)" .
       
   967           ultimately show False
       
   968           proof -
       
   969             show " \<lbrakk>(Cs cs, Th th) \<in> (depend s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
       
   970               by (ind_cases "step s (P th cs)", simp)
       
   971           qed
       
   972         qed
       
   973         with acyclic_insert ih eq_r show ?thesis by auto
       
   974       qed
       
   975       ultimately show ?thesis by simp
       
   976     next
       
   977       case (Set thread prio)
       
   978       with ih
       
   979       thm depend_set_unchanged
       
   980       show ?thesis by (simp add:depend_set_unchanged)
       
   981     qed
       
   982   next
       
   983     case vt_nil
       
   984     show "acyclic (depend ([]::state))"
       
   985       by (auto simp: s_depend_def cs_waiting_def 
       
   986                       cs_holding_def wq_def acyclic_def)
       
   987   qed
       
   988 qed
       
   989 
       
   990 lemma finite_depend: 
       
   991   fixes s
       
   992   assumes vt: "vt step s"
       
   993   shows "finite (depend s)"
       
   994 proof -
       
   995   from vt show ?thesis
       
   996   proof(induct)
       
   997     case (vt_cons s e)
       
   998     assume ih: "finite (depend s)"
       
   999       and stp: "step s e"
       
  1000       and vt: "vt step s"
       
  1001     show ?case
       
  1002     proof(cases e)
       
  1003       case (Create th prio)
       
  1004       with ih
       
  1005       show ?thesis by (simp add:depend_create_unchanged)
       
  1006     next
       
  1007       case (Exit th)
       
  1008       with ih show ?thesis by (simp add:depend_exit_unchanged)
       
  1009     next
       
  1010       case (V th cs)
       
  1011       from V vt stp have vtt: "vt step (V th cs#s)" by auto
       
  1012       from step_depend_v [OF this]
       
  1013       have eq_de: "depend (e # s) = 
       
  1014                    depend s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
       
  1015                       {(Cs cs, Th th') |th'. next_th s th cs th'}
       
  1016 "
       
  1017         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
       
  1018       moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
       
  1019       moreover have "finite ?D"
       
  1020       proof -
       
  1021         have "?D = {} \<or> (\<exists> a. ?D = {a})" 
       
  1022           by (unfold next_th_def, auto)
       
  1023         thus ?thesis
       
  1024         proof
       
  1025           assume h: "?D = {}"
       
  1026           show ?thesis by (unfold h, simp)
       
  1027         next
       
  1028           assume "\<exists> a. ?D = {a}"
       
  1029           thus ?thesis by auto
       
  1030         qed
       
  1031       qed
       
  1032       ultimately show ?thesis by simp
       
  1033     next
       
  1034       case (P th cs)
       
  1035       from P vt stp have vtt: "vt step (P th cs#s)" by auto
       
  1036       from step_depend_p [OF this] P
       
  1037       have "depend (e # s) = 
       
  1038               (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
       
  1039                                     depend s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
       
  1040         by simp
       
  1041       moreover have "finite ?R"
       
  1042       proof(cases "wq s cs = []")
       
  1043         case True
       
  1044         hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
       
  1045         with True and ih show ?thesis by auto
       
  1046       next
       
  1047         case False
       
  1048         hence "?R = depend s \<union> {(Th th, Cs cs)}" by simp
       
  1049         with False and ih show ?thesis by auto
       
  1050       qed
       
  1051       ultimately show ?thesis by auto
       
  1052     next
       
  1053       case (Set thread prio)
       
  1054       with ih
       
  1055       show ?thesis by (simp add:depend_set_unchanged)
       
  1056     qed
       
  1057   next
       
  1058     case vt_nil
       
  1059     show "finite (depend ([]::state))"
       
  1060       by (auto simp: s_depend_def cs_waiting_def 
       
  1061                    cs_holding_def wq_def acyclic_def)
       
  1062   qed
       
  1063 qed
       
  1064 
       
  1065 text {* Several useful lemmas *}
       
  1066 
       
  1067 lemma wf_dep_converse: 
       
  1068   fixes s
       
  1069   assumes vt: "vt step s"
       
  1070   shows "wf ((depend s)^-1)"
       
  1071 proof(rule finite_acyclic_wf_converse)
       
  1072   from finite_depend [OF vt]
       
  1073   show "finite (depend s)" .
       
  1074 next
       
  1075   from acyclic_depend[OF vt]
       
  1076   show "acyclic (depend s)" .
       
  1077 qed
       
  1078 
       
  1079 lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
       
  1080 by (induct l, auto)
       
  1081 
       
  1082 lemma th_chasing: "(Th th, Cs cs) \<in> depend (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> depend s"
       
  1083   by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1084 
       
  1085 lemma wq_threads: 
       
  1086   fixes s cs
       
  1087   assumes vt: "vt step s"
       
  1088   and h: "th \<in> set (wq s cs)"
       
  1089   shows "th \<in> threads s"
       
  1090 proof -
       
  1091  from vt and h show ?thesis
       
  1092   proof(induct arbitrary: th cs)
       
  1093     case (vt_cons s e)
       
  1094     assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
       
  1095       and stp: "step s e"
       
  1096       and vt: "vt step s"
       
  1097       and h: "th \<in> set (wq (e # s) cs)"
       
  1098     show ?case
       
  1099     proof(cases e)
       
  1100       case (Create th' prio)
       
  1101       with ih h show ?thesis
       
  1102         by (auto simp:wq_def Let_def)
       
  1103     next
       
  1104       case (Exit th')
       
  1105       with stp ih h show ?thesis
       
  1106         apply (auto simp:wq_def Let_def)
       
  1107         apply (ind_cases "step s (Exit th')")
       
  1108         apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
       
  1109                s_depend_def s_holding_def cs_holding_def)
       
  1110         by (fold wq_def, auto)
       
  1111     next
       
  1112       case (V th' cs')
       
  1113       show ?thesis
       
  1114       proof(cases "cs' = cs")
       
  1115         case False
       
  1116         with h
       
  1117         show ?thesis
       
  1118           apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
       
  1119           by (drule_tac ih, simp)
       
  1120       next
       
  1121         case True
       
  1122         from h
       
  1123         show ?thesis
       
  1124         proof(unfold V wq_def)
       
  1125           assume th_in: "th \<in> set (waiting_queue (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
       
  1126           show "th \<in> threads (V th' cs' # s)"
       
  1127           proof(cases "cs = cs'")
       
  1128             case False
       
  1129             hence "?l = waiting_queue (schs s) cs" by (simp add:Let_def)
       
  1130             with th_in have " th \<in> set (wq s cs)" 
       
  1131               by (fold wq_def, simp)
       
  1132             from ih [OF this] show ?thesis by simp
       
  1133           next
       
  1134             case True
       
  1135             show ?thesis
       
  1136             proof(cases "waiting_queue (schs s) cs'")
       
  1137               case Nil
       
  1138               with h V show ?thesis
       
  1139                 apply (auto simp:wq_def Let_def split:if_splits)
       
  1140                 by (fold wq_def, drule_tac ih, simp)
       
  1141             next
       
  1142               case (Cons a rest)
       
  1143               assume eq_wq: "waiting_queue (schs s) cs' = a # rest"
       
  1144               with h V show ?thesis
       
  1145                 apply (auto simp:Let_def wq_def split:if_splits)
       
  1146               proof -
       
  1147                 assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1148                 have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
       
  1149                 proof(rule someI2)
       
  1150                   from wq_distinct[OF vt, of cs'] and eq_wq[folded wq_def]
       
  1151                   show "distinct rest \<and> set rest = set rest" by auto
       
  1152                 next
       
  1153                   show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
       
  1154                     by auto
       
  1155                 qed
       
  1156                 with eq_wq th_in have "th \<in> set (waiting_queue (schs s) cs')" by auto
       
  1157                 from ih[OF this[folded wq_def]] show "th \<in> threads s" .
       
  1158               next
       
  1159                 assume th_in: "th \<in> set (waiting_queue (schs s) cs)"
       
  1160                 from ih[OF this[folded wq_def]]
       
  1161                 show "th \<in> threads s" .
       
  1162               qed
       
  1163             qed
       
  1164           qed
       
  1165         qed
       
  1166       qed
       
  1167     next
       
  1168       case (P th' cs')
       
  1169       from h stp
       
  1170       show ?thesis
       
  1171         apply (unfold P wq_def)
       
  1172         apply (auto simp:Let_def split:if_splits, fold wq_def)
       
  1173         apply (auto intro:ih)
       
  1174         apply(ind_cases "step s (P th' cs')")
       
  1175         by (unfold runing_def readys_def, auto)
       
  1176     next
       
  1177       case (Set thread prio)
       
  1178       with ih h show ?thesis
       
  1179         by (auto simp:wq_def Let_def)
       
  1180     qed
       
  1181   next
       
  1182     case vt_nil
       
  1183     thus ?case by (auto simp:wq_def)
       
  1184   qed
       
  1185 qed
       
  1186 
       
  1187 lemma range_in: "\<lbrakk>vt step s; (Th th) \<in> Range (depend (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
       
  1188   apply(unfold s_depend_def cs_waiting_def cs_holding_def)
       
  1189   by (auto intro:wq_threads)
       
  1190 
       
  1191 lemma readys_v_eq:
       
  1192   fixes th thread cs rest
       
  1193   assumes vt: "vt step s"
       
  1194   and neq_th: "th \<noteq> thread"
       
  1195   and eq_wq: "wq s cs = thread#rest"
       
  1196   and not_in: "th \<notin>  set rest"
       
  1197   shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
       
  1198 proof -
       
  1199   from prems show ?thesis
       
  1200     apply (auto simp:readys_def)
       
  1201     apply (case_tac "cs = csa", simp add:s_waiting_def)
       
  1202     apply (erule_tac x = csa in allE)
       
  1203     apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
       
  1204     apply (case_tac "csa = cs", simp)
       
  1205     apply (erule_tac x = cs in allE)
       
  1206     apply (auto simp:s_waiting_def wq_def Let_def split:list.splits)
       
  1207     proof -
       
  1208       assume th_nin: "th \<notin> set rest"
       
  1209         and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1210         and eq_wq: "waiting_queue (schs s) cs = thread # rest"
       
  1211       have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1212       proof(rule someI2)
       
  1213         from wq_distinct[OF vt, of cs] and eq_wq[folded wq_def]
       
  1214         show "distinct rest \<and> set rest = set rest" by auto
       
  1215       next
       
  1216         show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1217       qed
       
  1218       with th_nin th_in show False by auto
       
  1219     qed
       
  1220 qed
       
  1221 
       
  1222 lemma chain_building:
       
  1223   assumes vt: "vt step s"
       
  1224   shows "node \<in> Domain (depend s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (depend s)^+)"
       
  1225 proof -
       
  1226   from wf_dep_converse [OF vt]
       
  1227   have h: "wf ((depend s)\<inverse>)" .
       
  1228   show ?thesis
       
  1229   proof(induct rule:wf_induct [OF h])
       
  1230     fix x
       
  1231     assume ih [rule_format]: 
       
  1232       "\<forall>y. (y, x) \<in> (depend s)\<inverse> \<longrightarrow> 
       
  1233            y \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (depend s)\<^sup>+)"
       
  1234     show "x \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+)"
       
  1235     proof
       
  1236       assume x_d: "x \<in> Domain (depend s)"
       
  1237       show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+"
       
  1238       proof(cases x)
       
  1239         case (Th th)
       
  1240         from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> depend s" by (auto simp:s_depend_def)
       
  1241         with Th have x_in_r: "(Cs cs, x) \<in> (depend s)^-1" by simp
       
  1242         from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> depend s" by blast
       
  1243         hence "Cs cs \<in> Domain (depend s)" by auto
       
  1244         from ih [OF x_in_r this] obtain th'
       
  1245           where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (depend s)\<^sup>+" by auto
       
  1246         have "(x, Th th') \<in> (depend s)\<^sup>+" using Th x_in cs_in by auto
       
  1247         with th'_ready show ?thesis by auto
       
  1248       next
       
  1249         case (Cs cs)
       
  1250         from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (depend s)^-1" by (auto simp:s_depend_def)
       
  1251         show ?thesis
       
  1252         proof(cases "th' \<in> readys s")
       
  1253           case True
       
  1254           from True and th'_d show ?thesis by auto
       
  1255         next
       
  1256           case False
       
  1257           from th'_d and range_in [OF vt] have "th' \<in> threads s" by auto
       
  1258           with False have "Th th' \<in> Domain (depend s)" 
       
  1259             by (auto simp:readys_def s_waiting_def s_depend_def cs_waiting_def Domain_def)
       
  1260           from ih [OF th'_d this]
       
  1261           obtain th'' where 
       
  1262             th''_r: "th'' \<in> readys s" and 
       
  1263             th''_in: "(Th th', Th th'') \<in> (depend s)\<^sup>+" by auto
       
  1264           from th'_d and th''_in 
       
  1265           have "(x, Th th'') \<in> (depend s)\<^sup>+" by auto
       
  1266           with th''_r show ?thesis by auto
       
  1267         qed
       
  1268       qed
       
  1269     qed
       
  1270   qed
       
  1271 qed
       
  1272 
       
  1273 lemma th_chain_to_ready:
       
  1274   fixes s th
       
  1275   assumes vt: "vt step s"
       
  1276   and th_in: "th \<in> threads s"
       
  1277   shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (depend s)^+)"
       
  1278 proof(cases "th \<in> readys s")
       
  1279   case True
       
  1280   thus ?thesis by auto
       
  1281 next
       
  1282   case False
       
  1283   from False and th_in have "Th th \<in> Domain (depend s)" 
       
  1284     by (auto simp:readys_def s_waiting_def s_depend_def cs_waiting_def Domain_def)
       
  1285   from chain_building [rule_format, OF vt this]
       
  1286   show ?thesis by auto
       
  1287 qed
       
  1288 
       
  1289 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
       
  1290   by  (unfold s_waiting_def cs_waiting_def, auto)
       
  1291 
       
  1292 lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
       
  1293   by (unfold s_holding_def cs_holding_def, simp)
       
  1294 
       
  1295 lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
       
  1296   by (unfold s_holding_def cs_holding_def, auto)
       
  1297 
       
  1298 lemma unique_depend: "\<lbrakk>vt step s; (n, n1) \<in> depend s; (n, n2) \<in> depend s\<rbrakk> \<Longrightarrow> n1 = n2"
       
  1299   apply(unfold s_depend_def, auto, fold waiting_eq holding_eq)
       
  1300   by(auto elim:waiting_unique holding_unique)
       
  1301 
       
  1302 lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
       
  1303 by (induct rule:trancl_induct, auto)
       
  1304 
       
  1305 lemma dchain_unique:
       
  1306   assumes vt: "vt step s"
       
  1307   and th1_d: "(n, Th th1) \<in> (depend s)^+"
       
  1308   and th1_r: "th1 \<in> readys s"
       
  1309   and th2_d: "(n, Th th2) \<in> (depend s)^+"
       
  1310   and th2_r: "th2 \<in> readys s"
       
  1311   shows "th1 = th2"
       
  1312 proof -
       
  1313   { assume neq: "th1 \<noteq> th2"
       
  1314     hence "Th th1 \<noteq> Th th2" by simp
       
  1315     from unique_chain [OF _ th1_d th2_d this] and unique_depend [OF vt]
       
  1316     have "(Th th1, Th th2) \<in> (depend s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (depend s)\<^sup>+" by auto
       
  1317     hence "False"
       
  1318     proof
       
  1319       assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+"
       
  1320       from trancl_split [OF this]
       
  1321       obtain n where dd: "(Th th1, n) \<in> depend s" by auto
       
  1322       then obtain cs where eq_n: "n = Cs cs"
       
  1323         by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1324       from dd eq_n have "th1 \<notin> readys s"
       
  1325         by (auto simp:readys_def s_depend_def s_waiting_def cs_waiting_def)
       
  1326       with th1_r show ?thesis by auto
       
  1327     next
       
  1328       assume "(Th th2, Th th1) \<in> (depend s)\<^sup>+"
       
  1329       from trancl_split [OF this]
       
  1330       obtain n where dd: "(Th th2, n) \<in> depend s" by auto
       
  1331       then obtain cs where eq_n: "n = Cs cs"
       
  1332         by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1333       from dd eq_n have "th2 \<notin> readys s"
       
  1334         by (auto simp:readys_def s_depend_def s_waiting_def cs_waiting_def)
       
  1335       with th2_r show ?thesis by auto
       
  1336     qed
       
  1337   } thus ?thesis by auto
       
  1338 qed
       
  1339              
       
  1340 
       
  1341 lemma step_holdents_p_add:
       
  1342   fixes th cs s
       
  1343   assumes vt: "vt step (P th cs#s)"
       
  1344   and "wq s cs = []"
       
  1345   shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
       
  1346 proof -
       
  1347   from prems show ?thesis
       
  1348   unfolding  holdents_def step_depend_p[OF vt] by auto
       
  1349 qed
       
  1350 
       
  1351 lemma step_holdents_p_eq:
       
  1352   fixes th cs s
       
  1353   assumes vt: "vt step (P th cs#s)"
       
  1354   and "wq s cs \<noteq> []"
       
  1355   shows "holdents (P th cs#s) th = holdents s th"
       
  1356 proof -
       
  1357   from prems show ?thesis
       
  1358   unfolding  holdents_def step_depend_p[OF vt] by auto
       
  1359 qed
       
  1360 
       
  1361 
       
  1362 lemma finite_holding:
       
  1363   fixes s th cs
       
  1364   assumes vt: "vt step s"
       
  1365   shows "finite (holdents s th)"
       
  1366 proof -
       
  1367   let ?F = "\<lambda> (x, y). the_cs x"
       
  1368   from finite_depend [OF vt]
       
  1369   have "finite (depend s)" .
       
  1370   hence "finite (?F `(depend s))" by simp
       
  1371   moreover have "{cs . (Cs cs, Th th) \<in> depend s} \<subseteq> \<dots>" 
       
  1372   proof -
       
  1373     { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
       
  1374       fix x assume "(Cs x, Th th) \<in> depend s"
       
  1375       hence "?F (Cs x, Th th) \<in> ?F `(depend s)" by (rule h)
       
  1376       moreover have "?F (Cs x, Th th) = x" by simp
       
  1377       ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` depend s" by simp 
       
  1378     } thus ?thesis by auto
       
  1379   qed
       
  1380   ultimately show ?thesis by (unfold holdents_def, auto intro:finite_subset)
       
  1381 qed
       
  1382 
       
  1383 lemma cntCS_v_dec: 
       
  1384   fixes s thread cs
       
  1385   assumes vtv: "vt step (V thread cs#s)"
       
  1386   shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
       
  1387 proof -
       
  1388   from step_back_step[OF vtv]
       
  1389   have cs_in: "cs \<in> holdents s thread" 
       
  1390     apply (cases, unfold holdents_def s_depend_def, simp)
       
  1391     by (unfold cs_holding_def s_holding_def, auto)
       
  1392   moreover have cs_not_in: 
       
  1393     "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
       
  1394     apply (insert wq_distinct[OF step_back_vt[OF vtv], of cs])
       
  1395     apply (unfold holdents_def, unfold step_depend_v[OF vtv],
       
  1396             auto simp:next_th_def)
       
  1397   proof -
       
  1398     fix rest
       
  1399     assume dst: "distinct (rest::thread list)"
       
  1400       and ne: "rest \<noteq> []"
       
  1401     and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1402     moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1403     proof(rule someI2)
       
  1404       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1405     next
       
  1406       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1407     qed
       
  1408     ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1409                      set (SOME q. distinct q \<and> set q = set rest)" by simp
       
  1410     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1411     proof(rule someI2)
       
  1412       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1413     next
       
  1414       fix x assume " distinct x \<and> set x = set rest" with ne
       
  1415       show "x \<noteq> []" by auto
       
  1416     qed
       
  1417     ultimately 
       
  1418     show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> depend s"
       
  1419       by auto
       
  1420   next
       
  1421     fix rest
       
  1422     assume dst: "distinct (rest::thread list)"
       
  1423       and ne: "rest \<noteq> []"
       
  1424     and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1425     moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1426     proof(rule someI2)
       
  1427       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1428     next
       
  1429       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1430     qed
       
  1431     ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1432                      set (SOME q. distinct q \<and> set q = set rest)" by simp
       
  1433     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1434     proof(rule someI2)
       
  1435       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1436     next
       
  1437       fix x assume " distinct x \<and> set x = set rest" with ne
       
  1438       show "x \<noteq> []" by auto
       
  1439     qed
       
  1440     ultimately show "False" by auto 
       
  1441   qed
       
  1442   ultimately 
       
  1443   have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
       
  1444     by auto
       
  1445   moreover have "card \<dots> = 
       
  1446                     Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
       
  1447   proof(rule card_insert)
       
  1448     from finite_holding [OF vtv]
       
  1449     show " finite (holdents (V thread cs # s) thread)" .
       
  1450   qed
       
  1451   moreover from cs_not_in 
       
  1452   have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
       
  1453   ultimately show ?thesis by (simp add:cntCS_def)
       
  1454 qed 
       
  1455 
       
  1456 lemma cnp_cnv_cncs:
       
  1457   fixes s th
       
  1458   assumes vt: "vt step s"
       
  1459   shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
       
  1460                                        then cntCS s th else cntCS s th + 1)"
       
  1461 proof -
       
  1462   from vt show ?thesis
       
  1463   proof(induct arbitrary:th)
       
  1464     case (vt_cons s e)
       
  1465     assume vt: "vt step s"
       
  1466     and ih: "\<And>th. cntP s th  = cntV s th +
       
  1467                (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
       
  1468     and stp: "step s e"
       
  1469     from stp show ?case
       
  1470     proof(cases)
       
  1471       case (thread_create thread prio)
       
  1472       assume eq_e: "e = Create thread prio"
       
  1473         and not_in: "thread \<notin> threads s"
       
  1474       show ?thesis
       
  1475       proof -
       
  1476         { fix cs 
       
  1477           assume "thread \<in> set (wq s cs)"
       
  1478           from wq_threads [OF vt this] have "thread \<in> threads s" .
       
  1479           with not_in have "False" by simp
       
  1480         } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
       
  1481           by (auto simp:readys_def threads.simps s_waiting_def 
       
  1482             wq_def cs_waiting_def Let_def)
       
  1483         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
       
  1484         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
       
  1485         have eq_cncs: "cntCS (e#s) th = cntCS s th"
       
  1486           unfolding cntCS_def holdents_def
       
  1487           by (simp add:depend_create_unchanged eq_e)
       
  1488         { assume "th \<noteq> thread"
       
  1489           with eq_readys eq_e
       
  1490           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
       
  1491                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
       
  1492             by (simp add:threads.simps)
       
  1493           with eq_cnp eq_cnv eq_cncs ih not_in
       
  1494           have ?thesis by simp
       
  1495         } moreover {
       
  1496           assume eq_th: "th = thread"
       
  1497           with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
       
  1498           moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
       
  1499           moreover note eq_cnp eq_cnv eq_cncs
       
  1500           ultimately have ?thesis by auto
       
  1501         } ultimately show ?thesis by blast
       
  1502       qed
       
  1503     next
       
  1504       case (thread_exit thread)
       
  1505       assume eq_e: "e = Exit thread" 
       
  1506       and is_runing: "thread \<in> runing s"
       
  1507       and no_hold: "holdents s thread = {}"
       
  1508       from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
       
  1509       from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
       
  1510       have eq_cncs: "cntCS (e#s) th = cntCS s th"
       
  1511         unfolding cntCS_def holdents_def
       
  1512         by (simp add:depend_exit_unchanged eq_e)
       
  1513       { assume "th \<noteq> thread"
       
  1514         with eq_e
       
  1515         have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
       
  1516           (th \<in> readys (s) \<or> th \<notin> threads (s))" 
       
  1517           apply (simp add:threads.simps readys_def)
       
  1518           apply (subst s_waiting_def)
       
  1519           apply (subst (1 2) wq_def)
       
  1520           apply (simp add:Let_def)
       
  1521           apply (subst s_waiting_def, simp)
       
  1522           by (fold wq_def, simp)
       
  1523         with eq_cnp eq_cnv eq_cncs ih
       
  1524         have ?thesis by simp
       
  1525       } moreover {
       
  1526         assume eq_th: "th = thread"
       
  1527         with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
       
  1528           by (simp add:runing_def)
       
  1529         moreover from eq_th eq_e have "th \<notin> threads (e#s)"
       
  1530           by simp
       
  1531         moreover note eq_cnp eq_cnv eq_cncs
       
  1532         ultimately have ?thesis by auto
       
  1533       } ultimately show ?thesis by blast
       
  1534     next
       
  1535       case (thread_P thread cs)
       
  1536       assume eq_e: "e = P thread cs"
       
  1537         and is_runing: "thread \<in> runing s"
       
  1538         and no_dep: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
       
  1539       from prems have vtp: "vt step (P thread cs#s)" by auto
       
  1540       show ?thesis 
       
  1541       proof -
       
  1542         { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
       
  1543           assume neq_th: "th \<noteq> thread"
       
  1544           with eq_e
       
  1545           have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
       
  1546             apply (simp add:readys_def s_waiting_def wq_def Let_def)
       
  1547             apply (rule_tac hh, clarify)
       
  1548             apply (intro iffI allI, clarify)
       
  1549             apply (erule_tac x = csa in allE, auto)
       
  1550             apply (subgoal_tac "waiting_queue (schs s) cs \<noteq> []", auto)
       
  1551             apply (erule_tac x = cs in allE, auto)
       
  1552             by (case_tac "(waiting_queue (schs s) cs)", auto)
       
  1553           moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
       
  1554             apply (simp add:cntCS_def holdents_def)
       
  1555             by (unfold  step_depend_p [OF vtp], auto)
       
  1556           moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
       
  1557             by (simp add:cntP_def count_def)
       
  1558           moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
       
  1559             by (simp add:cntV_def count_def)
       
  1560           moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
       
  1561           moreover note ih [of th] 
       
  1562           ultimately have ?thesis by simp
       
  1563         } moreover {
       
  1564           assume eq_th: "th = thread"
       
  1565           have ?thesis
       
  1566           proof -
       
  1567             from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
       
  1568               by (simp add:cntP_def count_def)
       
  1569             from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
       
  1570               by (simp add:cntV_def count_def)
       
  1571             show ?thesis
       
  1572             proof (cases "wq s cs = []")
       
  1573               case True
       
  1574               with is_runing
       
  1575               have "th \<in> readys (e#s)"
       
  1576                 apply (unfold eq_e wq_def, unfold readys_def s_depend_def)
       
  1577                 apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
       
  1578                 by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
       
  1579               moreover have "cntCS (e # s) th = 1 + cntCS s th"
       
  1580               proof -
       
  1581                 have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> depend s} =
       
  1582                   Suc (card {cs. (Cs cs, Th thread) \<in> depend s})" (is "card ?L = Suc (card ?R)")
       
  1583                 proof -
       
  1584                   have "?L = insert cs ?R" by auto
       
  1585                   moreover have "card \<dots> = Suc (card (?R - {cs}))" 
       
  1586                   proof(rule card_insert)
       
  1587                     from finite_holding [OF vt, of thread]
       
  1588                     show " finite {cs. (Cs cs, Th thread) \<in> depend s}"
       
  1589                       by (unfold holdents_def, simp)
       
  1590                   qed
       
  1591                   moreover have "?R - {cs} = ?R"
       
  1592                   proof -
       
  1593                     have "cs \<notin> ?R"
       
  1594                     proof
       
  1595                       assume "cs \<in> {cs. (Cs cs, Th thread) \<in> depend s}"
       
  1596                       with no_dep show False by auto
       
  1597                     qed
       
  1598                     thus ?thesis by auto
       
  1599                   qed
       
  1600                   ultimately show ?thesis by auto
       
  1601                 qed
       
  1602                 thus ?thesis
       
  1603                   apply (unfold eq_e eq_th cntCS_def)
       
  1604                   apply (simp add: holdents_def)
       
  1605                   by (unfold step_depend_p [OF vtp], auto simp:True)
       
  1606               qed
       
  1607               moreover from is_runing have "th \<in> readys s"
       
  1608                 by (simp add:runing_def eq_th)
       
  1609               moreover note eq_cnp eq_cnv ih [of th]
       
  1610               ultimately show ?thesis by auto
       
  1611             next
       
  1612               case False
       
  1613               have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
       
  1614                     by (unfold eq_th eq_e wq_def, auto simp:Let_def)
       
  1615               have "th \<notin> readys (e#s)"
       
  1616               proof
       
  1617                 assume "th \<in> readys (e#s)"
       
  1618                 hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
       
  1619                 from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
       
  1620                 hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
       
  1621                   by (simp add:s_waiting_def)
       
  1622                 moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
       
  1623                 ultimately have "th = hd (wq (e#s) cs)" by blast
       
  1624                 with eq_wq have "th = hd (wq s cs @ [th])" by simp
       
  1625                 hence "th = hd (wq s cs)" using False by auto
       
  1626                 with False eq_wq wq_distinct [OF vtp, of cs]
       
  1627                 show False by (fold eq_e, auto)
       
  1628               qed
       
  1629               moreover from is_runing have "th \<in> threads (e#s)" 
       
  1630                 by (unfold eq_e, auto simp:runing_def readys_def eq_th)
       
  1631               moreover have "cntCS (e # s) th = cntCS s th"
       
  1632                 apply (unfold cntCS_def holdents_def eq_e step_depend_p[OF vtp])
       
  1633                 by (auto simp:False)
       
  1634               moreover note eq_cnp eq_cnv ih[of th]
       
  1635               moreover from is_runing have "th \<in> readys s"
       
  1636                 by (simp add:runing_def eq_th)
       
  1637               ultimately show ?thesis by auto
       
  1638             qed
       
  1639           qed
       
  1640         } ultimately show ?thesis by blast
       
  1641       qed
       
  1642     next
       
  1643       case (thread_V thread cs)
       
  1644       from prems have vtv: "vt step (V thread cs # s)" by auto
       
  1645       assume eq_e: "e = V thread cs"
       
  1646         and is_runing: "thread \<in> runing s"
       
  1647         and hold: "holding s thread cs"
       
  1648       from hold obtain rest 
       
  1649         where eq_wq: "wq s cs = thread # rest"
       
  1650         by (case_tac "wq s cs", auto simp:s_holding_def)
       
  1651       have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
       
  1652       have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1653       proof(rule someI2)
       
  1654         from wq_distinct[OF step_back_vt[OF vtv], of cs] and eq_wq
       
  1655         show "distinct rest \<and> set rest = set rest" by auto
       
  1656       next
       
  1657         show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
       
  1658           by auto
       
  1659       qed
       
  1660       show ?thesis
       
  1661       proof -
       
  1662         { assume eq_th: "th = thread"
       
  1663           from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
       
  1664             by (unfold eq_e, simp add:cntP_def count_def)
       
  1665           moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
       
  1666             by (unfold eq_e, simp add:cntV_def count_def)
       
  1667           moreover from cntCS_v_dec [OF vtv] 
       
  1668           have "cntCS (e # s) thread + 1 = cntCS s thread"
       
  1669             by (simp add:eq_e)
       
  1670           moreover from is_runing have rd_before: "thread \<in> readys s"
       
  1671             by (unfold runing_def, simp)
       
  1672           moreover have "thread \<in> readys (e # s)"
       
  1673           proof -
       
  1674             from is_runing
       
  1675             have "thread \<in> threads (e#s)" 
       
  1676               by (unfold eq_e, auto simp:runing_def readys_def)
       
  1677             moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
       
  1678             proof
       
  1679               fix cs1
       
  1680               { assume eq_cs: "cs1 = cs" 
       
  1681                 have "\<not> waiting (e # s) thread cs1"
       
  1682                 proof -
       
  1683                   from eq_wq
       
  1684                   have "thread \<notin> set (wq (e#s) cs1)"
       
  1685                     apply(unfold eq_e wq_def eq_cs s_holding_def)
       
  1686                     apply (auto simp:Let_def)
       
  1687                   proof -
       
  1688                     assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1689                     with eq_set have "thread \<in> set rest" by simp
       
  1690                     with wq_distinct[OF step_back_vt[OF vtv], of cs]
       
  1691                     and eq_wq show False by auto
       
  1692                   qed
       
  1693                   thus ?thesis by (simp add:s_waiting_def)
       
  1694                 qed
       
  1695               } moreover {
       
  1696                 assume neq_cs: "cs1 \<noteq> cs"
       
  1697                   have "\<not> waiting (e # s) thread cs1" 
       
  1698                   proof -
       
  1699                     from wq_v_neq [OF neq_cs[symmetric]]
       
  1700                     have "wq (V thread cs # s) cs1 = wq s cs1" .
       
  1701                     moreover have "\<not> waiting s thread cs1" 
       
  1702                     proof -
       
  1703                       from runing_ready and is_runing
       
  1704                       have "thread \<in> readys s" by auto
       
  1705                       thus ?thesis by (simp add:readys_def)
       
  1706                     qed
       
  1707                     ultimately show ?thesis 
       
  1708                       by (auto simp:s_waiting_def eq_e)
       
  1709                   qed
       
  1710               } ultimately show "\<not> waiting (e # s) thread cs1" by blast
       
  1711             qed
       
  1712             ultimately show ?thesis by (simp add:readys_def)
       
  1713           qed
       
  1714           moreover note eq_th ih
       
  1715           ultimately have ?thesis by auto
       
  1716         } moreover {
       
  1717           assume neq_th: "th \<noteq> thread"
       
  1718           from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
       
  1719             by (simp add:cntP_def count_def)
       
  1720           from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
       
  1721             by (simp add:cntV_def count_def)
       
  1722           have ?thesis
       
  1723           proof(cases "th \<in> set rest")
       
  1724             case False
       
  1725             have "(th \<in> readys (e # s)) = (th \<in> readys s)"
       
  1726               apply (insert step_back_vt[OF vtv])
       
  1727               by (unfold eq_e, rule readys_v_eq [OF _ neq_th eq_wq False], auto)
       
  1728             moreover have "cntCS (e#s) th = cntCS s th"
       
  1729               apply (insert neq_th, unfold eq_e cntCS_def holdents_def step_depend_v[OF vtv], auto)
       
  1730               proof -
       
  1731                 have "{csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
       
  1732                       {cs. (Cs cs, Th th) \<in> depend s}"
       
  1733                 proof -
       
  1734                   from False eq_wq
       
  1735                   have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> depend s"
       
  1736                     apply (unfold next_th_def, auto)
       
  1737                   proof -
       
  1738                     assume ne: "rest \<noteq> []"
       
  1739                       and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1740                       and eq_wq: "wq s cs = thread # rest"
       
  1741                     from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1742                                   set (SOME q. distinct q \<and> set q = set rest)
       
  1743                                   " by simp
       
  1744                     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1745                     proof(rule someI2)
       
  1746                       from wq_distinct[OF step_back_vt[OF vtv], of cs] and eq_wq
       
  1747                       show "distinct rest \<and> set rest = set rest" by auto
       
  1748                     next
       
  1749                       fix x assume "distinct x \<and> set x = set rest"
       
  1750                       with ne show "x \<noteq> []" by auto
       
  1751                     qed
       
  1752                     ultimately show 
       
  1753                       "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> depend s"
       
  1754                       by auto
       
  1755                   qed    
       
  1756                   thus ?thesis by auto
       
  1757                 qed
       
  1758                 thus "card {csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs \<and> next_th s thread cs th} =
       
  1759                              card {cs. (Cs cs, Th th) \<in> depend s}" by simp 
       
  1760               qed
       
  1761             moreover note ih eq_cnp eq_cnv eq_threads
       
  1762             ultimately show ?thesis by auto
       
  1763           next
       
  1764             case True
       
  1765             assume th_in: "th \<in> set rest"
       
  1766             show ?thesis
       
  1767             proof(cases "next_th s thread cs th")
       
  1768               case False
       
  1769               with eq_wq and th_in have 
       
  1770                 neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest")
       
  1771                 by (auto simp:next_th_def)
       
  1772               have "(th \<in> readys (e # s)) = (th \<in> readys s)"
       
  1773               proof -
       
  1774                 from eq_wq and th_in
       
  1775                 have "\<not> th \<in> readys s"
       
  1776                   apply (auto simp:readys_def s_waiting_def)
       
  1777                   apply (rule_tac x = cs in exI, auto)
       
  1778                   by (insert wq_distinct[OF step_back_vt[OF vtv], of cs], auto)
       
  1779                 moreover 
       
  1780                 from eq_wq and th_in and neq_hd
       
  1781                 have "\<not> (th \<in> readys (e # s))"
       
  1782                   apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits)
       
  1783                   by (rule_tac x = cs in exI, auto simp:eq_set)
       
  1784                 ultimately show ?thesis by auto
       
  1785               qed
       
  1786               moreover have "cntCS (e#s) th = cntCS s th" 
       
  1787               proof -
       
  1788                 from eq_wq and  th_in and neq_hd
       
  1789                 have "(holdents (e # s) th) = (holdents s th)"
       
  1790                   apply (unfold eq_e step_depend_v[OF vtv], 
       
  1791                          auto simp:next_th_def eq_set s_depend_def holdents_def wq_def
       
  1792                                    Let_def cs_holding_def)
       
  1793                   by (insert wq_distinct[OF step_back_vt[OF vtv], of cs], auto simp:wq_def)
       
  1794                 thus ?thesis by (simp add:cntCS_def)
       
  1795               qed
       
  1796               moreover note ih eq_cnp eq_cnv eq_threads
       
  1797               ultimately show ?thesis by auto
       
  1798             next
       
  1799               case True
       
  1800               let ?rest = " (SOME q. distinct q \<and> set q = set rest)"
       
  1801               let ?t = "hd ?rest"
       
  1802               from True eq_wq th_in neq_th
       
  1803               have "th \<in> readys (e # s)"
       
  1804                 apply (auto simp:eq_e readys_def s_waiting_def wq_def
       
  1805                         Let_def next_th_def)
       
  1806               proof -
       
  1807                 assume eq_wq: "waiting_queue (schs s) cs = thread # rest"
       
  1808                   and t_in: "?t \<in> set rest"
       
  1809                 show "?t \<in> threads s"
       
  1810                 proof(rule wq_threads[OF step_back_vt[OF vtv]])
       
  1811                   from eq_wq and t_in
       
  1812                   show "?t \<in> set (wq s cs)" by (auto simp:wq_def)
       
  1813                 qed
       
  1814               next
       
  1815                 fix csa
       
  1816                 assume eq_wq: "waiting_queue (schs s) cs = thread # rest"
       
  1817                   and t_in: "?t \<in> set rest"
       
  1818                   and neq_cs: "csa \<noteq> cs"
       
  1819                   and t_in': "?t \<in>  set (waiting_queue (schs s) csa)"
       
  1820                 show "?t = hd (waiting_queue (schs s) csa)"
       
  1821                 proof -
       
  1822                   { assume neq_hd': "?t \<noteq> hd (waiting_queue (schs s) csa)"
       
  1823                     from wq_distinct[OF step_back_vt[OF vtv], of cs] and 
       
  1824                     eq_wq[folded wq_def] and t_in eq_wq
       
  1825                     have "?t \<noteq> thread" by auto
       
  1826                     with eq_wq and t_in
       
  1827                     have w1: "waiting s ?t cs"
       
  1828                       by (auto simp:s_waiting_def wq_def)
       
  1829                     from t_in' neq_hd'
       
  1830                     have w2: "waiting s ?t csa"
       
  1831                       by (auto simp:s_waiting_def wq_def)
       
  1832                     from waiting_unique[OF step_back_vt[OF vtv] w1 w2]
       
  1833                     and neq_cs have "False" by auto
       
  1834                   } thus ?thesis by auto
       
  1835                 qed
       
  1836               qed
       
  1837               moreover have "cntP s th = cntV s th + cntCS s th + 1"
       
  1838               proof -
       
  1839                 have "th \<notin> readys s" 
       
  1840                 proof -
       
  1841                   from True eq_wq neq_th th_in
       
  1842                   show ?thesis
       
  1843                     apply (unfold readys_def s_waiting_def, auto)
       
  1844                     by (rule_tac x = cs in exI, auto)
       
  1845                 qed
       
  1846                 moreover have "th \<in> threads s"
       
  1847                 proof -
       
  1848                   from th_in eq_wq
       
  1849                   have "th \<in> set (wq s cs)" by simp
       
  1850                   from wq_threads [OF step_back_vt[OF vtv] this] 
       
  1851                   show ?thesis .
       
  1852                 qed
       
  1853                 ultimately show ?thesis using ih by auto
       
  1854               qed
       
  1855               moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
       
  1856                 apply (unfold cntCS_def holdents_def eq_e step_depend_v[OF vtv], auto)
       
  1857               proof -
       
  1858                 show "card {csa. (Cs csa, Th th) \<in> depend s \<or> csa = cs} =
       
  1859                                Suc (card {cs. (Cs cs, Th th) \<in> depend s})"
       
  1860                   (is "card ?A = Suc (card ?B)")
       
  1861                 proof -
       
  1862                   have "?A = insert cs ?B" by auto
       
  1863                   hence "card ?A = card (insert cs ?B)" by simp
       
  1864                   also have "\<dots> = Suc (card ?B)"
       
  1865                   proof(rule card_insert_disjoint)
       
  1866                     have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` depend s)" 
       
  1867                       apply (auto simp:image_def)
       
  1868                       by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
       
  1869                     with finite_depend[OF step_back_vt[OF vtv]]
       
  1870                     show "finite {cs. (Cs cs, Th th) \<in> depend s}" by (auto intro:finite_subset)
       
  1871                   next
       
  1872                     show "cs \<notin> {cs. (Cs cs, Th th) \<in> depend s}"
       
  1873                     proof
       
  1874                       assume "cs \<in> {cs. (Cs cs, Th th) \<in> depend s}"
       
  1875                       hence "(Cs cs, Th th) \<in> depend s" by simp
       
  1876                       with True neq_th eq_wq show False
       
  1877                         by (auto simp:next_th_def s_depend_def cs_holding_def)
       
  1878                     qed
       
  1879                   qed
       
  1880                   finally show ?thesis .
       
  1881                 qed
       
  1882               qed
       
  1883               moreover note eq_cnp eq_cnv
       
  1884               ultimately show ?thesis by simp
       
  1885             qed
       
  1886           qed
       
  1887         } ultimately show ?thesis by blast
       
  1888       qed
       
  1889     next
       
  1890       case (thread_set thread prio)
       
  1891       assume eq_e: "e = Set thread prio"
       
  1892         and is_runing: "thread \<in> runing s"
       
  1893       show ?thesis
       
  1894       proof -
       
  1895         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
       
  1896         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
       
  1897         have eq_cncs: "cntCS (e#s) th = cntCS s th"
       
  1898           unfolding cntCS_def holdents_def
       
  1899           by (simp add:depend_set_unchanged eq_e)
       
  1900         from eq_e have eq_readys: "readys (e#s) = readys s" 
       
  1901           by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
       
  1902                   auto simp:Let_def)
       
  1903         { assume "th \<noteq> thread"
       
  1904           with eq_readys eq_e
       
  1905           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
       
  1906                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
       
  1907             by (simp add:threads.simps)
       
  1908           with eq_cnp eq_cnv eq_cncs ih is_runing
       
  1909           have ?thesis by simp
       
  1910         } moreover {
       
  1911           assume eq_th: "th = thread"
       
  1912           with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
       
  1913             by (unfold runing_def, auto)
       
  1914           moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
       
  1915             by (simp add:runing_def)
       
  1916           moreover note eq_cnp eq_cnv eq_cncs
       
  1917           ultimately have ?thesis by auto
       
  1918         } ultimately show ?thesis by blast
       
  1919       qed   
       
  1920     qed
       
  1921   next
       
  1922     case vt_nil
       
  1923     show ?case 
       
  1924       by (unfold cntP_def cntV_def cntCS_def, 
       
  1925         auto simp:count_def holdents_def s_depend_def wq_def cs_holding_def)
       
  1926   qed
       
  1927 qed
       
  1928 
       
  1929 lemma not_thread_cncs:
       
  1930   fixes th s
       
  1931   assumes vt: "vt step s"
       
  1932   and not_in: "th \<notin> threads s" 
       
  1933   shows "cntCS s th = 0"
       
  1934 proof -
       
  1935   from vt not_in show ?thesis
       
  1936   proof(induct arbitrary:th)
       
  1937     case (vt_cons s e th)
       
  1938     assume vt: "vt step s"
       
  1939       and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
       
  1940       and stp: "step s e"
       
  1941       and not_in: "th \<notin> threads (e # s)"
       
  1942     from stp show ?case
       
  1943     proof(cases)
       
  1944       case (thread_create thread prio)
       
  1945       assume eq_e: "e = Create thread prio"
       
  1946         and not_in': "thread \<notin> threads s"
       
  1947       have "cntCS (e # s) th = cntCS s th"
       
  1948         apply (unfold eq_e cntCS_def holdents_def)
       
  1949         by (simp add:depend_create_unchanged)
       
  1950       moreover have "th \<notin> threads s" 
       
  1951       proof -
       
  1952         from not_in eq_e show ?thesis by simp
       
  1953       qed
       
  1954       moreover note ih ultimately show ?thesis by auto
       
  1955     next
       
  1956       case (thread_exit thread)
       
  1957       assume eq_e: "e = Exit thread"
       
  1958       and nh: "holdents s thread = {}"
       
  1959       have eq_cns: "cntCS (e # s) th = cntCS s th"
       
  1960         apply (unfold eq_e cntCS_def holdents_def)
       
  1961         by (simp add:depend_exit_unchanged)
       
  1962       show ?thesis
       
  1963       proof(cases "th = thread")
       
  1964         case True
       
  1965         have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
       
  1966         with eq_cns show ?thesis by simp
       
  1967       next
       
  1968         case False
       
  1969         with not_in and eq_e
       
  1970         have "th \<notin> threads s" by simp
       
  1971         from ih[OF this] and eq_cns show ?thesis by simp
       
  1972       qed
       
  1973     next
       
  1974       case (thread_P thread cs)
       
  1975       assume eq_e: "e = P thread cs"
       
  1976       and is_runing: "thread \<in> runing s"
       
  1977       from prems have vtp: "vt step (P thread cs#s)" by auto
       
  1978       have neq_th: "th \<noteq> thread" 
       
  1979       proof -
       
  1980         from not_in eq_e have "th \<notin> threads s" by simp
       
  1981         moreover from is_runing have "thread \<in> threads s"
       
  1982           by (simp add:runing_def readys_def)
       
  1983         ultimately show ?thesis by auto
       
  1984       qed
       
  1985       hence "cntCS (e # s) th  = cntCS s th "
       
  1986         apply (unfold cntCS_def holdents_def eq_e)
       
  1987         by (unfold step_depend_p[OF vtp], auto)
       
  1988       moreover have "cntCS s th = 0"
       
  1989       proof(rule ih)
       
  1990         from not_in eq_e show "th \<notin> threads s" by simp
       
  1991       qed
       
  1992       ultimately show ?thesis by simp
       
  1993     next
       
  1994       case (thread_V thread cs)
       
  1995       assume eq_e: "e = V thread cs"
       
  1996         and is_runing: "thread \<in> runing s"
       
  1997         and hold: "holding s thread cs"
       
  1998       have neq_th: "th \<noteq> thread" 
       
  1999       proof -
       
  2000         from not_in eq_e have "th \<notin> threads s" by simp
       
  2001         moreover from is_runing have "thread \<in> threads s"
       
  2002           by (simp add:runing_def readys_def)
       
  2003         ultimately show ?thesis by auto
       
  2004       qed
       
  2005       from prems have vtv: "vt step (V thread cs#s)" by auto
       
  2006       from hold obtain rest 
       
  2007         where eq_wq: "wq s cs = thread # rest"
       
  2008         by (case_tac "wq s cs", auto simp:s_holding_def)
       
  2009       from not_in eq_e eq_wq
       
  2010       have "\<not> next_th s thread cs th"
       
  2011         apply (auto simp:next_th_def)
       
  2012       proof -
       
  2013         assume ne: "rest \<noteq> []"
       
  2014           and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s")
       
  2015         have "?t \<in> set rest"
       
  2016         proof(rule someI2)
       
  2017           from wq_distinct[OF step_back_vt[OF vtv], of cs] and eq_wq
       
  2018           show "distinct rest \<and> set rest = set rest" by auto
       
  2019         next
       
  2020           fix x assume "distinct x \<and> set x = set rest" with ne
       
  2021           show "hd x \<in> set rest" by (cases x, auto)
       
  2022         qed
       
  2023         with eq_wq have "?t \<in> set (wq s cs)" by simp
       
  2024         from wq_threads[OF step_back_vt[OF vtv], OF this] and ni
       
  2025         show False by auto
       
  2026       qed
       
  2027       moreover note neq_th eq_wq
       
  2028       ultimately have "cntCS (e # s) th  = cntCS s th"
       
  2029         by (unfold eq_e cntCS_def holdents_def step_depend_v[OF vtv], auto)
       
  2030       moreover have "cntCS s th = 0"
       
  2031       proof(rule ih)
       
  2032         from not_in eq_e show "th \<notin> threads s" by simp
       
  2033       qed
       
  2034       ultimately show ?thesis by simp
       
  2035     next
       
  2036       case (thread_set thread prio)
       
  2037       print_facts
       
  2038       assume eq_e: "e = Set thread prio"
       
  2039         and is_runing: "thread \<in> runing s"
       
  2040       from not_in and eq_e have "th \<notin> threads s" by auto
       
  2041       from ih [OF this] and eq_e
       
  2042       show ?thesis 
       
  2043         apply (unfold eq_e cntCS_def holdents_def)
       
  2044         by (simp add:depend_set_unchanged)
       
  2045     qed
       
  2046     next
       
  2047       case vt_nil
       
  2048       show ?case
       
  2049       by (unfold cntCS_def, 
       
  2050         auto simp:count_def holdents_def s_depend_def wq_def cs_holding_def)
       
  2051   qed
       
  2052 qed
       
  2053 
       
  2054 lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
       
  2055   by (auto simp:s_waiting_def cs_waiting_def)
       
  2056 
       
  2057 lemma dm_depend_threads:
       
  2058   fixes th s
       
  2059   assumes vt: "vt step s"
       
  2060   and in_dom: "(Th th) \<in> Domain (depend s)"
       
  2061   shows "th \<in> threads s"
       
  2062 proof -
       
  2063   from in_dom obtain n where "(Th th, n) \<in> depend s" by auto
       
  2064   moreover from depend_target_th[OF this] obtain cs where "n = Cs cs" by auto
       
  2065   ultimately have "(Th th, Cs cs) \<in> depend s" by simp
       
  2066   hence "th \<in> set (wq s cs)"
       
  2067     by (unfold s_depend_def, auto simp:cs_waiting_def)
       
  2068   from wq_threads [OF vt this] show ?thesis .
       
  2069 qed
       
  2070 
       
  2071 lemma cp_eq_cpreced: "cp s th = cpreced s (wq s) th"
       
  2072 proof(unfold cp_def wq_def, induct s)
       
  2073   case (Cons e s')
       
  2074   show ?case
       
  2075     by (auto simp:Let_def)
       
  2076 next
       
  2077   case Nil
       
  2078   show ?case by (auto simp:Let_def)
       
  2079 qed
       
  2080 
       
  2081 
       
  2082 lemma runing_unique:
       
  2083   fixes th1 th2 s
       
  2084   assumes vt: "vt step s"
       
  2085   and runing_1: "th1 \<in> runing s"
       
  2086   and runing_2: "th2 \<in> runing s"
       
  2087   shows "th1 = th2"
       
  2088 proof -
       
  2089   from runing_1 and runing_2 have "cp s th1 = cp s th2"
       
  2090     by (unfold runing_def, simp)
       
  2091   hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1)) =
       
  2092                  Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependents (wq s) th2))"
       
  2093     (is "Max (?f ` ?A) = Max (?f ` ?B)")
       
  2094     by (unfold cp_eq_cpreced cpreced_def)
       
  2095   obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
       
  2096   proof -
       
  2097     have h1: "finite (?f ` ?A)"
       
  2098     proof -
       
  2099       have "finite ?A" 
       
  2100       proof -
       
  2101         have "finite (dependents (wq s) th1)"
       
  2102         proof-
       
  2103           have "finite {th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+}"
       
  2104           proof -
       
  2105             let ?F = "\<lambda> (x, y). the_th x"
       
  2106             have "{th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
       
  2107               apply (auto simp:image_def)
       
  2108               by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
       
  2109             moreover have "finite \<dots>"
       
  2110             proof -
       
  2111               from finite_depend[OF vt] have "finite (depend s)" .
       
  2112               hence "finite ((depend (wq s))\<^sup>+)"
       
  2113                 apply (unfold finite_trancl)
       
  2114                 by (auto simp: s_depend_def cs_depend_def wq_def)
       
  2115               thus ?thesis by auto
       
  2116             qed
       
  2117             ultimately show ?thesis by (auto intro:finite_subset)
       
  2118           qed
       
  2119           thus ?thesis by (simp add:cs_dependents_def)
       
  2120         qed
       
  2121         thus ?thesis by simp
       
  2122       qed
       
  2123       thus ?thesis by auto
       
  2124     qed
       
  2125     moreover have h2: "(?f ` ?A) \<noteq> {}"
       
  2126     proof -
       
  2127       have "?A \<noteq> {}" by simp
       
  2128       thus ?thesis by simp
       
  2129     qed
       
  2130     from Max_in [OF h1 h2]
       
  2131     have "Max (?f ` ?A) \<in> (?f ` ?A)" .
       
  2132     thus ?thesis by (auto intro:that)
       
  2133   qed
       
  2134   obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
       
  2135   proof -
       
  2136     have h1: "finite (?f ` ?B)"
       
  2137     proof -
       
  2138       have "finite ?B" 
       
  2139       proof -
       
  2140         have "finite (dependents (wq s) th2)"
       
  2141         proof-
       
  2142           have "finite {th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+}"
       
  2143           proof -
       
  2144             let ?F = "\<lambda> (x, y). the_th x"
       
  2145             have "{th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
       
  2146               apply (auto simp:image_def)
       
  2147               by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
       
  2148             moreover have "finite \<dots>"
       
  2149             proof -
       
  2150               from finite_depend[OF vt] have "finite (depend s)" .
       
  2151               hence "finite ((depend (wq s))\<^sup>+)"
       
  2152                 apply (unfold finite_trancl)
       
  2153                 by (auto simp: s_depend_def cs_depend_def wq_def)
       
  2154               thus ?thesis by auto
       
  2155             qed
       
  2156             ultimately show ?thesis by (auto intro:finite_subset)
       
  2157           qed
       
  2158           thus ?thesis by (simp add:cs_dependents_def)
       
  2159         qed
       
  2160         thus ?thesis by simp
       
  2161       qed
       
  2162       thus ?thesis by auto
       
  2163     qed
       
  2164     moreover have h2: "(?f ` ?B) \<noteq> {}"
       
  2165     proof -
       
  2166       have "?B \<noteq> {}" by simp
       
  2167       thus ?thesis by simp
       
  2168     qed
       
  2169     from Max_in [OF h1 h2]
       
  2170     have "Max (?f ` ?B) \<in> (?f ` ?B)" .
       
  2171     thus ?thesis by (auto intro:that)
       
  2172   qed
       
  2173   from eq_f_th1 eq_f_th2 eq_max 
       
  2174   have eq_preced: "preced th1' s = preced th2' s" by auto
       
  2175   hence eq_th12: "th1' = th2'"
       
  2176   proof (rule preced_unique)
       
  2177     from th1_in have "th1' = th1 \<or> (th1' \<in> dependents (wq s) th1)" by simp
       
  2178     thus "th1' \<in> threads s"
       
  2179     proof
       
  2180       assume "th1' \<in> dependents (wq s) th1"
       
  2181       hence "(Th th1') \<in> Domain ((depend s)^+)"
       
  2182         apply (unfold cs_dependents_def cs_depend_def s_depend_def)
       
  2183         by (auto simp:Domain_def)
       
  2184       hence "(Th th1') \<in> Domain (depend s)" by (simp add:trancl_domain)
       
  2185       from dm_depend_threads[OF vt this] show ?thesis .
       
  2186     next
       
  2187       assume "th1' = th1"
       
  2188       with runing_1 show ?thesis
       
  2189         by (unfold runing_def readys_def, auto)
       
  2190     qed
       
  2191   next
       
  2192     from th2_in have "th2' = th2 \<or> (th2' \<in> dependents (wq s) th2)" by simp
       
  2193     thus "th2' \<in> threads s"
       
  2194     proof
       
  2195       assume "th2' \<in> dependents (wq s) th2"
       
  2196       hence "(Th th2') \<in> Domain ((depend s)^+)"
       
  2197         apply (unfold cs_dependents_def cs_depend_def s_depend_def)
       
  2198         by (auto simp:Domain_def)
       
  2199       hence "(Th th2') \<in> Domain (depend s)" by (simp add:trancl_domain)
       
  2200       from dm_depend_threads[OF vt this] show ?thesis .
       
  2201     next
       
  2202       assume "th2' = th2"
       
  2203       with runing_2 show ?thesis
       
  2204         by (unfold runing_def readys_def, auto)
       
  2205     qed
       
  2206   qed
       
  2207   from th1_in have "th1' = th1 \<or> th1' \<in> dependents (wq s) th1" by simp
       
  2208   thus ?thesis
       
  2209   proof
       
  2210     assume eq_th': "th1' = th1"
       
  2211     from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
       
  2212     thus ?thesis
       
  2213     proof
       
  2214       assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
       
  2215     next
       
  2216       assume "th2' \<in> dependents (wq s) th2"
       
  2217       with eq_th12 eq_th' have "th1 \<in> dependents (wq s) th2" by simp
       
  2218       hence "(Th th1, Th th2) \<in> (depend s)^+"
       
  2219         by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
       
  2220       hence "Th th1 \<in> Domain ((depend s)^+)" using Domain_def [of "(depend s)^+"]
       
  2221         by auto
       
  2222       hence "Th th1 \<in> Domain (depend s)" by (simp add:trancl_domain)
       
  2223       then obtain n where d: "(Th th1, n) \<in> depend s" by (auto simp:Domain_def)
       
  2224       from depend_target_th [OF this]
       
  2225       obtain cs' where "n = Cs cs'" by auto
       
  2226       with d have "(Th th1, Cs cs') \<in> depend s" by simp
       
  2227       with runing_1 have "False"
       
  2228         apply (unfold runing_def readys_def s_depend_def)
       
  2229         by (auto simp:eq_waiting)
       
  2230       thus ?thesis by simp
       
  2231     qed
       
  2232   next
       
  2233     assume th1'_in: "th1' \<in> dependents (wq s) th1"
       
  2234     from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
       
  2235     thus ?thesis 
       
  2236     proof
       
  2237       assume "th2' = th2"
       
  2238       with th1'_in eq_th12 have "th2 \<in> dependents (wq s) th1" by simp
       
  2239       hence "(Th th2, Th th1) \<in> (depend s)^+"
       
  2240         by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
       
  2241       hence "Th th2 \<in> Domain ((depend s)^+)" using Domain_def [of "(depend s)^+"]
       
  2242         by auto
       
  2243       hence "Th th2 \<in> Domain (depend s)" by (simp add:trancl_domain)
       
  2244       then obtain n where d: "(Th th2, n) \<in> depend s" by (auto simp:Domain_def)
       
  2245       from depend_target_th [OF this]
       
  2246       obtain cs' where "n = Cs cs'" by auto
       
  2247       with d have "(Th th2, Cs cs') \<in> depend s" by simp
       
  2248       with runing_2 have "False"
       
  2249         apply (unfold runing_def readys_def s_depend_def)
       
  2250         by (auto simp:eq_waiting)
       
  2251       thus ?thesis by simp
       
  2252     next
       
  2253       assume "th2' \<in> dependents (wq s) th2"
       
  2254       with eq_th12 have "th1' \<in> dependents (wq s) th2" by simp
       
  2255       hence h1: "(Th th1', Th th2) \<in> (depend s)^+"
       
  2256         by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
       
  2257       from th1'_in have h2: "(Th th1', Th th1) \<in> (depend s)^+"
       
  2258         by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
       
  2259       show ?thesis
       
  2260       proof(rule dchain_unique[OF vt h1 _ h2, symmetric])
       
  2261         from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
       
  2262         from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
       
  2263       qed
       
  2264     qed
       
  2265   qed
       
  2266 qed
       
  2267 
       
  2268 lemma create_pre:
       
  2269   assumes stp: "step s e"
       
  2270   and not_in: "th \<notin> threads s"
       
  2271   and is_in: "th \<in> threads (e#s)"
       
  2272   obtains prio where "e = Create th prio"
       
  2273 proof -
       
  2274   from assms  
       
  2275   show ?thesis
       
  2276   proof(cases)
       
  2277     case (thread_create thread prio)
       
  2278     with is_in not_in have "e = Create th prio" by simp
       
  2279     from that[OF this] show ?thesis .
       
  2280   next
       
  2281     case (thread_exit thread)
       
  2282     with assms show ?thesis by (auto intro!:that)
       
  2283   next
       
  2284     case (thread_P thread)
       
  2285     with assms show ?thesis by (auto intro!:that)
       
  2286   next
       
  2287     case (thread_V thread)
       
  2288     with assms show ?thesis by (auto intro!:that)
       
  2289   next 
       
  2290     case (thread_set thread)
       
  2291     with assms show ?thesis by (auto intro!:that)
       
  2292   qed
       
  2293 qed
       
  2294 
       
  2295 lemma length_down_to_in: 
       
  2296   assumes le_ij: "i \<le> j"
       
  2297     and le_js: "j \<le> length s"
       
  2298   shows "length (down_to j i s) = j - i"
       
  2299 proof -
       
  2300   have "length (down_to j i s) = length (from_to i j (rev s))"
       
  2301     by (unfold down_to_def, auto)
       
  2302   also have "\<dots> = j - i"
       
  2303   proof(rule length_from_to_in[OF le_ij])
       
  2304     from le_js show "j \<le> length (rev s)" by simp
       
  2305   qed
       
  2306   finally show ?thesis .
       
  2307 qed
       
  2308 
       
  2309 
       
  2310 lemma moment_head: 
       
  2311   assumes le_it: "Suc i \<le> length t"
       
  2312   obtains e where "moment (Suc i) t = e#moment i t"
       
  2313 proof -
       
  2314   have "i \<le> Suc i" by simp
       
  2315   from length_down_to_in [OF this le_it]
       
  2316   have "length (down_to (Suc i) i t) = 1" by auto
       
  2317   then obtain e where "down_to (Suc i) i t = [e]"
       
  2318     apply (cases "(down_to (Suc i) i t)") by auto
       
  2319   moreover have "down_to (Suc i) 0 t = down_to (Suc i) i t @ down_to i 0 t"
       
  2320     by (rule down_to_conc[symmetric], auto)
       
  2321   ultimately have eq_me: "moment (Suc i) t = e#(moment i t)"
       
  2322     by (auto simp:down_to_moment)
       
  2323   from that [OF this] show ?thesis .
       
  2324 qed
       
  2325 
       
  2326 lemma cnp_cnv_eq:
       
  2327   fixes th s
       
  2328   assumes "vt step s"
       
  2329   and "th \<notin> threads s"
       
  2330   shows "cntP s th = cntV s th"
       
  2331 proof -
       
  2332   from assms show ?thesis
       
  2333   proof(induct)
       
  2334     case (vt_cons s e)
       
  2335     have ih: "th \<notin> threads s \<Longrightarrow> cntP s th = cntV s th" by fact
       
  2336     have not_in: "th \<notin> threads (e # s)" by fact
       
  2337     have "step s e" by fact
       
  2338     thus ?case proof(cases)
       
  2339       case (thread_create thread prio)
       
  2340       assume eq_e: "e = Create thread prio"
       
  2341       hence "thread \<in> threads (e#s)" by simp
       
  2342       with not_in and eq_e have "th \<notin> threads s" by auto
       
  2343       from ih [OF this] show ?thesis using eq_e
       
  2344         by (auto simp:cntP_def cntV_def count_def)
       
  2345     next
       
  2346       case (thread_exit thread)
       
  2347       assume eq_e: "e = Exit thread"
       
  2348         and not_holding: "holdents s thread = {}"
       
  2349       have vt_s: "vt step s" by fact
       
  2350       from finite_holding[OF vt_s] have "finite (holdents s thread)" .
       
  2351       with not_holding have "cntCS s thread = 0" by (unfold cntCS_def, auto)
       
  2352       moreover have "thread \<in> readys s" using thread_exit by (auto simp:runing_def)
       
  2353       moreover note cnp_cnv_cncs[OF vt_s, of thread]
       
  2354       ultimately have eq_thread: "cntP s thread = cntV s thread" by auto
       
  2355       show ?thesis
       
  2356       proof(cases "th = thread")
       
  2357         case True
       
  2358         with eq_thread eq_e show ?thesis 
       
  2359           by (auto simp:cntP_def cntV_def count_def)
       
  2360       next
       
  2361         case False
       
  2362         with not_in and eq_e have "th \<notin> threads s" by simp
       
  2363         from ih[OF this] and eq_e show ?thesis 
       
  2364            by (auto simp:cntP_def cntV_def count_def)
       
  2365       qed
       
  2366     next
       
  2367       case (thread_P thread cs)
       
  2368       assume eq_e: "e = P thread cs"
       
  2369       have "thread \<in> runing s" by fact
       
  2370       with not_in eq_e have neq_th: "thread \<noteq> th" 
       
  2371         by (auto simp:runing_def readys_def)
       
  2372       from not_in eq_e have "th \<notin> threads s" by simp
       
  2373       from ih[OF this] and neq_th and eq_e show ?thesis
       
  2374         by (auto simp:cntP_def cntV_def count_def)
       
  2375     next
       
  2376       case (thread_V thread cs)
       
  2377       assume eq_e: "e = V thread cs"
       
  2378       have "thread \<in> runing s" by fact
       
  2379       with not_in eq_e have neq_th: "thread \<noteq> th" 
       
  2380         by (auto simp:runing_def readys_def)
       
  2381       from not_in eq_e have "th \<notin> threads s" by simp
       
  2382       from ih[OF this] and neq_th and eq_e show ?thesis
       
  2383         by (auto simp:cntP_def cntV_def count_def)
       
  2384     next
       
  2385       case (thread_set thread prio)
       
  2386       assume eq_e: "e = Set thread prio"
       
  2387         and "thread \<in> runing s"
       
  2388       hence "thread \<in> threads (e#s)" 
       
  2389         by (simp add:runing_def readys_def)
       
  2390       with not_in and eq_e have "th \<notin> threads s" by auto
       
  2391       from ih [OF this] show ?thesis using eq_e
       
  2392         by (auto simp:cntP_def cntV_def count_def)  
       
  2393     qed
       
  2394   next
       
  2395     case vt_nil
       
  2396     show ?case by (auto simp:cntP_def cntV_def count_def)
       
  2397   qed
       
  2398 qed
       
  2399 
       
  2400 lemma eq_depend: 
       
  2401   "depend (wq s) = depend s"
       
  2402 by (unfold cs_depend_def s_depend_def, auto)
       
  2403 
       
  2404 lemma count_eq_dependents:
       
  2405   assumes vt: "vt step s"
       
  2406   and eq_pv: "cntP s th = cntV s th"
       
  2407   shows "dependents (wq s) th = {}"
       
  2408 proof -
       
  2409   from cnp_cnv_cncs[OF vt] and eq_pv
       
  2410   have "cntCS s th = 0" 
       
  2411     by (auto split:if_splits)
       
  2412   moreover have "finite {cs. (Cs cs, Th th) \<in> depend s}"
       
  2413   proof -
       
  2414     from finite_holding[OF vt, of th] show ?thesis
       
  2415       by (simp add:holdents_def)
       
  2416   qed
       
  2417   ultimately have h: "{cs. (Cs cs, Th th) \<in> depend s} = {}"
       
  2418     by (unfold cntCS_def holdents_def cs_dependents_def, auto)
       
  2419   show ?thesis
       
  2420   proof(unfold cs_dependents_def)
       
  2421     { assume "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}"
       
  2422       then obtain th' where "(Th th', Th th) \<in> (depend (wq s))\<^sup>+" by auto
       
  2423       hence "False"
       
  2424       proof(cases)
       
  2425         assume "(Th th', Th th) \<in> depend (wq s)"
       
  2426         thus "False" by (auto simp:cs_depend_def)
       
  2427       next
       
  2428         fix c
       
  2429         assume "(c, Th th) \<in> depend (wq s)"
       
  2430         with h and eq_depend show "False"
       
  2431           by (cases c, auto simp:cs_depend_def)
       
  2432       qed
       
  2433     } thus "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} = {}" by auto
       
  2434   qed
       
  2435 qed
       
  2436 
       
  2437 lemma dependents_threads:
       
  2438   fixes s th
       
  2439   assumes vt: "vt step s"
       
  2440   shows "dependents (wq s) th \<subseteq> threads s"
       
  2441 proof
       
  2442   { fix th th'
       
  2443     assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (depend (wq s))\<^sup>+}"
       
  2444     have "Th th \<in> Domain (depend s)"
       
  2445     proof -
       
  2446       from h obtain th' where "(Th th, Th th') \<in> (depend (wq s))\<^sup>+" by auto
       
  2447       hence "(Th th) \<in> Domain ( (depend (wq s))\<^sup>+)" by (auto simp:Domain_def)
       
  2448       with trancl_domain have "(Th th) \<in> Domain (depend (wq s))" by simp
       
  2449       thus ?thesis using eq_depend by simp
       
  2450     qed
       
  2451     from dm_depend_threads[OF vt this]
       
  2452     have "th \<in> threads s" .
       
  2453   } note hh = this
       
  2454   fix th1 
       
  2455   assume "th1 \<in> dependents (wq s) th"
       
  2456   hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (depend (wq s))\<^sup>+}"
       
  2457     by (unfold cs_dependents_def, simp)
       
  2458   from hh [OF this] show "th1 \<in> threads s" .
       
  2459 qed
       
  2460 
       
  2461 lemma finite_threads:
       
  2462   assumes vt: "vt step s"
       
  2463   shows "finite (threads s)"
       
  2464 proof -
       
  2465   from vt show ?thesis
       
  2466   proof(induct)
       
  2467     case (vt_cons s e)
       
  2468     assume vt: "vt step s"
       
  2469     and step: "step s e"
       
  2470     and ih: "finite (threads s)"
       
  2471     from step
       
  2472     show ?case
       
  2473     proof(cases)
       
  2474       case (thread_create thread prio)
       
  2475       assume eq_e: "e = Create thread prio"
       
  2476       with ih
       
  2477       show ?thesis by (unfold eq_e, auto)
       
  2478     next
       
  2479       case (thread_exit thread)
       
  2480       assume eq_e: "e = Exit thread"
       
  2481       with ih show ?thesis 
       
  2482         by (unfold eq_e, auto)
       
  2483     next
       
  2484       case (thread_P thread cs)
       
  2485       assume eq_e: "e = P thread cs"
       
  2486       with ih show ?thesis by (unfold eq_e, auto)
       
  2487     next
       
  2488       case (thread_V thread cs)
       
  2489       assume eq_e: "e = V thread cs"
       
  2490       with ih show ?thesis by (unfold eq_e, auto)
       
  2491     next 
       
  2492       case (thread_set thread prio)
       
  2493       from vt_cons thread_set show ?thesis by simp
       
  2494     qed
       
  2495   next
       
  2496     case vt_nil
       
  2497     show ?case by (auto)
       
  2498   qed
       
  2499 qed
       
  2500 
       
  2501 lemma Max_f_mono:
       
  2502   assumes seq: "A \<subseteq> B"
       
  2503   and np: "A \<noteq> {}"
       
  2504   and fnt: "finite B"
       
  2505   shows "Max (f ` A) \<le> Max (f ` B)"
       
  2506 proof(rule Max_mono)
       
  2507   from seq show "f ` A \<subseteq> f ` B" by auto
       
  2508 next
       
  2509   from np show "f ` A \<noteq> {}" by auto
       
  2510 next
       
  2511   from fnt and seq show "finite (f ` B)" by auto
       
  2512 qed
       
  2513 
       
  2514 lemma cp_le:
       
  2515   assumes vt: "vt step s"
       
  2516   and th_in: "th \<in> threads s"
       
  2517   shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
       
  2518 proof(unfold cp_eq_cpreced cpreced_def cs_dependents_def)
       
  2519   show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}))
       
  2520          \<le> Max ((\<lambda>th. preced th s) ` threads s)"
       
  2521     (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
       
  2522   proof(rule Max_f_mono)
       
  2523     show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}" by simp
       
  2524   next
       
  2525     from finite_threads [OF vt]
       
  2526     show "finite (threads s)" .
       
  2527   next
       
  2528     from th_in
       
  2529     show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> threads s"
       
  2530       apply (auto simp:Domain_def)
       
  2531       apply (rule_tac dm_depend_threads[OF vt])
       
  2532       apply (unfold trancl_domain [of "depend s", symmetric])
       
  2533       by (unfold cs_depend_def s_depend_def, auto simp:Domain_def)
       
  2534   qed
       
  2535 qed
       
  2536 
       
  2537 lemma le_cp:
       
  2538   assumes vt: "vt step s"
       
  2539   shows "preced th s \<le> cp s th"
       
  2540 proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
       
  2541   show "Prc (original_priority th s) (birthtime th s)
       
  2542     \<le> Max (insert (Prc (original_priority th s) (birthtime th s))
       
  2543             ((\<lambda>th. Prc (original_priority th s) (birthtime th s)) ` dependents (wq s) th))"
       
  2544     (is "?l \<le> Max (insert ?l ?A)")
       
  2545   proof(cases "?A = {}")
       
  2546     case False
       
  2547     have "finite ?A" (is "finite (?f ` ?B)")
       
  2548     proof -
       
  2549       have "finite ?B" 
       
  2550       proof-
       
  2551         have "finite {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}"
       
  2552         proof -
       
  2553           let ?F = "\<lambda> (x, y). the_th x"
       
  2554           have "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
       
  2555             apply (auto simp:image_def)
       
  2556             by (rule_tac x = "(Th x, Th th)" in bexI, auto)
       
  2557           moreover have "finite \<dots>"
       
  2558           proof -
       
  2559             from finite_depend[OF vt] have "finite (depend s)" .
       
  2560             hence "finite ((depend (wq s))\<^sup>+)"
       
  2561               apply (unfold finite_trancl)
       
  2562               by (auto simp: s_depend_def cs_depend_def wq_def)
       
  2563             thus ?thesis by auto
       
  2564           qed
       
  2565           ultimately show ?thesis by (auto intro:finite_subset)
       
  2566         qed
       
  2567         thus ?thesis by (simp add:cs_dependents_def)
       
  2568       qed
       
  2569       thus ?thesis by simp
       
  2570     qed
       
  2571     from Max_insert [OF this False, of ?l] show ?thesis by auto
       
  2572   next
       
  2573     case True
       
  2574     thus ?thesis by auto
       
  2575   qed
       
  2576 qed
       
  2577 
       
  2578 lemma max_cp_eq: 
       
  2579   assumes vt: "vt step s"
       
  2580   shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)"
       
  2581   (is "?l = ?r")
       
  2582 proof(cases "threads s = {}")
       
  2583   case True
       
  2584   thus ?thesis by auto
       
  2585 next
       
  2586   case False
       
  2587   have "?l \<in> ((cp s) ` threads s)"
       
  2588   proof(rule Max_in)
       
  2589     from finite_threads[OF vt] 
       
  2590     show "finite (cp s ` threads s)" by auto
       
  2591   next
       
  2592     from False show "cp s ` threads s \<noteq> {}" by auto
       
  2593   qed
       
  2594   then obtain th 
       
  2595     where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto
       
  2596   have "\<dots> \<le> ?r" by (rule cp_le[OF vt th_in])
       
  2597   moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th")
       
  2598   proof -
       
  2599     have "?r \<in> (?f ` ?A)"
       
  2600     proof(rule Max_in)
       
  2601       from finite_threads[OF vt]
       
  2602       show " finite ((\<lambda>th. preced th s) ` threads s)" by auto
       
  2603     next
       
  2604       from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
       
  2605     qed
       
  2606     then obtain th' where 
       
  2607       th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto
       
  2608     from le_cp [OF vt, of th']  eq_r
       
  2609     have "?r \<le> cp s th'" by auto
       
  2610     moreover have "\<dots> \<le> cp s th"
       
  2611     proof(fold eq_l)
       
  2612       show " cp s th' \<le> Max (cp s ` threads s)"
       
  2613       proof(rule Max_ge)
       
  2614         from th_in' show "cp s th' \<in> cp s ` threads s"
       
  2615           by auto
       
  2616       next
       
  2617         from finite_threads[OF vt]
       
  2618         show "finite (cp s ` threads s)" by auto
       
  2619       qed
       
  2620     qed
       
  2621     ultimately show ?thesis by auto
       
  2622   qed
       
  2623   ultimately show ?thesis using eq_l by auto
       
  2624 qed
       
  2625 
       
  2626 lemma max_cp_readys_threads_pre:
       
  2627   assumes vt: "vt step s"
       
  2628   and np: "threads s \<noteq> {}"
       
  2629   shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
       
  2630 proof(unfold max_cp_eq[OF vt])
       
  2631   show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
       
  2632   proof -
       
  2633     let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
       
  2634     let ?f = "(\<lambda>th. preced th s)"
       
  2635     have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
       
  2636     proof(rule Max_in)
       
  2637       from finite_threads[OF vt] show "finite (?f ` threads s)" by simp
       
  2638     next
       
  2639       from np show "?f ` threads s \<noteq> {}" by simp
       
  2640     qed
       
  2641     then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
       
  2642       by (auto simp:Image_def)
       
  2643     from th_chain_to_ready [OF vt tm_in]
       
  2644     have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+)" .
       
  2645     thus ?thesis
       
  2646     proof
       
  2647       assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+ "
       
  2648       then obtain th' where th'_in: "th' \<in> readys s" 
       
  2649         and tm_chain:"(Th tm, Th th') \<in> (depend s)\<^sup>+" by auto
       
  2650       have "cp s th' = ?f tm"
       
  2651       proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
       
  2652         from dependents_threads[OF vt] finite_threads[OF vt]
       
  2653         show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th'))" 
       
  2654           by (auto intro:finite_subset)
       
  2655       next
       
  2656         fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')"
       
  2657         from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
       
  2658         moreover have "p \<le> \<dots>"
       
  2659         proof(rule Max_ge)
       
  2660           from finite_threads[OF vt]
       
  2661           show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2662         next
       
  2663           from p_in and th'_in and dependents_threads[OF vt, of th']
       
  2664           show "p \<in> (\<lambda>th. preced th s) ` threads s"
       
  2665             by (auto simp:readys_def)
       
  2666         qed
       
  2667         ultimately show "p \<le> preced tm s" by auto
       
  2668       next
       
  2669         show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')"
       
  2670         proof -
       
  2671           from tm_chain
       
  2672           have "tm \<in> dependents (wq s) th'"
       
  2673             by (unfold cs_dependents_def s_depend_def cs_depend_def, auto)
       
  2674           thus ?thesis by auto
       
  2675         qed
       
  2676       qed
       
  2677       with tm_max
       
  2678       have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2679       show ?thesis
       
  2680       proof (fold h, rule Max_eqI)
       
  2681         fix q 
       
  2682         assume "q \<in> cp s ` readys s"
       
  2683         then obtain th1 where th1_in: "th1 \<in> readys s"
       
  2684           and eq_q: "q = cp s th1" by auto
       
  2685         show "q \<le> cp s th'"
       
  2686           apply (unfold h eq_q)
       
  2687           apply (unfold cp_eq_cpreced cpreced_def)
       
  2688           apply (rule Max_mono)
       
  2689         proof -
       
  2690           from dependents_threads [OF vt, of th1] th1_in
       
  2691           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<subseteq> 
       
  2692                  (\<lambda>th. preced th s) ` threads s"
       
  2693             by (auto simp:readys_def)
       
  2694         next
       
  2695           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<noteq> {}" by simp
       
  2696         next
       
  2697           from finite_threads[OF vt] 
       
  2698           show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2699         qed
       
  2700       next
       
  2701         from finite_threads[OF vt]
       
  2702         show "finite (cp s ` readys s)" by (auto simp:readys_def)
       
  2703       next
       
  2704         from th'_in
       
  2705         show "cp s th' \<in> cp s ` readys s" by simp
       
  2706       qed
       
  2707     next
       
  2708       assume tm_ready: "tm \<in> readys s"
       
  2709       show ?thesis
       
  2710       proof(fold tm_max)
       
  2711         have cp_eq_p: "cp s tm = preced tm s"
       
  2712         proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
       
  2713           fix y 
       
  2714           assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm)"
       
  2715           show "y \<le> preced tm s"
       
  2716           proof -
       
  2717             { fix y'
       
  2718               assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependents (wq s) tm)"
       
  2719               have "y' \<le> preced tm s"
       
  2720               proof(unfold tm_max, rule Max_ge)
       
  2721                 from hy' dependents_threads[OF vt, of tm]
       
  2722                 show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
       
  2723               next
       
  2724                 from finite_threads[OF vt] 
       
  2725                 show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2726               qed
       
  2727             } with hy show ?thesis by auto
       
  2728           qed
       
  2729         next
       
  2730           from dependents_threads[OF vt, of tm] finite_threads[OF vt]
       
  2731           show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm))"
       
  2732             by (auto intro:finite_subset)
       
  2733         next
       
  2734           show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm)"
       
  2735             by simp
       
  2736         qed 
       
  2737         moreover have "Max (cp s ` readys s) = cp s tm"
       
  2738         proof(rule Max_eqI)
       
  2739           from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
       
  2740         next
       
  2741           from finite_threads[OF vt]
       
  2742           show "finite (cp s ` readys s)" by (auto simp:readys_def)
       
  2743         next
       
  2744           fix y assume "y \<in> cp s ` readys s"
       
  2745           then obtain th1 where th1_readys: "th1 \<in> readys s"
       
  2746             and h: "y = cp s th1" by auto
       
  2747           show "y \<le> cp s tm"
       
  2748             apply(unfold cp_eq_p h)
       
  2749             apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
       
  2750           proof -
       
  2751             from finite_threads[OF vt]
       
  2752             show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2753           next
       
  2754             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<noteq> {}"
       
  2755               by simp
       
  2756           next
       
  2757             from dependents_threads[OF vt, of th1] th1_readys
       
  2758             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) 
       
  2759                     \<subseteq> (\<lambda>th. preced th s) ` threads s"
       
  2760               by (auto simp:readys_def)
       
  2761           qed
       
  2762         qed
       
  2763         ultimately show " Max (cp s ` readys s) = preced tm s" by simp
       
  2764       qed 
       
  2765     qed
       
  2766   qed
       
  2767 qed
       
  2768 
       
  2769 lemma max_cp_readys_threads:
       
  2770   assumes vt: "vt step s"
       
  2771   shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
       
  2772 proof(cases "threads s = {}")
       
  2773   case True
       
  2774   thus ?thesis 
       
  2775     by (auto simp:readys_def)
       
  2776 next
       
  2777   case False
       
  2778   show ?thesis by (rule max_cp_readys_threads_pre[OF vt False])
       
  2779 qed
       
  2780 
       
  2781 lemma readys_threads:
       
  2782   shows "readys s \<subseteq> threads s"
       
  2783 proof
       
  2784   fix th
       
  2785   assume "th \<in> readys s"
       
  2786   thus "th \<in> threads s"
       
  2787     by (unfold readys_def, auto)
       
  2788 qed
       
  2789 
       
  2790 lemma eq_holding: "holding (wq s) th cs = holding s th cs"
       
  2791   apply (unfold s_holding_def cs_holding_def, simp)
       
  2792   done
       
  2793 
       
  2794 lemma f_image_eq:
       
  2795   assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
       
  2796   shows "f ` A = g ` A"
       
  2797 proof
       
  2798   show "f ` A \<subseteq> g ` A"
       
  2799     by(rule image_subsetI, auto intro:h)
       
  2800 next
       
  2801   show "g ` A \<subseteq> f ` A"
       
  2802    by(rule image_subsetI, auto intro:h[symmetric])
       
  2803 qed
       
  2804 
       
  2805 end