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