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