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