PIPBasics.thy
changeset 92 4763aa246dbd
parent 90 ed938e2246b9
child 93 524bd3caa6b6
equal deleted inserted replaced
91:0525670d8e6a 92:4763aa246dbd
     1 theory PIPBasics
     1 theory CpsG
     2 imports PIPDefs 
     2 imports PIPDefs
     3 begin
     3 begin
       
     4 
       
     5 lemma f_image_eq:
       
     6   assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
       
     7   shows "f ` A = g ` A"
       
     8 proof
       
     9   show "f ` A \<subseteq> g ` A"
       
    10     by(rule image_subsetI, auto intro:h)
       
    11 next
       
    12   show "g ` A \<subseteq> f ` A"
       
    13    by (rule image_subsetI, auto intro:h[symmetric])
       
    14 qed
       
    15 
       
    16 lemma Max_fg_mono:
       
    17   assumes "finite A"
       
    18   and "\<forall> a \<in> A. f a \<le> g a"
       
    19   shows "Max (f ` A) \<le> Max (g ` A)"
       
    20 proof(cases "A = {}")
       
    21   case True
       
    22   thus ?thesis by auto
       
    23 next
       
    24   case False
       
    25   show ?thesis
       
    26   proof(rule Max.boundedI)
       
    27     from assms show "finite (f ` A)" by auto
       
    28   next
       
    29     from False show "f ` A \<noteq> {}" by auto
       
    30   next
       
    31     fix fa
       
    32     assume "fa \<in> f ` A"
       
    33     then obtain a where h_fa: "a \<in> A" "fa = f a" by auto
       
    34     show "fa \<le> Max (g ` A)"
       
    35     proof(rule Max_ge_iff[THEN iffD2])
       
    36       from assms show "finite (g ` A)" by auto
       
    37     next
       
    38       from False show "g ` A \<noteq> {}" by auto
       
    39     next
       
    40       from h_fa have "g a \<in> g ` A" by auto
       
    41       moreover have "fa \<le> g a" using h_fa assms(2) by auto
       
    42       ultimately show "\<exists>a\<in>g ` A. fa \<le> a" by auto
       
    43     qed
       
    44   qed
       
    45 qed 
       
    46 
       
    47 lemma Max_f_mono:
       
    48   assumes seq: "A \<subseteq> B"
       
    49   and np: "A \<noteq> {}"
       
    50   and fnt: "finite B"
       
    51   shows "Max (f ` A) \<le> Max (f ` B)"
       
    52 proof(rule Max_mono)
       
    53   from seq show "f ` A \<subseteq> f ` B" by auto
       
    54 next
       
    55   from np show "f ` A \<noteq> {}" by auto
       
    56 next
       
    57   from fnt and seq show "finite (f ` B)" by auto
       
    58 qed
       
    59 
       
    60 lemma Max_UNION: 
       
    61   assumes "finite A"
       
    62   and "A \<noteq> {}"
       
    63   and "\<forall> M \<in> f ` A. finite M"
       
    64   and "\<forall> M \<in> f ` A. M \<noteq> {}"
       
    65   shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
       
    66   using assms[simp]
       
    67 proof -
       
    68   have "?L = Max (\<Union>(f ` A))"
       
    69     by (fold Union_image_eq, simp)
       
    70   also have "... = ?R"
       
    71     by (subst Max_Union, simp+)
       
    72   finally show ?thesis .
       
    73 qed
       
    74 
       
    75 lemma max_Max_eq:
       
    76   assumes "finite A"
       
    77     and "A \<noteq> {}"
       
    78     and "x = y"
       
    79   shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
       
    80 proof -
       
    81   have "?R = Max (insert y A)" by simp
       
    82   also from assms have "... = ?L"
       
    83       by (subst Max.insert, simp+)
       
    84   finally show ?thesis by simp
       
    85 qed
       
    86 
       
    87 lemma birth_time_lt:  
       
    88   assumes "s \<noteq> []"
       
    89   shows "last_set th s < length s"
       
    90   using assms
       
    91 proof(induct s)
       
    92   case (Cons a s)
       
    93   show ?case
       
    94   proof(cases "s \<noteq> []")
       
    95     case False
       
    96     thus ?thesis
       
    97       by (cases a, auto)
       
    98   next
       
    99     case True
       
   100     show ?thesis using Cons(1)[OF True]
       
   101       by (cases a, auto)
       
   102   qed
       
   103 qed simp
       
   104 
       
   105 lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
       
   106   by (induct s, auto)
       
   107 
       
   108 lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
       
   109   by (drule_tac th_in_ne, unfold preced_def, auto intro: birth_time_lt)
       
   110 
       
   111 lemma eq_RAG: 
       
   112   "RAG (wq s) = RAG s"
       
   113   by (unfold cs_RAG_def s_RAG_def, auto)
       
   114 
       
   115 lemma waiting_holding:
       
   116   assumes "waiting (s::state) th cs"
       
   117   obtains th' where "holding s th' cs"
       
   118 proof -
       
   119   from assms[unfolded s_waiting_def, folded wq_def]
       
   120   obtain th' where "th' \<in> set (wq s cs)" "th' = hd (wq s cs)"
       
   121     by (metis empty_iff hd_in_set list.set(1))
       
   122   hence "holding s th' cs" 
       
   123     by (unfold s_holding_def, fold wq_def, auto)
       
   124   from that[OF this] show ?thesis .
       
   125 qed
       
   126 
       
   127 lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
       
   128 unfolding cp_def wq_def
       
   129 apply(induct s rule: schs.induct)
       
   130 apply(simp add: Let_def cpreced_initial)
       
   131 apply(simp add: Let_def)
       
   132 apply(simp add: Let_def)
       
   133 apply(simp add: Let_def)
       
   134 apply(subst (2) schs.simps)
       
   135 apply(simp add: Let_def)
       
   136 apply(subst (2) schs.simps)
       
   137 apply(simp add: Let_def)
       
   138 done
       
   139 
       
   140 lemma cp_alt_def:
       
   141   "cp s th =  
       
   142            Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})"
       
   143 proof -
       
   144   have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) =
       
   145         Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" 
       
   146           (is "Max (_ ` ?L) = Max (_ ` ?R)")
       
   147   proof -
       
   148     have "?L = ?R" 
       
   149     by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def)
       
   150     thus ?thesis by simp
       
   151   qed
       
   152   thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
       
   153 qed
       
   154 
       
   155 (* ccc *)
       
   156 
     4 
   157 
     5 locale valid_trace = 
   158 locale valid_trace = 
     6   fixes s
   159   fixes s
     7   assumes vt : "vt s"
   160   assumes vt : "vt s"
     8 
   161 
    14 lemma pip_e: "PIP s e"
   167 lemma pip_e: "PIP s e"
    15   using vt_e by (cases, simp)  
   168   using vt_e by (cases, simp)  
    16 
   169 
    17 end
   170 end
    18 
   171 
       
   172 locale valid_trace_create = valid_trace_e + 
       
   173   fixes th prio
       
   174   assumes is_create: "e = Create th prio"
       
   175 
       
   176 locale valid_trace_exit = valid_trace_e + 
       
   177   fixes th
       
   178   assumes is_exit: "e = Exit th"
       
   179 
       
   180 locale valid_trace_p = valid_trace_e + 
       
   181   fixes th cs
       
   182   assumes is_p: "e = P th cs"
       
   183 
       
   184 locale valid_trace_v = valid_trace_e + 
       
   185   fixes th cs
       
   186   assumes is_v: "e = V th cs"
       
   187 begin
       
   188   definition "rest = tl (wq s cs)"
       
   189   definition "wq' = (SOME q. distinct q \<and> set q = set rest)"
       
   190 end
       
   191 
       
   192 locale valid_trace_v_n = valid_trace_v +
       
   193   assumes rest_nnl: "rest \<noteq> []"
       
   194 
       
   195 locale valid_trace_v_e = valid_trace_v +
       
   196   assumes rest_nil: "rest = []"
       
   197 
       
   198 locale valid_trace_set= valid_trace_e + 
       
   199   fixes th prio
       
   200   assumes is_set: "e = Set th prio"
       
   201 
       
   202 context valid_trace
       
   203 begin
       
   204 
       
   205 lemma ind [consumes 0, case_names Nil Cons, induct type]:
       
   206   assumes "PP []"
       
   207      and "(\<And>s e. valid_trace_e s e \<Longrightarrow>
       
   208                    PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
       
   209      shows "PP s"
       
   210 proof(induct rule:vt.induct[OF vt, case_names Init Step])
       
   211   case Init
       
   212   from assms(1) show ?case .
       
   213 next
       
   214   case (Step s e)
       
   215   show ?case
       
   216   proof(rule assms(2))
       
   217     show "valid_trace_e s e" using Step by (unfold_locales, auto)
       
   218   next
       
   219     show "PP s" using Step by simp
       
   220   next
       
   221     show "PIP s e" using Step by simp
       
   222   qed
       
   223 qed
       
   224 
       
   225 lemma  vt_moment: "\<And> t. vt (moment t s)"
       
   226 proof(induct rule:ind)
       
   227   case Nil
       
   228   thus ?case by (simp add:vt_nil)
       
   229 next
       
   230   case (Cons s e t)
       
   231   show ?case
       
   232   proof(cases "t \<ge> length (e#s)")
       
   233     case True
       
   234     from True have "moment t (e#s) = e#s" by simp
       
   235     thus ?thesis using Cons
       
   236       by (simp add:valid_trace_def valid_trace_e_def, auto)
       
   237   next
       
   238     case False
       
   239     from Cons have "vt (moment t s)" by simp
       
   240     moreover have "moment t (e#s) = moment t s"
       
   241     proof -
       
   242       from False have "t \<le> length s" by simp
       
   243       from moment_app [OF this, of "[e]"] 
       
   244       show ?thesis by simp
       
   245     qed
       
   246     ultimately show ?thesis by simp
       
   247   qed
       
   248 qed
       
   249 
       
   250 lemma finite_threads:
       
   251   shows "finite (threads s)"
       
   252 using vt by (induct) (auto elim: step.cases)
       
   253 
       
   254 end
       
   255 
       
   256 lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
       
   257   by (unfold s_RAG_def, auto)
       
   258 
       
   259 locale valid_moment = valid_trace + 
       
   260   fixes i :: nat
       
   261 
       
   262 sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
       
   263   by (unfold_locales, insert vt_moment, auto)
       
   264 
       
   265 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
       
   266   by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
       
   267 
       
   268 lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
       
   269   by (unfold s_holding_def wq_def cs_holding_def, simp)
       
   270 
    19 lemma runing_ready: 
   271 lemma runing_ready: 
    20   shows "runing s \<subseteq> readys s"
   272   shows "runing s \<subseteq> readys s"
    21   unfolding runing_def readys_def
   273   unfolding runing_def readys_def
    22   by auto 
   274   by auto 
    23 
   275 
    24 lemma readys_threads:
   276 lemma readys_threads:
    25   shows "readys s \<subseteq> threads s"
   277   shows "readys s \<subseteq> threads s"
    26   unfolding readys_def
   278   unfolding readys_def
    27   by auto
   279   by auto
    28 
   280 
    29 lemma wq_v_neq:
   281 lemma wq_v_neq [simp]:
    30    "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
   282    "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)
   283   by (auto simp:wq_def Let_def cp_def split:list.splits)
    32 
   284 
    33 lemma runing_head:
   285 lemma runing_head:
    34   assumes "th \<in> runing s"
   286   assumes "th \<in> runing s"
    38   by (simp add:runing_def readys_def s_waiting_def wq_def)
   290   by (simp add:runing_def readys_def s_waiting_def wq_def)
    39 
   291 
    40 context valid_trace
   292 context valid_trace
    41 begin
   293 begin
    42 
   294 
       
   295 lemma runing_wqE:
       
   296   assumes "th \<in> runing s"
       
   297   and "th \<in> set (wq s cs)"
       
   298   obtains rest where "wq s cs = th#rest"
       
   299 proof -
       
   300   from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest"
       
   301     by (meson list.set_cases)
       
   302   have "th' = th"
       
   303   proof(rule ccontr)
       
   304     assume "th' \<noteq> th"
       
   305     hence "th \<noteq> hd (wq s cs)" using eq_wq by auto 
       
   306     with assms(2)
       
   307     have "waiting s th cs" 
       
   308       by (unfold s_waiting_def, fold wq_def, auto)
       
   309     with assms show False 
       
   310       by (unfold runing_def readys_def, auto)
       
   311   qed
       
   312   with eq_wq that show ?thesis by metis
       
   313 qed
       
   314 
       
   315 end
       
   316 
       
   317 context valid_trace_create
       
   318 begin
       
   319 
       
   320 lemma wq_neq_simp [simp]:
       
   321   shows "wq (e#s) cs' = wq s cs'"
       
   322     using assms unfolding is_create wq_def
       
   323   by (auto simp:Let_def)
       
   324 
       
   325 lemma wq_distinct_kept:
       
   326   assumes "distinct (wq s cs')"
       
   327   shows "distinct (wq (e#s) cs')"
       
   328   using assms by simp
       
   329 end
       
   330 
       
   331 context valid_trace_exit
       
   332 begin
       
   333 
       
   334 lemma wq_neq_simp [simp]:
       
   335   shows "wq (e#s) cs' = wq s cs'"
       
   336     using assms unfolding is_exit wq_def
       
   337   by (auto simp:Let_def)
       
   338 
       
   339 lemma wq_distinct_kept:
       
   340   assumes "distinct (wq s cs')"
       
   341   shows "distinct (wq (e#s) cs')"
       
   342   using assms by simp
       
   343 end
       
   344 
       
   345 context valid_trace_p
       
   346 begin
       
   347 
       
   348 lemma wq_neq_simp [simp]:
       
   349   assumes "cs' \<noteq> cs"
       
   350   shows "wq (e#s) cs' = wq s cs'"
       
   351     using assms unfolding is_p wq_def
       
   352   by (auto simp:Let_def)
       
   353 
       
   354 lemma runing_th_s:
       
   355   shows "th \<in> runing s"
       
   356 proof -
       
   357   from pip_e[unfolded is_p]
       
   358   show ?thesis by (cases, simp)
       
   359 qed
       
   360 
       
   361 lemma ready_th_s: "th \<in> readys s"
       
   362   using runing_th_s
       
   363   by (unfold runing_def, auto)
       
   364 
       
   365 lemma live_th_s: "th \<in> threads s"
       
   366   using readys_threads ready_th_s by auto
       
   367 
       
   368 lemma live_th_es: "th \<in> threads (e#s)"
       
   369   using live_th_s 
       
   370   by (unfold is_p, simp)
       
   371 
       
   372 lemma th_not_waiting: 
       
   373   "\<not> waiting s th c"
       
   374 proof -
       
   375   have "th \<in> readys s"
       
   376     using runing_ready runing_th_s by blast 
       
   377   thus ?thesis
       
   378     by (unfold readys_def, auto)
       
   379 qed
       
   380 
       
   381 lemma waiting_neq_th: 
       
   382   assumes "waiting s t c"
       
   383   shows "t \<noteq> th"
       
   384   using assms using th_not_waiting by blast 
       
   385 
       
   386 lemma th_not_in_wq: 
       
   387   shows "th \<notin> set (wq s cs)"
       
   388 proof
       
   389   assume otherwise: "th \<in> set (wq s cs)"
       
   390   from runing_wqE[OF runing_th_s this]
       
   391   obtain rest where eq_wq: "wq s cs = th#rest" by blast
       
   392   with otherwise
       
   393   have "holding s th cs"
       
   394     by (unfold s_holding_def, fold wq_def, simp)
       
   395   hence cs_th_RAG: "(Cs cs, Th th) \<in> RAG s"
       
   396     by (unfold s_RAG_def, fold holding_eq, auto)
       
   397   from pip_e[unfolded is_p]
       
   398   show False
       
   399   proof(cases)
       
   400     case (thread_P)
       
   401     with cs_th_RAG show ?thesis by auto
       
   402   qed
       
   403 qed
       
   404 
       
   405 lemma wq_es_cs: 
       
   406   "wq (e#s) cs =  wq s cs @ [th]"
       
   407   by (unfold is_p wq_def, auto simp:Let_def)
       
   408 
       
   409 lemma wq_distinct_kept:
       
   410   assumes "distinct (wq s cs')"
       
   411   shows "distinct (wq (e#s) cs')"
       
   412 proof(cases "cs' = cs")
       
   413   case True
       
   414   show ?thesis using True assms th_not_in_wq
       
   415     by (unfold True wq_es_cs, auto)
       
   416 qed (insert assms, simp)
       
   417 
       
   418 end
       
   419 
       
   420 context valid_trace_v
       
   421 begin
       
   422 
       
   423 lemma wq_neq_simp [simp]:
       
   424   assumes "cs' \<noteq> cs"
       
   425   shows "wq (e#s) cs' = wq s cs'"
       
   426     using assms unfolding is_v wq_def
       
   427   by (auto simp:Let_def)
       
   428 
       
   429 lemma runing_th_s:
       
   430   shows "th \<in> runing s"
       
   431 proof -
       
   432   from pip_e[unfolded is_v]
       
   433   show ?thesis by (cases, simp)
       
   434 qed
       
   435 
       
   436 lemma th_not_waiting: 
       
   437   "\<not> waiting s th c"
       
   438 proof -
       
   439   have "th \<in> readys s"
       
   440     using runing_ready runing_th_s by blast 
       
   441   thus ?thesis
       
   442     by (unfold readys_def, auto)
       
   443 qed
       
   444 
       
   445 lemma waiting_neq_th: 
       
   446   assumes "waiting s t c"
       
   447   shows "t \<noteq> th"
       
   448   using assms using th_not_waiting by blast 
       
   449 
       
   450 lemma wq_s_cs:
       
   451   "wq s cs = th#rest"
       
   452 proof -
       
   453   from pip_e[unfolded is_v]
       
   454   show ?thesis
       
   455   proof(cases)
       
   456     case (thread_V)
       
   457     from this(2) show ?thesis
       
   458       by (unfold rest_def s_holding_def, fold wq_def,
       
   459                  metis empty_iff list.collapse list.set(1))
       
   460   qed
       
   461 qed
       
   462 
       
   463 lemma wq_es_cs:
       
   464   "wq (e#s) cs = wq'"
       
   465  using wq_s_cs[unfolded wq_def]
       
   466  by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) 
       
   467 
       
   468 lemma wq_distinct_kept:
       
   469   assumes "distinct (wq s cs')"
       
   470   shows "distinct (wq (e#s) cs')"
       
   471 proof(cases "cs' = cs")
       
   472   case True
       
   473   show ?thesis
       
   474   proof(unfold True wq_es_cs wq'_def, rule someI2)
       
   475     show "distinct rest \<and> set rest = set rest"
       
   476         using assms[unfolded True wq_s_cs] by auto
       
   477   qed simp
       
   478 qed (insert assms, simp)
       
   479 
       
   480 end
       
   481 
       
   482 context valid_trace_set
       
   483 begin
       
   484 
       
   485 lemma wq_neq_simp [simp]:
       
   486   shows "wq (e#s) cs' = wq s cs'"
       
   487     using assms unfolding is_set wq_def
       
   488   by (auto simp:Let_def)
       
   489 
       
   490 lemma wq_distinct_kept:
       
   491   assumes "distinct (wq s cs')"
       
   492   shows "distinct (wq (e#s) cs')"
       
   493   using assms by simp
       
   494 end
       
   495 
       
   496 context valid_trace
       
   497 begin
       
   498 
    43 lemma actor_inv: 
   499 lemma actor_inv: 
    44   assumes "PIP s e"
   500   assumes "PIP s e"
    45   and "\<not> isCreate e"
   501   and "\<not> isCreate e"
    46   shows "actor e \<in> runing s"
   502   shows "actor e \<in> runing s"
    47   using assms
   503   using assms
    48   by (induct, auto)
   504   by (induct, auto)
    49 
   505 
    50 lemma ind [consumes 0, case_names Nil Cons, induct type]:
   506 lemma isP_E:
    51   assumes "PP []"
   507   assumes "isP e"
    52      and "(\<And>s e. valid_trace s \<Longrightarrow> valid_trace (e#s) \<Longrightarrow>
   508   obtains cs where "e = P (actor e) cs"
    53                    PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))"
   509   using assms by (cases e, auto)
    54      shows "PP s"
   510 
    55 proof(rule vt.induct[OF vt])
   511 lemma isV_E:
    56   from assms(1) show "PP []" .
   512   assumes "isV e"
    57 next
   513   obtains cs where "e = V (actor e) cs"
    58   fix s e
   514   using assms by (cases e, auto) 
    59   assume h: "vt s" "PP s" "PIP s e"
       
    60   show "PP (e # s)"
       
    61   proof(cases rule:assms(2))
       
    62     from h(1) show v1: "valid_trace s" by (unfold_locales, simp)
       
    63   next
       
    64     from h(1,3) have "vt (e#s)" by auto
       
    65     thus "valid_trace (e # s)" by (unfold_locales, simp)
       
    66   qed (insert h, auto)
       
    67 qed
       
    68 
   515 
    69 lemma wq_distinct: "distinct (wq s cs)"
   516 lemma wq_distinct: "distinct (wq s cs)"
    70 proof(induct rule:ind)
   517 proof(induct rule:ind)
    71   case (Cons s e)
   518   case (Cons s e)
    72   from Cons(4,3)
   519   interpret vt_e: valid_trace_e s e using Cons by simp
    73   show ?case 
   520   show ?case 
    74   proof(induct)
   521   proof(cases e)
    75     case (thread_P th s cs1)
   522     case (Create th prio)
    76     show ?case 
   523     interpret vt_create: valid_trace_create s e th prio 
    77     proof(cases "cs = cs1")
   524       using Create by (unfold_locales, simp)
    78       case True
   525     show ?thesis using Cons by (simp add: vt_create.wq_distinct_kept) 
    79       thus ?thesis (is "distinct ?L")
   526   next
    80       proof - 
   527     case (Exit th)
    81         have "?L = wq_fun (schs s) cs1 @ [th]" using True
   528     interpret vt_exit: valid_trace_exit s e th  
    82           by (simp add:wq_def wf_def Let_def split:list.splits)
   529         using Exit by (unfold_locales, simp)
    83         moreover have "distinct ..."
   530     show ?thesis using Cons by (simp add: vt_exit.wq_distinct_kept) 
    84         proof -
   531   next
    85           have "th \<notin> set (wq_fun (schs s) cs1)"
   532     case (P th cs)
    86           proof
   533     interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp)
    87             assume otherwise: "th \<in> set (wq_fun (schs s) cs1)"
   534     show ?thesis using Cons by (simp add: vt_p.wq_distinct_kept) 
    88             from runing_head[OF thread_P(1) this]
   535   next
    89             have "th = hd (wq_fun (schs s) cs1)" .
   536     case (V th cs)
    90             hence "(Cs cs1, Th th) \<in> (RAG s)" using otherwise
   537     interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp)
    91               by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def)
   538     show ?thesis using Cons by (simp add: vt_v.wq_distinct_kept) 
    92             with thread_P(2) show False by auto
   539   next
    93           qed
   540     case (Set th prio)
    94           moreover have "distinct (wq_fun (schs s) cs1)"
   541     interpret vt_set: valid_trace_set s e th prio
    95               using True thread_P wq_def by auto 
   542         using Set by (unfold_locales, simp)
    96           ultimately show ?thesis by auto
   543     show ?thesis using Cons by (simp add: vt_set.wq_distinct_kept) 
    97         qed
   544   qed
    98         ultimately show ?thesis by simp
       
    99       qed
       
   100     next
       
   101       case False
       
   102       with thread_P(3)
       
   103       show ?thesis
       
   104         by (auto simp:wq_def wf_def Let_def split:list.splits)
       
   105     qed
       
   106   next
       
   107     case (thread_V th s cs1)
       
   108     thus ?case
       
   109     proof(cases "cs = cs1")
       
   110       case True
       
   111       show ?thesis (is "distinct ?L")
       
   112       proof(cases "(wq s cs)")
       
   113         case Nil
       
   114         thus ?thesis
       
   115           by (auto simp:wq_def wf_def Let_def split:list.splits)
       
   116       next
       
   117         case (Cons w_hd w_tl)
       
   118         moreover have "distinct (SOME q. distinct q \<and> set q = set w_tl)"
       
   119         proof(rule someI2)
       
   120           from thread_V(3)[unfolded Cons]
       
   121           show  "distinct w_tl \<and> set w_tl = set w_tl" by auto
       
   122         qed auto
       
   123         ultimately show ?thesis
       
   124           by (auto simp:wq_def wf_def Let_def True split:list.splits)
       
   125       qed 
       
   126     next
       
   127       case False
       
   128       with thread_V(3)
       
   129       show ?thesis
       
   130         by (auto simp:wq_def wf_def Let_def split:list.splits)
       
   131     qed
       
   132   qed (insert Cons, auto simp: wq_def Let_def split:list.splits)
       
   133 qed (unfold wq_def Let_def, simp)
   545 qed (unfold wq_def Let_def, simp)
   134 
   546 
   135 end
   547 end
   136 
       
   137 
   548 
   138 context valid_trace_e
   549 context valid_trace_e
   139 begin
   550 begin
   140 
   551 
   141 text {*
   552 text {*
   143   operation can add new thread into waiting queues. 
   554   operation can add new thread into waiting queues. 
   144   Such kind of lemmas are very obvious, but need to be checked formally.
   555   Such kind of lemmas are very obvious, but need to be checked formally.
   145   This is a kind of confirmation that our modelling is correct.
   556   This is a kind of confirmation that our modelling is correct.
   146 *}
   557 *}
   147 
   558 
   148 lemma block_pre: 
   559 lemma wq_in_inv: 
   149   assumes s_ni: "thread \<notin> set (wq s cs)"
   560   assumes s_ni: "thread \<notin> set (wq s cs)"
   150   and s_i: "thread \<in> set (wq (e#s) cs)"
   561   and s_i: "thread \<in> set (wq (e#s) cs)"
   151   shows "e = P thread cs"
   562   shows "e = P thread cs"
   152 proof(cases e)
   563 proof(cases e)
   153   -- {* This is the only non-trivial case: *}
   564   -- {* This is the only non-trivial case: *}
   173     qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
   584     qed (insert assms V True, auto simp:wq_def Let_def split:if_splits)
   174   qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
   585   qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
   175   thus ?thesis by auto
   586   thus ?thesis by auto
   176 qed (insert assms, auto simp:wq_def Let_def split:if_splits)
   587 qed (insert assms, auto simp:wq_def Let_def split:if_splits)
   177 
   588 
       
   589 lemma wq_out_inv: 
       
   590   assumes s_in: "thread \<in> set (wq s cs)"
       
   591   and s_hd: "thread = hd (wq s cs)"
       
   592   and s_i: "thread \<noteq> hd (wq (e#s) cs)"
       
   593   shows "e = V thread cs"
       
   594 proof(cases e)
       
   595 -- {* There are only two non-trivial cases: *}
       
   596   case (V th cs1)
       
   597   show ?thesis
       
   598   proof(cases "cs1 = cs")
       
   599     case True
       
   600     have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] .
       
   601     thus ?thesis
       
   602     proof(cases)
       
   603       case (thread_V)
       
   604       moreover have "th = thread" using thread_V(2) s_hd
       
   605           by (unfold s_holding_def wq_def, simp)
       
   606       ultimately show ?thesis using V True by simp
       
   607     qed
       
   608   qed (insert assms V, auto simp:wq_def Let_def split:if_splits)
       
   609 next
       
   610   case (P th cs1)
       
   611   show ?thesis
       
   612   proof(cases "cs1 = cs")
       
   613     case True
       
   614     with P have "wq (e#s) cs = wq_fun (schs s) cs @ [th]"
       
   615       by (auto simp:wq_def Let_def split:if_splits)
       
   616     with s_i s_hd s_in have False
       
   617       by (metis empty_iff hd_append2 list.set(1) wq_def) 
       
   618     thus ?thesis by simp
       
   619   qed (insert assms P, auto simp:wq_def Let_def split:if_splits)
       
   620 qed (insert assms, auto simp:wq_def Let_def split:if_splits)
       
   621 
   178 end
   622 end
   179 
   623 
   180 text {*
       
   181   The following lemmas is also obvious and shallow. It says
       
   182   that only running thread can request for a critical resource 
       
   183   and that the requested resource must be one which is
       
   184   not current held by the thread.
       
   185 *}
       
   186 
       
   187 lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
       
   188   thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (RAG s)^+"
       
   189 apply (ind_cases "vt ((P thread cs)#s)")
       
   190 apply (ind_cases "step s (P thread cs)")
       
   191 by auto
       
   192 
       
   193 lemma abs1:
       
   194   assumes ein: "e \<in> set es"
       
   195   and neq: "hd es \<noteq> hd (es @ [x])"
       
   196   shows "False"
       
   197 proof -
       
   198   from ein have "es \<noteq> []" by auto
       
   199   then obtain e ess where "es = e # ess" by (cases es, auto)
       
   200   with neq show ?thesis by auto
       
   201 qed
       
   202 
       
   203 lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]"
       
   204   by (cases es, auto)
       
   205 
       
   206 inductive_cases evt_cons: "vt (a#s)"
       
   207 
       
   208 context valid_trace_e
       
   209 begin
       
   210 
       
   211 lemma abs2:
       
   212   assumes inq: "thread \<in> set (wq s cs)"
       
   213   and nh: "thread = hd (wq s cs)"
       
   214   and qt: "thread \<noteq> hd (wq (e#s) cs)"
       
   215   and inq': "thread \<in> set (wq (e#s) cs)"
       
   216   shows "False"
       
   217 proof -
       
   218   from vt_e assms show "False"
       
   219     apply (cases e)
       
   220     apply ((simp split:if_splits add:Let_def wq_def)[1])+
       
   221     apply (insert abs1, fast)[1]
       
   222     apply (auto simp:wq_def simp:Let_def split:if_splits list.splits)
       
   223   proof -
       
   224     fix th qs
       
   225     assume vt: "vt (V th cs # s)"
       
   226       and th_in: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)"
       
   227       and eq_wq: "wq_fun (schs s) cs = thread # qs"
       
   228     show "False"
       
   229     proof -
       
   230       from wq_distinct[of cs]
       
   231         and eq_wq[folded wq_def] have "distinct (thread#qs)" by simp
       
   232       moreover have "thread \<in> set qs"
       
   233       proof -
       
   234         have "set (SOME q. distinct q \<and> set q = set qs) = set qs"
       
   235         proof(rule someI2)
       
   236           from wq_distinct [of cs]
       
   237           and eq_wq [folded wq_def]
       
   238           show "distinct qs \<and> set qs = set qs" by auto
       
   239         next
       
   240           fix x assume "distinct x \<and> set x = set qs"
       
   241           thus "set x = set qs" by auto
       
   242         qed
       
   243         with th_in show ?thesis by auto
       
   244       qed
       
   245       ultimately show ?thesis by auto
       
   246     qed
       
   247   qed
       
   248 qed
       
   249 
       
   250 end
       
   251 
       
   252 
       
   253 context valid_trace
       
   254 begin
       
   255 lemma  vt_moment: "\<And> t. vt (moment t s)"
       
   256 proof(induct rule:ind)
       
   257   case Nil
       
   258   thus ?case by (simp add:vt_nil)
       
   259 next
       
   260   case (Cons s e t)
       
   261   show ?case
       
   262   proof(cases "t \<ge> length (e#s)")
       
   263     case True
       
   264     from True have "moment t (e#s) = e#s" by simp
       
   265     thus ?thesis using Cons
       
   266       by (simp add:valid_trace_def)
       
   267   next
       
   268     case False
       
   269     from Cons have "vt (moment t s)" by simp
       
   270     moreover have "moment t (e#s) = moment t s"
       
   271     proof -
       
   272       from False have "t \<le> length s" by simp
       
   273       from moment_app [OF this, of "[e]"] 
       
   274       show ?thesis by simp
       
   275     qed
       
   276     ultimately show ?thesis by simp
       
   277   qed
       
   278 qed
       
   279 end
       
   280 
       
   281 locale valid_moment = valid_trace + 
       
   282   fixes i :: nat
       
   283 
       
   284 sublocale valid_moment < vat_moment: valid_trace "(moment i s)"
       
   285   by (unfold_locales, insert vt_moment, auto)
       
   286 
   624 
   287 context valid_trace
   625 context valid_trace
   288 begin
   626 begin
   289 
   627 
   290 
   628 
   319   However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
   657   However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still
   320   in blocked state at moment @{text "t2"} and could not
   658   in blocked state at moment @{text "t2"} and could not
   321   make any request and get blocked the second time: Contradiction.
   659   make any request and get blocked the second time: Contradiction.
   322 *}
   660 *}
   323 
   661 
   324 lemma waiting_unique_pre: (* ccc *)
   662 lemma waiting_unique_pre: (* ddd *)
   325   assumes h11: "thread \<in> set (wq s cs1)"
   663   assumes h11: "thread \<in> set (wq s cs1)"
   326   and h12: "thread \<noteq> hd (wq s cs1)"
   664   and h12: "thread \<noteq> hd (wq s cs1)"
   327   assumes h21: "thread \<in> set (wq s cs2)"
   665   assumes h21: "thread \<in> set (wq s cs2)"
   328   and h22: "thread \<noteq> hd (wq s cs2)"
   666   and h22: "thread \<noteq> hd (wq s cs2)"
   329   and neq12: "cs1 \<noteq> cs2"
   667   and neq12: "cs1 \<noteq> cs2"
   330   shows "False"
   668   shows "False"
   331 proof -
   669 proof -
   332   let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
   670   let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
   333   from h11 and h12 have q1: "?Q cs1 s" by simp
   671   from h11 and h12 have q1: "?Q cs1 s" by simp
   334   from h21 and h22 have q2: "?Q cs2 s" by simp
   672   from h21 and h22 have q2: "?Q cs2 s" by simp
   335   have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
   673   have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
   336   have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
   674   have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
   337   from p_split [of "?Q cs1", OF q1 nq1]
   675   from p_split [of "?Q cs1", OF q1 nq1]
   338   obtain t1 where lt1: "t1 < length s"
   676   obtain t1 where lt1: "t1 < length s"
   339     and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and>
   677     and np1: "\<not> ?Q cs1 (moment t1 s)"
   340         thread \<noteq> hd (wq (moment t1 s) cs1))"
   678     and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto
   341     and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and>
       
   342              thread \<noteq> hd (wq (moment i' s) cs1))" by auto
       
   343   from p_split [of "?Q cs2", OF q2 nq2]
   679   from p_split [of "?Q cs2", OF q2 nq2]
   344   obtain t2 where lt2: "t2 < length s"
   680   obtain t2 where lt2: "t2 < length s"
   345     and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and>
   681     and np2: "\<not> ?Q cs2 (moment t2 s)"
   346         thread \<noteq> hd (wq (moment t2 s) cs2))"
   682     and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto
   347     and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and>
   683   { fix s cs
   348              thread \<noteq> hd (wq (moment i' s) cs2))" by auto
   684     assume q: "?Q cs s"
       
   685     have "thread \<notin> runing s"
       
   686     proof
       
   687       assume "thread \<in> runing s"
       
   688       hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> 
       
   689                  thread \<noteq> hd (wq_fun (schs s) cs))"
       
   690         by (unfold runing_def s_waiting_def readys_def, auto)
       
   691       from this[rule_format, of cs] q 
       
   692       show False by (simp add: wq_def) 
       
   693     qed
       
   694   } note q_not_runing = this
       
   695   { fix t1 t2 cs1 cs2
       
   696     assume  lt1: "t1 < length s"
       
   697     and np1: "\<not> ?Q cs1 (moment t1 s)"
       
   698     and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))"
       
   699     and lt2: "t2 < length s"
       
   700     and np2: "\<not> ?Q cs2 (moment t2 s)"
       
   701     and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))"
       
   702     and lt12: "t1 < t2"
       
   703     let ?t3 = "Suc t2"
       
   704     from lt2 have le_t3: "?t3 \<le> length s" by auto
       
   705     from moment_plus [OF this] 
       
   706     obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
       
   707     have "t2 < ?t3" by simp
       
   708     from nn2 [rule_format, OF this] and eq_m
       
   709     have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   710          h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   711     have "vt (e#moment t2 s)"
       
   712     proof -
       
   713       from vt_moment 
       
   714       have "vt (moment ?t3 s)" .
       
   715       with eq_m show ?thesis by simp
       
   716     qed
       
   717     then interpret vt_e: valid_trace_e "moment t2 s" "e"
       
   718         by (unfold_locales, auto, cases, simp)
       
   719     have ?thesis
       
   720     proof -
       
   721       have "thread \<in> runing (moment t2 s)"
       
   722       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   723         case True
       
   724         have "e = V thread cs2"
       
   725         proof -
       
   726           have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
       
   727               using True and np2  by auto 
       
   728           from vt_e.wq_out_inv[OF True this h2]
       
   729           show ?thesis .
       
   730         qed
       
   731         thus ?thesis using vt_e.actor_inv[OF vt_e.pip_e] by auto
       
   732       next
       
   733         case False
       
   734         have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
       
   735         with vt_e.actor_inv[OF vt_e.pip_e]
       
   736         show ?thesis by auto
       
   737       qed
       
   738       moreover have "thread \<notin> runing (moment t2 s)"
       
   739         by (rule q_not_runing[OF nn1[rule_format, OF lt12]])
       
   740       ultimately show ?thesis by simp
       
   741     qed
       
   742   } note lt_case = this
   349   show ?thesis
   743   show ?thesis
   350   proof -
   744   proof -
   351     { 
   745     { assume "t1 < t2"
   352       assume lt12: "t1 < t2"
   746       from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this]
       
   747       have ?thesis .
       
   748     } moreover {
       
   749       assume "t2 < t1"
       
   750       from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this]
       
   751       have ?thesis .
       
   752     } moreover {
       
   753       assume eq_12: "t1 = t2"
   353       let ?t3 = "Suc t2"
   754       let ?t3 = "Suc t2"
   354       from lt2 have le_t3: "?t3 \<le> length s" by auto
   755       from lt2 have le_t3: "?t3 \<le> length s" by auto
   355       from moment_plus [OF this] 
   756       from moment_plus [OF this] 
   356       obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
   757       obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
   357       have "t2 < ?t3" by simp
   758       have lt_2: "t2 < ?t3" by simp
   358       from nn2 [rule_format, OF this] and eq_m
   759       from nn2 [rule_format, OF this] and eq_m
   359       have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
   760       have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
   360         h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
   761            h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   762       from nn1[rule_format, OF lt_2[folded eq_12]] eq_m[folded eq_12]
       
   763       have g1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   764            g2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
   361       have "vt (e#moment t2 s)"
   765       have "vt (e#moment t2 s)"
   362       proof -
   766       proof -
   363         from vt_moment 
   767         from vt_moment 
   364         have "vt (moment ?t3 s)" .
   768         have "vt (moment ?t3 s)" .
   365         with eq_m show ?thesis by simp
   769         with eq_m show ?thesis by simp
   366       qed
   770       qed
   367       then interpret vt_e: valid_trace_e "moment t2 s" "e"
   771       then interpret vt_e: valid_trace_e "moment t2 s" "e"
   368         by (unfold_locales, auto, cases, simp)
   772           by (unfold_locales, auto, cases, simp)
   369       have ?thesis
   773       have "e = V thread cs2 \<or> e = P thread cs2"
   370       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
   774       proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
   371         case True
   775         case True
   372         from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
   776         have "e = V thread cs2"
   373           by auto 
   777         proof -
   374         from vt_e.abs2 [OF True eq_th h2 h1]
   778           have eq_th: "thread = hd (wq (moment t2 s) cs2)" 
   375         show ?thesis by auto
   779               using True and np2  by auto 
       
   780           from vt_e.wq_out_inv[OF True this h2]
       
   781           show ?thesis .
       
   782         qed
       
   783         thus ?thesis by auto
   376       next
   784       next
   377         case False
   785         case False
   378         from vt_e.block_pre[OF False h1]
   786         have "e = P thread cs2" using vt_e.wq_in_inv[OF False h1] .
   379         have "e = P thread cs2" .
   787         thus ?thesis by auto
   380         with vt_e.vt_e have "vt ((P thread cs2)# moment t2 s)" by simp
       
   381         from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp
       
   382         with runing_ready have "thread \<in> readys (moment t2 s)" by auto
       
   383         with nn1 [rule_format, OF lt12]
       
   384         show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
       
   385       qed
   788       qed
   386     } moreover {
   789       moreover have "e = V thread cs1 \<or> e = P thread cs1"
   387       assume lt12: "t2 < t1"
       
   388       let ?t3 = "Suc t1"
       
   389       from lt1 have le_t3: "?t3 \<le> length s" by auto
       
   390       from moment_plus [OF this] 
       
   391       obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
       
   392       have lt_t3: "t1 < ?t3" by simp
       
   393       from nn1 [rule_format, OF this] and eq_m
       
   394       have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   395         h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
       
   396       have "vt  (e#moment t1 s)"
       
   397       proof -
       
   398         from vt_moment
       
   399         have "vt (moment ?t3 s)" .
       
   400         with eq_m show ?thesis by simp
       
   401       qed
       
   402       then interpret vt_e: valid_trace_e "moment t1 s" e
       
   403         by (unfold_locales, auto, cases, auto)
       
   404       have ?thesis
       
   405       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
   790       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
   406         case True
   791         case True
   407         from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
   792         have eq_th: "thread = hd (wq (moment t1 s) cs1)" 
   408           by auto
   793               using True and np1  by auto 
   409         from vt_e.abs2 True eq_th h2 h1
   794         from vt_e.wq_out_inv[folded eq_12, OF True this g2]
   410         show ?thesis by auto
   795         have "e = V thread cs1" .
       
   796         thus ?thesis by auto
   411       next
   797       next
   412         case False
   798         case False
   413         from vt_e.block_pre [OF False h1]
   799         have "e = P thread cs1" using vt_e.wq_in_inv[folded eq_12, OF False g1] .
   414         have "e = P thread cs1" .
   800         thus ?thesis by auto
   415         with vt_e.vt_e have "vt ((P thread cs1)# moment t1 s)" by simp
       
   416         from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp
       
   417         with runing_ready have "thread \<in> readys (moment t1 s)" by auto
       
   418         with nn2 [rule_format, OF lt12]
       
   419         show ?thesis  by (simp add:readys_def wq_def s_waiting_def, auto)
       
   420       qed
   801       qed
   421     } moreover {
   802       ultimately have ?thesis using neq12 by auto
   422       assume eqt12: "t1 = t2"
   803     } ultimately show ?thesis using nat_neq_iff by blast 
   423       let ?t3 = "Suc t1"
       
   424       from lt1 have le_t3: "?t3 \<le> length s" by auto
       
   425       from moment_plus [OF this] 
       
   426       obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
       
   427       have lt_t3: "t1 < ?t3" by simp
       
   428       from nn1 [rule_format, OF this] and eq_m
       
   429       have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
       
   430         h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
       
   431       have vt_e: "vt (e#moment t1 s)"
       
   432       proof -
       
   433         from vt_moment
       
   434         have "vt (moment ?t3 s)" .
       
   435         with eq_m show ?thesis by simp
       
   436       qed
       
   437       then interpret vt_e: valid_trace_e "moment t1 s" e
       
   438         by (unfold_locales, auto, cases, auto)
       
   439       have ?thesis
       
   440       proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
       
   441         case True
       
   442         from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
       
   443           by auto
       
   444         from vt_e.abs2 [OF True eq_th h2 h1]
       
   445         show ?thesis by auto
       
   446       next
       
   447         case False
       
   448         from vt_e.block_pre [OF False h1]
       
   449         have eq_e1: "e = P thread cs1" .
       
   450         have lt_t3: "t1 < ?t3" by simp
       
   451         with eqt12 have "t2 < ?t3" by simp
       
   452         from nn2 [rule_format, OF this] and eq_m and eqt12
       
   453         have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
       
   454           h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
       
   455         show ?thesis
       
   456         proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
       
   457           case True
       
   458           from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
       
   459             by auto
       
   460           from vt_e and eqt12 have "vt (e#moment t2 s)" by simp 
       
   461           then interpret vt_e2: valid_trace_e "moment t2 s" e
       
   462             by (unfold_locales, auto, cases, auto)
       
   463           from vt_e2.abs2 [OF True eq_th h2 h1]
       
   464           show ?thesis .
       
   465         next
       
   466           case False
       
   467           have "vt (e#moment t2 s)"
       
   468           proof -
       
   469             from vt_moment eqt12
       
   470             have "vt (moment (Suc t2) s)" by auto
       
   471             with eq_m eqt12 show ?thesis by simp
       
   472           qed
       
   473           then interpret vt_e2: valid_trace_e "moment t2 s" e
       
   474             by (unfold_locales, auto, cases, auto)
       
   475           from vt_e2.block_pre [OF False h1]
       
   476           have "e = P thread cs2" .
       
   477           with eq_e1 neq12 show ?thesis by auto
       
   478         qed
       
   479       qed
       
   480     } ultimately show ?thesis by arith
       
   481   qed
   804   qed
   482 qed
   805 qed
   483 
   806 
   484 text {*
   807 text {*
   485   This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
   808   This lemma is a simple corrolary of @{text "waiting_unique_pre"}.
   487 
   810 
   488 lemma waiting_unique:
   811 lemma waiting_unique:
   489   assumes "waiting s th cs1"
   812   assumes "waiting s th cs1"
   490   and "waiting s th cs2"
   813   and "waiting s th cs2"
   491   shows "cs1 = cs2"
   814   shows "cs1 = cs2"
   492 using waiting_unique_pre assms
   815   using waiting_unique_pre assms
   493 unfolding wq_def s_waiting_def
   816   unfolding wq_def s_waiting_def
   494 by auto
   817   by auto
   495 
   818 
   496 end
   819 end
   497 
   820 
   498 (* not used *)
   821 (* not used *)
   499 text {*
   822 text {*
   505   assumes "holding (s::event list) th1 cs"
   828   assumes "holding (s::event list) th1 cs"
   506   and "holding s th2 cs"
   829   and "holding s th2 cs"
   507   shows "th1 = th2"
   830   shows "th1 = th2"
   508  by (insert assms, unfold s_holding_def, auto)
   831  by (insert assms, unfold s_holding_def, auto)
   509 
   832 
   510 
       
   511 lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
   833 lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s"
   512   apply (induct s, auto)
   834   apply (induct s, auto)
   513   by (case_tac a, auto split:if_splits)
   835   by (case_tac a, auto split:if_splits)
   514 
   836 
   515 lemma last_set_unique: 
   837 lemma last_set_unique: 
   526 proof -
   848 proof -
   527   from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
   849   from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def)
   528   from last_set_unique [OF this th_in1 th_in2]
   850   from last_set_unique [OF this th_in1 th_in2]
   529   show ?thesis .
   851   show ?thesis .
   530 qed
   852 qed
   531 
   853                       
   532 lemma preced_linorder: 
   854 lemma preced_linorder: 
   533   assumes neq_12: "th1 \<noteq> th2"
   855   assumes neq_12: "th1 \<noteq> th2"
   534   and th_in1: "th1 \<in> threads s"
   856   and th_in1: "th1 \<in> threads s"
   535   and th_in2: " th2 \<in> threads s"
   857   and th_in2: " th2 \<in> threads s"
   536   shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
   858   shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
   538   from preced_unique [OF _ th_in1 th_in2] and neq_12 
   860   from preced_unique [OF _ th_in1 th_in2] and neq_12 
   539   have "preced th1 s \<noteq> preced th2 s" by auto
   861   have "preced th1 s \<noteq> preced th2 s" by auto
   540   thus ?thesis by auto
   862   thus ?thesis by auto
   541 qed
   863 qed
   542 
   864 
   543 (* An aux lemma used later *)
       
   544 lemma unique_minus:
       
   545   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   546   and xy: "(x, y) \<in> r"
       
   547   and xz: "(x, z) \<in> r^+"
       
   548   and neq: "y \<noteq> z"
       
   549   shows "(y, z) \<in> r^+"
       
   550 proof -
       
   551  from xz and neq show ?thesis
       
   552  proof(induct)
       
   553    case (base ya)
       
   554    have "(x, ya) \<in> r" by fact
       
   555    from unique [OF xy this] have "y = ya" .
       
   556    with base show ?case by auto
       
   557  next
       
   558    case (step ya z)
       
   559    show ?case
       
   560    proof(cases "y = ya")
       
   561      case True
       
   562      from step True show ?thesis by simp
       
   563    next
       
   564      case False
       
   565      from step False
       
   566      show ?thesis by auto
       
   567    qed
       
   568  qed
       
   569 qed
       
   570 
       
   571 lemma unique_base:
       
   572   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   573   and xy: "(x, y) \<in> r"
       
   574   and xz: "(x, z) \<in> r^+"
       
   575   and neq_yz: "y \<noteq> z"
       
   576   shows "(y, z) \<in> r^+"
       
   577 proof -
       
   578   from xz neq_yz show ?thesis
       
   579   proof(induct)
       
   580     case (base ya)
       
   581     from xy unique base show ?case by auto
       
   582   next
       
   583     case (step ya z)
       
   584     show ?case
       
   585     proof(cases "y = ya")
       
   586       case True
       
   587       from True step show ?thesis by auto
       
   588     next
       
   589       case False
       
   590       from False step 
       
   591       have "(y, ya) \<in> r\<^sup>+" by auto
       
   592       with step show ?thesis by auto
       
   593     qed
       
   594   qed
       
   595 qed
       
   596 
       
   597 lemma unique_chain:
       
   598   assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
       
   599   and xy: "(x, y) \<in> r^+"
       
   600   and xz: "(x, z) \<in> r^+"
       
   601   and neq_yz: "y \<noteq> z"
       
   602   shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
       
   603 proof -
       
   604   from xy xz neq_yz show ?thesis
       
   605   proof(induct)
       
   606     case (base y)
       
   607     have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
       
   608     from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
       
   609   next
       
   610     case (step y za)
       
   611     show ?case
       
   612     proof(cases "y = z")
       
   613       case True
       
   614       from True step show ?thesis by auto
       
   615     next
       
   616       case False
       
   617       from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
       
   618       thus ?thesis
       
   619       proof
       
   620         assume "(z, y) \<in> r\<^sup>+"
       
   621         with step have "(z, za) \<in> r\<^sup>+" by auto
       
   622         thus ?thesis by auto
       
   623       next
       
   624         assume h: "(y, z) \<in> r\<^sup>+"
       
   625         from step have yza: "(y, za) \<in> r" by simp
       
   626         from step have "za \<noteq> z" by simp
       
   627         from unique_minus [OF _ yza h this] and unique
       
   628         have "(za, z) \<in> r\<^sup>+" by auto
       
   629         thus ?thesis by auto
       
   630       qed
       
   631     qed
       
   632   qed
       
   633 qed
       
   634 
       
   635 text {*
   865 text {*
   636   The following three lemmas show that @{text "RAG"} does not change
   866   The following three lemmas show that @{text "RAG"} does not change
   637   by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
   867   by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"}
   638   events, respectively.
   868   events, respectively.
   639 *}
   869 *}
   640 
   870 
   641 lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
   871 lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s"
   642 apply (unfold s_RAG_def s_waiting_def wq_def)
   872 apply (unfold s_RAG_def s_waiting_def wq_def)
   643 by (simp add:Let_def)
   873 by (simp add:Let_def)
   644 
   874 
       
   875 lemma (in valid_trace_set)
       
   876    RAG_unchanged: "(RAG (e # s)) = RAG s"
       
   877    by (unfold is_set RAG_set_unchanged, simp)
       
   878 
   645 lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
   879 lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s"
   646 apply (unfold s_RAG_def s_waiting_def wq_def)
   880 apply (unfold s_RAG_def s_waiting_def wq_def)
   647 by (simp add:Let_def)
   881 by (simp add:Let_def)
   648 
   882 
       
   883 lemma (in valid_trace_create)
       
   884    RAG_unchanged: "(RAG (e # s)) = RAG s"
       
   885    by (unfold is_create RAG_create_unchanged, simp)
       
   886 
   649 lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
   887 lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s"
   650 apply (unfold s_RAG_def s_waiting_def wq_def)
   888 apply (unfold s_RAG_def s_waiting_def wq_def)
   651 by (simp add:Let_def)
   889 by (simp add:Let_def)
   652 
   890 
   653 
   891 lemma (in valid_trace_exit)
   654 text {* 
   892    RAG_unchanged: "(RAG (e # s)) = RAG s"
   655   The following lemmas are used in the proof of 
   893    by (unfold is_exit RAG_exit_unchanged, simp)
   656   lemma @{text "step_RAG_v"}, which characterizes how the @{text "RAG"} is changed
   894 
   657   by @{text "V"}-events. 
   895 context valid_trace_v
   658   However, since our model is very concise, such  seemingly obvious lemmas need to be derived from scratch,
   896 begin
   659   starting from the model definitions.
   897 
   660 *}
   898 lemma distinct_rest: "distinct rest"
   661 lemma step_v_hold_inv[elim_format]:
   899   by (simp add: distinct_tl rest_def wq_distinct)
   662   "\<And>c t. \<lbrakk>vt (V th cs # s); 
   900 
   663           \<not> holding (wq s) t c; holding (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> 
   901 lemma holding_cs_eq_th:
   664             next_th s th cs t \<and> c = cs"
   902   assumes "holding s t cs"
   665 proof -
   903   shows "t = th"
   666   fix c t
   904 proof -
   667   assume vt: "vt (V th cs # s)"
   905   from pip_e[unfolded is_v]
   668     and nhd: "\<not> holding (wq s) t c"
   906   show ?thesis
   669     and hd: "holding (wq (V th cs # s)) t c"
   907   proof(cases)
   670   show "next_th s th cs t \<and> c = cs"
   908     case (thread_V)
       
   909     from held_unique[OF this(2) assms]
       
   910     show ?thesis by simp
       
   911   qed
       
   912 qed
       
   913 
       
   914 lemma distinct_wq': "distinct wq'"
       
   915   by (metis (mono_tags, lifting) distinct_rest  some_eq_ex wq'_def)
       
   916   
       
   917 lemma set_wq': "set wq' = set rest"
       
   918   by (metis (mono_tags, lifting) distinct_rest rest_def 
       
   919       some_eq_ex wq'_def)
       
   920     
       
   921 lemma th'_in_inv:
       
   922   assumes "th' \<in> set wq'"
       
   923   shows "th' \<in> set rest"
       
   924   using assms set_wq' by simp
       
   925 
       
   926 lemma neq_t_th: 
       
   927   assumes "waiting (e#s) t c"
       
   928   shows "t \<noteq> th"
       
   929 proof
       
   930   assume otherwise: "t = th"
       
   931   show False
   671   proof(cases "c = cs")
   932   proof(cases "c = cs")
       
   933     case True
       
   934     have "t \<in> set wq'" 
       
   935      using assms[unfolded True s_waiting_def, folded wq_def, unfolded wq_es_cs]
       
   936      by simp 
       
   937     from th'_in_inv[OF this] have "t \<in> set rest" .
       
   938     with wq_s_cs[folded otherwise] wq_distinct[of cs]
       
   939     show ?thesis by simp
       
   940   next
   672     case False
   941     case False
   673     with nhd hd show ?thesis
   942     have "wq (e#s) c = wq s c" using False
   674       by (unfold cs_holding_def wq_def, auto simp:Let_def)
   943         by (unfold is_v, simp)
   675   next
   944     hence "waiting s t c" using assms 
   676     case True
   945         by (simp add: cs_waiting_def waiting_eq)
   677     with step_back_step [OF vt] 
   946     hence "t \<notin> readys s" by (unfold readys_def, auto)
   678     have "step s (V th c)" by simp
   947     hence "t \<notin> runing s" using runing_ready by auto 
   679     hence "next_th s th cs t"
   948     with runing_th_s[folded otherwise] show ?thesis by auto
   680     proof(cases)
   949   qed
   681       assume "holding s th c"
   950 qed
   682       with nhd hd show ?thesis
   951 
   683         apply (unfold s_holding_def cs_holding_def wq_def next_th_def,
   952 lemma waiting_esI1:
   684                auto simp:Let_def split:list.splits if_splits)
   953   assumes "waiting s t c"
   685         proof -
   954       and "c \<noteq> cs" 
   686           assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
   955   shows "waiting (e#s) t c" 
   687           moreover have "\<dots> = set []"
   956 proof -
   688           proof(rule someI2)
   957   have "wq (e#s) c = wq s c" 
   689             show "distinct [] \<and> [] = []" by auto
   958     using assms(2) is_v by auto
   690           next
   959   with assms(1) show ?thesis 
   691             fix x assume "distinct x \<and> x = []"
   960     using cs_waiting_def waiting_eq by auto 
   692             thus "set x = set []" by auto
   961 qed
   693           qed
   962 
   694           ultimately show False by auto
   963 lemma holding_esI2:
       
   964   assumes "c \<noteq> cs" 
       
   965   and "holding s t c"
       
   966   shows "holding (e#s) t c"
       
   967 proof -
       
   968   from assms(1) have "wq (e#s) c = wq s c" using is_v by auto
       
   969   from assms(2)[unfolded s_holding_def, folded wq_def, 
       
   970                 folded this, unfolded wq_def, folded s_holding_def]
       
   971   show ?thesis .
       
   972 qed
       
   973 
       
   974 lemma holding_esI1:
       
   975   assumes "holding s t c"
       
   976   and "t \<noteq> th"
       
   977   shows "holding (e#s) t c"
       
   978 proof -
       
   979   have "c \<noteq> cs" using assms using holding_cs_eq_th by blast 
       
   980   from holding_esI2[OF this assms(1)]
       
   981   show ?thesis .
       
   982 qed
       
   983 
       
   984 end
       
   985 
       
   986 context valid_trace_v_n
       
   987 begin
       
   988 
       
   989 lemma neq_wq': "wq' \<noteq> []" 
       
   990 proof (unfold wq'_def, rule someI2)
       
   991   show "distinct rest \<and> set rest = set rest"
       
   992     by (simp add: distinct_rest) 
       
   993 next
       
   994   fix x
       
   995   assume " distinct x \<and> set x = set rest" 
       
   996   thus "x \<noteq> []" using rest_nnl by auto
       
   997 qed 
       
   998 
       
   999 definition "taker = hd wq'"
       
  1000 
       
  1001 definition "rest' = tl wq'"
       
  1002 
       
  1003 lemma eq_wq': "wq' = taker # rest'"
       
  1004   by (simp add: neq_wq' rest'_def taker_def)
       
  1005 
       
  1006 lemma next_th_taker: 
       
  1007   shows "next_th s th cs taker"
       
  1008   using rest_nnl taker_def wq'_def wq_s_cs 
       
  1009   by (auto simp:next_th_def)
       
  1010 
       
  1011 lemma taker_unique: 
       
  1012   assumes "next_th s th cs taker'"
       
  1013   shows "taker' = taker"
       
  1014 proof -
       
  1015   from assms
       
  1016   obtain rest' where 
       
  1017     h: "wq s cs = th # rest'" 
       
  1018        "taker' = hd (SOME q. distinct q \<and> set q = set rest')"
       
  1019           by (unfold next_th_def, auto)
       
  1020   with wq_s_cs have "rest' = rest" by auto
       
  1021   thus ?thesis using h(2) taker_def wq'_def by auto 
       
  1022 qed
       
  1023 
       
  1024 lemma waiting_set_eq:
       
  1025   "{(Th th', Cs cs) |th'. next_th s th cs th'} = {(Th taker, Cs cs)}"
       
  1026   by (smt all_not_in_conv bot.extremum insertI1 insert_subset 
       
  1027       mem_Collect_eq next_th_taker subsetI subset_antisym taker_def taker_unique)
       
  1028 
       
  1029 lemma holding_set_eq:
       
  1030   "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {(Cs cs, Th taker)}"
       
  1031   using next_th_taker taker_def waiting_set_eq 
       
  1032   by fastforce
       
  1033    
       
  1034 lemma holding_taker:
       
  1035   shows "holding (e#s) taker cs"
       
  1036     by (unfold s_holding_def, fold wq_def, unfold wq_es_cs, 
       
  1037         auto simp:neq_wq' taker_def)
       
  1038 
       
  1039 lemma waiting_esI2:
       
  1040   assumes "waiting s t cs"
       
  1041       and "t \<noteq> taker"
       
  1042   shows "waiting (e#s) t cs" 
       
  1043 proof -
       
  1044   have "t \<in> set wq'" 
       
  1045   proof(unfold wq'_def, rule someI2)
       
  1046     show "distinct rest \<and> set rest = set rest"
       
  1047           by (simp add: distinct_rest)
       
  1048   next
       
  1049     fix x
       
  1050     assume "distinct x \<and> set x = set rest"
       
  1051     moreover have "t \<in> set rest"
       
  1052         using assms(1) cs_waiting_def waiting_eq wq_s_cs by auto 
       
  1053     ultimately show "t \<in> set x" by simp
       
  1054   qed
       
  1055   moreover have "t \<noteq> hd wq'"
       
  1056     using assms(2) taker_def by auto 
       
  1057   ultimately show ?thesis
       
  1058     by (unfold s_waiting_def, fold wq_def, unfold wq_es_cs, simp)
       
  1059 qed
       
  1060 
       
  1061 lemma waiting_esE:
       
  1062   assumes "waiting (e#s) t c" 
       
  1063   obtains "c \<noteq> cs" "waiting s t c"
       
  1064      |    "c = cs" "t \<noteq> taker" "waiting s t cs" "t \<in> set rest'"
       
  1065 proof(cases "c = cs")
       
  1066   case False
       
  1067   hence "wq (e#s) c = wq s c" using is_v by auto
       
  1068   with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
       
  1069   from that(1)[OF False this] show ?thesis .
       
  1070 next
       
  1071   case True
       
  1072   from assms[unfolded s_waiting_def True, folded wq_def, unfolded wq_es_cs]
       
  1073   have "t \<noteq> hd wq'" "t \<in> set wq'" by auto
       
  1074   hence "t \<noteq> taker" by (simp add: taker_def) 
       
  1075   moreover hence "t \<noteq> th" using assms neq_t_th by blast 
       
  1076   moreover have "t \<in> set rest" by (simp add: `t \<in> set wq'` th'_in_inv) 
       
  1077   ultimately have "waiting s t cs"
       
  1078     by (metis cs_waiting_def list.distinct(2) list.sel(1) 
       
  1079                 list.set_sel(2) rest_def waiting_eq wq_s_cs)  
       
  1080   show ?thesis using that(2)
       
  1081   using True `t \<in> set wq'` `t \<noteq> taker` `waiting s t cs` eq_wq' by auto   
       
  1082 qed
       
  1083 
       
  1084 lemma holding_esI1:
       
  1085   assumes "c = cs"
       
  1086   and "t = taker"
       
  1087   shows "holding (e#s) t c"
       
  1088   by (unfold assms, simp add: holding_taker)
       
  1089 
       
  1090 lemma holding_esE:
       
  1091   assumes "holding (e#s) t c" 
       
  1092   obtains "c = cs" "t = taker"
       
  1093       | "c \<noteq> cs" "holding s t c"
       
  1094 proof(cases "c = cs")
       
  1095   case True
       
  1096   from assms[unfolded True, unfolded s_holding_def, 
       
  1097              folded wq_def, unfolded wq_es_cs]
       
  1098   have "t = taker" by (simp add: taker_def) 
       
  1099   from that(1)[OF True this] show ?thesis .
       
  1100 next
       
  1101   case False
       
  1102   hence "wq (e#s) c = wq s c" using is_v by auto
       
  1103   from assms[unfolded s_holding_def, folded wq_def, 
       
  1104              unfolded this, unfolded wq_def, folded s_holding_def]
       
  1105   have "holding s t c"  .
       
  1106   from that(2)[OF False this] show ?thesis .
       
  1107 qed
       
  1108 
       
  1109 end 
       
  1110 
       
  1111 
       
  1112 context valid_trace_v_e
       
  1113 begin
       
  1114 
       
  1115 lemma nil_wq': "wq' = []" 
       
  1116 proof (unfold wq'_def, rule someI2)
       
  1117   show "distinct rest \<and> set rest = set rest"
       
  1118     by (simp add: distinct_rest) 
       
  1119 next
       
  1120   fix x
       
  1121   assume " distinct x \<and> set x = set rest" 
       
  1122   thus "x = []" using rest_nil by auto
       
  1123 qed 
       
  1124 
       
  1125 lemma no_taker: 
       
  1126   assumes "next_th s th cs taker"
       
  1127   shows "False"
       
  1128 proof -
       
  1129   from assms[unfolded next_th_def]
       
  1130   obtain rest' where "wq s cs = th # rest'" "rest' \<noteq> []"
       
  1131     by auto
       
  1132   thus ?thesis using rest_def rest_nil by auto 
       
  1133 qed
       
  1134 
       
  1135 lemma waiting_set_eq:
       
  1136   "{(Th th', Cs cs) |th'. next_th s th cs th'} = {}"
       
  1137   using no_taker by auto
       
  1138 
       
  1139 lemma holding_set_eq:
       
  1140   "{(Cs cs, Th th') |th'.  next_th s th cs th'} = {}"
       
  1141   using no_taker by auto
       
  1142    
       
  1143 lemma no_holding:
       
  1144   assumes "holding (e#s) taker cs"
       
  1145   shows False
       
  1146 proof -
       
  1147   from wq_es_cs[unfolded nil_wq']
       
  1148   have " wq (e # s) cs = []" .
       
  1149   from assms[unfolded s_holding_def, folded wq_def, unfolded this]
       
  1150   show ?thesis by auto
       
  1151 qed
       
  1152 
       
  1153 lemma no_waiting:
       
  1154   assumes "waiting (e#s) t cs"
       
  1155   shows False
       
  1156 proof -
       
  1157   from wq_es_cs[unfolded nil_wq']
       
  1158   have " wq (e # s) cs = []" .
       
  1159   from assms[unfolded s_waiting_def, folded wq_def, unfolded this]
       
  1160   show ?thesis by auto
       
  1161 qed
       
  1162 
       
  1163 lemma waiting_esI2:
       
  1164   assumes "waiting s t c"
       
  1165   shows "waiting (e#s) t c"
       
  1166 proof -
       
  1167   have "c \<noteq> cs" using assms
       
  1168     using cs_waiting_def rest_nil waiting_eq wq_s_cs by auto 
       
  1169   from waiting_esI1[OF assms this]
       
  1170   show ?thesis .
       
  1171 qed
       
  1172 
       
  1173 lemma waiting_esE:
       
  1174   assumes "waiting (e#s) t c" 
       
  1175   obtains "c \<noteq> cs" "waiting s t c"
       
  1176 proof(cases "c = cs")
       
  1177   case False
       
  1178   hence "wq (e#s) c = wq s c" using is_v by auto
       
  1179   with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto 
       
  1180   from that(1)[OF False this] show ?thesis .
       
  1181 next
       
  1182   case True
       
  1183   from no_waiting[OF assms[unfolded True]]
       
  1184   show ?thesis by auto
       
  1185 qed
       
  1186 
       
  1187 lemma holding_esE:
       
  1188   assumes "holding (e#s) t c" 
       
  1189   obtains "c \<noteq> cs" "holding s t c"
       
  1190 proof(cases "c = cs")
       
  1191   case True
       
  1192   from no_holding[OF assms[unfolded True]] 
       
  1193   show ?thesis by auto
       
  1194 next
       
  1195   case False
       
  1196   hence "wq (e#s) c = wq s c" using is_v by auto
       
  1197   from assms[unfolded s_holding_def, folded wq_def, 
       
  1198              unfolded this, unfolded wq_def, folded s_holding_def]
       
  1199   have "holding s t c"  .
       
  1200   from that[OF False this] show ?thesis .
       
  1201 qed
       
  1202 
       
  1203 end 
       
  1204 
       
  1205 lemma rel_eqI:
       
  1206   assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B"
       
  1207   and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A"
       
  1208   shows "A = B"
       
  1209   using assms by auto
       
  1210 
       
  1211 lemma in_RAG_E:
       
  1212   assumes "(n1, n2) \<in> RAG (s::state)"
       
  1213   obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs"
       
  1214       | (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs"
       
  1215   using assms[unfolded s_RAG_def, folded waiting_eq holding_eq]
       
  1216   by auto
       
  1217   
       
  1218 context valid_trace_v
       
  1219 begin
       
  1220 
       
  1221 lemma RAG_es:
       
  1222   "RAG (e # s) =
       
  1223    RAG s - {(Cs cs, Th th)} -
       
  1224      {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
       
  1225      {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
       
  1226 proof(rule rel_eqI)
       
  1227   fix n1 n2
       
  1228   assume "(n1, n2) \<in> ?L"
       
  1229   thus "(n1, n2) \<in> ?R"
       
  1230   proof(cases rule:in_RAG_E)
       
  1231     case (waiting th' cs')
       
  1232     show ?thesis
       
  1233     proof(cases "rest = []")
       
  1234       case False
       
  1235       interpret h_n: valid_trace_v_n s e th cs
       
  1236         by (unfold_locales, insert False, simp)
       
  1237       from waiting(3)
       
  1238       show ?thesis
       
  1239       proof(cases rule:h_n.waiting_esE)
       
  1240         case 1
       
  1241         with waiting(1,2)
       
  1242         show ?thesis
       
  1243         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
       
  1244              fold waiting_eq, auto)
       
  1245       next
       
  1246         case 2
       
  1247         with waiting(1,2)
       
  1248         show ?thesis
       
  1249          by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
       
  1250              fold waiting_eq, auto)
       
  1251       qed
       
  1252     next
       
  1253       case True
       
  1254       interpret h_e: valid_trace_v_e s e th cs
       
  1255         by (unfold_locales, insert True, simp)
       
  1256       from waiting(3)
       
  1257       show ?thesis
       
  1258       proof(cases rule:h_e.waiting_esE)
       
  1259         case 1
       
  1260         with waiting(1,2)
       
  1261         show ?thesis
       
  1262         by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
       
  1263              fold waiting_eq, auto)
       
  1264       qed
       
  1265     qed
       
  1266   next
       
  1267     case (holding th' cs')
       
  1268     show ?thesis
       
  1269     proof(cases "rest = []")
       
  1270       case False
       
  1271       interpret h_n: valid_trace_v_n s e th cs
       
  1272         by (unfold_locales, insert False, simp)
       
  1273       from holding(3)
       
  1274       show ?thesis
       
  1275       proof(cases rule:h_n.holding_esE)
       
  1276         case 1
       
  1277         with holding(1,2)
       
  1278         show ?thesis
       
  1279         by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
       
  1280              fold waiting_eq, auto)
       
  1281       next
       
  1282         case 2
       
  1283         with holding(1,2)
       
  1284         show ?thesis
       
  1285          by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, 
       
  1286              fold holding_eq, auto)
       
  1287       qed
       
  1288     next
       
  1289       case True
       
  1290       interpret h_e: valid_trace_v_e s e th cs
       
  1291         by (unfold_locales, insert True, simp)
       
  1292       from holding(3)
       
  1293       show ?thesis
       
  1294       proof(cases rule:h_e.holding_esE)
       
  1295         case 1
       
  1296         with holding(1,2)
       
  1297         show ?thesis
       
  1298         by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, 
       
  1299              fold holding_eq, auto)
       
  1300       qed
       
  1301     qed
       
  1302   qed
       
  1303 next
       
  1304   fix n1 n2
       
  1305   assume h: "(n1, n2) \<in> ?R"
       
  1306   show "(n1, n2) \<in> ?L"
       
  1307   proof(cases "rest = []")
       
  1308     case False
       
  1309     interpret h_n: valid_trace_v_n s e th cs
       
  1310         by (unfold_locales, insert False, simp)
       
  1311     from h[unfolded h_n.waiting_set_eq h_n.holding_set_eq]
       
  1312     have "((n1, n2) \<in> RAG s \<and> (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th)
       
  1313                             \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)) \<or> 
       
  1314           (n2 = Th h_n.taker \<and> n1 = Cs cs)" 
       
  1315       by auto
       
  1316    thus ?thesis
       
  1317    proof
       
  1318       assume "n2 = Th h_n.taker \<and> n1 = Cs cs"
       
  1319       with h_n.holding_taker
       
  1320       show ?thesis 
       
  1321         by (unfold s_RAG_def, fold holding_eq, auto)
       
  1322    next
       
  1323     assume h: "(n1, n2) \<in> RAG s \<and>
       
  1324         (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)"
       
  1325     hence "(n1, n2) \<in> RAG s" by simp
       
  1326     thus ?thesis
       
  1327     proof(cases rule:in_RAG_E)
       
  1328       case (waiting th' cs')
       
  1329       from h and this(1,2)
       
  1330       have "th' \<noteq> h_n.taker \<or> cs' \<noteq> cs" by auto
       
  1331       hence "waiting (e#s) th' cs'" 
       
  1332       proof
       
  1333         assume "cs' \<noteq> cs"
       
  1334         from waiting_esI1[OF waiting(3) this] 
       
  1335         show ?thesis .
       
  1336       next
       
  1337         assume neq_th': "th' \<noteq> h_n.taker"
       
  1338         show ?thesis
       
  1339         proof(cases "cs' = cs")
       
  1340           case False
       
  1341           from waiting_esI1[OF waiting(3) this] 
       
  1342           show ?thesis .
   695         next
  1343         next
   696           assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])"
  1344           case True
   697           moreover have "\<dots> = set []"
  1345           from h_n.waiting_esI2[OF waiting(3)[unfolded True] neq_th', folded True]
   698           proof(rule someI2)
  1346           show ?thesis .
   699             show "distinct [] \<and> [] = []" by auto
       
   700           next
       
   701             fix x assume "distinct x \<and> x = []"
       
   702             thus "set x = set []" by auto
       
   703           qed
       
   704           ultimately show False by auto
       
   705         qed
  1347         qed
   706     qed
  1348       qed
   707     with True show ?thesis by auto
  1349       thus ?thesis using waiting(1,2)
   708   qed
  1350         by (unfold s_RAG_def, fold waiting_eq, auto)
   709 qed
  1351     next
   710 
  1352       case (holding th' cs')
   711 text {* 
  1353       from h this(1,2)
   712   The following @{text "step_v_wait_inv"} is also an obvious lemma, which, however, needs to be
  1354       have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
   713   derived from scratch, which confirms the correctness of the definition of @{text "next_th"}.
  1355       hence "holding (e#s) th' cs'"
   714 *}
  1356       proof
   715 lemma step_v_wait_inv[elim_format]:
  1357         assume "cs' \<noteq> cs"
   716     "\<And>t c. \<lbrakk>vt (V th cs # s); \<not> waiting (wq (V th cs # s)) t c; waiting (wq s) t c
  1358         from holding_esI2[OF this holding(3)] 
   717            \<rbrakk>
  1359         show ?thesis .
   718           \<Longrightarrow> (next_th s th cs t \<and> cs = c)"
       
   719 proof -
       
   720   fix t c 
       
   721   assume vt: "vt (V th cs # s)"
       
   722     and nw: "\<not> waiting (wq (V th cs # s)) t c"
       
   723     and wt: "waiting (wq s) t c"
       
   724   from vt interpret vt_v: valid_trace_e s "V th cs" 
       
   725     by  (cases, unfold_locales, simp)
       
   726   show "next_th s th cs t \<and> cs = c"
       
   727   proof(cases "cs = c")
       
   728     case False
       
   729     with nw wt show ?thesis
       
   730       by (auto simp:cs_waiting_def wq_def Let_def)
       
   731   next
       
   732     case True
       
   733     from nw[folded True] wt[folded True]
       
   734     have "next_th s th cs t"
       
   735       apply (unfold next_th_def, auto simp:cs_waiting_def wq_def Let_def split:list.splits)
       
   736     proof -
       
   737       fix a list
       
   738       assume t_in: "t \<in> set list"
       
   739         and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
       
   740         and eq_wq: "wq_fun (schs s) cs = a # list"
       
   741       have " set (SOME q. distinct q \<and> set q = set list) = set list"
       
   742       proof(rule someI2)
       
   743         from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
       
   744         show "distinct list \<and> set list = set list" by auto
       
   745       next
  1360       next
   746         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
  1361         assume "th' \<noteq> th"
   747           by auto
  1362         from holding_esI1[OF holding(3) this]
       
  1363         show ?thesis .
   748       qed
  1364       qed
   749       with t_ni and t_in show "a = th" by auto
  1365       thus ?thesis using holding(1,2)
       
  1366         by (unfold s_RAG_def, fold holding_eq, auto)
       
  1367     qed
       
  1368    qed
       
  1369  next
       
  1370    case True
       
  1371    interpret h_e: valid_trace_v_e s e th cs
       
  1372         by (unfold_locales, insert True, simp)
       
  1373    from h[unfolded h_e.waiting_set_eq h_e.holding_set_eq]
       
  1374    have h_s: "(n1, n2) \<in> RAG s" "(n1, n2) \<noteq> (Cs cs, Th th)" 
       
  1375       by auto
       
  1376    from h_s(1)
       
  1377    show ?thesis
       
  1378    proof(cases rule:in_RAG_E)
       
  1379     case (waiting th' cs')
       
  1380     from h_e.waiting_esI2[OF this(3)]
       
  1381     show ?thesis using waiting(1,2)
       
  1382       by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1383    next
       
  1384     case (holding th' cs')
       
  1385     with h_s(2)
       
  1386     have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto
       
  1387     thus ?thesis
       
  1388     proof
       
  1389       assume neq_cs: "cs' \<noteq> cs"
       
  1390       from holding_esI2[OF this holding(3)]
       
  1391       show ?thesis using holding(1,2)
       
  1392         by (unfold s_RAG_def, fold holding_eq, auto)
   750     next
  1393     next
   751       fix a list
  1394       assume "th' \<noteq> th"
   752       assume t_in: "t \<in> set list"
  1395       from holding_esI1[OF holding(3) this]
   753         and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)"
  1396       show ?thesis using holding(1,2)
   754         and eq_wq: "wq_fun (schs s) cs = a # list"
  1397         by (unfold s_RAG_def, fold holding_eq, auto)
   755       have " set (SOME q. distinct q \<and> set q = set list) = set list"
  1398     qed
   756       proof(rule someI2)
  1399    qed
   757         from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def]
  1400  qed
   758         show "distinct list \<and> set list = set list" by auto
  1401 qed
   759       next
  1402 
   760         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
  1403 end
   761           by auto
  1404 
   762       qed
  1405 lemma step_RAG_v: 
   763       with t_ni and t_in show "t = hd (SOME q. distinct q \<and> set q = set list)" by auto
       
   764     next
       
   765       fix a list
       
   766       assume eq_wq: "wq_fun (schs s) cs = a # list"
       
   767       from step_back_step[OF vt]
       
   768       show "a = th"
       
   769       proof(cases)
       
   770         assume "holding s th cs"
       
   771         with eq_wq show ?thesis
       
   772           by (unfold s_holding_def wq_def, auto)
       
   773       qed
       
   774     qed
       
   775     with True show ?thesis by simp
       
   776   qed
       
   777 qed
       
   778 
       
   779 lemma step_v_not_wait[consumes 3]:
       
   780   "\<lbrakk>vt (V th cs # s); next_th s th cs t; waiting (wq (V th cs # s)) t cs\<rbrakk> \<Longrightarrow> False"
       
   781   by (unfold next_th_def cs_waiting_def wq_def, auto simp:Let_def)
       
   782 
       
   783 lemma step_v_release:
       
   784   "\<lbrakk>vt (V th cs # s); holding (wq (V th cs # s)) th cs\<rbrakk> \<Longrightarrow> False"
       
   785 proof -
       
   786   assume vt: "vt (V th cs # s)"
       
   787     and hd: "holding (wq (V th cs # s)) th cs"
       
   788   from vt interpret vt_v: valid_trace_e s "V th cs"
       
   789     by (cases, unfold_locales, simp+)
       
   790   from step_back_step [OF vt] and hd
       
   791   show "False"
       
   792   proof(cases)
       
   793     assume "holding (wq (V th cs # s)) th cs" and "holding s th cs"
       
   794     thus ?thesis
       
   795       apply (unfold s_holding_def wq_def cs_holding_def)
       
   796       apply (auto simp:Let_def split:list.splits)
       
   797     proof -
       
   798       fix list
       
   799       assume eq_wq[folded wq_def]: 
       
   800         "wq_fun (schs s) cs = hd (SOME q. distinct q \<and> set q = set list) # list"
       
   801       and hd_in: "hd (SOME q. distinct q \<and> set q = set list)
       
   802             \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   803       have "set (SOME q. distinct q \<and> set q = set list) = set list"
       
   804       proof(rule someI2)
       
   805         from vt_v.wq_distinct[of cs] and eq_wq
       
   806         show "distinct list \<and> set list = set list" by auto
       
   807       next
       
   808         show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list"
       
   809           by auto
       
   810       qed
       
   811       moreover have "distinct  (hd (SOME q. distinct q \<and> set q = set list) # list)"
       
   812       proof -
       
   813         from vt_v.wq_distinct[of cs] and eq_wq
       
   814         show ?thesis by auto
       
   815       qed
       
   816       moreover note eq_wq and hd_in
       
   817       ultimately show "False" by auto
       
   818     qed
       
   819   qed
       
   820 qed
       
   821 
       
   822 lemma step_v_get_hold:
       
   823   "\<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"
       
   824   apply (unfold cs_holding_def next_th_def wq_def,
       
   825          auto simp:Let_def)
       
   826 proof -
       
   827   fix rest
       
   828   assume vt: "vt (V th cs # s)"
       
   829     and eq_wq[folded wq_def]: " wq_fun (schs s) cs = th # rest"
       
   830     and nrest: "rest \<noteq> []"
       
   831     and ni: "hd (SOME q. distinct q \<and> set q = set rest)
       
   832             \<notin> set (SOME q. distinct q \<and> set q = set rest)"
       
   833   from vt interpret vt_v: valid_trace_e s "V th cs"
       
   834     by (cases, unfold_locales, simp+)
       
   835   have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
   836   proof(rule someI2)
       
   837     from vt_v.wq_distinct[of cs] and eq_wq
       
   838     show "distinct rest \<and> set rest = set rest" by auto
       
   839   next
       
   840     fix x assume "distinct x \<and> set x = set rest"
       
   841     hence "set x = set rest" by auto
       
   842     with nrest
       
   843     show "x \<noteq> []" by (case_tac x, auto)
       
   844   qed
       
   845   with ni show "False" by auto
       
   846 qed
       
   847 
       
   848 lemma step_v_release_inv[elim_format]:
       
   849 "\<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> 
       
   850   c = cs \<and> t = th"
       
   851   apply (unfold cs_holding_def wq_def, auto simp:Let_def split:if_splits list.splits)
       
   852   proof -
       
   853     fix a list
       
   854     assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
       
   855     from step_back_step [OF vt] show "a = th"
       
   856     proof(cases)
       
   857       assume "holding s th cs" with eq_wq
       
   858       show ?thesis
       
   859         by (unfold s_holding_def wq_def, auto)
       
   860     qed
       
   861   next
       
   862     fix a list
       
   863     assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list"
       
   864     from step_back_step [OF vt] show "a = th"
       
   865     proof(cases)
       
   866       assume "holding s th cs" with eq_wq
       
   867       show ?thesis
       
   868         by (unfold s_holding_def wq_def, auto)
       
   869     qed
       
   870   qed
       
   871 
       
   872 lemma step_v_waiting_mono:
       
   873   "\<And>t c. \<lbrakk>vt (V th cs # s); waiting (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> waiting (wq s) t c"
       
   874 proof -
       
   875   fix t c
       
   876   let ?s' = "(V th cs # s)"
       
   877   assume vt: "vt ?s'" 
       
   878     and wt: "waiting (wq ?s') t c"
       
   879   from vt interpret vt_v: valid_trace_e s "V th cs"
       
   880     by (cases, unfold_locales, simp+)
       
   881   show "waiting (wq s) t c"
       
   882   proof(cases "c = cs")
       
   883     case False
       
   884     assume neq_cs: "c \<noteq> cs"
       
   885     hence "waiting (wq ?s') t c = waiting (wq s) t c"
       
   886       by (unfold cs_waiting_def wq_def, auto simp:Let_def)
       
   887     with wt show ?thesis by simp
       
   888   next
       
   889     case True
       
   890     with wt show ?thesis
       
   891       apply (unfold cs_waiting_def wq_def, auto simp:Let_def split:list.splits)
       
   892     proof -
       
   893       fix a list
       
   894       assume not_in: "t \<notin> set list"
       
   895         and is_in: "t \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   896         and eq_wq: "wq_fun (schs s) cs = a # list"
       
   897       have "set (SOME q. distinct q \<and> set q = set list) = set list"
       
   898       proof(rule someI2)
       
   899         from vt_v.wq_distinct [of cs]
       
   900         and eq_wq[folded wq_def]
       
   901         show "distinct list \<and> set list = set list" by auto
       
   902       next
       
   903         fix x assume "distinct x \<and> set x = set list"
       
   904         thus "set x = set list" by auto
       
   905       qed
       
   906       with not_in is_in show "t = a" by auto
       
   907     next
       
   908       fix list
       
   909       assume is_waiting: "waiting (wq (V th cs # s)) t cs"
       
   910       and eq_wq: "wq_fun (schs s) cs = t # list"
       
   911       hence "t \<in> set list"
       
   912         apply (unfold wq_def, auto simp:Let_def cs_waiting_def)
       
   913       proof -
       
   914         assume " t \<in> set (SOME q. distinct q \<and> set q = set list)"
       
   915         moreover have "\<dots> = set list" 
       
   916         proof(rule someI2)
       
   917           from vt_v.wq_distinct [of cs]
       
   918             and eq_wq[folded wq_def]
       
   919           show "distinct list \<and> set list = set list" by auto
       
   920         next
       
   921           fix x assume "distinct x \<and> set x = set list" 
       
   922           thus "set x = set list" by auto
       
   923         qed
       
   924         ultimately show "t \<in> set list" by simp
       
   925       qed
       
   926       with eq_wq and vt_v.wq_distinct [of cs, unfolded wq_def]
       
   927       show False by auto
       
   928     qed
       
   929   qed
       
   930 qed
       
   931 
       
   932 text {* (* ddd *) 
       
   933   The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed
       
   934   with the happening of @{text "V"}-events:
       
   935 *}
       
   936 lemma step_RAG_v:
       
   937 assumes vt:
  1406 assumes vt:
   938   "vt (V th cs#s)"
  1407   "vt (V th cs#s)"
   939 shows "
  1408 shows "
   940   RAG (V th cs # s) =
  1409   RAG (V th cs # s) =
   941   RAG s - {(Cs cs, Th th)} -
  1410   RAG s - {(Cs cs, Th th)} -
   942   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
  1411   {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
   943   {(Cs cs, Th th') |th'.  next_th s th cs th'}"
  1412   {(Cs cs, Th th') |th'.  next_th s th cs th'}" (is "?L = ?R")
   944   apply (insert vt, unfold s_RAG_def) 
  1413 proof -
   945   apply (auto split:if_splits list.splits simp:Let_def)
  1414   interpret vt_v: valid_trace_v s "V th cs"
   946   apply (auto elim: step_v_waiting_mono step_v_hold_inv 
  1415     using assms step_back_vt by (unfold_locales, auto) 
   947               step_v_release step_v_wait_inv
  1416   show ?thesis using vt_v.RAG_es .
   948               step_v_get_hold step_v_release_inv)
  1417 qed
   949   apply (erule_tac step_v_not_wait, auto)
  1418 
   950   done
  1419 lemma (in valid_trace_create)
   951 
  1420   th_not_in_threads: "th \<notin> threads s"
   952 text {* 
  1421 proof -
   953   The following @{text "step_RAG_p"} lemma charaterizes how @{text "RAG"} is changed
  1422   from pip_e[unfolded is_create]
   954   with the happening of @{text "P"}-events:
  1423   show ?thesis by (cases, simp)
   955 *}
  1424 qed
   956 lemma step_RAG_p:
  1425 
   957   "vt (P th cs#s) \<Longrightarrow>
  1426 lemma (in valid_trace_create)
   958   RAG (P th cs # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
  1427   threads_es [simp]: "threads (e#s) = threads s \<union> {th}"
   959                                              else RAG s \<union> {(Th th, Cs cs)})"
  1428   by (unfold is_create, simp)
   960   apply(simp only: s_RAG_def wq_def)
  1429 
   961   apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def)
  1430 lemma (in valid_trace_exit)
   962   apply(case_tac "csa = cs", auto)
  1431   threads_es [simp]: "threads (e#s) = threads s - {th}"
   963   apply(fold wq_def)
  1432   by (unfold is_exit, simp)
   964   apply(drule_tac step_back_step)
  1433 
   965   apply(ind_cases " step s (P (hd (wq s cs)) cs)")
  1434 lemma (in valid_trace_p)
   966   apply(simp add:s_RAG_def wq_def cs_holding_def)
  1435   threads_es [simp]: "threads (e#s) = threads s"
   967   apply(auto)
  1436   by (unfold is_p, simp)
   968   done
  1437 
   969 
  1438 lemma (in valid_trace_v)
   970 
  1439   threads_es [simp]: "threads (e#s) = threads s"
   971 lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
  1440   by (unfold is_v, simp)
   972   by (unfold s_RAG_def, auto)
  1441 
       
  1442 lemma (in valid_trace_v)
       
  1443   th_not_in_rest[simp]: "th \<notin> set rest"
       
  1444 proof
       
  1445   assume otherwise: "th \<in> set rest"
       
  1446   have "distinct (wq s cs)" by (simp add: wq_distinct)
       
  1447   from this[unfolded wq_s_cs] and otherwise
       
  1448   show False by auto
       
  1449 qed
       
  1450 
       
  1451 lemma (in valid_trace_v)
       
  1452   set_wq_es_cs [simp]: "set (wq (e#s) cs) = set (wq s cs) - {th}"
       
  1453 proof(unfold wq_es_cs wq'_def, rule someI2)
       
  1454   show "distinct rest \<and> set rest = set rest"
       
  1455     by (simp add: distinct_rest)
       
  1456 next
       
  1457   fix x
       
  1458   assume "distinct x \<and> set x = set rest"
       
  1459   thus "set x = set (wq s cs) - {th}" 
       
  1460       by (unfold wq_s_cs, simp)
       
  1461 qed
       
  1462 
       
  1463 lemma (in valid_trace_exit)
       
  1464   th_not_in_wq: "th \<notin> set (wq s cs)"
       
  1465 proof -
       
  1466   from pip_e[unfolded is_exit]
       
  1467   show ?thesis
       
  1468   by (cases, unfold holdents_def s_holding_def, fold wq_def, 
       
  1469              auto elim!:runing_wqE)
       
  1470 qed
       
  1471 
       
  1472 lemma (in valid_trace) wq_threads: 
       
  1473   assumes "th \<in> set (wq s cs)"
       
  1474   shows "th \<in> threads s"
       
  1475   using assms
       
  1476 proof(induct rule:ind)
       
  1477   case (Nil)
       
  1478   thus ?case by (auto simp:wq_def)
       
  1479 next
       
  1480   case (Cons s e)
       
  1481   interpret vt_e: valid_trace_e s e using Cons by simp
       
  1482   show ?case
       
  1483   proof(cases e)
       
  1484     case (Create th' prio')
       
  1485     interpret vt: valid_trace_create s e th' prio'
       
  1486       using Create by (unfold_locales, simp)
       
  1487     show ?thesis
       
  1488       using Cons.hyps(2) Cons.prems by auto
       
  1489   next
       
  1490     case (Exit th')
       
  1491     interpret vt: valid_trace_exit s e th'
       
  1492       using Exit by (unfold_locales, simp)
       
  1493     show ?thesis
       
  1494       using Cons.hyps(2) Cons.prems vt.th_not_in_wq by auto 
       
  1495   next
       
  1496     case (P th' cs')
       
  1497     interpret vt: valid_trace_p s e th' cs'
       
  1498       using P by (unfold_locales, simp)
       
  1499     show ?thesis
       
  1500       using Cons.hyps(2) Cons.prems readys_threads 
       
  1501         runing_ready vt.is_p vt.runing_th_s vt_e.wq_in_inv 
       
  1502         by fastforce 
       
  1503   next
       
  1504     case (V th' cs')
       
  1505     interpret vt: valid_trace_v s e th' cs'
       
  1506       using V by (unfold_locales, simp)
       
  1507     show ?thesis using Cons
       
  1508       using vt.is_v vt.threads_es vt_e.wq_in_inv by blast
       
  1509   next
       
  1510     case (Set th' prio)
       
  1511     interpret vt: valid_trace_set s e th' prio
       
  1512       using Set by (unfold_locales, simp)
       
  1513     show ?thesis using Cons.hyps(2) Cons.prems vt.is_set 
       
  1514         by (auto simp:wq_def Let_def)
       
  1515   qed
       
  1516 qed 
   973 
  1517 
   974 context valid_trace
  1518 context valid_trace
   975 begin
  1519 begin
   976 
  1520 
   977 text {*
  1521 lemma  dm_RAG_threads:
   978   The following lemma shows that @{text "RAG"} is acyclic.
  1522   assumes in_dom: "(Th th) \<in> Domain (RAG s)"
   979   The overall structure is by induction on the formation of @{text "vt s"}
  1523   shows "th \<in> threads s"
   980   and then case analysis on event @{text "e"}, where the non-trivial cases 
  1524 proof -
   981   for those for @{text "V"} and @{text "P"} events.
  1525   from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
   982 *}
  1526   moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
   983 lemma acyclic_RAG:
  1527   ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
   984   shows "acyclic (RAG s)"
  1528   hence "th \<in> set (wq s cs)"
   985 using vt
  1529     by (unfold s_RAG_def, auto simp:cs_waiting_def)
   986 proof(induct)
  1530   from wq_threads [OF this] show ?thesis .
   987   case (vt_cons s e)
  1531 qed
   988   interpret vt_s: valid_trace s using vt_cons(1)
  1532 
       
  1533 lemma rg_RAG_threads: 
       
  1534   assumes "(Th th) \<in> Range (RAG s)"
       
  1535   shows "th \<in> threads s"
       
  1536   using assms
       
  1537   by (unfold s_RAG_def cs_waiting_def cs_holding_def, 
       
  1538        auto intro:wq_threads)
       
  1539 
       
  1540 lemma RAG_threads:
       
  1541   assumes "(Th th) \<in> Field (RAG s)"
       
  1542   shows "th \<in> threads s"
       
  1543   using assms
       
  1544   by (metis Field_def UnE dm_RAG_threads rg_RAG_threads)
       
  1545 
       
  1546 end
       
  1547 
       
  1548 lemma (in valid_trace_v)
       
  1549   preced_es [simp]: "preced th (e#s) = preced th s"
       
  1550   by (unfold is_v preced_def, simp)
       
  1551 
       
  1552 lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s"
       
  1553 proof
       
  1554   fix th'
       
  1555   show "the_preced (V th cs # s) th' = the_preced s th'"
       
  1556     by (unfold the_preced_def preced_def, simp)
       
  1557 qed
       
  1558 
       
  1559 lemma (in valid_trace_v)
       
  1560   the_preced_es: "the_preced (e#s) = the_preced s"
       
  1561   by (unfold is_v preced_def, simp)
       
  1562 
       
  1563 context valid_trace_p
       
  1564 begin
       
  1565 
       
  1566 lemma not_holding_s_th_cs: "\<not> holding s th cs"
       
  1567 proof
       
  1568   assume otherwise: "holding s th cs"
       
  1569   from pip_e[unfolded is_p]
       
  1570   show False
       
  1571   proof(cases)
       
  1572     case (thread_P)
       
  1573     moreover have "(Cs cs, Th th) \<in> RAG s"
       
  1574       using otherwise cs_holding_def 
       
  1575             holding_eq th_not_in_wq by auto
       
  1576     ultimately show ?thesis by auto
       
  1577   qed
       
  1578 qed
       
  1579 
       
  1580 lemma waiting_kept:
       
  1581   assumes "waiting s th' cs'"
       
  1582   shows "waiting (e#s) th' cs'"
       
  1583   using assms
       
  1584   by (metis cs_waiting_def hd_append2 list.sel(1) list.set_intros(2) 
       
  1585       rotate1.simps(2) self_append_conv2 set_rotate1 
       
  1586         th_not_in_wq waiting_eq wq_es_cs wq_neq_simp)
       
  1587   
       
  1588 lemma holding_kept:
       
  1589   assumes "holding s th' cs'"
       
  1590   shows "holding (e#s) th' cs'"
       
  1591 proof(cases "cs' = cs")
       
  1592   case False
       
  1593   hence "wq (e#s) cs' = wq s cs'" by simp
       
  1594   with assms show ?thesis using cs_holding_def holding_eq by auto 
       
  1595 next
       
  1596   case True
       
  1597   from assms[unfolded s_holding_def, folded wq_def]
       
  1598   obtain rest where eq_wq: "wq s cs' = th'#rest"
       
  1599     by (metis empty_iff list.collapse list.set(1)) 
       
  1600   hence "wq (e#s) cs' = th'#(rest@[th])"
       
  1601     by (simp add: True wq_es_cs) 
       
  1602   thus ?thesis
       
  1603     by (simp add: cs_holding_def holding_eq) 
       
  1604 qed
       
  1605 
       
  1606 end
       
  1607 
       
  1608 locale valid_trace_p_h = valid_trace_p +
       
  1609   assumes we: "wq s cs = []"
       
  1610 
       
  1611 locale valid_trace_p_w = valid_trace_p +
       
  1612   assumes wne: "wq s cs \<noteq> []"
       
  1613 begin
       
  1614 
       
  1615 definition "holder = hd (wq s cs)"
       
  1616 definition "waiters = tl (wq s cs)"
       
  1617 definition "waiters' = waiters @ [th]"
       
  1618 
       
  1619 lemma wq_s_cs: "wq s cs = holder#waiters"
       
  1620     by (simp add: holder_def waiters_def wne)
       
  1621     
       
  1622 lemma wq_es_cs': "wq (e#s) cs = holder#waiters@[th]"
       
  1623   by (simp add: wq_es_cs wq_s_cs)
       
  1624 
       
  1625 lemma waiting_es_th_cs: "waiting (e#s) th cs"
       
  1626   using cs_waiting_def th_not_in_wq waiting_eq wq_es_cs' wq_s_cs by auto
       
  1627 
       
  1628 lemma RAG_edge: "(Th th, Cs cs) \<in> RAG (e#s)"
       
  1629    by (unfold s_RAG_def, fold waiting_eq, insert waiting_es_th_cs, auto)
       
  1630 
       
  1631 lemma holding_esE:
       
  1632   assumes "holding (e#s) th' cs'"
       
  1633   obtains "holding s th' cs'"
       
  1634   using assms 
       
  1635 proof(cases "cs' = cs")
       
  1636   case False
       
  1637   hence "wq (e#s) cs' = wq s cs'" by simp
       
  1638   with assms show ?thesis
       
  1639     using cs_holding_def holding_eq that by auto 
       
  1640 next
       
  1641   case True
       
  1642   with assms show ?thesis
       
  1643   by (metis cs_holding_def holding_eq list.sel(1) list.set_intros(1) that 
       
  1644         wq_es_cs' wq_s_cs) 
       
  1645 qed
       
  1646 
       
  1647 lemma waiting_esE:
       
  1648   assumes "waiting (e#s) th' cs'"
       
  1649   obtains "th' \<noteq> th" "waiting s th' cs'"
       
  1650      |  "th' = th" "cs' = cs"
       
  1651 proof(cases "waiting s th' cs'")
       
  1652   case True
       
  1653   have "th' \<noteq> th"
       
  1654   proof
       
  1655     assume otherwise: "th' = th"
       
  1656     from True[unfolded this]
       
  1657     show False by (simp add: th_not_waiting) 
       
  1658   qed
       
  1659   from that(1)[OF this True] show ?thesis .
       
  1660 next
       
  1661   case False
       
  1662   hence "th' = th \<and> cs' = cs"
       
  1663       by (metis assms cs_waiting_def holder_def list.sel(1) rotate1.simps(2) 
       
  1664         set_ConsD set_rotate1 waiting_eq wq_es_cs wq_es_cs' wq_neq_simp)
       
  1665   with that(2) show ?thesis by metis
       
  1666 qed
       
  1667 
       
  1668 lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Th th, Cs cs)}" (is "?L = ?R")
       
  1669 proof(rule rel_eqI)
       
  1670   fix n1 n2
       
  1671   assume "(n1, n2) \<in> ?L"
       
  1672   thus "(n1, n2) \<in> ?R" 
       
  1673   proof(cases rule:in_RAG_E)
       
  1674     case (waiting th' cs')
       
  1675     from this(3)
       
  1676     show ?thesis
       
  1677     proof(cases rule:waiting_esE)
       
  1678       case 1
       
  1679       thus ?thesis using waiting(1,2)
       
  1680         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1681     next
       
  1682       case 2
       
  1683       thus ?thesis using waiting(1,2) by auto
       
  1684     qed
       
  1685   next
       
  1686     case (holding th' cs')
       
  1687     from this(3)
       
  1688     show ?thesis
       
  1689     proof(cases rule:holding_esE)
       
  1690       case 1
       
  1691       with holding(1,2)
       
  1692       show ?thesis by (unfold s_RAG_def, fold holding_eq, auto)
       
  1693     qed
       
  1694   qed
       
  1695 next
       
  1696   fix n1 n2
       
  1697   assume "(n1, n2) \<in> ?R"
       
  1698   hence "(n1, n2) \<in> RAG s \<or> (n1 = Th th \<and> n2 = Cs cs)" by auto
       
  1699   thus "(n1, n2) \<in> ?L"
       
  1700   proof
       
  1701     assume "(n1, n2) \<in> RAG s"
       
  1702     thus ?thesis
       
  1703     proof(cases rule:in_RAG_E)
       
  1704       case (waiting th' cs')
       
  1705       from waiting_kept[OF this(3)]
       
  1706       show ?thesis using waiting(1,2)
       
  1707          by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1708     next
       
  1709       case (holding th' cs')
       
  1710       from holding_kept[OF this(3)]
       
  1711       show ?thesis using holding(1,2)
       
  1712          by (unfold s_RAG_def, fold holding_eq, auto)
       
  1713     qed
       
  1714   next
       
  1715     assume "n1 = Th th \<and> n2 = Cs cs"
       
  1716     thus ?thesis using RAG_edge by auto
       
  1717   qed
       
  1718 qed
       
  1719 
       
  1720 end
       
  1721 
       
  1722 context valid_trace_p_h
       
  1723 begin
       
  1724 
       
  1725 lemma wq_es_cs': "wq (e#s) cs = [th]"
       
  1726   using wq_es_cs[unfolded we] by simp
       
  1727 
       
  1728 lemma holding_es_th_cs: 
       
  1729   shows "holding (e#s) th cs"
       
  1730 proof -
       
  1731   from wq_es_cs'
       
  1732   have "th \<in> set (wq (e#s) cs)" "th = hd (wq (e#s) cs)" by auto
       
  1733   thus ?thesis using cs_holding_def holding_eq by blast 
       
  1734 qed
       
  1735 
       
  1736 lemma RAG_edge: "(Cs cs, Th th) \<in> RAG (e#s)"
       
  1737   by (unfold s_RAG_def, fold holding_eq, insert holding_es_th_cs, auto)
       
  1738 
       
  1739 lemma waiting_esE:
       
  1740   assumes "waiting (e#s) th' cs'"
       
  1741   obtains "waiting s th' cs'"
       
  1742   using assms
       
  1743   by (metis cs_waiting_def event.distinct(15) is_p list.sel(1) 
       
  1744         set_ConsD waiting_eq we wq_es_cs' wq_neq_simp wq_out_inv)
       
  1745   
       
  1746 lemma holding_esE:
       
  1747   assumes "holding (e#s) th' cs'"
       
  1748   obtains "cs' \<noteq> cs" "holding s th' cs'"
       
  1749     | "cs' = cs" "th' = th"
       
  1750 proof(cases "cs' = cs")
       
  1751   case True
       
  1752   from held_unique[OF holding_es_th_cs assms[unfolded True]]
       
  1753   have "th' = th" by simp
       
  1754   from that(2)[OF True this] show ?thesis .
       
  1755 next
       
  1756   case False
       
  1757   have "holding s th' cs'" using assms
       
  1758     using False cs_holding_def holding_eq by auto
       
  1759   from that(1)[OF False this] show ?thesis .
       
  1760 qed
       
  1761 
       
  1762 lemma RAG_es: "RAG (e # s) =  RAG s \<union> {(Cs cs, Th th)}" (is "?L = ?R")
       
  1763 proof(rule rel_eqI)
       
  1764   fix n1 n2
       
  1765   assume "(n1, n2) \<in> ?L"
       
  1766   thus "(n1, n2) \<in> ?R" 
       
  1767   proof(cases rule:in_RAG_E)
       
  1768     case (waiting th' cs')
       
  1769     from this(3)
       
  1770     show ?thesis
       
  1771     proof(cases rule:waiting_esE)
       
  1772       case 1
       
  1773       thus ?thesis using waiting(1,2)
       
  1774         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1775     qed
       
  1776   next
       
  1777     case (holding th' cs')
       
  1778     from this(3)
       
  1779     show ?thesis
       
  1780     proof(cases rule:holding_esE)
       
  1781       case 1
       
  1782       with holding(1,2)
       
  1783       show ?thesis by (unfold s_RAG_def, fold holding_eq, auto)
       
  1784     next
       
  1785       case 2
       
  1786       with holding(1,2) show ?thesis by auto
       
  1787     qed
       
  1788   qed
       
  1789 next
       
  1790   fix n1 n2
       
  1791   assume "(n1, n2) \<in> ?R"
       
  1792   hence "(n1, n2) \<in> RAG s \<or> (n1 = Cs cs \<and> n2 = Th th)" by auto
       
  1793   thus "(n1, n2) \<in> ?L"
       
  1794   proof
       
  1795     assume "(n1, n2) \<in> RAG s"
       
  1796     thus ?thesis
       
  1797     proof(cases rule:in_RAG_E)
       
  1798       case (waiting th' cs')
       
  1799       from waiting_kept[OF this(3)]
       
  1800       show ?thesis using waiting(1,2)
       
  1801          by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1802     next
       
  1803       case (holding th' cs')
       
  1804       from holding_kept[OF this(3)]
       
  1805       show ?thesis using holding(1,2)
       
  1806          by (unfold s_RAG_def, fold holding_eq, auto)
       
  1807     qed
       
  1808   next
       
  1809     assume "n1 = Cs cs \<and> n2 = Th th"
       
  1810     with holding_es_th_cs
       
  1811     show ?thesis 
       
  1812       by (unfold s_RAG_def, fold holding_eq, auto)
       
  1813   qed
       
  1814 qed
       
  1815 
       
  1816 end
       
  1817 
       
  1818 context valid_trace_p
       
  1819 begin
       
  1820 
       
  1821 lemma RAG_es': "RAG (e # s) =  (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)}
       
  1822                                                   else RAG s \<union> {(Th th, Cs cs)})"
       
  1823 proof(cases "wq s cs = []")
       
  1824   case True
       
  1825   interpret vt_p: valid_trace_p_h using True
   989     by (unfold_locales, simp)
  1826     by (unfold_locales, simp)
   990   assume ih: "acyclic (RAG s)"
  1827   show ?thesis by (simp add: vt_p.RAG_es vt_p.we) 
   991     and stp: "step s e"
  1828 next
   992     and vt: "vt s"
  1829   case False
       
  1830   interpret vt_p: valid_trace_p_w using False
       
  1831     by (unfold_locales, simp)
       
  1832   show ?thesis by (simp add: vt_p.RAG_es vt_p.wne) 
       
  1833 qed
       
  1834 
       
  1835 end
       
  1836 
       
  1837 lemma (in valid_trace_v_n) finite_waiting_set:
       
  1838   "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
       
  1839     by (simp add: waiting_set_eq)
       
  1840 
       
  1841 lemma (in valid_trace_v_n) finite_holding_set:
       
  1842   "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
       
  1843     by (simp add: holding_set_eq)
       
  1844 
       
  1845 lemma (in valid_trace_v_e) finite_waiting_set:
       
  1846   "finite {(Th th', Cs cs) |th'. next_th s th cs th'}"
       
  1847     by (simp add: waiting_set_eq)
       
  1848 
       
  1849 lemma (in valid_trace_v_e) finite_holding_set:
       
  1850   "finite {(Cs cs, Th th') |th'. next_th s th cs th'}"
       
  1851     by (simp add: holding_set_eq)
       
  1852 
       
  1853 context valid_trace_v
       
  1854 begin
       
  1855 
       
  1856 lemma 
       
  1857   finite_RAG_kept:
       
  1858   assumes "finite (RAG s)"
       
  1859   shows "finite (RAG (e#s))"
       
  1860 proof(cases "rest = []")
       
  1861   case True
       
  1862   interpret vt: valid_trace_v_e using True
       
  1863     by (unfold_locales, simp)
       
  1864   show ?thesis using assms
       
  1865     by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
       
  1866 next
       
  1867   case False
       
  1868   interpret vt: valid_trace_v_n using False
       
  1869     by (unfold_locales, simp)
       
  1870   show ?thesis using assms
       
  1871     by  (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp)
       
  1872 qed
       
  1873 
       
  1874 end
       
  1875 
       
  1876 context valid_trace_v_e
       
  1877 begin 
       
  1878 
       
  1879 lemma 
       
  1880   acylic_RAG_kept:
       
  1881   assumes "acyclic (RAG s)"
       
  1882   shows "acyclic (RAG (e#s))"
       
  1883 proof(rule acyclic_subset[OF assms])
       
  1884   show "RAG (e # s) \<subseteq> RAG s"
       
  1885       by (unfold RAG_es waiting_set_eq holding_set_eq, auto)
       
  1886 qed
       
  1887 
       
  1888 end
       
  1889 
       
  1890 context valid_trace_v_n
       
  1891 begin 
       
  1892 
       
  1893 lemma waiting_taker: "waiting s taker cs"
       
  1894   apply (unfold s_waiting_def, fold wq_def, unfold wq_s_cs taker_def)
       
  1895   using eq_wq' th'_in_inv wq'_def by fastforce
       
  1896 
       
  1897 lemma 
       
  1898   acylic_RAG_kept:
       
  1899   assumes "acyclic (RAG s)"
       
  1900   shows "acyclic (RAG (e#s))"
       
  1901 proof -
       
  1902   have "acyclic ((RAG s - {(Cs cs, Th th)} - {(Th taker, Cs cs)}) \<union> 
       
  1903                  {(Cs cs, Th taker)})" (is "acyclic (?A \<union> _)")
       
  1904   proof -
       
  1905     from assms
       
  1906     have "acyclic ?A"
       
  1907        by (rule acyclic_subset, auto)
       
  1908     moreover have "(Th taker, Cs cs) \<notin> ?A^*"
       
  1909     proof
       
  1910       assume otherwise: "(Th taker, Cs cs) \<in> ?A^*"
       
  1911       hence "(Th taker, Cs cs) \<in> ?A^+"
       
  1912         by (unfold rtrancl_eq_or_trancl, auto)
       
  1913       from tranclD[OF this]
       
  1914       obtain cs' where h: "(Th taker, Cs cs') \<in> ?A" 
       
  1915                           "(Th taker, Cs cs') \<in> RAG s"
       
  1916         by (unfold s_RAG_def, auto)
       
  1917       from this(2) have "waiting s taker cs'" 
       
  1918         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1919       from waiting_unique[OF this waiting_taker]
       
  1920       have "cs' = cs" .
       
  1921       from h(1)[unfolded this] show False by auto
       
  1922     qed
       
  1923     ultimately show ?thesis by auto
       
  1924   qed
       
  1925   thus ?thesis 
       
  1926     by (unfold RAG_es waiting_set_eq holding_set_eq, simp)
       
  1927 qed
       
  1928 
       
  1929 end
       
  1930 
       
  1931 context valid_trace_p_h
       
  1932 begin
       
  1933 
       
  1934 lemma 
       
  1935   acylic_RAG_kept:
       
  1936   assumes "acyclic (RAG s)"
       
  1937   shows "acyclic (RAG (e#s))"
       
  1938 proof -
       
  1939   have "acyclic (RAG s \<union> {(Cs cs, Th th)})" (is "acyclic (?A \<union> _)") 
       
  1940   proof -
       
  1941     from assms
       
  1942     have "acyclic ?A"
       
  1943        by (rule acyclic_subset, auto)
       
  1944     moreover have "(Th th, Cs cs) \<notin> ?A^*"
       
  1945     proof
       
  1946       assume otherwise: "(Th th, Cs cs) \<in> ?A^*"
       
  1947       hence "(Th th, Cs cs) \<in> ?A^+"
       
  1948         by (unfold rtrancl_eq_or_trancl, auto)
       
  1949       from tranclD[OF this]
       
  1950       obtain cs' where h: "(Th th, Cs cs') \<in> RAG s"
       
  1951         by (unfold s_RAG_def, auto)
       
  1952       hence "waiting s th cs'" 
       
  1953         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  1954       with th_not_waiting show False by auto
       
  1955     qed
       
  1956     ultimately show ?thesis by auto
       
  1957   qed
       
  1958   thus ?thesis by (unfold RAG_es, simp)
       
  1959 qed
       
  1960 
       
  1961 end
       
  1962 
       
  1963 context valid_trace_p_w
       
  1964 begin
       
  1965 
       
  1966 lemma 
       
  1967   acylic_RAG_kept:
       
  1968   assumes "acyclic (RAG s)"
       
  1969   shows "acyclic (RAG (e#s))"
       
  1970 proof -
       
  1971   have "acyclic (RAG s \<union> {(Th th, Cs cs)})" (is "acyclic (?A \<union> _)") 
       
  1972   proof -
       
  1973     from assms
       
  1974     have "acyclic ?A"
       
  1975        by (rule acyclic_subset, auto)
       
  1976     moreover have "(Cs cs, Th th) \<notin> ?A^*"
       
  1977     proof
       
  1978       assume otherwise: "(Cs cs, Th th) \<in> ?A^*"
       
  1979       from pip_e[unfolded is_p]
       
  1980       show False
       
  1981       proof(cases)
       
  1982         case (thread_P)
       
  1983         moreover from otherwise have "(Cs cs, Th th) \<in> ?A^+"
       
  1984             by (unfold rtrancl_eq_or_trancl, auto)
       
  1985         ultimately show ?thesis by auto
       
  1986       qed
       
  1987     qed
       
  1988     ultimately show ?thesis by auto
       
  1989   qed
       
  1990   thus ?thesis by (unfold RAG_es, simp)
       
  1991 qed
       
  1992 
       
  1993 end
       
  1994 
       
  1995 context valid_trace
       
  1996 begin
       
  1997 
       
  1998 lemma finite_RAG:
       
  1999   shows "finite (RAG s)"
       
  2000 proof(induct rule:ind)
       
  2001   case Nil
       
  2002   show ?case 
       
  2003   by (auto simp: s_RAG_def cs_waiting_def 
       
  2004                    cs_holding_def wq_def acyclic_def)
       
  2005 next
       
  2006   case (Cons s e)
       
  2007   interpret vt_e: valid_trace_e s e using Cons by simp
   993   show ?case
  2008   show ?case
   994   proof(cases e)
  2009   proof(cases e)
   995     case (Create th prio)
  2010     case (Create th prio)
   996     with ih
  2011     interpret vt: valid_trace_create s e th prio using Create
   997     show ?thesis by (simp add:RAG_create_unchanged)
  2012       by (unfold_locales, simp)
       
  2013     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
   998   next
  2014   next
   999     case (Exit th)
  2015     case (Exit th)
  1000     with ih show ?thesis by (simp add:RAG_exit_unchanged)
  2016     interpret vt: valid_trace_exit s e th using Exit
       
  2017       by (unfold_locales, simp)
       
  2018     show ?thesis using Cons by (simp add: vt.RAG_unchanged)
       
  2019   next
       
  2020     case (P th cs)
       
  2021     interpret vt: valid_trace_p s e th cs using P
       
  2022       by (unfold_locales, simp)
       
  2023     show ?thesis using Cons using vt.RAG_es' by auto 
  1001   next
  2024   next
  1002     case (V th cs)
  2025     case (V th cs)
  1003     from V vt stp have vtt: "vt (V th cs#s)" by auto
  2026     interpret vt: valid_trace_v s e th cs using V
  1004     from step_RAG_v [OF this]
  2027       by (unfold_locales, simp)
  1005     have eq_de: 
  2028     show ?thesis using Cons by (simp add: vt.finite_RAG_kept) 
  1006       "RAG (e # s) = 
  2029   next
  1007       RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
  2030     case (Set th prio)
  1008       {(Cs cs, Th th') |th'. next_th s th cs th'}"
  2031     interpret vt: valid_trace_set s e th prio using Set
  1009       (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
  2032       by (unfold_locales, simp)
  1010     from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
  2033     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
  1011     from step_back_step [OF vtt]
  2034   qed
  1012     have "step s (V th cs)" .
  2035 qed
  1013     thus ?thesis
  2036 
  1014     proof(cases)
  2037 lemma acyclic_RAG:
  1015       assume "holding s th cs"
  2038   shows "acyclic (RAG s)"
  1016       hence th_in: "th \<in> set (wq s cs)" and
  2039 proof(induct rule:ind)
  1017         eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto
  2040   case Nil
  1018       then obtain rest where
  2041   show ?case 
  1019         eq_wq: "wq s cs = th#rest"
  2042   by (auto simp: s_RAG_def cs_waiting_def 
  1020         by (cases "wq s cs", auto)
  2043                    cs_holding_def wq_def acyclic_def)
  1021       show ?thesis
  2044 next
  1022       proof(cases "rest = []")
  2045   case (Cons s e)
  1023         case False
  2046   interpret vt_e: valid_trace_e s e using Cons by simp
  1024         let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
  2047   show ?case
  1025         from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}"
  2048   proof(cases e)
  1026           by (unfold next_th_def, auto)
  2049     case (Create th prio)
  1027         let ?E = "(?A - ?B - ?C)"
  2050     interpret vt: valid_trace_create s e th prio using Create
  1028         have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*"
  2051       by (unfold_locales, simp)
  1029         proof
  2052     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
  1030           assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*"
  2053   next
  1031           hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
  2054     case (Exit th)
  1032           from tranclD [OF this]
  2055     interpret vt: valid_trace_exit s e th using Exit
  1033           obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast
  2056       by (unfold_locales, simp)
  1034           hence th_d: "(Th ?th', x) \<in> ?A" by simp
  2057     show ?thesis using Cons by (simp add: vt.RAG_unchanged)
  1035           from RAG_target_th [OF this]
       
  1036           obtain cs' where eq_x: "x = Cs cs'" by auto
       
  1037           with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp
       
  1038           hence wt_th': "waiting s ?th' cs'"
       
  1039             unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp
       
  1040           hence "cs' = cs"
       
  1041           proof(rule vt_s.waiting_unique)
       
  1042             from eq_wq vt_s.wq_distinct[of cs]
       
  1043             show "waiting s ?th' cs" 
       
  1044               apply (unfold s_waiting_def wq_def, auto)
       
  1045             proof -
       
  1046               assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1047                 and eq_wq: "wq_fun (schs s) cs = th # rest"
       
  1048               have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1049               proof(rule someI2)
       
  1050                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1051                 show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
       
  1052               next
       
  1053                 fix x assume "distinct x \<and> set x = set rest"
       
  1054                 with False show "x \<noteq> []" by auto
       
  1055               qed
       
  1056               hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
       
  1057                 set (SOME q. distinct q \<and> set q = set rest)" by auto
       
  1058               moreover have "\<dots> = set rest" 
       
  1059               proof(rule someI2)
       
  1060                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1061                 show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto
       
  1062               next
       
  1063                 show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1064               qed
       
  1065               moreover note hd_in
       
  1066               ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
       
  1067             next
       
  1068               assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1069                 and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
       
  1070               have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1071               proof(rule someI2)
       
  1072                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1073                 show "distinct rest \<and> set rest = set rest" by auto
       
  1074               next
       
  1075                 fix x assume "distinct x \<and> set x = set rest"
       
  1076                 with False show "x \<noteq> []" by auto
       
  1077               qed
       
  1078               hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> 
       
  1079                 set (SOME q. distinct q \<and> set q = set rest)" by auto
       
  1080               moreover have "\<dots> = set rest" 
       
  1081               proof(rule someI2)
       
  1082                 from vt_s.wq_distinct[of cs] and eq_wq
       
  1083                 show "distinct rest \<and> set rest = set rest" by auto
       
  1084               next
       
  1085                 show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1086               qed
       
  1087               moreover note hd_in
       
  1088               ultimately show False by auto
       
  1089             qed
       
  1090           qed
       
  1091           with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp
       
  1092           with False
       
  1093           show "False" by (auto simp: next_th_def eq_wq)
       
  1094         qed
       
  1095         with acyclic_insert[symmetric] and ac
       
  1096           and eq_de eq_D show ?thesis by auto
       
  1097       next
       
  1098         case True
       
  1099         with eq_wq
       
  1100         have eq_D: "?D = {}"
       
  1101           by (unfold next_th_def, auto)
       
  1102         with eq_de ac
       
  1103         show ?thesis by auto
       
  1104       qed 
       
  1105     qed
       
  1106   next
  2058   next
  1107     case (P th cs)
  2059     case (P th cs)
  1108     from P vt stp have vtt: "vt (P th cs#s)" by auto
  2060     interpret vt: valid_trace_p s e th cs using P
  1109     from step_RAG_p [OF this] P
  2061       by (unfold_locales, simp)
  1110     have "RAG (e # s) = 
  2062     show ?thesis
  1111       (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
       
  1112       RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
       
  1113       by simp
       
  1114     moreover have "acyclic ?R"
       
  1115     proof(cases "wq s cs = []")
  2063     proof(cases "wq s cs = []")
  1116       case True
  2064       case True
  1117       hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
  2065       then interpret vt_h: valid_trace_p_h s e th cs
  1118       have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*"
  2066         by (unfold_locales, simp)
  1119       proof
  2067       show ?thesis using Cons by (simp add: vt_h.acylic_RAG_kept) 
  1120         assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*"
       
  1121         hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
       
  1122         from tranclD2 [OF this]
       
  1123         obtain x where "(x, Cs cs) \<in> RAG s" by auto
       
  1124         with True show False by (auto simp:s_RAG_def cs_waiting_def)
       
  1125       qed
       
  1126       with acyclic_insert ih eq_r show ?thesis by auto
       
  1127     next
  2068     next
  1128       case False
  2069       case False
  1129       hence eq_r: "?R =  RAG s \<union> {(Th th, Cs cs)}" by simp
  2070       then interpret vt_w: valid_trace_p_w s e th cs
  1130       have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*"
  2071         by (unfold_locales, simp)
  1131       proof
  2072       show ?thesis using Cons by (simp add: vt_w.acylic_RAG_kept) 
  1132         assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*"
  2073     qed
  1133         hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
  2074   next
  1134         moreover from step_back_step [OF vtt] have "step s (P th cs)" .
  2075     case (V th cs)
  1135         ultimately show False
  2076     interpret vt: valid_trace_v s e th cs using V
  1136         proof -
  2077       by (unfold_locales, simp)
  1137           show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
  2078     show ?thesis
  1138             by (ind_cases "step s (P th cs)", simp)
  2079     proof(cases "vt.rest = []")
       
  2080       case True
       
  2081       then interpret vt_e: valid_trace_v_e s e th cs
       
  2082         by (unfold_locales, simp)
       
  2083       show ?thesis by (simp add: Cons.hyps(2) vt_e.acylic_RAG_kept) 
       
  2084     next
       
  2085       case False
       
  2086       then interpret vt_n: valid_trace_v_n s e th cs
       
  2087         by (unfold_locales, simp)
       
  2088       show ?thesis by (simp add: Cons.hyps(2) vt_n.acylic_RAG_kept) 
       
  2089     qed
       
  2090   next
       
  2091     case (Set th prio)
       
  2092     interpret vt: valid_trace_set s e th prio using Set
       
  2093       by (unfold_locales, simp)
       
  2094     show ?thesis using Cons by (simp add: vt.RAG_unchanged) 
       
  2095   qed
       
  2096 qed
       
  2097 
       
  2098 lemma wf_RAG: "wf (RAG s)"
       
  2099 proof(rule finite_acyclic_wf)
       
  2100   from finite_RAG show "finite (RAG s)" .
       
  2101 next
       
  2102   from acyclic_RAG show "acyclic (RAG s)" .
       
  2103 qed
       
  2104 
       
  2105 lemma sgv_wRAG: "single_valued (wRAG s)"
       
  2106   using waiting_unique
       
  2107   by (unfold single_valued_def wRAG_def, auto)
       
  2108 
       
  2109 lemma sgv_hRAG: "single_valued (hRAG s)"
       
  2110   using held_unique 
       
  2111   by (unfold single_valued_def hRAG_def, auto)
       
  2112 
       
  2113 lemma sgv_tRAG: "single_valued (tRAG s)"
       
  2114   by (unfold tRAG_def, rule single_valued_relcomp, 
       
  2115               insert sgv_wRAG sgv_hRAG, auto)
       
  2116 
       
  2117 lemma acyclic_tRAG: "acyclic (tRAG s)"
       
  2118 proof(unfold tRAG_def, rule acyclic_compose)
       
  2119   show "acyclic (RAG s)" using acyclic_RAG .
       
  2120 next
       
  2121   show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
  2122 next
       
  2123   show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
  2124 qed
       
  2125 
       
  2126 lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
       
  2127   apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
       
  2128   by(auto elim:waiting_unique held_unique)
       
  2129 
       
  2130 lemma sgv_RAG: "single_valued (RAG s)"
       
  2131   using unique_RAG by (auto simp:single_valued_def)
       
  2132 
       
  2133 lemma rtree_RAG: "rtree (RAG s)"
       
  2134   using sgv_RAG acyclic_RAG
       
  2135   by (unfold rtree_def rtree_axioms_def sgv_def, auto)
       
  2136 
       
  2137 end
       
  2138 
       
  2139 sublocale valid_trace < rtree_RAG: rtree "RAG s"
       
  2140 proof
       
  2141   show "single_valued (RAG s)"
       
  2142   apply (intro_locales)
       
  2143     by (unfold single_valued_def, 
       
  2144         auto intro:unique_RAG)
       
  2145 
       
  2146   show "acyclic (RAG s)"
       
  2147      by (rule acyclic_RAG)
       
  2148 qed
       
  2149 
       
  2150 sublocale valid_trace < rtree_s: rtree "tRAG s"
       
  2151 proof(unfold_locales)
       
  2152   from sgv_tRAG show "single_valued (tRAG s)" .
       
  2153 next
       
  2154   from acyclic_tRAG show "acyclic (tRAG s)" .
       
  2155 qed
       
  2156 
       
  2157 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
       
  2158 proof -
       
  2159   show "fsubtree (RAG s)"
       
  2160   proof(intro_locales)
       
  2161     show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
       
  2162   next
       
  2163     show "fsubtree_axioms (RAG s)"
       
  2164     proof(unfold fsubtree_axioms_def)
       
  2165       from wf_RAG show "wf (RAG s)" .
       
  2166     qed
       
  2167   qed
       
  2168 qed
       
  2169 
       
  2170 lemma tRAG_alt_def: 
       
  2171   "tRAG s = {(Th th1, Th th2) | th1 th2. 
       
  2172                   \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
       
  2173  by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
       
  2174 
       
  2175 sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
       
  2176 proof -
       
  2177   have "fsubtree (tRAG s)"
       
  2178   proof -
       
  2179     have "fbranch (tRAG s)"
       
  2180     proof(unfold tRAG_def, rule fbranch_compose)
       
  2181         show "fbranch (wRAG s)"
       
  2182         proof(rule finite_fbranchI)
       
  2183            from finite_RAG show "finite (wRAG s)"
       
  2184            by (unfold RAG_split, auto)
  1139         qed
  2185         qed
  1140       qed
       
  1141       with acyclic_insert ih eq_r show ?thesis by auto
       
  1142       qed
       
  1143       ultimately show ?thesis by simp
       
  1144     next
  2186     next
  1145       case (Set thread prio)
  2187         show "fbranch (hRAG s)"
  1146       with ih
  2188         proof(rule finite_fbranchI)
  1147       thm RAG_set_unchanged
  2189            from finite_RAG 
  1148       show ?thesis by (simp add:RAG_set_unchanged)
  2190            show "finite (hRAG s)" by (unfold RAG_split, auto)
  1149     qed
  2191         qed
  1150   next
  2192     qed
  1151     case vt_nil
  2193     moreover have "wf (tRAG s)"
  1152     show "acyclic (RAG ([]::state))"
  2194     proof(rule wf_subset)
  1153       by (auto simp: s_RAG_def cs_waiting_def 
  2195       show "wf (RAG s O RAG s)" using wf_RAG
  1154         cs_holding_def wq_def acyclic_def)
  2196         by (fold wf_comp_self, simp)
  1155 qed
       
  1156 
       
  1157 
       
  1158 lemma finite_RAG:
       
  1159   shows "finite (RAG s)"
       
  1160 proof -
       
  1161   from vt show ?thesis
       
  1162   proof(induct)
       
  1163     case (vt_cons s e)
       
  1164     interpret vt_s: valid_trace s using vt_cons(1)
       
  1165       by (unfold_locales, simp)
       
  1166     assume ih: "finite (RAG s)"
       
  1167       and stp: "step s e"
       
  1168       and vt: "vt s"
       
  1169     show ?case
       
  1170     proof(cases e)
       
  1171       case (Create th prio)
       
  1172       with ih
       
  1173       show ?thesis by (simp add:RAG_create_unchanged)
       
  1174     next
  2197     next
  1175       case (Exit th)
  2198       show "tRAG s \<subseteq> (RAG s O RAG s)"
  1176       with ih show ?thesis by (simp add:RAG_exit_unchanged)
  2199         by (unfold tRAG_alt_def, auto)
  1177     next
  2200     qed
  1178       case (V th cs)
  2201     ultimately show ?thesis
  1179       from V vt stp have vtt: "vt (V th cs#s)" by auto
  2202       by (unfold fsubtree_def fsubtree_axioms_def,auto)
  1180       from step_RAG_v [OF this]
  2203   qed
  1181       have eq_de: "RAG (e # s) = 
  2204   from this[folded tRAG_def] show "fsubtree (tRAG s)" .
  1182                    RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union>
  2205 qed
  1183                       {(Cs cs, Th th') |th'. next_th s th cs th'}
  2206 
  1184 "
  2207 
  1185         (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
  2208 context valid_trace
  1186       moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
  2209 begin
  1187       moreover have "finite ?D"
  2210 
  1188       proof -
  2211 lemma finite_subtree_threads:
  1189         have "?D = {} \<or> (\<exists> a. ?D = {a})" 
  2212     "finite {th'. Th th' \<in> subtree (RAG s) (Th th)}" (is "finite ?A")
  1190           by (unfold next_th_def, auto)
  2213 proof -
  1191         thus ?thesis
  2214   have "?A = the_thread ` {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
  1192         proof
  2215         by (auto, insert image_iff, fastforce)
  1193           assume h: "?D = {}"
  2216   moreover have "finite {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}"
  1194           show ?thesis by (unfold h, simp)
  2217         (is "finite ?B")
  1195         next
  2218   proof -
  1196           assume "\<exists> a. ?D = {a}"
  2219      have "?B = (subtree (RAG s) (Th th)) \<inter> {Th th' | th'. True}"
  1197           thus ?thesis
  2220       by auto
  1198             by (metis finite.simps)
  2221      moreover have "... \<subseteq> (subtree (RAG s) (Th th))" by auto
  1199         qed
  2222      moreover have "finite ..." by (simp add: fsbtRAGs.finite_subtree) 
  1200       qed
  2223      ultimately show ?thesis by auto
  1201       ultimately show ?thesis by simp
  2224   qed
  1202     next
  2225   ultimately show ?thesis by auto
  1203       case (P th cs)
  2226 qed
  1204       from P vt stp have vtt: "vt (P th cs#s)" by auto
  2227 
  1205       from step_RAG_p [OF this] P
  2228 lemma le_cp:
  1206       have "RAG (e # s) = 
  2229   shows "preced th s \<le> cp s th"
  1207               (if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else 
  2230   proof(unfold cp_alt_def, rule Max_ge)
  1208                                     RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
  2231     show "finite (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})"
  1209         by simp
  2232       by (simp add: finite_subtree_threads)
  1210       moreover have "finite ?R"
  2233   next
  1211       proof(cases "wq s cs = []")
  2234     show "preced th s \<in> the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}"
  1212         case True
  2235       by (simp add: subtree_def the_preced_def)   
  1213         hence eq_r: "?R =  RAG s \<union> {(Cs cs, Th th)}" by simp
  2236   qed
  1214         with True and ih show ?thesis by auto
  2237 
  1215       next
  2238 lemma cp_le:
  1216         case False
  2239   assumes th_in: "th \<in> threads s"
  1217         hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp
  2240   shows "cp s th \<le> Max (the_preced s ` threads s)"
  1218         with False and ih show ?thesis by auto
  2241 proof(unfold cp_alt_def, rule Max_f_mono)
  1219       qed
  2242   show "finite (threads s)" by (simp add: finite_threads) 
  1220       ultimately show ?thesis by auto
  2243 next
  1221     next
  2244   show " {th'. Th th' \<in> subtree (RAG s) (Th th)} \<noteq> {}"
  1222       case (Set thread prio)
  2245     using subtree_def by fastforce
  1223       with ih
  2246 next
  1224       show ?thesis by (simp add:RAG_set_unchanged)
  2247   show "{th'. Th th' \<in> subtree (RAG s) (Th th)} \<subseteq> threads s"
  1225     qed
  2248     using assms
  1226   next
  2249     by (smt Domain.DomainI dm_RAG_threads mem_Collect_eq 
  1227     case vt_nil
  2250         node.inject(1) rtranclD subsetI subtree_def trancl_domain) 
  1228     show "finite (RAG ([]::state))"
  2251 qed
  1229       by (auto simp: s_RAG_def cs_waiting_def 
  2252 
  1230                    cs_holding_def wq_def acyclic_def)
  2253 lemma max_cp_eq: 
  1231   qed
  2254   shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)"
  1232 qed
  2255   (is "?L = ?R")
  1233 
  2256 proof -
  1234 text {* Several useful lemmas *}
  2257   have "?L \<le> ?R" 
  1235 
  2258   proof(cases "threads s = {}")
  1236 lemma wf_dep_converse: 
  2259     case False
       
  2260     show ?thesis 
       
  2261       by (rule Max.boundedI, 
       
  2262           insert cp_le, 
       
  2263           auto simp:finite_threads False)
       
  2264   qed auto
       
  2265   moreover have "?R \<le> ?L"
       
  2266     by (rule Max_fg_mono, 
       
  2267         simp add: finite_threads,
       
  2268         simp add: le_cp the_preced_def)
       
  2269   ultimately show ?thesis by auto
       
  2270 qed
       
  2271 
       
  2272 lemma wf_RAG_converse: 
  1237   shows "wf ((RAG s)^-1)"
  2273   shows "wf ((RAG s)^-1)"
  1238 proof(rule finite_acyclic_wf_converse)
  2274 proof(rule finite_acyclic_wf_converse)
  1239   from finite_RAG 
  2275   from finite_RAG 
  1240   show "finite (RAG s)" .
  2276   show "finite (RAG s)" .
  1241 next
  2277 next
  1242   from acyclic_RAG
  2278   from acyclic_RAG
  1243   show "acyclic (RAG s)" .
  2279   show "acyclic (RAG s)" .
  1244 qed
  2280 qed
  1245 
  2281 
  1246 end
       
  1247 
       
  1248 lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
       
  1249   by (induct l, auto)
       
  1250 
       
  1251 lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s"
       
  1252   by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1253 
       
  1254 context valid_trace
       
  1255 begin
       
  1256 
       
  1257 lemma wq_threads: 
       
  1258   assumes h: "th \<in> set (wq s cs)"
       
  1259   shows "th \<in> threads s"
       
  1260 proof -
       
  1261  from vt and h show ?thesis
       
  1262   proof(induct arbitrary: th cs)
       
  1263     case (vt_cons s e)
       
  1264     interpret vt_s: valid_trace s
       
  1265       using vt_cons(1) by (unfold_locales, auto)
       
  1266     assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
       
  1267       and stp: "step s e"
       
  1268       and vt: "vt s"
       
  1269       and h: "th \<in> set (wq (e # s) cs)"
       
  1270     show ?case
       
  1271     proof(cases e)
       
  1272       case (Create th' prio)
       
  1273       with ih h show ?thesis
       
  1274         by (auto simp:wq_def Let_def)
       
  1275     next
       
  1276       case (Exit th')
       
  1277       with stp ih h show ?thesis
       
  1278         apply (auto simp:wq_def Let_def)
       
  1279         apply (ind_cases "step s (Exit th')")
       
  1280         apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
       
  1281                s_RAG_def s_holding_def cs_holding_def)
       
  1282         done
       
  1283     next
       
  1284       case (V th' cs')
       
  1285       show ?thesis
       
  1286       proof(cases "cs' = cs")
       
  1287         case False
       
  1288         with h
       
  1289         show ?thesis
       
  1290           apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
       
  1291           by (drule_tac ih, simp)
       
  1292       next
       
  1293         case True
       
  1294         from h
       
  1295         show ?thesis
       
  1296         proof(unfold V wq_def)
       
  1297           assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
       
  1298           show "th \<in> threads (V th' cs' # s)"
       
  1299           proof(cases "cs = cs'")
       
  1300             case False
       
  1301             hence "?l = wq_fun (schs s) cs" by (simp add:Let_def)
       
  1302             with th_in have " th \<in> set (wq s cs)" 
       
  1303               by (fold wq_def, simp)
       
  1304             from ih [OF this] show ?thesis by simp
       
  1305           next
       
  1306             case True
       
  1307             show ?thesis
       
  1308             proof(cases "wq_fun (schs s) cs'")
       
  1309               case Nil
       
  1310               with h V show ?thesis
       
  1311                 apply (auto simp:wq_def Let_def split:if_splits)
       
  1312                 by (fold wq_def, drule_tac ih, simp)
       
  1313             next
       
  1314               case (Cons a rest)
       
  1315               assume eq_wq: "wq_fun (schs s) cs' = a # rest"
       
  1316               with h V show ?thesis
       
  1317                 apply (auto simp:Let_def wq_def split:if_splits)
       
  1318               proof -
       
  1319                 assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1320                 have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
       
  1321                 proof(rule someI2)
       
  1322                   from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def]
       
  1323                   show "distinct rest \<and> set rest = set rest" by auto
       
  1324                 next
       
  1325                   show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
       
  1326                     by auto
       
  1327                 qed
       
  1328                 with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto
       
  1329                 from ih[OF this[folded wq_def]] show "th \<in> threads s" .
       
  1330               next
       
  1331                 assume th_in: "th \<in> set (wq_fun (schs s) cs)"
       
  1332                 from ih[OF this[folded wq_def]]
       
  1333                 show "th \<in> threads s" .
       
  1334               qed
       
  1335             qed
       
  1336           qed
       
  1337         qed
       
  1338       qed
       
  1339     next
       
  1340       case (P th' cs')
       
  1341       from h stp
       
  1342       show ?thesis
       
  1343         apply (unfold P wq_def)
       
  1344         apply (auto simp:Let_def split:if_splits, fold wq_def)
       
  1345         apply (auto intro:ih)
       
  1346         apply(ind_cases "step s (P th' cs')")
       
  1347         by (unfold runing_def readys_def, auto)
       
  1348     next
       
  1349       case (Set thread prio)
       
  1350       with ih h show ?thesis
       
  1351         by (auto simp:wq_def Let_def)
       
  1352     qed
       
  1353   next
       
  1354     case vt_nil
       
  1355     thus ?case by (auto simp:wq_def)
       
  1356   qed
       
  1357 qed
       
  1358 
       
  1359 lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
       
  1360   apply(unfold s_RAG_def cs_waiting_def cs_holding_def)
       
  1361   by (auto intro:wq_threads)
       
  1362 
       
  1363 lemma readys_v_eq:
       
  1364   assumes neq_th: "th \<noteq> thread"
       
  1365   and eq_wq: "wq s cs = thread#rest"
       
  1366   and not_in: "th \<notin>  set rest"
       
  1367   shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
       
  1368 proof -
       
  1369   from assms show ?thesis
       
  1370     apply (auto simp:readys_def)
       
  1371     apply(simp add:s_waiting_def[folded wq_def])
       
  1372     apply (erule_tac x = csa in allE)
       
  1373     apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
       
  1374     apply (case_tac "csa = cs", simp)
       
  1375     apply (erule_tac x = cs in allE)
       
  1376     apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits)
       
  1377     apply(auto simp add: wq_def)
       
  1378     apply (auto simp:s_waiting_def wq_def Let_def split:list.splits)
       
  1379     proof -
       
  1380        assume th_nin: "th \<notin> set rest"
       
  1381         and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1382         and eq_wq: "wq_fun (schs s) cs = thread # rest"
       
  1383       have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1384       proof(rule someI2)
       
  1385         from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def]
       
  1386         show "distinct rest \<and> set rest = set rest" by auto
       
  1387       next
       
  1388         show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1389       qed
       
  1390       with th_nin th_in show False by auto
       
  1391     qed
       
  1392 qed
       
  1393 
       
  1394 text {* \noindent
       
  1395   The following lemmas shows that: starting from any node in @{text "RAG"}, 
       
  1396   by chasing out-going edges, it is always possible to reach a node representing a ready
       
  1397   thread. In this lemma, it is the @{text "th'"}.
       
  1398 *}
       
  1399 
       
  1400 lemma chain_building:
  2282 lemma chain_building:
  1401   shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)"
  2283   assumes "node \<in> Domain (RAG s)"
  1402 proof -
  2284   obtains th' where "th' \<in> readys s" "(node, Th th') \<in> (RAG s)^+"
  1403   from wf_dep_converse
  2285 proof -
  1404   have h: "wf ((RAG s)\<inverse>)" .
  2286   from assms have "node \<in> Range ((RAG s)^-1)" by auto
  1405   show ?thesis
  2287   from wf_base[OF wf_RAG_converse this]
  1406   proof(induct rule:wf_induct [OF h])
  2288   obtain b where h_b: "(b, node) \<in> ((RAG s)\<inverse>)\<^sup>+" "\<forall>c. (c, b) \<notin> (RAG s)\<inverse>" by auto
  1407     fix x
  2289   obtain th' where eq_b: "b = Th th'"
  1408     assume ih [rule_format]: 
  2290   proof(cases b)
  1409       "\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> 
  2291     case (Cs cs)
  1410            y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)"
  2292     from h_b(1)[unfolded trancl_converse] 
  1411     show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)"
  2293     have "(node, b) \<in> ((RAG s)\<^sup>+)" by auto
  1412     proof
  2294     from tranclE[OF this]
  1413       assume x_d: "x \<in> Domain (RAG s)"
  2295     obtain n where "(n, b) \<in> RAG s" by auto
  1414       show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+"
  2296     from this[unfolded Cs]
  1415       proof(cases x)
  2297     obtain th1 where "waiting s th1 cs"
  1416         case (Th th)
  2298       by (unfold s_RAG_def, fold waiting_eq, auto)
  1417         from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def)
  2299     from waiting_holding[OF this]
  1418         with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp
  2300     obtain th2 where "holding s th2 cs" .
  1419         from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast
  2301     hence "(Cs cs, Th th2) \<in> RAG s"
  1420         hence "Cs cs \<in> Domain (RAG s)" by auto
  2302       by (unfold s_RAG_def, fold holding_eq, auto)
  1421         from ih [OF x_in_r this] obtain th'
  2303     with h_b(2)[unfolded Cs, rule_format]
  1422           where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto
  2304     have False by auto
  1423         have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto
  2305     thus ?thesis by auto
  1424         with th'_ready show ?thesis by auto
  2306   qed auto
  1425       next
  2307   have "th' \<in> readys s" 
  1426         case (Cs cs)
  2308   proof -
  1427         from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def)
  2309     from h_b(2)[unfolded eq_b]
  1428         show ?thesis
  2310     have "\<forall>cs. \<not> waiting s th' cs"
  1429         proof(cases "th' \<in> readys s")
  2311       by (unfold s_RAG_def, fold waiting_eq, auto)
  1430           case True
  2312     moreover have "th' \<in> threads s"
  1431           from True and th'_d show ?thesis by auto
  2313     proof(rule rg_RAG_threads)
  1432         next
  2314       from tranclD[OF h_b(1), unfolded eq_b]
  1433           case False
  2315       obtain z where "(z, Th th') \<in> (RAG s)" by auto
  1434           from th'_d and range_in  have "th' \<in> threads s" by auto
  2316       thus "Th th' \<in> Range (RAG s)" by auto
  1435           with False have "Th th' \<in> Domain (RAG s)" 
  2317     qed
  1436             by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def)
  2318     ultimately show ?thesis by (auto simp:readys_def)
  1437           from ih [OF th'_d this]
  2319   qed
  1438           obtain th'' where 
  2320   moreover have "(node, Th th') \<in> (RAG s)^+" 
  1439             th''_r: "th'' \<in> readys s" and 
  2321     using h_b(1)[unfolded trancl_converse] eq_b by auto
  1440             th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto
  2322   ultimately show ?thesis using that by metis
  1441           from th'_d and th''_in 
       
  1442           have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto
       
  1443           with th''_r show ?thesis by auto
       
  1444         qed
       
  1445       qed
       
  1446     qed
       
  1447   qed
       
  1448 qed
  2323 qed
  1449 
  2324 
  1450 text {* \noindent
  2325 text {* \noindent
  1451   The following is just an instance of @{text "chain_building"}.
  2326   The following is just an instance of @{text "chain_building"}.
  1452 *}
  2327 *}
  1464   show ?thesis by auto
  2339   show ?thesis by auto
  1465 qed
  2340 qed
  1466 
  2341 
  1467 end
  2342 end
  1468 
  2343 
  1469 lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
       
  1470   by  (unfold s_waiting_def cs_waiting_def wq_def, auto)
       
  1471 
       
  1472 lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
       
  1473   by (unfold s_holding_def wq_def cs_holding_def, simp)
       
  1474 
       
  1475 lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
       
  1476   by (unfold s_holding_def cs_holding_def, auto)
       
  1477 
       
  1478 context valid_trace
       
  1479 begin
       
  1480 
       
  1481 lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2"
       
  1482   apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq)
       
  1483   by(auto elim:waiting_unique holding_unique)
       
  1484 
       
  1485 end
       
  1486 
       
  1487 
       
  1488 lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
       
  1489 by (induct rule:trancl_induct, auto)
       
  1490 
       
  1491 context valid_trace
       
  1492 begin
       
  1493 
       
  1494 lemma dchain_unique:
       
  1495   assumes th1_d: "(n, Th th1) \<in> (RAG s)^+"
       
  1496   and th1_r: "th1 \<in> readys s"
       
  1497   and th2_d: "(n, Th th2) \<in> (RAG s)^+"
       
  1498   and th2_r: "th2 \<in> readys s"
       
  1499   shows "th1 = th2"
       
  1500 proof -
       
  1501   { assume neq: "th1 \<noteq> th2"
       
  1502     hence "Th th1 \<noteq> Th th2" by simp
       
  1503     from unique_chain [OF _ th1_d th2_d this] and unique_RAG 
       
  1504     have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto
       
  1505     hence "False"
       
  1506     proof
       
  1507       assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+"
       
  1508       from trancl_split [OF this]
       
  1509       obtain n where dd: "(Th th1, n) \<in> RAG s" by auto
       
  1510       then obtain cs where eq_n: "n = Cs cs"
       
  1511         by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1512       from dd eq_n have "th1 \<notin> readys s"
       
  1513         by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def)
       
  1514       with th1_r show ?thesis by auto
       
  1515     next
       
  1516       assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+"
       
  1517       from trancl_split [OF this]
       
  1518       obtain n where dd: "(Th th2, n) \<in> RAG s" by auto
       
  1519       then obtain cs where eq_n: "n = Cs cs"
       
  1520         by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
       
  1521       from dd eq_n have "th2 \<notin> readys s"
       
  1522         by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
       
  1523       with th2_r show ?thesis by auto
       
  1524     qed
       
  1525   } thus ?thesis by auto
       
  1526 qed
       
  1527 
       
  1528 end
       
  1529              
       
  1530 
       
  1531 lemma step_holdents_p_add:
       
  1532   assumes vt: "vt (P th cs#s)"
       
  1533   and "wq s cs = []"
       
  1534   shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
       
  1535 proof -
       
  1536   from assms show ?thesis
       
  1537   unfolding  holdents_test step_RAG_p[OF vt] by (auto)
       
  1538 qed
       
  1539 
       
  1540 lemma step_holdents_p_eq:
       
  1541   assumes vt: "vt (P th cs#s)"
       
  1542   and "wq s cs \<noteq> []"
       
  1543   shows "holdents (P th cs#s) th = holdents s th"
       
  1544 proof -
       
  1545   from assms show ?thesis
       
  1546   unfolding  holdents_test step_RAG_p[OF vt] by auto
       
  1547 qed
       
  1548 
       
  1549 
       
  1550 lemma (in valid_trace) finite_holding :
       
  1551   shows "finite (holdents s th)"
       
  1552 proof -
       
  1553   let ?F = "\<lambda> (x, y). the_cs x"
       
  1554   from finite_RAG 
       
  1555   have "finite (RAG s)" .
       
  1556   hence "finite (?F `(RAG s))" by simp
       
  1557   moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" 
       
  1558   proof -
       
  1559     { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
       
  1560       fix x assume "(Cs x, Th th) \<in> RAG s"
       
  1561       hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h)
       
  1562       moreover have "?F (Cs x, Th th) = x" by simp
       
  1563       ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp 
       
  1564     } thus ?thesis by auto
       
  1565   qed
       
  1566   ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset)
       
  1567 qed
       
  1568 
       
  1569 lemma cntCS_v_dec: 
       
  1570   assumes vtv: "vt (V thread cs#s)"
       
  1571   shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
       
  1572 proof -
       
  1573   from vtv interpret vt_s: valid_trace s
       
  1574     by (cases, unfold_locales, simp)
       
  1575   from vtv interpret vt_v: valid_trace "V thread cs#s"
       
  1576      by (unfold_locales, simp)
       
  1577   from step_back_step[OF vtv]
       
  1578   have cs_in: "cs \<in> holdents s thread" 
       
  1579     apply (cases, unfold holdents_test s_RAG_def, simp)
       
  1580     by (unfold cs_holding_def s_holding_def wq_def, auto)
       
  1581   moreover have cs_not_in: 
       
  1582     "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
       
  1583     apply (insert vt_s.wq_distinct[of cs])
       
  1584     apply (unfold holdents_test, unfold step_RAG_v[OF vtv],
       
  1585             auto simp:next_th_def)
       
  1586   proof -
       
  1587     fix rest
       
  1588     assume dst: "distinct (rest::thread list)"
       
  1589       and ne: "rest \<noteq> []"
       
  1590     and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1591     moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1592     proof(rule someI2)
       
  1593       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1594     next
       
  1595       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1596     qed
       
  1597     ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1598                      set (SOME q. distinct q \<and> set q = set rest)" by simp
       
  1599     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1600     proof(rule someI2)
       
  1601       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1602     next
       
  1603       fix x assume " distinct x \<and> set x = set rest" with ne
       
  1604       show "x \<noteq> []" by auto
       
  1605     qed
       
  1606     ultimately 
       
  1607     show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
       
  1608       by auto
       
  1609   next
       
  1610     fix rest
       
  1611     assume dst: "distinct (rest::thread list)"
       
  1612       and ne: "rest \<noteq> []"
       
  1613     and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1614     moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest"
       
  1615     proof(rule someI2)
       
  1616       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1617     next
       
  1618       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
  1619     qed
       
  1620     ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1621                      set (SOME q. distinct q \<and> set q = set rest)" by simp
       
  1622     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1623     proof(rule someI2)
       
  1624       from dst show "distinct rest \<and> set rest = set rest" by auto
       
  1625     next
       
  1626       fix x assume " distinct x \<and> set x = set rest" with ne
       
  1627       show "x \<noteq> []" by auto
       
  1628     qed
       
  1629     ultimately show "False" by auto 
       
  1630   qed
       
  1631   ultimately 
       
  1632   have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
       
  1633     by auto
       
  1634   moreover have "card \<dots> = 
       
  1635                     Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
       
  1636   proof(rule card_insert)
       
  1637     from vt_v.finite_holding
       
  1638     show " finite (holdents (V thread cs # s) thread)" .
       
  1639   qed
       
  1640   moreover from cs_not_in 
       
  1641   have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
       
  1642   ultimately show ?thesis by (simp add:cntCS_def)
       
  1643 qed 
       
  1644 
       
  1645 lemma count_rec1 [simp]: 
  2344 lemma count_rec1 [simp]: 
  1646   assumes "Q e"
  2345   assumes "Q e"
  1647   shows "count Q (e#es) = Suc (count Q es)"
  2346   shows "count Q (e#es) = Suc (count Q es)"
  1648   using assms
  2347   using assms
  1649   by (unfold count_def, auto)
  2348   by (unfold count_def, auto)
  1655   by (unfold count_def, auto)
  2354   by (unfold count_def, auto)
  1656 
  2355 
  1657 lemma count_rec3 [simp]: 
  2356 lemma count_rec3 [simp]: 
  1658   shows "count Q [] =  0"
  2357   shows "count Q [] =  0"
  1659   by (unfold count_def, auto)
  2358   by (unfold count_def, auto)
  1660   
  2359 
       
  2360 lemma cntP_simp1[simp]:
       
  2361   "cntP (P th cs'#s) th = cntP s th + 1"
       
  2362   by (unfold cntP_def, simp)
       
  2363 
       
  2364 lemma cntP_simp2[simp]:
       
  2365   assumes "th' \<noteq> th"
       
  2366   shows "cntP (P th cs'#s) th' = cntP s th'"
       
  2367   using assms
       
  2368   by (unfold cntP_def, simp)
       
  2369 
       
  2370 lemma cntP_simp3[simp]:
       
  2371   assumes "\<not> isP e"
       
  2372   shows "cntP (e#s) th' = cntP s th'"
       
  2373   using assms
       
  2374   by (unfold cntP_def, cases e, simp+)
       
  2375 
       
  2376 lemma cntV_simp1[simp]:
       
  2377   "cntV (V th cs'#s) th = cntV s th + 1"
       
  2378   by (unfold cntV_def, simp)
       
  2379 
       
  2380 lemma cntV_simp2[simp]:
       
  2381   assumes "th' \<noteq> th"
       
  2382   shows "cntV (V th cs'#s) th' = cntV s th'"
       
  2383   using assms
       
  2384   by (unfold cntV_def, simp)
       
  2385 
       
  2386 lemma cntV_simp3[simp]:
       
  2387   assumes "\<not> isV e"
       
  2388   shows "cntV (e#s) th' = cntV s th'"
       
  2389   using assms
       
  2390   by (unfold cntV_def, cases e, simp+)
       
  2391 
  1661 lemma cntP_diff_inv:
  2392 lemma cntP_diff_inv:
  1662   assumes "cntP (e#s) th \<noteq> cntP s th"
  2393   assumes "cntP (e#s) th \<noteq> cntP s th"
  1663   shows "isP e \<and> actor e = th"
  2394   shows "isP e \<and> actor e = th"
  1664 proof(cases e)
  2395 proof(cases e)
  1665   case (P th' pty)
  2396   case (P th' pty)
  1666   show ?thesis
  2397   show ?thesis
  1667   by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
  2398   by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", 
  1668         insert assms P, auto simp:cntP_def)
  2399         insert assms P, auto simp:cntP_def)
  1669 qed (insert assms, auto simp:cntP_def)
  2400 qed (insert assms, auto simp:cntP_def)
  1670 
  2401   
  1671 lemma isP_E:
       
  1672   assumes "isP e"
       
  1673   obtains cs where "e = P (actor e) cs"
       
  1674   using assms by (cases e, auto)
       
  1675 
       
  1676 lemma isV_E:
       
  1677   assumes "isV e"
       
  1678   obtains cs where "e = V (actor e) cs"
       
  1679   using assms by (cases e, auto) (* ccc *)
       
  1680 
       
  1681 lemma cntV_diff_inv:
  2402 lemma cntV_diff_inv:
  1682   assumes "cntV (e#s) th \<noteq> cntV s th"
  2403   assumes "cntV (e#s) th \<noteq> cntV s th"
  1683   shows "isV e \<and> actor e = th"
  2404   shows "isV e \<and> actor e = th"
  1684 proof(cases e)
  2405 proof(cases e)
  1685   case (V th' pty)
  2406   case (V th' pty)
  1686   show ?thesis
  2407   show ?thesis
  1687   by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
  2408   by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", 
  1688         insert assms V, auto simp:cntV_def)
  2409         insert assms V, auto simp:cntV_def)
  1689 qed (insert assms, auto simp:cntV_def)
  2410 qed (insert assms, auto simp:cntV_def)
  1690 
  2411 
       
  2412 lemma children_RAG_alt_def:
       
  2413   "children (RAG (s::state)) (Th th) = Cs ` {cs. holding s th cs}"
       
  2414   by (unfold s_RAG_def, auto simp:children_def holding_eq)
       
  2415 
       
  2416 lemma holdents_alt_def:
       
  2417   "holdents s th = the_cs ` (children (RAG (s::state)) (Th th))"
       
  2418   by (unfold children_RAG_alt_def holdents_def, simp add: image_image)
       
  2419 
       
  2420 lemma cntCS_alt_def:
       
  2421   "cntCS s th = card (children (RAG s) (Th th))"
       
  2422   apply (unfold children_RAG_alt_def cntCS_def holdents_def)
       
  2423   by (rule card_image[symmetric], auto simp:inj_on_def)
       
  2424 
  1691 context valid_trace
  2425 context valid_trace
  1692 begin
  2426 begin
  1693 
  2427 
  1694 text {* (* ddd *) \noindent
  2428 lemma finite_holdents: "finite (holdents s th)"
  1695   The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} 
  2429   by (unfold holdents_alt_def, insert fsbtRAGs.finite_children, auto)
  1696   of one particular thread. 
  2430   
  1697 *} 
  2431 end
  1698 
  2432 
  1699 lemma cnp_cnv_cncs:
  2433 context valid_trace_p_w
  1700   shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
  2434 begin
  1701                                        then cntCS s th else cntCS s th + 1)"
  2435 
  1702 proof -
  2436 lemma holding_s_holder: "holding s holder cs"
  1703   from vt show ?thesis
  2437   by (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
  1704   proof(induct arbitrary:th)
  2438 
  1705     case (vt_cons s e)
  2439 lemma holding_es_holder: "holding (e#s) holder cs"
  1706     interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp)
  2440   by (unfold s_holding_def, fold wq_def, unfold wq_es_cs wq_s_cs, auto)
  1707     assume vt: "vt s"
  2441 
  1708     and ih: "\<And>th. cntP s th  = cntV s th +
  2442 lemma holdents_es:
  1709                (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
  2443   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") 
  1710     and stp: "step s e"
  2444 proof -
  1711     from stp show ?case
  2445   { fix cs'
  1712     proof(cases)
  2446     assume "cs' \<in> ?L"
  1713       case (thread_create thread prio)
  2447     hence h: "holding (e#s) th' cs'" by (auto simp:holdents_def)
  1714       assume eq_e: "e = Create thread prio"
  2448     have "holding s th' cs'"
  1715         and not_in: "thread \<notin> threads s"
  2449     proof(cases "cs' = cs")
       
  2450       case True
       
  2451       from held_unique[OF h[unfolded True] holding_es_holder]
       
  2452       have "th' = holder" .
       
  2453       thus ?thesis 
       
  2454         by (unfold True holdents_def, insert holding_s_holder, simp)
       
  2455     next
       
  2456       case False
       
  2457       hence "wq (e#s) cs' = wq s cs'" by simp
       
  2458       from h[unfolded s_holding_def, folded wq_def, unfolded this]
  1716       show ?thesis
  2459       show ?thesis
  1717       proof -
  2460        by (unfold s_holding_def, fold wq_def, auto)
  1718         { fix cs 
  2461     qed 
  1719           assume "thread \<in> set (wq s cs)"
  2462     hence "cs' \<in> ?R" by (auto simp:holdents_def)
  1720           from vt_s.wq_threads [OF this] have "thread \<in> threads s" .
  2463   } moreover {
  1721           with not_in have "False" by simp
  2464     fix cs'
  1722         } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
  2465     assume "cs' \<in> ?R"
  1723           by (auto simp:readys_def threads.simps s_waiting_def 
  2466     hence h: "holding s th' cs'" by (auto simp:holdents_def)
  1724             wq_def cs_waiting_def Let_def)
  2467     have "holding (e#s) th' cs'"
  1725         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  2468     proof(cases "cs' = cs")
  1726         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  2469       case True
  1727         have eq_cncs: "cntCS (e#s) th = cntCS s th"
  2470       from held_unique[OF h[unfolded True] holding_s_holder]
  1728           unfolding cntCS_def holdents_test
  2471       have "th' = holder" .
  1729           by (simp add:RAG_create_unchanged eq_e)
  2472       thus ?thesis 
  1730         { assume "th \<noteq> thread"
  2473         by (unfold True holdents_def, insert holding_es_holder, simp)
  1731           with eq_readys eq_e
  2474     next
  1732           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
  2475       case False
  1733                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
  2476       hence "wq s cs' = wq (e#s) cs'" by simp
  1734             by (simp add:threads.simps)
  2477       from h[unfolded s_holding_def, folded wq_def, unfolded this]
  1735           with eq_cnp eq_cnv eq_cncs ih not_in
  2478       show ?thesis
  1736           have ?thesis by simp
  2479        by (unfold s_holding_def, fold wq_def, auto)
  1737         } moreover {
  2480     qed 
  1738           assume eq_th: "th = thread"
  2481     hence "cs' \<in> ?L" by (auto simp:holdents_def)
  1739           with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
  2482   } ultimately show ?thesis by auto
  1740           moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
  2483 qed
  1741           moreover note eq_cnp eq_cnv eq_cncs
  2484 
  1742           ultimately have ?thesis by auto
  2485 lemma cntCS_es_th[simp]: "cntCS (e#s) th' = cntCS s th'"
  1743         } ultimately show ?thesis by blast
  2486  by (unfold cntCS_def holdents_es, simp)
       
  2487 
       
  2488 lemma th_not_ready_es: 
       
  2489   shows "th \<notin> readys (e#s)"
       
  2490   using waiting_es_th_cs 
       
  2491   by (unfold readys_def, auto)
       
  2492 
       
  2493 end
       
  2494   
       
  2495 context valid_trace_p_h
       
  2496 begin
       
  2497 
       
  2498 lemma th_not_waiting':
       
  2499   "\<not> waiting (e#s) th cs'"
       
  2500 proof(cases "cs' = cs")
       
  2501   case True
       
  2502   show ?thesis
       
  2503     by (unfold True s_waiting_def, fold wq_def, unfold wq_es_cs', auto)
       
  2504 next
       
  2505   case False
       
  2506   from th_not_waiting[of cs', unfolded s_waiting_def, folded wq_def]
       
  2507   show ?thesis
       
  2508     by (unfold s_waiting_def, fold wq_def, insert False, simp)
       
  2509 qed
       
  2510 
       
  2511 lemma ready_th_es: 
       
  2512   shows "th \<in> readys (e#s)"
       
  2513   using th_not_waiting'
       
  2514   by (unfold readys_def, insert live_th_es, auto)
       
  2515 
       
  2516 lemma holdents_es_th:
       
  2517   "holdents (e#s) th = (holdents s th) \<union> {cs}" (is "?L = ?R")
       
  2518 proof -
       
  2519   { fix cs'
       
  2520     assume "cs' \<in> ?L" 
       
  2521     hence "holding (e#s) th cs'"
       
  2522       by (unfold holdents_def, auto)
       
  2523     hence "cs' \<in> ?R"
       
  2524      by (cases rule:holding_esE, auto simp:holdents_def)
       
  2525   } moreover {
       
  2526     fix cs'
       
  2527     assume "cs' \<in> ?R"
       
  2528     hence "holding s th cs' \<or> cs' = cs" 
       
  2529       by (auto simp:holdents_def)
       
  2530     hence "cs' \<in> ?L"
       
  2531     proof
       
  2532       assume "holding s th cs'"
       
  2533       from holding_kept[OF this]
       
  2534       show ?thesis by (auto simp:holdents_def)
       
  2535     next
       
  2536       assume "cs' = cs"
       
  2537       thus ?thesis using holding_es_th_cs
       
  2538         by (unfold holdents_def, auto)
       
  2539     qed
       
  2540   } ultimately show ?thesis by auto
       
  2541 qed
       
  2542 
       
  2543 lemma cntCS_es_th: "cntCS (e#s) th = cntCS s th + 1"
       
  2544 proof -
       
  2545   have "card (holdents s th \<union> {cs}) = card (holdents s th) + 1"
       
  2546   proof(subst card_Un_disjoint)
       
  2547     show "holdents s th \<inter> {cs} = {}"
       
  2548       using not_holding_s_th_cs by (auto simp:holdents_def)
       
  2549   qed (auto simp:finite_holdents)
       
  2550   thus ?thesis
       
  2551    by (unfold cntCS_def holdents_es_th, simp)
       
  2552 qed
       
  2553 
       
  2554 lemma no_holder: 
       
  2555   "\<not> holding s th' cs"
       
  2556 proof
       
  2557   assume otherwise: "holding s th' cs"
       
  2558   from this[unfolded s_holding_def, folded wq_def, unfolded we]
       
  2559   show False by auto
       
  2560 qed
       
  2561 
       
  2562 lemma holdents_es_th':
       
  2563   assumes "th' \<noteq> th"
       
  2564   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  2565 proof -
       
  2566   { fix cs'
       
  2567     assume "cs' \<in> ?L"
       
  2568     hence h_e: "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  2569     have "cs' \<noteq> cs"
       
  2570     proof
       
  2571       assume "cs' = cs"
       
  2572       from held_unique[OF h_e[unfolded this] holding_es_th_cs]
       
  2573       have "th' = th" .
       
  2574       with assms show False by simp
       
  2575     qed
       
  2576     from h_e[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp[OF this]]
       
  2577     have "th' \<in> set (wq s cs') \<and> th' = hd (wq s cs')" .
       
  2578     hence "cs' \<in> ?R" 
       
  2579       by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  2580   } moreover {
       
  2581     fix cs'
       
  2582     assume "cs' \<in> ?R"
       
  2583     hence "holding s th' cs'" by (auto simp:holdents_def)
       
  2584     from holding_kept[OF this]
       
  2585     have "holding (e # s) th' cs'" .
       
  2586     hence "cs' \<in> ?L"
       
  2587       by (unfold holdents_def, auto)
       
  2588   } ultimately show ?thesis by auto
       
  2589 qed
       
  2590 
       
  2591 lemma cntCS_es_th'[simp]: 
       
  2592   assumes "th' \<noteq> th"
       
  2593   shows "cntCS (e#s) th' = cntCS s th'"
       
  2594   by (unfold cntCS_def holdents_es_th'[OF assms], simp)
       
  2595 
       
  2596 end
       
  2597 
       
  2598 context valid_trace_p
       
  2599 begin
       
  2600 
       
  2601 lemma readys_kept1: 
       
  2602   assumes "th' \<noteq> th"
       
  2603   and "th' \<in> readys (e#s)"
       
  2604   shows "th' \<in> readys s"
       
  2605 proof -
       
  2606   { fix cs'
       
  2607     assume wait: "waiting s th' cs'"
       
  2608     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  2609         using assms(2)[unfolded readys_def] by auto
       
  2610     have False
       
  2611     proof(cases "cs' = cs")
       
  2612       case False
       
  2613       with n_wait wait
       
  2614       show ?thesis 
       
  2615         by (unfold s_waiting_def, fold wq_def, auto)
       
  2616     next
       
  2617       case True
       
  2618       show ?thesis
       
  2619       proof(cases "wq s cs = []")
       
  2620         case True
       
  2621         then interpret vt: valid_trace_p_h
       
  2622           by (unfold_locales, simp)
       
  2623         show ?thesis using n_wait wait waiting_kept by auto 
       
  2624       next
       
  2625         case False
       
  2626         then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  2627         show ?thesis using n_wait wait waiting_kept by blast 
  1744       qed
  2628       qed
       
  2629     qed
       
  2630   } with assms(2) show ?thesis  
       
  2631     by (unfold readys_def, auto)
       
  2632 qed
       
  2633 
       
  2634 lemma readys_kept2: 
       
  2635   assumes "th' \<noteq> th"
       
  2636   and "th' \<in> readys s"
       
  2637   shows "th' \<in> readys (e#s)"
       
  2638 proof -
       
  2639   { fix cs'
       
  2640     assume wait: "waiting (e#s) th' cs'"
       
  2641     have n_wait: "\<not> waiting s th' cs'" 
       
  2642         using assms(2)[unfolded readys_def] by auto
       
  2643     have False
       
  2644     proof(cases "cs' = cs")
       
  2645       case False
       
  2646       with n_wait wait
       
  2647       show ?thesis 
       
  2648         by (unfold s_waiting_def, fold wq_def, auto)
  1745     next
  2649     next
  1746       case (thread_exit thread)
  2650       case True
  1747       assume eq_e: "e = Exit thread" 
  2651       show ?thesis
  1748       and is_runing: "thread \<in> runing s"
  2652       proof(cases "wq s cs = []")
  1749       and no_hold: "holdents s thread = {}"
  2653         case True
  1750       from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  2654         then interpret vt: valid_trace_p_h
  1751       from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  2655           by (unfold_locales, simp)
  1752       have eq_cncs: "cntCS (e#s) th = cntCS s th"
  2656         show ?thesis using n_wait vt.waiting_esE wait by blast 
  1753         unfolding cntCS_def holdents_test
  2657       next
  1754         by (simp add:RAG_exit_unchanged eq_e)
  2658         case False
  1755       { assume "th \<noteq> thread"
  2659         then interpret vt: valid_trace_p_w by (unfold_locales, simp)
  1756         with eq_e
  2660         show ?thesis using assms(1) n_wait vt.waiting_esE wait by auto 
  1757         have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
  2661       qed
  1758           (th \<in> readys (s) \<or> th \<notin> threads (s))" 
  2662     qed
  1759           apply (simp add:threads.simps readys_def)
  2663   } with assms(2) show ?thesis  
  1760           apply (subst s_waiting_def)
  2664     by (unfold readys_def, auto)
  1761           apply (simp add:Let_def)
  2665 qed
  1762           apply (subst s_waiting_def, simp)
  2666 
  1763           done
  2667 lemma readys_simp [simp]:
  1764         with eq_cnp eq_cnv eq_cncs ih
  2668   assumes "th' \<noteq> th"
  1765         have ?thesis by simp
  2669   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
  1766       } moreover {
  2670   using readys_kept1[OF assms] readys_kept2[OF assms]
  1767         assume eq_th: "th = thread"
  2671   by metis
  1768         with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
  2672 
  1769           by (simp add:runing_def)
  2673 lemma cnp_cnv_cncs_kept: (* ddd *)
  1770         moreover from eq_th eq_e have "th \<notin> threads (e#s)"
  2674   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
  1771           by simp
  2675   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
  1772         moreover note eq_cnp eq_cnv eq_cncs
  2676 proof(cases "th' = th")
  1773         ultimately have ?thesis by auto
  2677   case True
  1774       } ultimately show ?thesis by blast
  2678   note eq_th' = this
       
  2679   show ?thesis
       
  2680   proof(cases "wq s cs = []")
       
  2681     case True
       
  2682     then interpret vt: valid_trace_p_h by (unfold_locales, simp)
       
  2683     show ?thesis
       
  2684       using assms eq_th' is_p ready_th_s vt.cntCS_es_th vt.ready_th_es pvD_def by auto 
       
  2685   next
       
  2686     case False
       
  2687     then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  2688     show ?thesis
       
  2689       using add.commute add.left_commute assms eq_th' is_p live_th_s 
       
  2690             ready_th_s vt.th_not_ready_es pvD_def
       
  2691       apply (auto)
       
  2692       by (fold is_p, simp)
       
  2693   qed
       
  2694 next
       
  2695   case False
       
  2696   note h_False = False
       
  2697   thus ?thesis
       
  2698   proof(cases "wq s cs = []")
       
  2699     case True
       
  2700     then interpret vt: valid_trace_p_h by (unfold_locales, simp)
       
  2701     show ?thesis using assms
       
  2702       by (insert True h_False pvD_def, auto split:if_splits,unfold is_p, auto)
       
  2703   next
       
  2704     case False
       
  2705     then interpret vt: valid_trace_p_w by (unfold_locales, simp)
       
  2706     show ?thesis using assms
       
  2707       by (insert False h_False pvD_def, auto split:if_splits,unfold is_p, auto)
       
  2708   qed
       
  2709 qed
       
  2710 
       
  2711 end
       
  2712 
       
  2713 
       
  2714 context valid_trace_v (* ccc *)
       
  2715 begin
       
  2716 
       
  2717 lemma holding_th_cs_s: 
       
  2718   "holding s th cs" 
       
  2719  by  (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto)
       
  2720 
       
  2721 lemma th_ready_s [simp]: "th \<in> readys s"
       
  2722   using runing_th_s
       
  2723   by (unfold runing_def readys_def, auto)
       
  2724 
       
  2725 lemma th_live_s [simp]: "th \<in> threads s"
       
  2726   using th_ready_s by (unfold readys_def, auto)
       
  2727 
       
  2728 lemma th_ready_es [simp]: "th \<in> readys (e#s)"
       
  2729   using runing_th_s neq_t_th
       
  2730   by (unfold is_v runing_def readys_def, auto)
       
  2731 
       
  2732 lemma th_live_es [simp]: "th \<in> threads (e#s)"
       
  2733   using th_ready_es by (unfold readys_def, auto)
       
  2734 
       
  2735 lemma pvD_th_s[simp]: "pvD s th = 0"
       
  2736   by (unfold pvD_def, simp)
       
  2737 
       
  2738 lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
       
  2739   by (unfold pvD_def, simp)
       
  2740 
       
  2741 lemma cntCS_s_th [simp]: "cntCS s th > 0"
       
  2742 proof -
       
  2743   have "cs \<in> holdents s th" using holding_th_cs_s
       
  2744     by (unfold holdents_def, simp)
       
  2745   moreover have "finite (holdents s th)" using finite_holdents
       
  2746     by simp
       
  2747   ultimately show ?thesis
       
  2748     by (unfold cntCS_def, 
       
  2749         auto intro!:card_gt_0_iff[symmetric, THEN iffD1])
       
  2750 qed
       
  2751 
       
  2752 end
       
  2753 
       
  2754 context valid_trace_v_n
       
  2755 begin
       
  2756 
       
  2757 lemma not_ready_taker_s[simp]: 
       
  2758   "taker \<notin> readys s"
       
  2759   using waiting_taker
       
  2760   by (unfold readys_def, auto)
       
  2761 
       
  2762 lemma taker_live_s [simp]: "taker \<in> threads s"
       
  2763 proof -
       
  2764   have "taker \<in> set wq'" by (simp add: eq_wq') 
       
  2765   from th'_in_inv[OF this]
       
  2766   have "taker \<in> set rest" .
       
  2767   hence "taker \<in> set (wq s cs)" by (simp add: wq_s_cs) 
       
  2768   thus ?thesis using wq_threads by auto 
       
  2769 qed
       
  2770 
       
  2771 lemma taker_live_es [simp]: "taker \<in> threads (e#s)"
       
  2772   using taker_live_s threads_es by blast
       
  2773 
       
  2774 lemma taker_ready_es [simp]:
       
  2775   shows "taker \<in> readys (e#s)"
       
  2776 proof -
       
  2777   { fix cs'
       
  2778     assume "waiting (e#s) taker cs'"
       
  2779     hence False
       
  2780     proof(cases rule:waiting_esE)
       
  2781       case 1
       
  2782       thus ?thesis using waiting_taker waiting_unique by auto 
       
  2783     qed simp
       
  2784   } thus ?thesis by (unfold readys_def, auto)
       
  2785 qed
       
  2786 
       
  2787 lemma neq_taker_th: "taker \<noteq> th"
       
  2788   using th_not_waiting waiting_taker by blast
       
  2789 
       
  2790 lemma not_holding_taker_s_cs:
       
  2791   shows "\<not> holding s taker cs"
       
  2792   using holding_cs_eq_th neq_taker_th by auto
       
  2793 
       
  2794 lemma holdents_es_taker:
       
  2795   "holdents (e#s) taker = holdents s taker \<union> {cs}" (is "?L = ?R")
       
  2796 proof -
       
  2797   { fix cs'
       
  2798     assume "cs' \<in> ?L"
       
  2799     hence "holding (e#s) taker cs'" by (auto simp:holdents_def)
       
  2800     hence "cs' \<in> ?R"
       
  2801     proof(cases rule:holding_esE)
       
  2802       case 2
       
  2803       thus ?thesis by (auto simp:holdents_def)
       
  2804     qed auto
       
  2805   } moreover {
       
  2806     fix cs'
       
  2807     assume "cs' \<in> ?R"
       
  2808     hence "holding s taker cs' \<or> cs' = cs" by (auto simp:holdents_def)
       
  2809     hence "cs' \<in> ?L" 
       
  2810     proof
       
  2811       assume "holding s taker cs'"
       
  2812       hence "holding (e#s) taker cs'" 
       
  2813           using holding_esI2 holding_taker by fastforce 
       
  2814       thus ?thesis by (auto simp:holdents_def)
  1775     next
  2815     next
  1776       case (thread_P thread cs)
  2816       assume "cs' = cs"
  1777       assume eq_e: "e = P thread cs"
  2817       with holding_taker
  1778         and is_runing: "thread \<in> runing s"
  2818       show ?thesis by (auto simp:holdents_def)
  1779         and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+"
  2819     qed
  1780       from thread_P vt stp ih  have vtp: "vt (P thread cs#s)" by auto
  2820   } ultimately show ?thesis by auto
  1781       then interpret vt_p: valid_trace "(P thread cs#s)"
  2821 qed
  1782         by (unfold_locales, simp)
  2822 
       
  2823 lemma cntCS_es_taker [simp]: "cntCS (e#s) taker = cntCS s taker + 1"
       
  2824 proof -
       
  2825   have "card (holdents s taker \<union> {cs}) = card (holdents s taker) + 1"
       
  2826   proof(subst card_Un_disjoint)
       
  2827     show "holdents s taker \<inter> {cs} = {}"
       
  2828       using not_holding_taker_s_cs by (auto simp:holdents_def)
       
  2829   qed (auto simp:finite_holdents)
       
  2830   thus ?thesis 
       
  2831     by (unfold cntCS_def, insert holdents_es_taker, simp)
       
  2832 qed
       
  2833 
       
  2834 lemma pvD_taker_s[simp]: "pvD s taker = 1"
       
  2835   by (unfold pvD_def, simp)
       
  2836 
       
  2837 lemma pvD_taker_es[simp]: "pvD (e#s) taker = 0"
       
  2838   by (unfold pvD_def, simp)  
       
  2839 
       
  2840 lemma pvD_th_s[simp]: "pvD s th = 0"
       
  2841   by (unfold pvD_def, simp)
       
  2842 
       
  2843 lemma pvD_th_es[simp]: "pvD (e#s) th = 0"
       
  2844   by (unfold pvD_def, simp)
       
  2845 
       
  2846 lemma holdents_es_th:
       
  2847   "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
       
  2848 proof -
       
  2849   { fix cs'
       
  2850     assume "cs' \<in> ?L"
       
  2851     hence "holding (e#s) th cs'" by (auto simp:holdents_def)
       
  2852     hence "cs' \<in> ?R"
       
  2853     proof(cases rule:holding_esE)
       
  2854       case 2
       
  2855       thus ?thesis by (auto simp:holdents_def)
       
  2856     qed (insert neq_taker_th, auto)
       
  2857   } moreover {
       
  2858     fix cs'
       
  2859     assume "cs' \<in> ?R"
       
  2860     hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
       
  2861     from holding_esI2[OF this]
       
  2862     have "cs' \<in> ?L" by (auto simp:holdents_def)
       
  2863   } ultimately show ?thesis by auto
       
  2864 qed
       
  2865 
       
  2866 lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
       
  2867 proof -
       
  2868   have "card (holdents s th - {cs}) = card (holdents s th) - 1"
       
  2869   proof -
       
  2870     have "cs \<in> holdents s th" using holding_th_cs_s
       
  2871       by (auto simp:holdents_def)
       
  2872     moreover have "finite (holdents s th)"
       
  2873         by (simp add: finite_holdents) 
       
  2874     ultimately show ?thesis by auto
       
  2875   qed
       
  2876   thus ?thesis by (unfold cntCS_def holdents_es_th)
       
  2877 qed
       
  2878 
       
  2879 lemma holdents_kept:
       
  2880   assumes "th' \<noteq> taker"
       
  2881   and "th' \<noteq> th"
       
  2882   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  2883 proof -
       
  2884   { fix cs'
       
  2885     assume h: "cs' \<in> ?L"
       
  2886     have "cs' \<in> ?R"
       
  2887     proof(cases "cs' = cs")
       
  2888       case False
       
  2889       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  2890       from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  2891       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  2892       show ?thesis
       
  2893         by (unfold holdents_def s_holding_def, fold wq_def, auto)
       
  2894     next
       
  2895       case True
       
  2896       from h[unfolded this]
       
  2897       have "holding (e#s) th' cs" by (auto simp:holdents_def)
       
  2898       from held_unique[OF this holding_taker]
       
  2899       have "th' = taker" .
       
  2900       with assms show ?thesis by auto
       
  2901     qed
       
  2902   } moreover {
       
  2903     fix cs'
       
  2904     assume h: "cs' \<in> ?R"
       
  2905     have "cs' \<in> ?L"
       
  2906     proof(cases "cs' = cs")
       
  2907       case False
       
  2908       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  2909       from h have "holding s th' cs'" by (auto simp:holdents_def)
       
  2910       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
       
  2911       show ?thesis
       
  2912         by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
       
  2913     next
       
  2914       case True
       
  2915       from h[unfolded this]
       
  2916       have "holding s th' cs" by (auto simp:holdents_def)
       
  2917       from held_unique[OF this holding_th_cs_s]
       
  2918       have "th' = th" .
       
  2919       with assms show ?thesis by auto
       
  2920     qed
       
  2921   } ultimately show ?thesis by auto
       
  2922 qed
       
  2923 
       
  2924 lemma cntCS_kept [simp]:
       
  2925   assumes "th' \<noteq> taker"
       
  2926   and "th' \<noteq> th"
       
  2927   shows "cntCS (e#s) th' = cntCS s th'"
       
  2928   by (unfold cntCS_def holdents_kept[OF assms], simp)
       
  2929 
       
  2930 lemma readys_kept1: 
       
  2931   assumes "th' \<noteq> taker"
       
  2932   and "th' \<in> readys (e#s)"
       
  2933   shows "th' \<in> readys s"
       
  2934 proof -
       
  2935   { fix cs'
       
  2936     assume wait: "waiting s th' cs'"
       
  2937     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  2938         using assms(2)[unfolded readys_def] by auto
       
  2939     have False
       
  2940     proof(cases "cs' = cs")
       
  2941       case False
       
  2942       with n_wait wait
  1783       show ?thesis 
  2943       show ?thesis 
  1784       proof -
  2944         by (unfold s_waiting_def, fold wq_def, auto)
  1785         { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
       
  1786           assume neq_th: "th \<noteq> thread"
       
  1787           with eq_e
       
  1788           have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
       
  1789             apply (simp add:readys_def s_waiting_def wq_def Let_def)
       
  1790             apply (rule_tac hh)
       
  1791              apply (intro iffI allI, clarify)
       
  1792             apply (erule_tac x = csa in allE, auto)
       
  1793             apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto)
       
  1794             apply (erule_tac x = cs in allE, auto)
       
  1795             by (case_tac "(wq_fun (schs s) cs)", auto)
       
  1796           moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
       
  1797             apply (simp add:cntCS_def holdents_test)
       
  1798             by (unfold  step_RAG_p [OF vtp], auto)
       
  1799           moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
       
  1800             by (simp add:cntP_def count_def)
       
  1801           moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
       
  1802             by (simp add:cntV_def count_def)
       
  1803           moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
       
  1804           moreover note ih [of th] 
       
  1805           ultimately have ?thesis by simp
       
  1806         } moreover {
       
  1807           assume eq_th: "th = thread"
       
  1808           have ?thesis
       
  1809           proof -
       
  1810             from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
       
  1811               by (simp add:cntP_def count_def)
       
  1812             from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
       
  1813               by (simp add:cntV_def count_def)
       
  1814             show ?thesis
       
  1815             proof (cases "wq s cs = []")
       
  1816               case True
       
  1817               with is_runing
       
  1818               have "th \<in> readys (e#s)"
       
  1819                 apply (unfold eq_e wq_def, unfold readys_def s_RAG_def)
       
  1820                 apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
       
  1821                 by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
       
  1822               moreover have "cntCS (e # s) th = 1 + cntCS s th"
       
  1823               proof -
       
  1824                 have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} =
       
  1825                   Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)")
       
  1826                 proof -
       
  1827                   have "?L = insert cs ?R" by auto
       
  1828                   moreover have "card \<dots> = Suc (card (?R - {cs}))" 
       
  1829                   proof(rule card_insert)
       
  1830                     from vt_s.finite_holding [of thread]
       
  1831                     show " finite {cs. (Cs cs, Th thread) \<in> RAG s}"
       
  1832                       by (unfold holdents_test, simp)
       
  1833                   qed
       
  1834                   moreover have "?R - {cs} = ?R"
       
  1835                   proof -
       
  1836                     have "cs \<notin> ?R"
       
  1837                     proof
       
  1838                       assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}"
       
  1839                       with no_dep show False by auto
       
  1840                     qed
       
  1841                     thus ?thesis by auto
       
  1842                   qed
       
  1843                   ultimately show ?thesis by auto
       
  1844                 qed
       
  1845                 thus ?thesis
       
  1846                   apply (unfold eq_e eq_th cntCS_def)
       
  1847                   apply (simp add: holdents_test)
       
  1848                   by (unfold step_RAG_p [OF vtp], auto simp:True)
       
  1849               qed
       
  1850               moreover from is_runing have "th \<in> readys s"
       
  1851                 by (simp add:runing_def eq_th)
       
  1852               moreover note eq_cnp eq_cnv ih [of th]
       
  1853               ultimately show ?thesis by auto
       
  1854             next
       
  1855               case False
       
  1856               have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
       
  1857                     by (unfold eq_th eq_e wq_def, auto simp:Let_def)
       
  1858               have "th \<notin> readys (e#s)"
       
  1859               proof
       
  1860                 assume "th \<in> readys (e#s)"
       
  1861                 hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
       
  1862                 from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
       
  1863                 hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
       
  1864                   by (simp add:s_waiting_def wq_def)
       
  1865                 moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
       
  1866                 ultimately have "th = hd (wq (e#s) cs)" by blast
       
  1867                 with eq_wq have "th = hd (wq s cs @ [th])" by simp
       
  1868                 hence "th = hd (wq s cs)" using False by auto
       
  1869                 with False eq_wq vt_p.wq_distinct [of cs]
       
  1870                 show False by (fold eq_e, auto)
       
  1871               qed
       
  1872               moreover from is_runing have "th \<in> threads (e#s)" 
       
  1873                 by (unfold eq_e, auto simp:runing_def readys_def eq_th)
       
  1874               moreover have "cntCS (e # s) th = cntCS s th"
       
  1875                 apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp])
       
  1876                 by (auto simp:False)
       
  1877               moreover note eq_cnp eq_cnv ih[of th]
       
  1878               moreover from is_runing have "th \<in> readys s"
       
  1879                 by (simp add:runing_def eq_th)
       
  1880               ultimately show ?thesis by auto
       
  1881             qed
       
  1882           qed
       
  1883         } ultimately show ?thesis by blast
       
  1884       qed
       
  1885     next
  2945     next
  1886       case (thread_V thread cs)
  2946       case True
  1887       from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto
  2947       have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
  1888       then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp)
  2948         using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
  1889       assume eq_e: "e = V thread cs"
  2949       moreover have "\<not> (th' \<in> set rest \<and> th' \<noteq> hd (taker # rest'))" 
  1890         and is_runing: "thread \<in> runing s"
  2950         using n_wait[unfolded True s_waiting_def, folded wq_def, 
  1891         and hold: "holding s thread cs"
  2951                     unfolded wq_es_cs set_wq', unfolded eq_wq'] .
  1892       from hold obtain rest 
  2952       ultimately have "th' = taker" by auto
  1893         where eq_wq: "wq s cs = thread # rest"
  2953       with assms(1)
  1894         by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
  2954       show ?thesis by simp
  1895       have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
  2955     qed
  1896       have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest"
  2956   } with assms(2) show ?thesis  
  1897       proof(rule someI2)
  2957     by (unfold readys_def, auto)
  1898         from vt_v.wq_distinct[of cs] and eq_wq
  2958 qed
  1899         show "distinct rest \<and> set rest = set rest"
  2959 
  1900           by (metis distinct.simps(2) vt_s.wq_distinct)
  2960 lemma readys_kept2: 
  1901       next
  2961   assumes "th' \<noteq> taker"
  1902         show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest"
  2962   and "th' \<in> readys s"
  1903           by auto
  2963   shows "th' \<in> readys (e#s)"
  1904       qed
  2964 proof -
       
  2965   { fix cs'
       
  2966     assume wait: "waiting (e#s) th' cs'"
       
  2967     have n_wait: "\<not> waiting s th' cs'" 
       
  2968         using assms(2)[unfolded readys_def] by auto
       
  2969     have False
       
  2970     proof(cases "cs' = cs")
       
  2971       case False
       
  2972       with n_wait wait
       
  2973       show ?thesis 
       
  2974         by (unfold s_waiting_def, fold wq_def, auto)
       
  2975     next
       
  2976       case True
       
  2977       have "th' \<in> set rest \<and> th' \<noteq> hd (taker # rest')"
       
  2978           using  wait [unfolded True s_waiting_def, folded wq_def, 
       
  2979                     unfolded wq_es_cs set_wq', unfolded eq_wq']  .
       
  2980       moreover have "\<not> (th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest))"
       
  2981           using n_wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] .
       
  2982       ultimately have "th' = taker" by auto
       
  2983       with assms(1)
       
  2984       show ?thesis by simp
       
  2985     qed
       
  2986   } with assms(2) show ?thesis  
       
  2987     by (unfold readys_def, auto)
       
  2988 qed
       
  2989 
       
  2990 lemma readys_simp [simp]:
       
  2991   assumes "th' \<noteq> taker"
       
  2992   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  2993   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  2994   by metis
       
  2995 
       
  2996 lemma cnp_cnv_cncs_kept:
       
  2997   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  2998   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  2999 proof -
       
  3000   { assume eq_th': "th' = taker"
       
  3001     have ?thesis
       
  3002       apply (unfold eq_th' pvD_taker_es cntCS_es_taker)
       
  3003       by (insert neq_taker_th assms[unfolded eq_th'], unfold is_v, simp)
       
  3004   } moreover {
       
  3005     assume eq_th': "th' = th"
       
  3006     have ?thesis 
       
  3007       apply (unfold eq_th' pvD_th_es cntCS_es_th)
       
  3008       by (insert assms[unfolded eq_th'], unfold is_v, simp)
       
  3009   } moreover {
       
  3010     assume h: "th' \<noteq> taker" "th' \<noteq> th"
       
  3011     have ?thesis using assms
       
  3012       apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
       
  3013       by (fold is_v, unfold pvD_def, simp)
       
  3014   } ultimately show ?thesis by metis
       
  3015 qed
       
  3016 
       
  3017 end
       
  3018 
       
  3019 context valid_trace_v_e
       
  3020 begin
       
  3021 
       
  3022 lemma holdents_es_th:
       
  3023   "holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R")
       
  3024 proof -
       
  3025   { fix cs'
       
  3026     assume "cs' \<in> ?L"
       
  3027     hence "holding (e#s) th cs'" by (auto simp:holdents_def)
       
  3028     hence "cs' \<in> ?R"
       
  3029     proof(cases rule:holding_esE)
       
  3030       case 1
       
  3031       thus ?thesis by (auto simp:holdents_def)
       
  3032     qed 
       
  3033   } moreover {
       
  3034     fix cs'
       
  3035     assume "cs' \<in> ?R"
       
  3036     hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def)
       
  3037     from holding_esI2[OF this]
       
  3038     have "cs' \<in> ?L" by (auto simp:holdents_def)
       
  3039   } ultimately show ?thesis by auto
       
  3040 qed
       
  3041 
       
  3042 lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1"
       
  3043 proof -
       
  3044   have "card (holdents s th - {cs}) = card (holdents s th) - 1"
       
  3045   proof -
       
  3046     have "cs \<in> holdents s th" using holding_th_cs_s
       
  3047       by (auto simp:holdents_def)
       
  3048     moreover have "finite (holdents s th)"
       
  3049         by (simp add: finite_holdents) 
       
  3050     ultimately show ?thesis by auto
       
  3051   qed
       
  3052   thus ?thesis by (unfold cntCS_def holdents_es_th)
       
  3053 qed
       
  3054 
       
  3055 lemma holdents_kept:
       
  3056   assumes "th' \<noteq> th"
       
  3057   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3058 proof -
       
  3059   { fix cs'
       
  3060     assume h: "cs' \<in> ?L"
       
  3061     have "cs' \<in> ?R"
       
  3062     proof(cases "cs' = cs")
       
  3063       case False
       
  3064       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3065       from h have "holding (e#s) th' cs'" by (auto simp:holdents_def)
       
  3066       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
  1905       show ?thesis
  3067       show ?thesis
  1906       proof -
  3068         by (unfold holdents_def s_holding_def, fold wq_def, auto)
  1907         { assume eq_th: "th = thread"
       
  1908           from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
       
  1909             by (unfold eq_e, simp add:cntP_def count_def)
       
  1910           moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
       
  1911             by (unfold eq_e, simp add:cntV_def count_def)
       
  1912           moreover from cntCS_v_dec [OF vtv] 
       
  1913           have "cntCS (e # s) thread + 1 = cntCS s thread"
       
  1914             by (simp add:eq_e)
       
  1915           moreover from is_runing have rd_before: "thread \<in> readys s"
       
  1916             by (unfold runing_def, simp)
       
  1917           moreover have "thread \<in> readys (e # s)"
       
  1918           proof -
       
  1919             from is_runing
       
  1920             have "thread \<in> threads (e#s)" 
       
  1921               by (unfold eq_e, auto simp:runing_def readys_def)
       
  1922             moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
       
  1923             proof
       
  1924               fix cs1
       
  1925               { assume eq_cs: "cs1 = cs" 
       
  1926                 have "\<not> waiting (e # s) thread cs1"
       
  1927                 proof -
       
  1928                   from eq_wq
       
  1929                   have "thread \<notin> set (wq (e#s) cs1)"
       
  1930                     apply(unfold eq_e wq_def eq_cs s_holding_def)
       
  1931                     apply (auto simp:Let_def)
       
  1932                   proof -
       
  1933                     assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)"
       
  1934                     with eq_set have "thread \<in> set rest" by simp
       
  1935                     with vt_v.wq_distinct[of cs]
       
  1936                     and eq_wq show False
       
  1937                         by (metis distinct.simps(2) vt_s.wq_distinct)
       
  1938                   qed
       
  1939                   thus ?thesis by (simp add:wq_def s_waiting_def)
       
  1940                 qed
       
  1941               } moreover {
       
  1942                 assume neq_cs: "cs1 \<noteq> cs"
       
  1943                   have "\<not> waiting (e # s) thread cs1" 
       
  1944                   proof -
       
  1945                     from wq_v_neq [OF neq_cs[symmetric]]
       
  1946                     have "wq (V thread cs # s) cs1 = wq s cs1" .
       
  1947                     moreover have "\<not> waiting s thread cs1" 
       
  1948                     proof -
       
  1949                       from runing_ready and is_runing
       
  1950                       have "thread \<in> readys s" by auto
       
  1951                       thus ?thesis by (simp add:readys_def)
       
  1952                     qed
       
  1953                     ultimately show ?thesis 
       
  1954                       by (auto simp:wq_def s_waiting_def eq_e)
       
  1955                   qed
       
  1956               } ultimately show "\<not> waiting (e # s) thread cs1" by blast
       
  1957             qed
       
  1958             ultimately show ?thesis by (simp add:readys_def)
       
  1959           qed
       
  1960           moreover note eq_th ih
       
  1961           ultimately have ?thesis by auto
       
  1962         } moreover {
       
  1963           assume neq_th: "th \<noteq> thread"
       
  1964           from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
       
  1965             by (simp add:cntP_def count_def)
       
  1966           from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
       
  1967             by (simp add:cntV_def count_def)
       
  1968           have ?thesis
       
  1969           proof(cases "th \<in> set rest")
       
  1970             case False
       
  1971             have "(th \<in> readys (e # s)) = (th \<in> readys s)"
       
  1972               apply (insert step_back_vt[OF vtv])
       
  1973               by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq)
       
  1974             moreover have "cntCS (e#s) th = cntCS s th"
       
  1975               apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
       
  1976               proof -
       
  1977                 have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
       
  1978                       {cs. (Cs cs, Th th) \<in> RAG s}"
       
  1979                 proof -
       
  1980                   from False eq_wq
       
  1981                   have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s"
       
  1982                     apply (unfold next_th_def, auto)
       
  1983                   proof -
       
  1984                     assume ne: "rest \<noteq> []"
       
  1985                       and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
       
  1986                       and eq_wq: "wq s cs = thread # rest"
       
  1987                     from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> 
       
  1988                                   set (SOME q. distinct q \<and> set q = set rest)
       
  1989                                   " by simp
       
  1990                     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
       
  1991                     proof(rule someI2)
       
  1992                       from vt_s.wq_distinct[ of cs] and eq_wq
       
  1993                       show "distinct rest \<and> set rest = set rest" by auto
       
  1994                     next
       
  1995                       fix x assume "distinct x \<and> set x = set rest"
       
  1996                       with ne show "x \<noteq> []" by auto
       
  1997                     qed
       
  1998                     ultimately show 
       
  1999                       "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s"
       
  2000                       by auto
       
  2001                   qed    
       
  2002                   thus ?thesis by auto
       
  2003                 qed
       
  2004                 thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} =
       
  2005                              card {cs. (Cs cs, Th th) \<in> RAG s}" by simp 
       
  2006               qed
       
  2007             moreover note ih eq_cnp eq_cnv eq_threads
       
  2008             ultimately show ?thesis by auto
       
  2009           next
       
  2010             case True
       
  2011             assume th_in: "th \<in> set rest"
       
  2012             show ?thesis
       
  2013             proof(cases "next_th s thread cs th")
       
  2014               case False
       
  2015               with eq_wq and th_in have 
       
  2016                 neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest")
       
  2017                 by (auto simp:next_th_def)
       
  2018               have "(th \<in> readys (e # s)) = (th \<in> readys s)"
       
  2019               proof -
       
  2020                 from eq_wq and th_in
       
  2021                 have "\<not> th \<in> readys s"
       
  2022                   apply (auto simp:readys_def s_waiting_def)
       
  2023                   apply (rule_tac x = cs in exI, auto)
       
  2024                   by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def)
       
  2025                 moreover 
       
  2026                 from eq_wq and th_in and neq_hd
       
  2027                 have "\<not> (th \<in> readys (e # s))"
       
  2028                   apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits)
       
  2029                   by (rule_tac x = cs in exI, auto simp:eq_set)
       
  2030                 ultimately show ?thesis by auto
       
  2031               qed
       
  2032               moreover have "cntCS (e#s) th = cntCS s th" 
       
  2033               proof -
       
  2034                 from eq_wq and  th_in and neq_hd
       
  2035                 have "(holdents (e # s) th) = (holdents s th)"
       
  2036                   apply (unfold eq_e step_RAG_v[OF vtv], 
       
  2037                          auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def
       
  2038                                    Let_def cs_holding_def)
       
  2039                   by (insert vt_s.wq_distinct[of cs], auto simp:wq_def)
       
  2040                 thus ?thesis by (simp add:cntCS_def)
       
  2041               qed
       
  2042               moreover note ih eq_cnp eq_cnv eq_threads
       
  2043               ultimately show ?thesis by auto
       
  2044             next
       
  2045               case True
       
  2046               let ?rest = " (SOME q. distinct q \<and> set q = set rest)"
       
  2047               let ?t = "hd ?rest"
       
  2048               from True eq_wq th_in neq_th
       
  2049               have "th \<in> readys (e # s)"
       
  2050                 apply (auto simp:eq_e readys_def s_waiting_def wq_def
       
  2051                         Let_def next_th_def)
       
  2052               proof -
       
  2053                 assume eq_wq: "wq_fun (schs s) cs = thread # rest"
       
  2054                   and t_in: "?t \<in> set rest"
       
  2055                 show "?t \<in> threads s"
       
  2056                 proof(rule vt_s.wq_threads)
       
  2057                   from eq_wq and t_in
       
  2058                   show "?t \<in> set (wq s cs)" by (auto simp:wq_def)
       
  2059                 qed
       
  2060               next
       
  2061                 fix csa
       
  2062                 assume eq_wq: "wq_fun (schs s) cs = thread # rest"
       
  2063                   and t_in: "?t \<in> set rest"
       
  2064                   and neq_cs: "csa \<noteq> cs"
       
  2065                   and t_in': "?t \<in>  set (wq_fun (schs s) csa)"
       
  2066                 show "?t = hd (wq_fun (schs s) csa)"
       
  2067                 proof -
       
  2068                   { assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)"
       
  2069                     from vt_s.wq_distinct[of cs] and 
       
  2070                     eq_wq[folded wq_def] and t_in eq_wq
       
  2071                     have "?t \<noteq> thread" by auto
       
  2072                     with eq_wq and t_in
       
  2073                     have w1: "waiting s ?t cs"
       
  2074                       by (auto simp:s_waiting_def wq_def)
       
  2075                     from t_in' neq_hd'
       
  2076                     have w2: "waiting s ?t csa"
       
  2077                       by (auto simp:s_waiting_def wq_def)
       
  2078                     from vt_s.waiting_unique[OF w1 w2]
       
  2079                     and neq_cs have "False" by auto
       
  2080                   } thus ?thesis by auto
       
  2081                 qed
       
  2082               qed
       
  2083               moreover have "cntP s th = cntV s th + cntCS s th + 1"
       
  2084               proof -
       
  2085                 have "th \<notin> readys s" 
       
  2086                 proof -
       
  2087                   from True eq_wq neq_th th_in
       
  2088                   show ?thesis
       
  2089                     apply (unfold readys_def s_waiting_def, auto)
       
  2090                     by (rule_tac x = cs in exI, auto simp add: wq_def)
       
  2091                 qed
       
  2092                 moreover have "th \<in> threads s"
       
  2093                 proof -
       
  2094                   from th_in eq_wq
       
  2095                   have "th \<in> set (wq s cs)" by simp
       
  2096                   from vt_s.wq_threads [OF this] 
       
  2097                   show ?thesis .
       
  2098                 qed
       
  2099                 ultimately show ?thesis using ih by auto
       
  2100               qed
       
  2101               moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th"
       
  2102                 apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto)
       
  2103               proof -
       
  2104                 show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} =
       
  2105                                Suc (card {cs. (Cs cs, Th th) \<in> RAG s})"
       
  2106                   (is "card ?A = Suc (card ?B)")
       
  2107                 proof -
       
  2108                   have "?A = insert cs ?B" by auto
       
  2109                   hence "card ?A = card (insert cs ?B)" by simp
       
  2110                   also have "\<dots> = Suc (card ?B)"
       
  2111                   proof(rule card_insert_disjoint)
       
  2112                     have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" 
       
  2113                       apply (auto simp:image_def)
       
  2114                       by (rule_tac x = "(Cs x, Th th)" in bexI, auto)
       
  2115                     with vt_s.finite_RAG
       
  2116                     show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset)
       
  2117                   next
       
  2118                     show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}"
       
  2119                     proof
       
  2120                       assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}"
       
  2121                       hence "(Cs cs, Th th) \<in> RAG s" by simp
       
  2122                       with True neq_th eq_wq show False
       
  2123                         by (auto simp:next_th_def s_RAG_def cs_holding_def)
       
  2124                     qed
       
  2125                   qed
       
  2126                   finally show ?thesis .
       
  2127                 qed
       
  2128               qed
       
  2129               moreover note eq_cnp eq_cnv
       
  2130               ultimately show ?thesis by simp
       
  2131             qed
       
  2132           qed
       
  2133         } ultimately show ?thesis by blast
       
  2134       qed
       
  2135     next
  3069     next
  2136       case (thread_set thread prio)
  3070       case True
  2137       assume eq_e: "e = Set thread prio"
  3071       from h[unfolded this]
  2138         and is_runing: "thread \<in> runing s"
  3072       have "holding (e#s) th' cs" by (auto simp:holdents_def)
       
  3073       from this[unfolded s_holding_def, folded wq_def, 
       
  3074             unfolded wq_es_cs nil_wq']
       
  3075       show ?thesis by auto
       
  3076     qed
       
  3077   } moreover {
       
  3078     fix cs'
       
  3079     assume h: "cs' \<in> ?R"
       
  3080     have "cs' \<in> ?L"
       
  3081     proof(cases "cs' = cs")
       
  3082       case False
       
  3083       hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp
       
  3084       from h have "holding s th' cs'" by (auto simp:holdents_def)
       
  3085       from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq]
  2139       show ?thesis
  3086       show ?thesis
  2140       proof -
  3087         by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp)
  2141         from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
  3088     next
  2142         from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
  3089       case True
  2143         have eq_cncs: "cntCS (e#s) th = cntCS s th"
  3090       from h[unfolded this]
  2144           unfolding cntCS_def holdents_test
  3091       have "holding s th' cs" by (auto simp:holdents_def)
  2145           by (simp add:RAG_set_unchanged eq_e)
  3092       from held_unique[OF this holding_th_cs_s]
  2146         from eq_e have eq_readys: "readys (e#s) = readys s" 
  3093       have "th' = th" .
  2147           by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
  3094       with assms show ?thesis by auto
  2148                   auto simp:Let_def)
  3095     qed
  2149         { assume "th \<noteq> thread"
  3096   } ultimately show ?thesis by auto
  2150           with eq_readys eq_e
  3097 qed
  2151           have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
  3098 
  2152                       (th \<in> readys (s) \<or> th \<notin> threads (s))" 
  3099 lemma cntCS_kept [simp]:
  2153             by (simp add:threads.simps)
  3100   assumes "th' \<noteq> th"
  2154           with eq_cnp eq_cnv eq_cncs ih is_runing
  3101   shows "cntCS (e#s) th' = cntCS s th'"
  2155           have ?thesis by simp
  3102   by (unfold cntCS_def holdents_kept[OF assms], simp)
  2156         } moreover {
  3103 
  2157           assume eq_th: "th = thread"
  3104 lemma readys_kept1: 
  2158           with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
  3105   assumes "th' \<in> readys (e#s)"
  2159             by (unfold runing_def, auto)
  3106   shows "th' \<in> readys s"
  2160           moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
  3107 proof -
  2161             by (simp add:runing_def)
  3108   { fix cs'
  2162           moreover note eq_cnp eq_cnv eq_cncs
  3109     assume wait: "waiting s th' cs'"
  2163           ultimately have ?thesis by auto
  3110     have n_wait: "\<not> waiting (e#s) th' cs'" 
  2164         } ultimately show ?thesis by blast
  3111         using assms(1)[unfolded readys_def] by auto
  2165       qed   
  3112     have False
  2166     qed
  3113     proof(cases "cs' = cs")
  2167   next
  3114       case False
  2168     case vt_nil
  3115       with n_wait wait
  2169     show ?case 
  3116       show ?thesis 
  2170       by (unfold cntP_def cntV_def cntCS_def, 
  3117         by (unfold s_waiting_def, fold wq_def, auto)
  2171         auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
  3118     next
  2172   qed
  3119       case True
       
  3120       have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" 
       
  3121         using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . 
       
  3122       hence "th' \<in> set rest" by auto
       
  3123       with set_wq' have "th' \<in> set wq'" by metis
       
  3124       with nil_wq' show ?thesis by simp
       
  3125     qed
       
  3126   } thus ?thesis using assms
       
  3127     by (unfold readys_def, auto)
       
  3128 qed
       
  3129 
       
  3130 lemma readys_kept2: 
       
  3131   assumes "th' \<in> readys s"
       
  3132   shows "th' \<in> readys (e#s)"
       
  3133 proof -
       
  3134   { fix cs'
       
  3135     assume wait: "waiting (e#s) th' cs'"
       
  3136     have n_wait: "\<not> waiting s th' cs'" 
       
  3137         using assms[unfolded readys_def] by auto
       
  3138     have False
       
  3139     proof(cases "cs' = cs")
       
  3140       case False
       
  3141       with n_wait wait
       
  3142       show ?thesis 
       
  3143         by (unfold s_waiting_def, fold wq_def, auto)
       
  3144     next
       
  3145       case True
       
  3146       have "th' \<in> set [] \<and> th' \<noteq> hd []"
       
  3147         using wait[unfolded True s_waiting_def, folded wq_def, 
       
  3148               unfolded wq_es_cs nil_wq'] .
       
  3149       thus ?thesis by simp
       
  3150     qed
       
  3151   } with assms show ?thesis  
       
  3152     by (unfold readys_def, auto)
       
  3153 qed
       
  3154 
       
  3155 lemma readys_simp [simp]:
       
  3156   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3157   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3158   by metis
       
  3159 
       
  3160 lemma cnp_cnv_cncs_kept:
       
  3161   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3162   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3163 proof -
       
  3164   {
       
  3165     assume eq_th': "th' = th"
       
  3166     have ?thesis 
       
  3167       apply (unfold eq_th' pvD_th_es cntCS_es_th)
       
  3168       by (insert assms[unfolded eq_th'], unfold is_v, simp)
       
  3169   } moreover {
       
  3170     assume h: "th' \<noteq> th"
       
  3171     have ?thesis using assms
       
  3172       apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp)
       
  3173       by (fold is_v, unfold pvD_def, simp)
       
  3174   } ultimately show ?thesis by metis
       
  3175 qed
       
  3176 
       
  3177 end
       
  3178 
       
  3179 context valid_trace_v
       
  3180 begin
       
  3181 
       
  3182 lemma cnp_cnv_cncs_kept:
       
  3183   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3184   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3185 proof(cases "rest = []")
       
  3186   case True
       
  3187   then interpret vt: valid_trace_v_e by (unfold_locales, simp)
       
  3188   show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
       
  3189 next
       
  3190   case False
       
  3191   then interpret vt: valid_trace_v_n by (unfold_locales, simp)
       
  3192   show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast 
       
  3193 qed
       
  3194 
       
  3195 end
       
  3196 
       
  3197 context valid_trace_create
       
  3198 begin
       
  3199 
       
  3200 lemma th_not_live_s [simp]: "th \<notin> threads s"
       
  3201 proof -
       
  3202   from pip_e[unfolded is_create]
       
  3203   show ?thesis by (cases, simp)
       
  3204 qed
       
  3205 
       
  3206 lemma th_not_ready_s [simp]: "th \<notin> readys s"
       
  3207   using th_not_live_s by (unfold readys_def, simp)
       
  3208 
       
  3209 lemma th_live_es [simp]: "th \<in> threads (e#s)"
       
  3210   by (unfold is_create, simp)
       
  3211 
       
  3212 lemma not_waiting_th_s [simp]: "\<not> waiting s th cs'"
       
  3213 proof
       
  3214   assume "waiting s th cs'"
       
  3215   from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3216   have "th \<in> set (wq s cs')" by auto
       
  3217   from wq_threads[OF this] have "th \<in> threads s" .
       
  3218   with th_not_live_s show False by simp
       
  3219 qed
       
  3220 
       
  3221 lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
       
  3222 proof
       
  3223   assume "holding s th cs'"
       
  3224   from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
       
  3225   have "th \<in> set (wq s cs')" by auto
       
  3226   from wq_threads[OF this] have "th \<in> threads s" .
       
  3227   with th_not_live_s show False by simp
       
  3228 qed
       
  3229 
       
  3230 lemma not_waiting_th_es [simp]: "\<not> waiting (e#s) th cs'"
       
  3231 proof
       
  3232   assume "waiting (e # s) th cs'"
       
  3233   from this[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3234   have "th \<in> set (wq s cs')" by auto
       
  3235   from wq_threads[OF this] have "th \<in> threads s" .
       
  3236   with th_not_live_s show False by simp
       
  3237 qed
       
  3238 
       
  3239 lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
       
  3240 proof
       
  3241   assume "holding (e # s) th cs'"
       
  3242   from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
       
  3243   have "th \<in> set (wq s cs')" by auto
       
  3244   from wq_threads[OF this] have "th \<in> threads s" .
       
  3245   with th_not_live_s show False by simp
       
  3246 qed
       
  3247 
       
  3248 lemma ready_th_es [simp]: "th \<in> readys (e#s)"
       
  3249   by (simp add:readys_def)
       
  3250 
       
  3251 lemma holdents_th_s: "holdents s th = {}"
       
  3252   by (unfold holdents_def, auto)
       
  3253 
       
  3254 lemma holdents_th_es: "holdents (e#s) th = {}"
       
  3255   by (unfold holdents_def, auto)
       
  3256 
       
  3257 lemma cntCS_th_s [simp]: "cntCS s th = 0"
       
  3258   by (unfold cntCS_def, simp add:holdents_th_s)
       
  3259 
       
  3260 lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
       
  3261   by (unfold cntCS_def, simp add:holdents_th_es)
       
  3262 
       
  3263 lemma pvD_th_s [simp]: "pvD s th = 0"
       
  3264   by (unfold pvD_def, simp)
       
  3265 
       
  3266 lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
       
  3267   by (unfold pvD_def, simp)
       
  3268 
       
  3269 lemma holdents_kept:
       
  3270   assumes "th' \<noteq> th"
       
  3271   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3272 proof -
       
  3273   { fix cs'
       
  3274     assume h: "cs' \<in> ?L"
       
  3275     hence "cs' \<in> ?R"
       
  3276       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3277              unfold wq_neq_simp, auto)
       
  3278   } moreover {
       
  3279     fix cs'
       
  3280     assume h: "cs' \<in> ?R"
       
  3281     hence "cs' \<in> ?L"
       
  3282       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3283              unfold wq_neq_simp, auto)
       
  3284   } ultimately show ?thesis by auto
       
  3285 qed
       
  3286 
       
  3287 lemma cntCS_kept [simp]:
       
  3288   assumes "th' \<noteq> th"
       
  3289   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  3290   using holdents_kept[OF assms]
       
  3291   by (unfold cntCS_def, simp)
       
  3292 
       
  3293 lemma readys_kept1: 
       
  3294   assumes "th' \<noteq> th"
       
  3295   and "th' \<in> readys (e#s)"
       
  3296   shows "th' \<in> readys s"
       
  3297 proof -
       
  3298   { fix cs'
       
  3299     assume wait: "waiting s th' cs'"
       
  3300     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3301       using assms by (auto simp:readys_def)
       
  3302     from wait[unfolded s_waiting_def, folded wq_def]
       
  3303          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3304     have False by auto
       
  3305   } thus ?thesis using assms
       
  3306     by (unfold readys_def, auto)
       
  3307 qed
       
  3308 
       
  3309 lemma readys_kept2: 
       
  3310   assumes "th' \<noteq> th"
       
  3311   and "th' \<in> readys s"
       
  3312   shows "th' \<in> readys (e#s)"
       
  3313 proof -
       
  3314   { fix cs'
       
  3315     assume wait: "waiting (e#s) th' cs'"
       
  3316     have n_wait: "\<not> waiting s th' cs'"
       
  3317       using assms(2) by (auto simp:readys_def)
       
  3318     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3319          n_wait[unfolded s_waiting_def, folded wq_def]
       
  3320     have False by auto
       
  3321   } with assms show ?thesis  
       
  3322     by (unfold readys_def, auto)
       
  3323 qed
       
  3324 
       
  3325 lemma readys_simp [simp]:
       
  3326   assumes "th' \<noteq> th"
       
  3327   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3328   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3329   by metis
       
  3330 
       
  3331 lemma pvD_kept [simp]:
       
  3332   assumes "th' \<noteq> th"
       
  3333   shows "pvD (e#s) th' = pvD s th'"
       
  3334   using assms
       
  3335   by (unfold pvD_def, simp)
       
  3336 
       
  3337 lemma cnp_cnv_cncs_kept:
       
  3338   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3339   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3340 proof -
       
  3341   {
       
  3342     assume eq_th': "th' = th"
       
  3343     have ?thesis using assms
       
  3344       by (unfold eq_th', simp, unfold is_create, simp)
       
  3345   } moreover {
       
  3346     assume h: "th' \<noteq> th"
       
  3347     hence ?thesis using assms
       
  3348       by (simp, simp add:is_create)
       
  3349   } ultimately show ?thesis by metis
       
  3350 qed
       
  3351 
       
  3352 end
       
  3353 
       
  3354 context valid_trace_exit
       
  3355 begin
       
  3356 
       
  3357 lemma th_live_s [simp]: "th \<in> threads s"
       
  3358 proof -
       
  3359   from pip_e[unfolded is_exit]
       
  3360   show ?thesis
       
  3361   by (cases, unfold runing_def readys_def, simp)
       
  3362 qed
       
  3363 
       
  3364 lemma th_ready_s [simp]: "th \<in> readys s"
       
  3365 proof -
       
  3366   from pip_e[unfolded is_exit]
       
  3367   show ?thesis
       
  3368   by (cases, unfold runing_def, simp)
       
  3369 qed
       
  3370 
       
  3371 lemma th_not_live_es [simp]: "th \<notin> threads (e#s)"
       
  3372   by (unfold is_exit, simp)
       
  3373 
       
  3374 lemma not_holding_th_s [simp]: "\<not> holding s th cs'"
       
  3375 proof -
       
  3376   from pip_e[unfolded is_exit]
       
  3377   show ?thesis 
       
  3378    by (cases, unfold holdents_def, auto)
       
  3379 qed
       
  3380 
       
  3381 lemma cntCS_th_s [simp]: "cntCS s th = 0"
       
  3382 proof -
       
  3383   from pip_e[unfolded is_exit]
       
  3384   show ?thesis 
       
  3385    by (cases, unfold cntCS_def, simp)
       
  3386 qed
       
  3387 
       
  3388 lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'"
       
  3389 proof
       
  3390   assume "holding (e # s) th cs'"
       
  3391   from this[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp]
       
  3392   have "holding s th cs'" 
       
  3393     by (unfold s_holding_def, fold wq_def, auto)
       
  3394   with not_holding_th_s 
       
  3395   show False by simp
       
  3396 qed
       
  3397 
       
  3398 lemma ready_th_es [simp]: "th \<notin> readys (e#s)"
       
  3399   by (simp add:readys_def)
       
  3400 
       
  3401 lemma holdents_th_s: "holdents s th = {}"
       
  3402   by (unfold holdents_def, auto)
       
  3403 
       
  3404 lemma holdents_th_es: "holdents (e#s) th = {}"
       
  3405   by (unfold holdents_def, auto)
       
  3406 
       
  3407 lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0"
       
  3408   by (unfold cntCS_def, simp add:holdents_th_es)
       
  3409 
       
  3410 lemma pvD_th_s [simp]: "pvD s th = 0"
       
  3411   by (unfold pvD_def, simp)
       
  3412 
       
  3413 lemma pvD_th_es [simp]: "pvD (e#s) th = 0"
       
  3414   by (unfold pvD_def, simp)
       
  3415 
       
  3416 lemma holdents_kept:
       
  3417   assumes "th' \<noteq> th"
       
  3418   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3419 proof -
       
  3420   { fix cs'
       
  3421     assume h: "cs' \<in> ?L"
       
  3422     hence "cs' \<in> ?R"
       
  3423       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3424              unfold wq_neq_simp, auto)
       
  3425   } moreover {
       
  3426     fix cs'
       
  3427     assume h: "cs' \<in> ?R"
       
  3428     hence "cs' \<in> ?L"
       
  3429       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3430              unfold wq_neq_simp, auto)
       
  3431   } ultimately show ?thesis by auto
       
  3432 qed
       
  3433 
       
  3434 lemma cntCS_kept [simp]:
       
  3435   assumes "th' \<noteq> th"
       
  3436   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  3437   using holdents_kept[OF assms]
       
  3438   by (unfold cntCS_def, simp)
       
  3439 
       
  3440 lemma readys_kept1: 
       
  3441   assumes "th' \<noteq> th"
       
  3442   and "th' \<in> readys (e#s)"
       
  3443   shows "th' \<in> readys s"
       
  3444 proof -
       
  3445   { fix cs'
       
  3446     assume wait: "waiting s th' cs'"
       
  3447     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3448       using assms by (auto simp:readys_def)
       
  3449     from wait[unfolded s_waiting_def, folded wq_def]
       
  3450          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3451     have False by auto
       
  3452   } thus ?thesis using assms
       
  3453     by (unfold readys_def, auto)
       
  3454 qed
       
  3455 
       
  3456 lemma readys_kept2: 
       
  3457   assumes "th' \<noteq> th"
       
  3458   and "th' \<in> readys s"
       
  3459   shows "th' \<in> readys (e#s)"
       
  3460 proof -
       
  3461   { fix cs'
       
  3462     assume wait: "waiting (e#s) th' cs'"
       
  3463     have n_wait: "\<not> waiting s th' cs'"
       
  3464       using assms(2) by (auto simp:readys_def)
       
  3465     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3466          n_wait[unfolded s_waiting_def, folded wq_def]
       
  3467     have False by auto
       
  3468   } with assms show ?thesis  
       
  3469     by (unfold readys_def, auto)
       
  3470 qed
       
  3471 
       
  3472 lemma readys_simp [simp]:
       
  3473   assumes "th' \<noteq> th"
       
  3474   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3475   using readys_kept1[OF assms] readys_kept2[OF assms]
       
  3476   by metis
       
  3477 
       
  3478 lemma pvD_kept [simp]:
       
  3479   assumes "th' \<noteq> th"
       
  3480   shows "pvD (e#s) th' = pvD s th'"
       
  3481   using assms
       
  3482   by (unfold pvD_def, simp)
       
  3483 
       
  3484 lemma cnp_cnv_cncs_kept:
       
  3485   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3486   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3487 proof -
       
  3488   {
       
  3489     assume eq_th': "th' = th"
       
  3490     have ?thesis using assms
       
  3491       by (unfold eq_th', simp, unfold is_exit, simp)
       
  3492   } moreover {
       
  3493     assume h: "th' \<noteq> th"
       
  3494     hence ?thesis using assms
       
  3495       by (simp, simp add:is_exit)
       
  3496   } ultimately show ?thesis by metis
       
  3497 qed
       
  3498 
       
  3499 end
       
  3500 
       
  3501 context valid_trace_set
       
  3502 begin
       
  3503 
       
  3504 lemma th_live_s [simp]: "th \<in> threads s"
       
  3505 proof -
       
  3506   from pip_e[unfolded is_set]
       
  3507   show ?thesis
       
  3508   by (cases, unfold runing_def readys_def, simp)
       
  3509 qed
       
  3510 
       
  3511 lemma th_ready_s [simp]: "th \<in> readys s"
       
  3512 proof -
       
  3513   from pip_e[unfolded is_set]
       
  3514   show ?thesis
       
  3515   by (cases, unfold runing_def, simp)
       
  3516 qed
       
  3517 
       
  3518 lemma th_not_live_es [simp]: "th \<in> threads (e#s)"
       
  3519   by (unfold is_set, simp)
       
  3520 
       
  3521 
       
  3522 lemma holdents_kept:
       
  3523   shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R")
       
  3524 proof -
       
  3525   { fix cs'
       
  3526     assume h: "cs' \<in> ?L"
       
  3527     hence "cs' \<in> ?R"
       
  3528       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3529              unfold wq_neq_simp, auto)
       
  3530   } moreover {
       
  3531     fix cs'
       
  3532     assume h: "cs' \<in> ?R"
       
  3533     hence "cs' \<in> ?L"
       
  3534       by (unfold holdents_def s_holding_def, fold wq_def, 
       
  3535              unfold wq_neq_simp, auto)
       
  3536   } ultimately show ?thesis by auto
       
  3537 qed
       
  3538 
       
  3539 lemma cntCS_kept [simp]:
       
  3540   shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R")
       
  3541   using holdents_kept
       
  3542   by (unfold cntCS_def, simp)
       
  3543 
       
  3544 lemma threads_kept[simp]:
       
  3545   "threads (e#s) = threads s"
       
  3546   by (unfold is_set, simp)
       
  3547 
       
  3548 lemma readys_kept1: 
       
  3549   assumes "th' \<in> readys (e#s)"
       
  3550   shows "th' \<in> readys s"
       
  3551 proof -
       
  3552   { fix cs'
       
  3553     assume wait: "waiting s th' cs'"
       
  3554     have n_wait: "\<not> waiting (e#s) th' cs'" 
       
  3555       using assms by (auto simp:readys_def)
       
  3556     from wait[unfolded s_waiting_def, folded wq_def]
       
  3557          n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3558     have False by auto
       
  3559   } moreover have "th' \<in> threads s" 
       
  3560     using assms[unfolded readys_def] by auto
       
  3561   ultimately show ?thesis 
       
  3562     by (unfold readys_def, auto)
       
  3563 qed
       
  3564 
       
  3565 lemma readys_kept2: 
       
  3566   assumes "th' \<in> readys s"
       
  3567   shows "th' \<in> readys (e#s)"
       
  3568 proof -
       
  3569   { fix cs'
       
  3570     assume wait: "waiting (e#s) th' cs'"
       
  3571     have n_wait: "\<not> waiting s th' cs'"
       
  3572       using assms by (auto simp:readys_def)
       
  3573     from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_neq_simp]
       
  3574          n_wait[unfolded s_waiting_def, folded wq_def]
       
  3575     have False by auto
       
  3576   } with assms show ?thesis  
       
  3577     by (unfold readys_def, auto)
       
  3578 qed
       
  3579 
       
  3580 lemma readys_simp [simp]:
       
  3581   shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)"
       
  3582   using readys_kept1 readys_kept2
       
  3583   by metis
       
  3584 
       
  3585 lemma pvD_kept [simp]:
       
  3586   shows "pvD (e#s) th' = pvD s th'"
       
  3587   by (unfold pvD_def, simp)
       
  3588 
       
  3589 lemma cnp_cnv_cncs_kept:
       
  3590   assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3591   shows "cntP (e#s) th' = cntV (e#s) th' +  cntCS (e#s) th' + pvD (e#s) th'"
       
  3592   using assms
       
  3593   by (unfold is_set, simp, fold is_set, simp)
       
  3594 
       
  3595 end
       
  3596 
       
  3597 context valid_trace
       
  3598 begin
       
  3599 
       
  3600 lemma cnp_cnv_cncs: "cntP s th' = cntV s th' + cntCS s th' + pvD s th'"
       
  3601 proof(induct rule:ind)
       
  3602   case Nil
       
  3603   thus ?case 
       
  3604     by (unfold cntP_def cntV_def pvD_def cntCS_def holdents_def 
       
  3605               s_holding_def, simp)
       
  3606 next
       
  3607   case (Cons s e)
       
  3608   interpret vt_e: valid_trace_e s e using Cons by simp
       
  3609   show ?case
       
  3610   proof(cases e)
       
  3611     case (Create th prio)
       
  3612     interpret vt_create: valid_trace_create s e th prio 
       
  3613       using Create by (unfold_locales, simp)
       
  3614     show ?thesis using Cons by (simp add: vt_create.cnp_cnv_cncs_kept) 
       
  3615   next
       
  3616     case (Exit th)
       
  3617     interpret vt_exit: valid_trace_exit s e th  
       
  3618         using Exit by (unfold_locales, simp)
       
  3619    show ?thesis using Cons by (simp add: vt_exit.cnp_cnv_cncs_kept) 
       
  3620   next
       
  3621     case (P th cs)
       
  3622     interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp)
       
  3623     show ?thesis using Cons by (simp add: vt_p.cnp_cnv_cncs_kept) 
       
  3624   next
       
  3625     case (V th cs)
       
  3626     interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp)
       
  3627     show ?thesis using Cons by (simp add: vt_v.cnp_cnv_cncs_kept) 
       
  3628   next
       
  3629     case (Set th prio)
       
  3630     interpret vt_set: valid_trace_set s e th prio
       
  3631         using Set by (unfold_locales, simp)
       
  3632     show ?thesis using Cons by (simp add: vt_set.cnp_cnv_cncs_kept) 
       
  3633   qed
       
  3634 qed
       
  3635 
       
  3636 lemma not_thread_holdents:
       
  3637   assumes not_in: "th \<notin> threads s" 
       
  3638   shows "holdents s th = {}"
       
  3639 proof -
       
  3640   { fix cs
       
  3641     assume "cs \<in> holdents s th"
       
  3642     hence "holding s th cs" by (auto simp:holdents_def)
       
  3643     from this[unfolded s_holding_def, folded wq_def]
       
  3644     have "th \<in> set (wq s cs)" by auto
       
  3645     with wq_threads have "th \<in> threads s" by auto
       
  3646     with assms
       
  3647     have False by simp
       
  3648   } thus ?thesis by auto
  2173 qed
  3649 qed
  2174 
  3650 
  2175 lemma not_thread_cncs:
  3651 lemma not_thread_cncs:
  2176   assumes not_in: "th \<notin> threads s" 
  3652   assumes not_in: "th \<notin> threads s" 
  2177   shows "cntCS s th = 0"
  3653   shows "cntCS s th = 0"
  2178 proof -
  3654   using not_thread_holdents[OF assms]
  2179   from vt not_in show ?thesis
  3655   by (simp add:cntCS_def)
  2180   proof(induct arbitrary:th)
  3656 
  2181     case (vt_cons s e th)
  3657 lemma cnp_cnv_eq:
  2182     interpret vt_s: valid_trace s using vt_cons(1)
  3658   assumes "th \<notin> threads s"
  2183        by (unfold_locales, simp)
  3659   shows "cntP s th = cntV s th"
  2184     assume vt: "vt s"
  3660   using assms cnp_cnv_cncs not_thread_cncs pvD_def
  2185       and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
  3661   by (auto)
  2186       and stp: "step s e"
       
  2187       and not_in: "th \<notin> threads (e # s)"
       
  2188     from stp show ?case
       
  2189     proof(cases)
       
  2190       case (thread_create thread prio)
       
  2191       assume eq_e: "e = Create thread prio"
       
  2192         and not_in': "thread \<notin> threads s"
       
  2193       have "cntCS (e # s) th = cntCS s th"
       
  2194         apply (unfold eq_e cntCS_def holdents_test)
       
  2195         by (simp add:RAG_create_unchanged)
       
  2196       moreover have "th \<notin> threads s" 
       
  2197       proof -
       
  2198         from not_in eq_e show ?thesis by simp
       
  2199       qed
       
  2200       moreover note ih ultimately show ?thesis by auto
       
  2201     next
       
  2202       case (thread_exit thread)
       
  2203       assume eq_e: "e = Exit thread"
       
  2204       and nh: "holdents s thread = {}"
       
  2205       have eq_cns: "cntCS (e # s) th = cntCS s th"
       
  2206         apply (unfold eq_e cntCS_def holdents_test)
       
  2207         by (simp add:RAG_exit_unchanged)
       
  2208       show ?thesis
       
  2209       proof(cases "th = thread")
       
  2210         case True
       
  2211         have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
       
  2212         with eq_cns show ?thesis by simp
       
  2213       next
       
  2214         case False
       
  2215         with not_in and eq_e
       
  2216         have "th \<notin> threads s" by simp
       
  2217         from ih[OF this] and eq_cns show ?thesis by simp
       
  2218       qed
       
  2219     next
       
  2220       case (thread_P thread cs)
       
  2221       assume eq_e: "e = P thread cs"
       
  2222       and is_runing: "thread \<in> runing s"
       
  2223       from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto
       
  2224       have neq_th: "th \<noteq> thread" 
       
  2225       proof -
       
  2226         from not_in eq_e have "th \<notin> threads s" by simp
       
  2227         moreover from is_runing have "thread \<in> threads s"
       
  2228           by (simp add:runing_def readys_def)
       
  2229         ultimately show ?thesis by auto
       
  2230       qed
       
  2231       hence "cntCS (e # s) th  = cntCS s th "
       
  2232         apply (unfold cntCS_def holdents_test eq_e)
       
  2233         by (unfold step_RAG_p[OF vtp], auto)
       
  2234       moreover have "cntCS s th = 0"
       
  2235       proof(rule ih)
       
  2236         from not_in eq_e show "th \<notin> threads s" by simp
       
  2237       qed
       
  2238       ultimately show ?thesis by simp
       
  2239     next
       
  2240       case (thread_V thread cs)
       
  2241       assume eq_e: "e = V thread cs"
       
  2242         and is_runing: "thread \<in> runing s"
       
  2243         and hold: "holding s thread cs"
       
  2244       have neq_th: "th \<noteq> thread" 
       
  2245       proof -
       
  2246         from not_in eq_e have "th \<notin> threads s" by simp
       
  2247         moreover from is_runing have "thread \<in> threads s"
       
  2248           by (simp add:runing_def readys_def)
       
  2249         ultimately show ?thesis by auto
       
  2250       qed
       
  2251       from assms thread_V vt stp ih 
       
  2252       have vtv: "vt (V thread cs#s)" by auto
       
  2253       then interpret vt_v: valid_trace "(V thread cs#s)"
       
  2254         by (unfold_locales, simp)
       
  2255       from hold obtain rest 
       
  2256         where eq_wq: "wq s cs = thread # rest"
       
  2257         by (case_tac "wq s cs", auto simp: wq_def s_holding_def)
       
  2258       from not_in eq_e eq_wq
       
  2259       have "\<not> next_th s thread cs th"
       
  2260         apply (auto simp:next_th_def)
       
  2261       proof -
       
  2262         assume ne: "rest \<noteq> []"
       
  2263           and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s")
       
  2264         have "?t \<in> set rest"
       
  2265         proof(rule someI2)
       
  2266           from vt_v.wq_distinct[of cs] and eq_wq
       
  2267           show "distinct rest \<and> set rest = set rest"
       
  2268             by (metis distinct.simps(2) vt_s.wq_distinct) 
       
  2269         next
       
  2270           fix x assume "distinct x \<and> set x = set rest" with ne
       
  2271           show "hd x \<in> set rest" by (cases x, auto)
       
  2272         qed
       
  2273         with eq_wq have "?t \<in> set (wq s cs)" by simp
       
  2274         from vt_s.wq_threads[OF this] and ni
       
  2275         show False
       
  2276           using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` 
       
  2277             ni vt_s.wq_threads by blast 
       
  2278       qed
       
  2279       moreover note neq_th eq_wq
       
  2280       ultimately have "cntCS (e # s) th  = cntCS s th"
       
  2281         by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto)
       
  2282       moreover have "cntCS s th = 0"
       
  2283       proof(rule ih)
       
  2284         from not_in eq_e show "th \<notin> threads s" by simp
       
  2285       qed
       
  2286       ultimately show ?thesis by simp
       
  2287     next
       
  2288       case (thread_set thread prio)
       
  2289       print_facts
       
  2290       assume eq_e: "e = Set thread prio"
       
  2291         and is_runing: "thread \<in> runing s"
       
  2292       from not_in and eq_e have "th \<notin> threads s" by auto
       
  2293       from ih [OF this] and eq_e
       
  2294       show ?thesis 
       
  2295         apply (unfold eq_e cntCS_def holdents_test)
       
  2296         by (simp add:RAG_set_unchanged)
       
  2297     qed
       
  2298     next
       
  2299       case vt_nil
       
  2300       show ?case
       
  2301       by (unfold cntCS_def, 
       
  2302         auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def)
       
  2303   qed
       
  2304 qed
       
  2305 
       
  2306 end
       
  2307 
       
  2308 lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
       
  2309   by (auto simp:s_waiting_def cs_waiting_def wq_def)
       
  2310 
       
  2311 context valid_trace
       
  2312 begin
       
  2313 
       
  2314 lemma dm_RAG_threads:
       
  2315   assumes in_dom: "(Th th) \<in> Domain (RAG s)"
       
  2316   shows "th \<in> threads s"
       
  2317 proof -
       
  2318   from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto
       
  2319   moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto
       
  2320   ultimately have "(Th th, Cs cs) \<in> RAG s" by simp
       
  2321   hence "th \<in> set (wq s cs)"
       
  2322     by (unfold s_RAG_def, auto simp:cs_waiting_def)
       
  2323   from wq_threads [OF this] show ?thesis .
       
  2324 qed
       
  2325 
       
  2326 end
       
  2327 
       
  2328 lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th"
       
  2329 unfolding cp_def wq_def
       
  2330 apply(induct s rule: schs.induct)
       
  2331 thm cpreced_initial
       
  2332 apply(simp add: Let_def cpreced_initial)
       
  2333 apply(simp add: Let_def)
       
  2334 apply(simp add: Let_def)
       
  2335 apply(simp add: Let_def)
       
  2336 apply(subst (2) schs.simps)
       
  2337 apply(simp add: Let_def)
       
  2338 apply(subst (2) schs.simps)
       
  2339 apply(simp add: Let_def)
       
  2340 done
       
  2341 
       
  2342 context valid_trace
       
  2343 begin
       
  2344 
  3662 
  2345 lemma runing_unique:
  3663 lemma runing_unique:
  2346   assumes runing_1: "th1 \<in> runing s"
  3664   assumes runing_1: "th1 \<in> runing s"
  2347   and runing_2: "th2 \<in> runing s"
  3665   and runing_2: "th2 \<in> runing s"
  2348   shows "th1 = th2"
  3666   shows "th1 = th2"
  2349 proof -
  3667 proof -
  2350   from runing_1 and runing_2 have "cp s th1 = cp s th2"
  3668   from runing_1 and runing_2 have "cp s th1 = cp s th2"
  2351     unfolding runing_def
  3669     unfolding runing_def by auto
  2352     apply(simp)
  3670   from this[unfolded cp_alt_def]
  2353     done
  3671   have eq_max: 
  2354   hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) =
  3672     "Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th1)}) =
  2355                  Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))"
  3673      Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th2)})" 
  2356     (is "Max (?f ` ?A) = Max (?f ` ?B)")
  3674         (is "Max ?L = Max ?R") .
  2357     unfolding cp_eq_cpreced 
  3675   have "Max ?L \<in> ?L"
  2358     unfolding cpreced_def .
  3676   proof(rule Max_in)
  2359   obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
  3677     show "finite ?L" by (simp add: finite_subtree_threads)
  2360   proof -
  3678   next
  2361     have h1: "finite (?f ` ?A)"
  3679     show "?L \<noteq> {}" using subtree_def by fastforce 
  2362     proof -
  3680   qed
  2363       have "finite ?A" 
  3681   then obtain th1' where 
  2364       proof -
  3682     h_1: "Th th1' \<in> subtree (RAG s) (Th th1)" "the_preced s th1' = Max ?L"
  2365         have "finite (dependants (wq s) th1)"
  3683     by auto
  2366         proof-
  3684   have "Max ?R \<in> ?R"
  2367           have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}"
  3685   proof(rule Max_in)
  2368           proof -
  3686     show "finite ?R" by (simp add: finite_subtree_threads)
  2369             let ?F = "\<lambda> (x, y). the_th x"
  3687   next
  2370             have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
  3688     show "?R \<noteq> {}" using subtree_def by fastforce 
  2371               apply (auto simp:image_def)
  3689   qed
  2372               by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
  3690   then obtain th2' where 
  2373             moreover have "finite \<dots>"
  3691     h_2: "Th th2' \<in> subtree (RAG s) (Th th2)" "the_preced s th2' = Max ?R"
  2374             proof -
  3692     by auto
  2375               from finite_RAG have "finite (RAG s)" .
  3693   have "th1' = th2'"
  2376               hence "finite ((RAG (wq s))\<^sup>+)"
  3694   proof(rule preced_unique)
  2377                 apply (unfold finite_trancl)
  3695     from h_1(1)
  2378                 by (auto simp: s_RAG_def cs_RAG_def wq_def)
  3696     show "th1' \<in> threads s"
  2379               thus ?thesis by auto
  3697     proof(cases rule:subtreeE)
  2380             qed
  3698       case 1
  2381             ultimately show ?thesis by (auto intro:finite_subset)
  3699       hence "th1' = th1" by simp
  2382           qed
  3700       with runing_1 show ?thesis by (auto simp:runing_def readys_def)
  2383           thus ?thesis by (simp add:cs_dependants_def)
  3701     next
  2384         qed
  3702       case 2
  2385         thus ?thesis by simp
  3703       from this(2)
  2386       qed
  3704       have "(Th th1', Th th1) \<in> (RAG s)^+" by (auto simp:ancestors_def)
  2387       thus ?thesis by auto
  3705       from tranclD[OF this]
  2388     qed
  3706       have "(Th th1') \<in> Domain (RAG s)" by auto
  2389     moreover have h2: "(?f ` ?A) \<noteq> {}"
       
  2390     proof -
       
  2391       have "?A \<noteq> {}" by simp
       
  2392       thus ?thesis by simp
       
  2393     qed
       
  2394     from Max_in [OF h1 h2]
       
  2395     have "Max (?f ` ?A) \<in> (?f ` ?A)" .
       
  2396     thus ?thesis 
       
  2397       thm cpreced_def
       
  2398       unfolding cpreced_def[symmetric] 
       
  2399       unfolding cp_eq_cpreced[symmetric] 
       
  2400       unfolding cpreced_def 
       
  2401       using that[intro] by (auto)
       
  2402   qed
       
  2403   obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
       
  2404   proof -
       
  2405     have h1: "finite (?f ` ?B)"
       
  2406     proof -
       
  2407       have "finite ?B" 
       
  2408       proof -
       
  2409         have "finite (dependants (wq s) th2)"
       
  2410         proof-
       
  2411           have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}"
       
  2412           proof -
       
  2413             let ?F = "\<lambda> (x, y). the_th x"
       
  2414             have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
       
  2415               apply (auto simp:image_def)
       
  2416               by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
       
  2417             moreover have "finite \<dots>"
       
  2418             proof -
       
  2419               from finite_RAG have "finite (RAG s)" .
       
  2420               hence "finite ((RAG (wq s))\<^sup>+)"
       
  2421                 apply (unfold finite_trancl)
       
  2422                 by (auto simp: s_RAG_def cs_RAG_def wq_def)
       
  2423               thus ?thesis by auto
       
  2424             qed
       
  2425             ultimately show ?thesis by (auto intro:finite_subset)
       
  2426           qed
       
  2427           thus ?thesis by (simp add:cs_dependants_def)
       
  2428         qed
       
  2429         thus ?thesis by simp
       
  2430       qed
       
  2431       thus ?thesis by auto
       
  2432     qed
       
  2433     moreover have h2: "(?f ` ?B) \<noteq> {}"
       
  2434     proof -
       
  2435       have "?B \<noteq> {}" by simp
       
  2436       thus ?thesis by simp
       
  2437     qed
       
  2438     from Max_in [OF h1 h2]
       
  2439     have "Max (?f ` ?B) \<in> (?f ` ?B)" .
       
  2440     thus ?thesis by (auto intro:that)
       
  2441   qed
       
  2442   from eq_f_th1 eq_f_th2 eq_max 
       
  2443   have eq_preced: "preced th1' s = preced th2' s" by auto
       
  2444   hence eq_th12: "th1' = th2'"
       
  2445   proof (rule preced_unique)
       
  2446     from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp
       
  2447     thus "th1' \<in> threads s"
       
  2448     proof
       
  2449       assume "th1' \<in> dependants (wq s) th1"
       
  2450       hence "(Th th1') \<in> Domain ((RAG s)^+)"
       
  2451         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
       
  2452         by (auto simp:Domain_def)
       
  2453       hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain)
       
  2454       from dm_RAG_threads[OF this] show ?thesis .
  3707       from dm_RAG_threads[OF this] show ?thesis .
       
  3708     qed
       
  3709   next
       
  3710     from h_2(1)
       
  3711     show "th2' \<in> threads s"
       
  3712     proof(cases rule:subtreeE)
       
  3713       case 1
       
  3714       hence "th2' = th2" by simp
       
  3715       with runing_2 show ?thesis by (auto simp:runing_def readys_def)
  2455     next
  3716     next
  2456       assume "th1' = th1"
  3717       case 2
  2457       with runing_1 show ?thesis
  3718       from this(2)
       
  3719       have "(Th th2', Th th2) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
  3720       from tranclD[OF this]
       
  3721       have "(Th th2') \<in> Domain (RAG s)" by auto
       
  3722       from dm_RAG_threads[OF this] show ?thesis .
       
  3723     qed
       
  3724   next
       
  3725     have "the_preced s th1' = the_preced s th2'" 
       
  3726      using eq_max h_1(2) h_2(2) by metis
       
  3727     thus "preced th1' s = preced th2' s" by (simp add:the_preced_def)
       
  3728   qed
       
  3729   from h_1(1)[unfolded this]
       
  3730   have star1: "(Th th2', Th th1) \<in> (RAG s)^*" by (auto simp:subtree_def)
       
  3731   from h_2(1)[unfolded this]
       
  3732   have star2: "(Th th2', Th th2) \<in> (RAG s)^*" by (auto simp:subtree_def)
       
  3733   from star_rpath[OF star1] obtain xs1 
       
  3734     where rp1: "rpath (RAG s) (Th th2') xs1 (Th th1)"
       
  3735     by auto
       
  3736   from star_rpath[OF star2] obtain xs2 
       
  3737     where rp2: "rpath (RAG s) (Th th2') xs2 (Th th2)"
       
  3738     by auto
       
  3739   from rp1 rp2
       
  3740   show ?thesis
       
  3741   proof(cases)
       
  3742     case (less_1 xs')
       
  3743     moreover have "xs' = []"
       
  3744     proof(rule ccontr)
       
  3745       assume otherwise: "xs' \<noteq> []"
       
  3746       from rpath_plus[OF less_1(3) this]
       
  3747       have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+" .
       
  3748       from tranclD[OF this]
       
  3749       obtain cs where "waiting s th1 cs"
       
  3750         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  3751       with runing_1 show False
  2458         by (unfold runing_def readys_def, auto)
  3752         by (unfold runing_def readys_def, auto)
  2459     qed
  3753     qed
  2460   next
  3754     ultimately have "xs2 = xs1" by simp
  2461     from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp
  3755     from rpath_dest_eq[OF rp1 rp2[unfolded this]]
  2462     thus "th2' \<in> threads s"
  3756     show ?thesis by simp
  2463     proof
  3757   next
  2464       assume "th2' \<in> dependants (wq s) th2"
  3758     case (less_2 xs')
  2465       hence "(Th th2') \<in> Domain ((RAG s)^+)"
  3759     moreover have "xs' = []"
  2466         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
  3760     proof(rule ccontr)
  2467         by (auto simp:Domain_def)
  3761       assume otherwise: "xs' \<noteq> []"
  2468       hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain)
  3762       from rpath_plus[OF less_2(3) this]
  2469       from dm_RAG_threads[OF this] show ?thesis .
  3763       have "(Th th2, Th th1) \<in> (RAG s)\<^sup>+" .
  2470     next
  3764       from tranclD[OF this]
  2471       assume "th2' = th2"
  3765       obtain cs where "waiting s th2 cs"
  2472       with runing_2 show ?thesis
  3766         by (unfold s_RAG_def, fold waiting_eq, auto)
       
  3767       with runing_2 show False
  2473         by (unfold runing_def readys_def, auto)
  3768         by (unfold runing_def readys_def, auto)
  2474     qed
  3769     qed
  2475   qed
  3770     ultimately have "xs2 = xs1" by simp
  2476   from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp
  3771     from rpath_dest_eq[OF rp1 rp2[unfolded this]]
  2477   thus ?thesis
  3772     show ?thesis by simp
  2478   proof
  3773   qed
  2479     assume eq_th': "th1' = th1"
  3774 qed
  2480     from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
  3775 
  2481     thus ?thesis
  3776 lemma card_runing: "card (runing s) \<le> 1"
  2482     proof
  3777 proof(cases "runing s = {}")
  2483       assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
  3778   case True
  2484     next
  3779   thus ?thesis by auto
  2485       assume "th2' \<in> dependants (wq s) th2"
  3780 next
  2486       with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp
  3781   case False
  2487       hence "(Th th1, Th th2) \<in> (RAG s)^+"
  3782   then obtain th where [simp]: "th \<in> runing s" by auto
  2488         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
  3783   from runing_unique[OF this]
  2489       hence "Th th1 \<in> Domain ((RAG s)^+)" 
  3784   have "runing s = {th}" by auto
  2490         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
  3785   thus ?thesis by auto
  2491         by (auto simp:Domain_def)
  3786 qed
  2492       hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain)
       
  2493       then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def)
       
  2494       from RAG_target_th [OF this]
       
  2495       obtain cs' where "n = Cs cs'" by auto
       
  2496       with d have "(Th th1, Cs cs') \<in> RAG s" by simp
       
  2497       with runing_1 have "False"
       
  2498         apply (unfold runing_def readys_def s_RAG_def)
       
  2499         by (auto simp:eq_waiting)
       
  2500       thus ?thesis by simp
       
  2501     qed
       
  2502   next
       
  2503     assume th1'_in: "th1' \<in> dependants (wq s) th1"
       
  2504     from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp
       
  2505     thus ?thesis 
       
  2506     proof
       
  2507       assume "th2' = th2"
       
  2508       with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp
       
  2509       hence "(Th th2, Th th1) \<in> (RAG s)^+"
       
  2510         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       
  2511       hence "Th th2 \<in> Domain ((RAG s)^+)" 
       
  2512         apply (unfold cs_dependants_def cs_RAG_def s_RAG_def)
       
  2513         by (auto simp:Domain_def)
       
  2514       hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain)
       
  2515       then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def)
       
  2516       from RAG_target_th [OF this]
       
  2517       obtain cs' where "n = Cs cs'" by auto
       
  2518       with d have "(Th th2, Cs cs') \<in> RAG s" by simp
       
  2519       with runing_2 have "False"
       
  2520         apply (unfold runing_def readys_def s_RAG_def)
       
  2521         by (auto simp:eq_waiting)
       
  2522       thus ?thesis by simp
       
  2523     next
       
  2524       assume "th2' \<in> dependants (wq s) th2"
       
  2525       with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp
       
  2526       hence h1: "(Th th1', Th th2) \<in> (RAG s)^+"
       
  2527         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       
  2528       from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+"
       
  2529         by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp)
       
  2530       show ?thesis
       
  2531       proof(rule dchain_unique[OF h1 _ h2, symmetric])
       
  2532         from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
       
  2533         from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
       
  2534       qed
       
  2535     qed
       
  2536   qed
       
  2537 qed
       
  2538 
       
  2539 
       
  2540 lemma "card (runing s) \<le> 1"
       
  2541 apply(subgoal_tac "finite (runing s)")
       
  2542 prefer 2
       
  2543 apply (metis finite_nat_set_iff_bounded lessI runing_unique)
       
  2544 apply(rule ccontr)
       
  2545 apply(simp)
       
  2546 apply(case_tac "Suc (Suc 0) \<le> card (runing s)")
       
  2547 apply(subst (asm) card_le_Suc_iff)
       
  2548 apply(simp)
       
  2549 apply(auto)[1]
       
  2550 apply (metis insertCI runing_unique)
       
  2551 apply(auto) 
       
  2552 done
       
  2553 
       
  2554 end
       
  2555 
       
  2556 
  3787 
  2557 lemma create_pre:
  3788 lemma create_pre:
  2558   assumes stp: "step s e"
  3789   assumes stp: "step s e"
  2559   and not_in: "th \<notin> threads s"
  3790   and not_in: "th \<notin> threads s"
  2560   and is_in: "th \<in> threads (e#s)"
  3791   and is_in: "th \<in> threads (e#s)"
  2579     case (thread_set thread)
  3810     case (thread_set thread)
  2580     with assms show ?thesis by (auto intro!:that)
  3811     with assms show ?thesis by (auto intro!:that)
  2581   qed
  3812   qed
  2582 qed
  3813 qed
  2583 
  3814 
  2584 
  3815 lemma eq_pv_children:
  2585 context valid_trace
  3816   assumes eq_pv: "cntP s th = cntV s th"
  2586 begin
  3817   shows "children (RAG s) (Th th) = {}"
  2587 
  3818 proof -
  2588 lemma cnp_cnv_eq:
  3819     from cnp_cnv_cncs and eq_pv
  2589   assumes "th \<notin> threads s"
  3820     have "cntCS s th = 0" 
  2590   shows "cntP s th = cntV s th"
  3821       by (auto split:if_splits)
  2591   using assms
  3822     from this[unfolded cntCS_def holdents_alt_def]
  2592   using cnp_cnv_cncs not_thread_cncs by auto
  3823     have card_0: "card (the_cs ` children (RAG s) (Th th)) = 0" .
       
  3824     have "finite (the_cs ` children (RAG s) (Th th))"
       
  3825       by (simp add: fsbtRAGs.finite_children)
       
  3826     from card_0[unfolded card_0_eq[OF this]]
       
  3827     show ?thesis by auto
       
  3828 qed
       
  3829 
       
  3830 lemma eq_pv_holdents:
       
  3831   assumes eq_pv: "cntP s th = cntV s th"
       
  3832   shows "holdents s th = {}"
       
  3833   by (unfold holdents_alt_def eq_pv_children[OF assms], simp)
       
  3834 
       
  3835 lemma eq_pv_subtree:
       
  3836   assumes eq_pv: "cntP s th = cntV s th"
       
  3837   shows "subtree (RAG s) (Th th) = {Th th}"
       
  3838   using eq_pv_children[OF assms]
       
  3839     by (unfold subtree_children, simp)
  2593 
  3840 
  2594 end
  3841 end
  2595 
       
  2596 
       
  2597 lemma eq_RAG: 
       
  2598   "RAG (wq s) = RAG s"
       
  2599 by (unfold cs_RAG_def s_RAG_def, auto)
       
  2600 
       
  2601 context valid_trace
       
  2602 begin
       
  2603 
       
  2604 lemma count_eq_dependants:
       
  2605   assumes eq_pv: "cntP s th = cntV s th"
       
  2606   shows "dependants (wq s) th = {}"
       
  2607 proof -
       
  2608   from cnp_cnv_cncs and eq_pv
       
  2609   have "cntCS s th = 0" 
       
  2610     by (auto split:if_splits)
       
  2611   moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}"
       
  2612   proof -
       
  2613     from finite_holding[of th] show ?thesis
       
  2614       by (simp add:holdents_test)
       
  2615   qed
       
  2616   ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}"
       
  2617     by (unfold cntCS_def holdents_test cs_dependants_def, auto)
       
  2618   show ?thesis
       
  2619   proof(unfold cs_dependants_def)
       
  2620     { assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}"
       
  2621       then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto
       
  2622       hence "False"
       
  2623       proof(cases)
       
  2624         assume "(Th th', Th th) \<in> RAG (wq s)"
       
  2625         thus "False" by (auto simp:cs_RAG_def)
       
  2626       next
       
  2627         fix c
       
  2628         assume "(c, Th th) \<in> RAG (wq s)"
       
  2629         with h and eq_RAG show "False"
       
  2630           by (cases c, auto simp:cs_RAG_def)
       
  2631       qed
       
  2632     } thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto
       
  2633   qed
       
  2634 qed
       
  2635 
       
  2636 lemma dependants_threads:
       
  2637   shows "dependants (wq s) th \<subseteq> threads s"
       
  2638 proof
       
  2639   { fix th th'
       
  2640     assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}"
       
  2641     have "Th th \<in> Domain (RAG s)"
       
  2642     proof -
       
  2643       from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto
       
  2644       hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def)
       
  2645       with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp
       
  2646       thus ?thesis using eq_RAG by simp
       
  2647     qed
       
  2648     from dm_RAG_threads[OF this]
       
  2649     have "th \<in> threads s" .
       
  2650   } note hh = this
       
  2651   fix th1 
       
  2652   assume "th1 \<in> dependants (wq s) th"
       
  2653   hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}"
       
  2654     by (unfold cs_dependants_def, simp)
       
  2655   from hh [OF this] show "th1 \<in> threads s" .
       
  2656 qed
       
  2657 
       
  2658 lemma finite_threads:
       
  2659   shows "finite (threads s)"
       
  2660 using vt by (induct) (auto elim: step.cases)
       
  2661 
       
  2662 end
       
  2663 
       
  2664 lemma Max_f_mono:
       
  2665   assumes seq: "A \<subseteq> B"
       
  2666   and np: "A \<noteq> {}"
       
  2667   and fnt: "finite B"
       
  2668   shows "Max (f ` A) \<le> Max (f ` B)"
       
  2669 proof(rule Max_mono)
       
  2670   from seq show "f ` A \<subseteq> f ` B" by auto
       
  2671 next
       
  2672   from np show "f ` A \<noteq> {}" by auto
       
  2673 next
       
  2674   from fnt and seq show "finite (f ` B)" by auto
       
  2675 qed
       
  2676 
       
  2677 context valid_trace
       
  2678 begin
       
  2679 
       
  2680 lemma cp_le:
       
  2681   assumes th_in: "th \<in> threads s"
       
  2682   shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
       
  2683 proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def)
       
  2684   show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}))
       
  2685          \<le> Max ((\<lambda>th. preced th s) ` threads s)"
       
  2686     (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
       
  2687   proof(rule Max_f_mono)
       
  2688     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp
       
  2689   next
       
  2690     from finite_threads
       
  2691     show "finite (threads s)" .
       
  2692   next
       
  2693     from th_in
       
  2694     show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s"
       
  2695       apply (auto simp:Domain_def)
       
  2696       apply (rule_tac dm_RAG_threads)
       
  2697       apply (unfold trancl_domain [of "RAG s", symmetric])
       
  2698       by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def)
       
  2699   qed
       
  2700 qed
       
  2701 
       
  2702 lemma le_cp:
       
  2703   shows "preced th s \<le> cp s th"
       
  2704 proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
       
  2705   show "Prc (priority th s) (last_set th s)
       
  2706     \<le> Max (insert (Prc (priority th s) (last_set th s))
       
  2707             ((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))"
       
  2708     (is "?l \<le> Max (insert ?l ?A)")
       
  2709   proof(cases "?A = {}")
       
  2710     case False
       
  2711     have "finite ?A" (is "finite (?f ` ?B)")
       
  2712     proof -
       
  2713       have "finite ?B" 
       
  2714       proof-
       
  2715         have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}"
       
  2716         proof -
       
  2717           let ?F = "\<lambda> (x, y). the_th x"
       
  2718           have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)"
       
  2719             apply (auto simp:image_def)
       
  2720             by (rule_tac x = "(Th x, Th th)" in bexI, auto)
       
  2721           moreover have "finite \<dots>"
       
  2722           proof -
       
  2723             from finite_RAG have "finite (RAG s)" .
       
  2724             hence "finite ((RAG (wq s))\<^sup>+)"
       
  2725               apply (unfold finite_trancl)
       
  2726               by (auto simp: s_RAG_def cs_RAG_def wq_def)
       
  2727             thus ?thesis by auto
       
  2728           qed
       
  2729           ultimately show ?thesis by (auto intro:finite_subset)
       
  2730         qed
       
  2731         thus ?thesis by (simp add:cs_dependants_def)
       
  2732       qed
       
  2733       thus ?thesis by simp
       
  2734     qed
       
  2735     from Max_insert [OF this False, of ?l] show ?thesis by auto
       
  2736   next
       
  2737     case True
       
  2738     thus ?thesis by auto
       
  2739   qed
       
  2740 qed
       
  2741 
       
  2742 lemma max_cp_eq: 
       
  2743   shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)"
       
  2744   (is "?l = ?r")
       
  2745 proof(cases "threads s = {}")
       
  2746   case True
       
  2747   thus ?thesis by auto
       
  2748 next
       
  2749   case False
       
  2750   have "?l \<in> ((cp s) ` threads s)"
       
  2751   proof(rule Max_in)
       
  2752     from finite_threads
       
  2753     show "finite (cp s ` threads s)" by auto
       
  2754   next
       
  2755     from False show "cp s ` threads s \<noteq> {}" by auto
       
  2756   qed
       
  2757   then obtain th 
       
  2758     where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto
       
  2759   have "\<dots> \<le> ?r" by (rule cp_le[OF th_in])
       
  2760   moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th")
       
  2761   proof -
       
  2762     have "?r \<in> (?f ` ?A)"
       
  2763     proof(rule Max_in)
       
  2764       from finite_threads
       
  2765       show " finite ((\<lambda>th. preced th s) ` threads s)" by auto
       
  2766     next
       
  2767       from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
       
  2768     qed
       
  2769     then obtain th' where 
       
  2770       th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto
       
  2771     from le_cp [of th']  eq_r
       
  2772     have "?r \<le> cp s th'" by auto
       
  2773     moreover have "\<dots> \<le> cp s th"
       
  2774     proof(fold eq_l)
       
  2775       show " cp s th' \<le> Max (cp s ` threads s)"
       
  2776       proof(rule Max_ge)
       
  2777         from th_in' show "cp s th' \<in> cp s ` threads s"
       
  2778           by auto
       
  2779       next
       
  2780         from finite_threads
       
  2781         show "finite (cp s ` threads s)" by auto
       
  2782       qed
       
  2783     qed
       
  2784     ultimately show ?thesis by auto
       
  2785   qed
       
  2786   ultimately show ?thesis using eq_l by auto
       
  2787 qed
       
  2788 
       
  2789 lemma max_cp_readys_threads_pre:
       
  2790   assumes np: "threads s \<noteq> {}"
       
  2791   shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
       
  2792 proof(unfold max_cp_eq)
       
  2793   show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
       
  2794   proof -
       
  2795     let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
       
  2796     let ?f = "(\<lambda>th. preced th s)"
       
  2797     have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
       
  2798     proof(rule Max_in)
       
  2799       from finite_threads show "finite (?f ` threads s)" by simp
       
  2800     next
       
  2801       from np show "?f ` threads s \<noteq> {}" by simp
       
  2802     qed
       
  2803     then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
       
  2804       by (auto simp:Image_def)
       
  2805     from th_chain_to_ready [OF tm_in]
       
  2806     have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" .
       
  2807     thus ?thesis
       
  2808     proof
       
  2809       assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ "
       
  2810       then obtain th' where th'_in: "th' \<in> readys s" 
       
  2811         and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto
       
  2812       have "cp s th' = ?f tm"
       
  2813       proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
       
  2814         from dependants_threads finite_threads
       
  2815         show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" 
       
  2816           by (auto intro:finite_subset)
       
  2817       next
       
  2818         fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
       
  2819         from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
       
  2820         moreover have "p \<le> \<dots>"
       
  2821         proof(rule Max_ge)
       
  2822           from finite_threads
       
  2823           show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2824         next
       
  2825           from p_in and th'_in and dependants_threads[of th']
       
  2826           show "p \<in> (\<lambda>th. preced th s) ` threads s"
       
  2827             by (auto simp:readys_def)
       
  2828         qed
       
  2829         ultimately show "p \<le> preced tm s" by auto
       
  2830       next
       
  2831         show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')"
       
  2832         proof -
       
  2833           from tm_chain
       
  2834           have "tm \<in> dependants (wq s) th'"
       
  2835             by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto)
       
  2836           thus ?thesis by auto
       
  2837         qed
       
  2838       qed
       
  2839       with tm_max
       
  2840       have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2841       show ?thesis
       
  2842       proof (fold h, rule Max_eqI)
       
  2843         fix q 
       
  2844         assume "q \<in> cp s ` readys s"
       
  2845         then obtain th1 where th1_in: "th1 \<in> readys s"
       
  2846           and eq_q: "q = cp s th1" by auto
       
  2847         show "q \<le> cp s th'"
       
  2848           apply (unfold h eq_q)
       
  2849           apply (unfold cp_eq_cpreced cpreced_def)
       
  2850           apply (rule Max_mono)
       
  2851         proof -
       
  2852           from dependants_threads [of th1] th1_in
       
  2853           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> 
       
  2854                  (\<lambda>th. preced th s) ` threads s"
       
  2855             by (auto simp:readys_def)
       
  2856         next
       
  2857           show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp
       
  2858         next
       
  2859           from finite_threads 
       
  2860           show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2861         qed
       
  2862       next
       
  2863         from finite_threads
       
  2864         show "finite (cp s ` readys s)" by (auto simp:readys_def)
       
  2865       next
       
  2866         from th'_in
       
  2867         show "cp s th' \<in> cp s ` readys s" by simp
       
  2868       qed
       
  2869     next
       
  2870       assume tm_ready: "tm \<in> readys s"
       
  2871       show ?thesis
       
  2872       proof(fold tm_max)
       
  2873         have cp_eq_p: "cp s tm = preced tm s"
       
  2874         proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
       
  2875           fix y 
       
  2876           assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
       
  2877           show "y \<le> preced tm s"
       
  2878           proof -
       
  2879             { fix y'
       
  2880               assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)"
       
  2881               have "y' \<le> preced tm s"
       
  2882               proof(unfold tm_max, rule Max_ge)
       
  2883                 from hy' dependants_threads[of tm]
       
  2884                 show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
       
  2885               next
       
  2886                 from finite_threads
       
  2887                 show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2888               qed
       
  2889             } with hy show ?thesis by auto
       
  2890           qed
       
  2891         next
       
  2892           from dependants_threads[of tm] finite_threads
       
  2893           show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))"
       
  2894             by (auto intro:finite_subset)
       
  2895         next
       
  2896           show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)"
       
  2897             by simp
       
  2898         qed 
       
  2899         moreover have "Max (cp s ` readys s) = cp s tm"
       
  2900         proof(rule Max_eqI)
       
  2901           from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
       
  2902         next
       
  2903           from finite_threads
       
  2904           show "finite (cp s ` readys s)" by (auto simp:readys_def)
       
  2905         next
       
  2906           fix y assume "y \<in> cp s ` readys s"
       
  2907           then obtain th1 where th1_readys: "th1 \<in> readys s"
       
  2908             and h: "y = cp s th1" by auto
       
  2909           show "y \<le> cp s tm"
       
  2910             apply(unfold cp_eq_p h)
       
  2911             apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
       
  2912           proof -
       
  2913             from finite_threads
       
  2914             show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
       
  2915           next
       
  2916             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}"
       
  2917               by simp
       
  2918           next
       
  2919             from dependants_threads[of th1] th1_readys
       
  2920             show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) 
       
  2921                     \<subseteq> (\<lambda>th. preced th s) ` threads s"
       
  2922               by (auto simp:readys_def)
       
  2923           qed
       
  2924         qed
       
  2925         ultimately show " Max (cp s ` readys s) = preced tm s" by simp
       
  2926       qed 
       
  2927     qed
       
  2928   qed
       
  2929 qed
       
  2930 
       
  2931 text {* (* ccc *) \noindent
       
  2932   Since the current precedence of the threads in ready queue will always be boosted,
       
  2933   there must be one inside it has the maximum precedence of the whole system. 
       
  2934 *}
       
  2935 lemma max_cp_readys_threads:
       
  2936   shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
       
  2937 proof(cases "threads s = {}")
       
  2938   case True
       
  2939   thus ?thesis 
       
  2940     by (auto simp:readys_def)
       
  2941 next
       
  2942   case False
       
  2943   show ?thesis by (rule max_cp_readys_threads_pre[OF False])
       
  2944 qed
       
  2945 
       
  2946 end
       
  2947 
       
  2948 lemma eq_holding: "holding (wq s) th cs = holding s th cs"
       
  2949   apply (unfold s_holding_def cs_holding_def wq_def, simp)
       
  2950   done
       
  2951 
       
  2952 lemma f_image_eq:
       
  2953   assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
       
  2954   shows "f ` A = g ` A"
       
  2955 proof
       
  2956   show "f ` A \<subseteq> g ` A"
       
  2957     by(rule image_subsetI, auto intro:h)
       
  2958 next
       
  2959   show "g ` A \<subseteq> f ` A"
       
  2960    by (rule image_subsetI, auto intro:h[symmetric])
       
  2961 qed
       
  2962 
       
  2963 
       
  2964 definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
       
  2965   where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
       
  2966 
       
  2967 
       
  2968 lemma detached_test:
       
  2969   shows "detached s th = (Th th \<notin> Field (RAG s))"
       
  2970 apply(simp add: detached_def Field_def)
       
  2971 apply(simp add: s_RAG_def)
       
  2972 apply(simp add: s_holding_abv s_waiting_abv)
       
  2973 apply(simp add: Domain_iff Range_iff)
       
  2974 apply(simp add: wq_def)
       
  2975 apply(auto)
       
  2976 done
       
  2977 
       
  2978 context valid_trace
       
  2979 begin
       
  2980 
       
  2981 lemma detached_intro:
       
  2982   assumes eq_pv: "cntP s th = cntV s th"
       
  2983   shows "detached s th"
       
  2984 proof -
       
  2985  from cnp_cnv_cncs
       
  2986   have eq_cnt: "cntP s th =
       
  2987     cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
       
  2988   hence cncs_zero: "cntCS s th = 0"
       
  2989     by (auto simp:eq_pv split:if_splits)
       
  2990   with eq_cnt
       
  2991   have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv)
       
  2992   thus ?thesis
       
  2993   proof
       
  2994     assume "th \<notin> threads s"
       
  2995     with range_in dm_RAG_threads
       
  2996     show ?thesis
       
  2997       by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff)
       
  2998   next
       
  2999     assume "th \<in> readys s"
       
  3000     moreover have "Th th \<notin> Range (RAG s)"
       
  3001     proof -
       
  3002       from card_0_eq [OF finite_holding] and cncs_zero
       
  3003       have "holdents s th = {}"
       
  3004         by (simp add:cntCS_def)
       
  3005       thus ?thesis
       
  3006         apply(auto simp:holdents_test)
       
  3007         apply(case_tac a)
       
  3008         apply(auto simp:holdents_test s_RAG_def)
       
  3009         done
       
  3010     qed
       
  3011     ultimately show ?thesis
       
  3012       by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def)
       
  3013   qed
       
  3014 qed
       
  3015 
       
  3016 lemma detached_elim:
       
  3017   assumes dtc: "detached s th"
       
  3018   shows "cntP s th = cntV s th"
       
  3019 proof -
       
  3020   from cnp_cnv_cncs
       
  3021   have eq_pv: " cntP s th =
       
  3022     cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" .
       
  3023   have cncs_z: "cntCS s th = 0"
       
  3024   proof -
       
  3025     from dtc have "holdents s th = {}"
       
  3026       unfolding detached_def holdents_test s_RAG_def
       
  3027       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
       
  3028     thus ?thesis by (auto simp:cntCS_def)
       
  3029   qed
       
  3030   show ?thesis
       
  3031   proof(cases "th \<in> threads s")
       
  3032     case True
       
  3033     with dtc 
       
  3034     have "th \<in> readys s"
       
  3035       by (unfold readys_def detached_def Field_def Domain_def Range_def, 
       
  3036            auto simp:eq_waiting s_RAG_def)
       
  3037     with cncs_z and eq_pv show ?thesis by simp
       
  3038   next
       
  3039     case False
       
  3040     with cncs_z and eq_pv show ?thesis by simp
       
  3041   qed
       
  3042 qed
       
  3043 
       
  3044 lemma detached_eq:
       
  3045   shows "(detached s th) = (cntP s th = cntV s th)"
       
  3046   by (insert vt, auto intro:detached_intro detached_elim)
       
  3047 
       
  3048 end
       
  3049 
       
  3050 text {* 
       
  3051   The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived
       
  3052   from the concise and miniature model of PIP given in PrioGDef.thy.
       
  3053 *}
       
  3054 
       
  3055 lemma eq_dependants: "dependants (wq s) = dependants s"
       
  3056   by (simp add: s_dependants_abv wq_def)
       
  3057 
       
  3058 lemma next_th_unique: 
       
  3059   assumes nt1: "next_th s th cs th1"
       
  3060   and nt2: "next_th s th cs th2"
       
  3061   shows "th1 = th2"
       
  3062 using assms by (unfold next_th_def, auto)
       
  3063 
       
  3064 lemma birth_time_lt:  "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
       
  3065   apply (induct s, simp)
       
  3066 proof -
       
  3067   fix a s
       
  3068   assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s"
       
  3069     and eq_as: "a # s \<noteq> []"
       
  3070   show "last_set th (a # s) < length (a # s)"
       
  3071   proof(cases "s \<noteq> []")
       
  3072     case False
       
  3073     from False show ?thesis
       
  3074       by (cases a, auto simp:last_set.simps)
       
  3075   next
       
  3076     case True
       
  3077     from ih [OF True] show ?thesis
       
  3078       by (cases a, auto simp:last_set.simps)
       
  3079   qed
       
  3080 qed
       
  3081 
       
  3082 lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []"
       
  3083   by (induct s, auto simp:threads.simps)
       
  3084 
       
  3085 lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s"
       
  3086   apply (drule_tac th_in_ne)
       
  3087   by (unfold preced_def, auto intro: birth_time_lt)
       
  3088 
       
  3089 lemma inj_the_preced: 
       
  3090   "inj_on (the_preced s) (threads s)"
       
  3091   by (metis inj_onI preced_unique the_preced_def)
       
  3092 
       
  3093 lemma tRAG_alt_def: 
       
  3094   "tRAG s = {(Th th1, Th th2) | th1 th2. 
       
  3095                   \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
       
  3096  by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
       
  3097 
       
  3098 lemma tRAG_Field:
       
  3099   "Field (tRAG s) \<subseteq> Field (RAG s)"
       
  3100   by (unfold tRAG_alt_def Field_def, auto)
       
  3101 
       
  3102 lemma tRAG_ancestorsE:
       
  3103   assumes "x \<in> ancestors (tRAG s) u"
       
  3104   obtains th where "x = Th th"
       
  3105 proof -
       
  3106   from assms have "(u, x) \<in> (tRAG s)^+" 
       
  3107       by (unfold ancestors_def, auto)
       
  3108   from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
       
  3109   then obtain th where "x = Th th"
       
  3110     by (unfold tRAG_alt_def, auto)
       
  3111   from that[OF this] show ?thesis .
       
  3112 qed
       
  3113 
       
  3114 lemma tRAG_mono:
       
  3115   assumes "RAG s' \<subseteq> RAG s"
       
  3116   shows "tRAG s' \<subseteq> tRAG s"
       
  3117   using assms 
       
  3118   by (unfold tRAG_alt_def, auto)
       
  3119 
       
  3120 lemma holding_next_thI:
       
  3121   assumes "holding s th cs"
       
  3122   and "length (wq s cs) > 1"
       
  3123   obtains th' where "next_th s th cs th'"
       
  3124 proof -
       
  3125   from assms(1)[folded eq_holding, unfolded cs_holding_def]
       
  3126   have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" .
       
  3127   then obtain rest where h1: "wq s cs = th#rest" 
       
  3128     by (cases "wq s cs", auto)
       
  3129   with assms(2) have h2: "rest \<noteq> []" by auto
       
  3130   let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
  3131   have "next_th s th cs ?th'" using  h1(1) h2 
       
  3132     by (unfold next_th_def, auto)
       
  3133   from that[OF this] show ?thesis .
       
  3134 qed
       
  3135 
       
  3136 lemma RAG_tRAG_transfer:
       
  3137   assumes "vt s'"
       
  3138   assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
       
  3139   and "(Cs cs, Th th'') \<in> RAG s'"
       
  3140   shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
       
  3141 proof -
       
  3142   interpret vt_s': valid_trace "s'" using assms(1)
       
  3143     by (unfold_locales, simp)
       
  3144   interpret rtree: rtree "RAG s'"
       
  3145   proof
       
  3146   show "single_valued (RAG s')"
       
  3147   apply (intro_locales)
       
  3148     by (unfold single_valued_def, 
       
  3149         auto intro:vt_s'.unique_RAG)
       
  3150 
       
  3151   show "acyclic (RAG s')"
       
  3152      by (rule vt_s'.acyclic_RAG)
       
  3153   qed
       
  3154   { fix n1 n2
       
  3155     assume "(n1, n2) \<in> ?L"
       
  3156     from this[unfolded tRAG_alt_def]
       
  3157     obtain th1 th2 cs' where 
       
  3158       h: "n1 = Th th1" "n2 = Th th2" 
       
  3159          "(Th th1, Cs cs') \<in> RAG s"
       
  3160          "(Cs cs', Th th2) \<in> RAG s" by auto
       
  3161     from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
       
  3162     from h(3) and assms(2) 
       
  3163     have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
       
  3164           (Th th1, Cs cs') \<in> RAG s'" by auto
       
  3165     hence "(n1, n2) \<in> ?R"
       
  3166     proof
       
  3167       assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
       
  3168       hence eq_th1: "th1 = th" by simp
       
  3169       moreover have "th2 = th''"
       
  3170       proof -
       
  3171         from h1 have "cs' = cs" by simp
       
  3172         from assms(3) cs_in[unfolded this] rtree.sgv
       
  3173         show ?thesis
       
  3174           by (unfold single_valued_def, auto)
       
  3175       qed
       
  3176       ultimately show ?thesis using h(1,2) by auto
       
  3177     next
       
  3178       assume "(Th th1, Cs cs') \<in> RAG s'"
       
  3179       with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
       
  3180         by (unfold tRAG_alt_def, auto)
       
  3181       from this[folded h(1, 2)] show ?thesis by auto
       
  3182     qed
       
  3183   } moreover {
       
  3184     fix n1 n2
       
  3185     assume "(n1, n2) \<in> ?R"
       
  3186     hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
       
  3187     hence "(n1, n2) \<in> ?L" 
       
  3188     proof
       
  3189       assume "(n1, n2) \<in> tRAG s'"
       
  3190       moreover have "... \<subseteq> ?L"
       
  3191       proof(rule tRAG_mono)
       
  3192         show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
       
  3193       qed
       
  3194       ultimately show ?thesis by auto
       
  3195     next
       
  3196       assume eq_n: "(n1, n2) = (Th th, Th th'')"
       
  3197       from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
       
  3198       moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
       
  3199       ultimately show ?thesis 
       
  3200         by (unfold eq_n tRAG_alt_def, auto)
       
  3201     qed
       
  3202   } ultimately show ?thesis by auto
       
  3203 qed
       
  3204 
       
  3205 context valid_trace
       
  3206 begin
       
  3207 
       
  3208 lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
       
  3209 
       
  3210 end
       
  3211 
       
  3212 lemma cp_alt_def:
       
  3213   "cp s th =  
       
  3214            Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})"
       
  3215 proof -
       
  3216   have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) =
       
  3217         Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" 
       
  3218           (is "Max (_ ` ?L) = Max (_ ` ?R)")
       
  3219   proof -
       
  3220     have "?L = ?R" 
       
  3221     by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def)
       
  3222     thus ?thesis by simp
       
  3223   qed
       
  3224   thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
       
  3225 qed
       
  3226 
  3842 
  3227 lemma cp_gen_alt_def:
  3843 lemma cp_gen_alt_def:
  3228   "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
  3844   "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
  3229     by (auto simp:cp_gen_def)
  3845     by (auto simp:cp_gen_def)
  3230 
  3846 
  3270 lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
  3886 lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x"
  3271 proof -
  3887 proof -
  3272   { fix a
  3888   { fix a
  3273     assume "a \<in> subtree (tRAG s) x"
  3889     assume "a \<in> subtree (tRAG s) x"
  3274     hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
  3890     hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def)
  3275     with tRAG_star_RAG[of s]
  3891     with tRAG_star_RAG
  3276     have "(a, x) \<in> (RAG s)^*" by auto
  3892     have "(a, x) \<in> (RAG s)^*" by auto
  3277     hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
  3893     hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def)
  3278   } thus ?thesis by auto
  3894   } thus ?thesis by auto
  3279 qed
  3895 qed
  3280 
  3896 
  3286   { fix th'
  3902   { fix th'
  3287     assume "th' \<in> ?L"
  3903     assume "th' \<in> ?L"
  3288     hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
  3904     hence "(Th th', Th th) \<in> (tRAG s)^+" by auto
  3289     from tranclD[OF this]
  3905     from tranclD[OF this]
  3290     obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
  3906     obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto
  3291     from tRAG_subtree_RAG[of s] and this(2)
  3907     from tRAG_subtree_RAG and this(2)
  3292     have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
  3908     have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) 
  3293     moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
  3909     moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto 
  3294     ultimately have "th' \<in> ?R"  by auto 
  3910     ultimately have "th' \<in> ?R"  by auto 
  3295   } moreover 
  3911   } moreover 
  3296   { fix th'
  3912   { fix th'
  3305       show ?case
  3921       show ?case
  3306       proof(cases "xs1")
  3922       proof(cases "xs1")
  3307         case Nil
  3923         case Nil
  3308         from 1(2)[unfolded Cons1 Nil]
  3924         from 1(2)[unfolded Cons1 Nil]
  3309         have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
  3925         have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
  3310         hence "(Th th', x1) \<in> (RAG s)" by (cases, simp)
  3926         hence "(Th th', x1) \<in> (RAG s)" 
       
  3927           by (cases, auto)
  3311         then obtain cs where "x1 = Cs cs" 
  3928         then obtain cs where "x1 = Cs cs" 
  3312               by (unfold s_RAG_def, auto)
  3929               by (unfold s_RAG_def, auto)
  3313         from rpath_nnl_lastE[OF rp[unfolded this]]
  3930         from rpath_nnl_lastE[OF rp[unfolded this]]
  3314         show ?thesis by auto
  3931         show ?thesis by auto
  3315       next
  3932       next
  3357     using tRAG_trancl_eq by auto
  3974     using tRAG_trancl_eq by auto
  3358 
  3975 
  3359 lemma dependants_alt_def:
  3976 lemma dependants_alt_def:
  3360   "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
  3977   "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}"
  3361   by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
  3978   by (metis eq_RAG s_dependants_def tRAG_trancl_eq)
  3362   
  3979 
       
  3980 lemma dependants_alt_def1:
       
  3981   "dependants (s::state) th = {th'. (Th th', Th th) \<in> (RAG s)^+}"
       
  3982   using dependants_alt_def tRAG_trancl_eq by auto
       
  3983 
       
  3984 context valid_trace
       
  3985 begin
       
  3986 lemma count_eq_RAG_plus:
       
  3987   assumes "cntP s th = cntV s th"
       
  3988   shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  3989 proof(rule ccontr)
       
  3990     assume otherwise: "{th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} \<noteq> {}"
       
  3991     then obtain th' where "(Th th', Th th) \<in> (RAG s)^+" by auto
       
  3992     from tranclD2[OF this]
       
  3993     obtain z where "z \<in> children (RAG s) (Th th)" 
       
  3994       by (auto simp:children_def)
       
  3995     with eq_pv_children[OF assms]
       
  3996     show False by simp
       
  3997 qed
       
  3998 
       
  3999 lemma eq_pv_dependants:
       
  4000   assumes eq_pv: "cntP s th = cntV s th"
       
  4001   shows "dependants s th = {}"
       
  4002 proof -
       
  4003   from count_eq_RAG_plus[OF assms, folded dependants_alt_def1]
       
  4004   show ?thesis .
       
  4005 qed
       
  4006 
       
  4007 end
       
  4008 
       
  4009 lemma eq_dependants: "dependants (wq s) = dependants s"
       
  4010   by (simp add: s_dependants_abv wq_def)
       
  4011 
  3363 context valid_trace
  4012 context valid_trace
  3364 begin
  4013 begin
  3365 
  4014 
  3366 lemma count_eq_tRAG_plus:
  4015 lemma count_eq_tRAG_plus:
  3367   assumes "cntP s th = cntV s th"
  4016   assumes "cntP s th = cntV s th"
  3368   shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
  4017   shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
  3369   using assms count_eq_dependants dependants_alt_def eq_dependants by auto 
  4018   using assms eq_pv_dependants dependants_alt_def eq_dependants by auto 
  3370 
       
  3371 lemma count_eq_RAG_plus:
       
  3372   assumes "cntP s th = cntV s th"
       
  3373   shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
       
  3374   using assms count_eq_dependants cs_dependants_def eq_RAG by auto
       
  3375 
  4019 
  3376 lemma count_eq_RAG_plus_Th:
  4020 lemma count_eq_RAG_plus_Th:
  3377   assumes "cntP s th = cntV s th"
  4021   assumes "cntP s th = cntV s th"
  3378   shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
  4022   shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}"
  3379   using count_eq_RAG_plus[OF assms] by auto
  4023   using count_eq_RAG_plus[OF assms] by auto
  3380 
  4024 
  3381 lemma count_eq_tRAG_plus_Th:
  4025 lemma count_eq_tRAG_plus_Th:
  3382   assumes "cntP s th = cntV s th"
  4026   assumes "cntP s th = cntV s th"
  3383   shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
  4027   shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}"
  3384    using count_eq_tRAG_plus[OF assms] by auto
  4028    using count_eq_tRAG_plus[OF assms] by auto
       
  4029 end
       
  4030 
       
  4031 lemma inj_the_preced: 
       
  4032   "inj_on (the_preced s) (threads s)"
       
  4033   by (metis inj_onI preced_unique the_preced_def)
       
  4034 
       
  4035 lemma tRAG_Field:
       
  4036   "Field (tRAG s) \<subseteq> Field (RAG s)"
       
  4037   by (unfold tRAG_alt_def Field_def, auto)
       
  4038 
       
  4039 lemma tRAG_ancestorsE:
       
  4040   assumes "x \<in> ancestors (tRAG s) u"
       
  4041   obtains th where "x = Th th"
       
  4042 proof -
       
  4043   from assms have "(u, x) \<in> (tRAG s)^+" 
       
  4044       by (unfold ancestors_def, auto)
       
  4045   from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
       
  4046   then obtain th where "x = Th th"
       
  4047     by (unfold tRAG_alt_def, auto)
       
  4048   from that[OF this] show ?thesis .
       
  4049 qed
       
  4050 
       
  4051 lemma tRAG_mono:
       
  4052   assumes "RAG s' \<subseteq> RAG s"
       
  4053   shows "tRAG s' \<subseteq> tRAG s"
       
  4054   using assms 
       
  4055   by (unfold tRAG_alt_def, auto)
       
  4056 
       
  4057 lemma holding_next_thI:
       
  4058   assumes "holding s th cs"
       
  4059   and "length (wq s cs) > 1"
       
  4060   obtains th' where "next_th s th cs th'"
       
  4061 proof -
       
  4062   from assms(1)[folded holding_eq, unfolded cs_holding_def]
       
  4063   have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" 
       
  4064     by (unfold s_holding_def, fold wq_def, auto)
       
  4065   then obtain rest where h1: "wq s cs = th#rest" 
       
  4066     by (cases "wq s cs", auto)
       
  4067   with assms(2) have h2: "rest \<noteq> []" by auto
       
  4068   let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
  4069   have "next_th s th cs ?th'" using  h1(1) h2 
       
  4070     by (unfold next_th_def, auto)
       
  4071   from that[OF this] show ?thesis .
       
  4072 qed
       
  4073 
       
  4074 lemma RAG_tRAG_transfer:
       
  4075   assumes "vt s'"
       
  4076   assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
       
  4077   and "(Cs cs, Th th'') \<in> RAG s'"
       
  4078   shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
       
  4079 proof -
       
  4080   interpret vt_s': valid_trace "s'" using assms(1)
       
  4081     by (unfold_locales, simp)
       
  4082   { fix n1 n2
       
  4083     assume "(n1, n2) \<in> ?L"
       
  4084     from this[unfolded tRAG_alt_def]
       
  4085     obtain th1 th2 cs' where 
       
  4086       h: "n1 = Th th1" "n2 = Th th2" 
       
  4087          "(Th th1, Cs cs') \<in> RAG s"
       
  4088          "(Cs cs', Th th2) \<in> RAG s" by auto
       
  4089     from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
       
  4090     from h(3) and assms(2) 
       
  4091     have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
       
  4092           (Th th1, Cs cs') \<in> RAG s'" by auto
       
  4093     hence "(n1, n2) \<in> ?R"
       
  4094     proof
       
  4095       assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
       
  4096       hence eq_th1: "th1 = th" by simp
       
  4097       moreover have "th2 = th''"
       
  4098       proof -
       
  4099         from h1 have "cs' = cs" by simp
       
  4100         from assms(3) cs_in[unfolded this]
       
  4101         show ?thesis using vt_s'.unique_RAG by auto 
       
  4102       qed
       
  4103       ultimately show ?thesis using h(1,2) by auto
       
  4104     next
       
  4105       assume "(Th th1, Cs cs') \<in> RAG s'"
       
  4106       with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
       
  4107         by (unfold tRAG_alt_def, auto)
       
  4108       from this[folded h(1, 2)] show ?thesis by auto
       
  4109     qed
       
  4110   } moreover {
       
  4111     fix n1 n2
       
  4112     assume "(n1, n2) \<in> ?R"
       
  4113     hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
       
  4114     hence "(n1, n2) \<in> ?L" 
       
  4115     proof
       
  4116       assume "(n1, n2) \<in> tRAG s'"
       
  4117       moreover have "... \<subseteq> ?L"
       
  4118       proof(rule tRAG_mono)
       
  4119         show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
       
  4120       qed
       
  4121       ultimately show ?thesis by auto
       
  4122     next
       
  4123       assume eq_n: "(n1, n2) = (Th th, Th th'')"
       
  4124       from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
       
  4125       moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
       
  4126       ultimately show ?thesis 
       
  4127         by (unfold eq_n tRAG_alt_def, auto)
       
  4128     qed
       
  4129   } ultimately show ?thesis by auto
       
  4130 qed
       
  4131 
       
  4132 context valid_trace
       
  4133 begin
       
  4134 
       
  4135 lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt]
  3385 
  4136 
  3386 end
  4137 end
  3387 
  4138 
  3388 lemma tRAG_subtree_eq: 
  4139 lemma tRAG_subtree_eq: 
  3389    "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
  4140    "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th'  \<in> (subtree (RAG s) (Th th))}"
  3437   obtain th where eq_a: "a = Th th" by auto
  4188   obtain th where eq_a: "a = Th th" by auto
  3438   show "cp_gen s a = (cp s \<circ> the_thread) a"
  4189   show "cp_gen s a = (cp s \<circ> the_thread) a"
  3439     by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
  4190     by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
  3440 qed
  4191 qed
  3441 
  4192 
  3442 
       
  3443 context valid_trace
  4193 context valid_trace
  3444 begin
  4194 begin
  3445 
       
  3446 lemma RAG_threads:
       
  3447   assumes "(Th th) \<in> Field (RAG s)"
       
  3448   shows "th \<in> threads s"
       
  3449   using assms
       
  3450   by (metis Field_def UnE dm_RAG_threads range_in vt)
       
  3451 
  4195 
  3452 lemma subtree_tRAG_thread:
  4196 lemma subtree_tRAG_thread:
  3453   assumes "th \<in> threads s"
  4197   assumes "th \<in> threads s"
  3454   shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
  4198   shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R")
  3455 proof -
  4199 proof -
  3508 lemma not_in_thread_isolated:
  4252 lemma not_in_thread_isolated:
  3509   assumes "th \<notin> threads s"
  4253   assumes "th \<notin> threads s"
  3510   shows "(Th th) \<notin> Field (RAG s)"
  4254   shows "(Th th) \<notin> Field (RAG s)"
  3511 proof
  4255 proof
  3512   assume "(Th th) \<in> Field (RAG s)"
  4256   assume "(Th th) \<in> Field (RAG s)"
  3513   with dm_RAG_threads and range_in assms
  4257   with dm_RAG_threads and rg_RAG_threads assms
  3514   show False by (unfold Field_def, blast)
  4258   show False by (unfold Field_def, blast)
  3515 qed
  4259 qed
  3516 
  4260 
  3517 lemma wf_RAG: "wf (RAG s)"
       
  3518 proof(rule finite_acyclic_wf)
       
  3519   from finite_RAG show "finite (RAG s)" .
       
  3520 next
       
  3521   from acyclic_RAG show "acyclic (RAG s)" .
       
  3522 qed
       
  3523 
       
  3524 lemma sgv_wRAG: "single_valued (wRAG s)"
       
  3525   using waiting_unique
       
  3526   by (unfold single_valued_def wRAG_def, auto)
       
  3527 
       
  3528 lemma sgv_hRAG: "single_valued (hRAG s)"
       
  3529   using holding_unique 
       
  3530   by (unfold single_valued_def hRAG_def, auto)
       
  3531 
       
  3532 lemma sgv_tRAG: "single_valued (tRAG s)"
       
  3533   by (unfold tRAG_def, rule single_valued_relcomp, 
       
  3534               insert sgv_wRAG sgv_hRAG, auto)
       
  3535 
       
  3536 lemma acyclic_tRAG: "acyclic (tRAG s)"
       
  3537 proof(unfold tRAG_def, rule acyclic_compose)
       
  3538   show "acyclic (RAG s)" using acyclic_RAG .
       
  3539 next
       
  3540   show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
  3541 next
       
  3542   show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
  3543 qed
       
  3544 
       
  3545 lemma sgv_RAG: "single_valued (RAG s)"
       
  3546   using unique_RAG by (auto simp:single_valued_def)
       
  3547 
       
  3548 lemma rtree_RAG: "rtree (RAG s)"
       
  3549   using sgv_RAG acyclic_RAG
       
  3550   by (unfold rtree_def rtree_axioms_def sgv_def, auto)
       
  3551 
       
  3552 end
  4261 end
  3553 
  4262 
  3554 sublocale valid_trace < rtree_RAG: rtree "RAG s"
  4263 definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool"
  3555 proof
  4264   where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))"
  3556   show "single_valued (RAG s)"
  4265 
  3557   apply (intro_locales)
  4266 
  3558     by (unfold single_valued_def, 
  4267 lemma detached_test:
  3559         auto intro:unique_RAG)
  4268   shows "detached s th = (Th th \<notin> Field (RAG s))"
  3560 
  4269 apply(simp add: detached_def Field_def)
  3561   show "acyclic (RAG s)"
  4270 apply(simp add: s_RAG_def)
  3562      by (rule acyclic_RAG)
  4271 apply(simp add: s_holding_abv s_waiting_abv)
  3563 qed
  4272 apply(simp add: Domain_iff Range_iff)
  3564 
  4273 apply(simp add: wq_def)
  3565 sublocale valid_trace < rtree_s: rtree "tRAG s"
  4274 apply(auto)
  3566 proof(unfold_locales)
  4275 done
  3567   from sgv_tRAG show "single_valued (tRAG s)" .
       
  3568 next
       
  3569   from acyclic_tRAG show "acyclic (tRAG s)" .
       
  3570 qed
       
  3571 
       
  3572 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
       
  3573 proof -
       
  3574   show "fsubtree (RAG s)"
       
  3575   proof(intro_locales)
       
  3576     show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] .
       
  3577   next
       
  3578     show "fsubtree_axioms (RAG s)"
       
  3579     proof(unfold fsubtree_axioms_def)
       
  3580       from wf_RAG show "wf (RAG s)" .
       
  3581     qed
       
  3582   qed
       
  3583 qed
       
  3584 
       
  3585 sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
       
  3586 proof -
       
  3587   have "fsubtree (tRAG s)"
       
  3588   proof -
       
  3589     have "fbranch (tRAG s)"
       
  3590     proof(unfold tRAG_def, rule fbranch_compose)
       
  3591         show "fbranch (wRAG s)"
       
  3592         proof(rule finite_fbranchI)
       
  3593            from finite_RAG show "finite (wRAG s)"
       
  3594            by (unfold RAG_split, auto)
       
  3595         qed
       
  3596     next
       
  3597         show "fbranch (hRAG s)"
       
  3598         proof(rule finite_fbranchI)
       
  3599            from finite_RAG 
       
  3600            show "finite (hRAG s)" by (unfold RAG_split, auto)
       
  3601         qed
       
  3602     qed
       
  3603     moreover have "wf (tRAG s)"
       
  3604     proof(rule wf_subset)
       
  3605       show "wf (RAG s O RAG s)" using wf_RAG
       
  3606         by (fold wf_comp_self, simp)
       
  3607     next
       
  3608       show "tRAG s \<subseteq> (RAG s O RAG s)"
       
  3609         by (unfold tRAG_alt_def, auto)
       
  3610     qed
       
  3611     ultimately show ?thesis
       
  3612       by (unfold fsubtree_def fsubtree_axioms_def,auto)
       
  3613   qed
       
  3614   from this[folded tRAG_def] show "fsubtree (tRAG s)" .
       
  3615 qed
       
  3616 
       
  3617 lemma Max_UNION: 
       
  3618   assumes "finite A"
       
  3619   and "A \<noteq> {}"
       
  3620   and "\<forall> M \<in> f ` A. finite M"
       
  3621   and "\<forall> M \<in> f ` A. M \<noteq> {}"
       
  3622   shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
       
  3623   using assms[simp]
       
  3624 proof -
       
  3625   have "?L = Max (\<Union>(f ` A))"
       
  3626     by (fold Union_image_eq, simp)
       
  3627   also have "... = ?R"
       
  3628     by (subst Max_Union, simp+)
       
  3629   finally show ?thesis .
       
  3630 qed
       
  3631 
       
  3632 lemma max_Max_eq:
       
  3633   assumes "finite A"
       
  3634     and "A \<noteq> {}"
       
  3635     and "x = y"
       
  3636   shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
       
  3637 proof -
       
  3638   have "?R = Max (insert y A)" by simp
       
  3639   also from assms have "... = ?L"
       
  3640       by (subst Max.insert, simp+)
       
  3641   finally show ?thesis by simp
       
  3642 qed
       
  3643 
  4276 
  3644 context valid_trace
  4277 context valid_trace
  3645 begin
  4278 begin
  3646 
  4279 
       
  4280 lemma detached_intro:
       
  4281   assumes eq_pv: "cntP s th = cntV s th"
       
  4282   shows "detached s th"
       
  4283 proof -
       
  4284   from eq_pv cnp_cnv_cncs
       
  4285   have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:pvD_def)
       
  4286   thus ?thesis
       
  4287   proof
       
  4288     assume "th \<notin> threads s"
       
  4289     with rg_RAG_threads dm_RAG_threads
       
  4290     show ?thesis
       
  4291       by (auto simp add: detached_def s_RAG_def s_waiting_abv 
       
  4292               s_holding_abv wq_def Domain_iff Range_iff)
       
  4293   next
       
  4294     assume "th \<in> readys s"
       
  4295     moreover have "Th th \<notin> Range (RAG s)"
       
  4296     proof -
       
  4297       from eq_pv_children[OF assms]
       
  4298       have "children (RAG s) (Th th) = {}" .
       
  4299       thus ?thesis
       
  4300       by (unfold children_def, auto)
       
  4301     qed
       
  4302     ultimately show ?thesis
       
  4303       by (auto simp add: detached_def s_RAG_def s_waiting_abv 
       
  4304               s_holding_abv wq_def readys_def)
       
  4305   qed
       
  4306 qed
       
  4307 
       
  4308 lemma detached_elim:
       
  4309   assumes dtc: "detached s th"
       
  4310   shows "cntP s th = cntV s th"
       
  4311 proof -
       
  4312   have cncs_z: "cntCS s th = 0"
       
  4313   proof -
       
  4314     from dtc have "holdents s th = {}"
       
  4315       unfolding detached_def holdents_test s_RAG_def
       
  4316       by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff)
       
  4317     thus ?thesis by (auto simp:cntCS_def)
       
  4318   qed
       
  4319   show ?thesis
       
  4320   proof(cases "th \<in> threads s")
       
  4321     case True
       
  4322     with dtc 
       
  4323     have "th \<in> readys s"
       
  4324       by (unfold readys_def detached_def Field_def Domain_def Range_def, 
       
  4325            auto simp:waiting_eq s_RAG_def)
       
  4326     with cncs_z  show ?thesis using cnp_cnv_cncs by (simp add:pvD_def)
       
  4327   next
       
  4328     case False
       
  4329     with cncs_z and cnp_cnv_cncs show ?thesis by (simp add:pvD_def)
       
  4330   qed
       
  4331 qed
       
  4332 
       
  4333 lemma detached_eq:
       
  4334   shows "(detached s th) = (cntP s th = cntV s th)"
       
  4335   by (insert vt, auto intro:detached_intro detached_elim)
       
  4336 
       
  4337 end
       
  4338 
       
  4339 context valid_trace
       
  4340 begin
  3647 (* ddd *)
  4341 (* ddd *)
  3648 lemma cp_gen_rec:
  4342 lemma cp_gen_rec:
  3649   assumes "x = Th th"
  4343   assumes "x = Th th"
  3650   shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
  4344   shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
  3651 proof(cases "children (tRAG s) x = {}")
  4345 proof(cases "children (tRAG s) x = {}")
  3718     qed
  4412     qed
  3719     thus ?thesis by (subst (1) h(1), unfold h(2), simp)
  4413     thus ?thesis by (subst (1) h(1), unfold h(2), simp)
  3720   qed
  4414   qed
  3721 qed
  4415 qed
  3722 
  4416 
  3723 end
       
  3724 
       
  3725 (* keep *)
       
  3726 lemma next_th_holding:
  4417 lemma next_th_holding:
  3727   assumes vt: "vt s"
  4418   assumes nxt: "next_th s th cs th'"
  3728   and nxt: "next_th s th cs th'"
       
  3729   shows "holding (wq s) th cs"
  4419   shows "holding (wq s) th cs"
  3730 proof -
  4420 proof -
  3731   from nxt[unfolded next_th_def]
  4421   from nxt[unfolded next_th_def]
  3732   obtain rest where h: "wq s cs = th # rest"
  4422   obtain rest where h: "wq s cs = th # rest"
  3733                        "rest \<noteq> []" 
  4423                        "rest \<noteq> []" 
  3734                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
  4424                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
  3735   thus ?thesis
  4425   thus ?thesis
  3736     by (unfold cs_holding_def, auto)
  4426     by (unfold cs_holding_def, auto)
  3737 qed
  4427 qed
  3738 
       
  3739 context valid_trace
       
  3740 begin
       
  3741 
  4428 
  3742 lemma next_th_waiting:
  4429 lemma next_th_waiting:
  3743   assumes nxt: "next_th s th cs th'"
  4430   assumes nxt: "next_th s th cs th'"
  3744   shows "waiting (wq s) th' cs"
  4431   shows "waiting (wq s) th' cs"
  3745 proof -
  4432 proof -
  3769   using vt assms next_th_holding next_th_waiting
  4456   using vt assms next_th_holding next_th_waiting
  3770   by (unfold s_RAG_def, simp)
  4457   by (unfold s_RAG_def, simp)
  3771 
  4458 
  3772 end
  4459 end
  3773 
  4460 
  3774 -- {* A useless definition *}
  4461 lemma next_th_unique: 
  3775 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
  4462   assumes nt1: "next_th s th cs th1"
  3776 where "cps s = {(th, cp s th) | th . th \<in> threads s}"
  4463   and nt2: "next_th s th cs th2"
  3777 
  4464   shows "th1 = th2"
  3778 
  4465 using assms by (unfold next_th_def, auto)
  3779 find_theorems "waiting" holding
  4466 
  3780 context valid_trace
  4467 context valid_trace
  3781 begin
  4468 begin
  3782 
  4469 
  3783 find_theorems "waiting" holding
  4470 thm th_chain_to_ready
       
  4471 
       
  4472 find_theorems subtree Th RAG
       
  4473 
       
  4474 lemma threads_alt_def:
       
  4475   "(threads s) = (\<Union> th \<in> readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)})"
       
  4476     (is "?L = ?R")
       
  4477 proof -
       
  4478   { fix th1
       
  4479     assume "th1 \<in> ?L"
       
  4480     from th_chain_to_ready[OF this]
       
  4481     have "th1 \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th th1, Th th') \<in> (RAG s)\<^sup>+)" .
       
  4482     hence "th1 \<in> ?R" by (auto simp:subtree_def)
       
  4483   } moreover 
       
  4484   { fix th'
       
  4485     assume "th' \<in> ?R"
       
  4486     then obtain th where h: "th \<in> readys s" " Th th' \<in> subtree (RAG s) (Th th)"
       
  4487       by auto
       
  4488     from this(2)
       
  4489     have "th' \<in> ?L" 
       
  4490     proof(cases rule:subtreeE)
       
  4491       case 1
       
  4492       with h(1) show ?thesis by (auto simp:readys_def)
       
  4493     next
       
  4494       case 2
       
  4495       from tranclD[OF this(2)[unfolded ancestors_def, simplified]]
       
  4496       have "Th th' \<in> Domain (RAG s)" by auto
       
  4497       from dm_RAG_threads[OF this]
       
  4498       show ?thesis .
       
  4499     qed
       
  4500   } ultimately show ?thesis by auto
       
  4501 qed
       
  4502 
       
  4503 lemma finite_readys [simp]: "finite (readys s)"
       
  4504   using finite_threads readys_threads rev_finite_subset by blast
       
  4505 
       
  4506 text {* (* ccc *) \noindent
       
  4507   Since the current precedence of the threads in ready queue will always be boosted,
       
  4508   there must be one inside it has the maximum precedence of the whole system. 
       
  4509 *}
       
  4510 lemma max_cp_readys_threads:
       
  4511   shows "Max (cp s ` readys s) = Max (cp s ` threads s)" (is "?L = ?R")
       
  4512 proof(cases "readys s = {}")
       
  4513   case False
       
  4514   have "?R = Max (the_preced s ` threads s)" by (unfold max_cp_eq, simp)
       
  4515   also have "... = 
       
  4516     Max (the_preced s ` (\<Union>th\<in>readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)}))" 
       
  4517          by (unfold threads_alt_def, simp)
       
  4518   also have "... = 
       
  4519     Max ((\<Union>th\<in>readys s. the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}))"
       
  4520           by (unfold image_UN, simp)
       
  4521   also have "... = 
       
  4522     Max (Max ` (\<lambda>th. the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}) ` readys s)" 
       
  4523   proof(rule Max_UNION)
       
  4524     show "\<forall>M\<in>(\<lambda>x. the_preced s ` 
       
  4525                     {th'. Th th' \<in> subtree (RAG s) (Th x)}) ` readys s. finite M"
       
  4526                         using finite_subtree_threads by auto
       
  4527   qed (auto simp:False subtree_def)
       
  4528   also have "... =  
       
  4529     Max ((Max \<circ> (\<lambda>th. the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})) ` readys s)" 
       
  4530       by (unfold image_comp, simp)
       
  4531   also have "... = ?L" (is "Max (?f ` ?A) = Max (?g ` ?A)")
       
  4532   proof -
       
  4533     have "(?f ` ?A) = (?g ` ?A)"
       
  4534     proof(rule f_image_eq)
       
  4535       fix th1 
       
  4536       assume "th1 \<in> ?A"
       
  4537       thus "?f th1 = ?g th1"
       
  4538         by (unfold cp_alt_def, simp)
       
  4539     qed
       
  4540     thus ?thesis by simp
       
  4541   qed
       
  4542   finally show ?thesis by simp
       
  4543 qed (auto simp:threads_alt_def)
  3784 
  4544 
  3785 end
  4545 end
  3786 
  4546 
  3787 end
  4547 end
       
  4548