CpsG.thy
changeset 58 ad57323fd4d6
parent 56 0fd478e14e87
child 59 0a069a667301
equal deleted inserted replaced
57:f1b39d77db00 58:ad57323fd4d6
     4 *}
     4 *}
     5 theory CpsG
     5 theory CpsG
     6 imports PrioG Max RTree
     6 imports PrioG Max RTree
     7 begin
     7 begin
     8 
     8 
     9 locale pip = 
     9 definition "wRAG (s::state) = {(Th th, Cs cs) | th cs. waiting s th cs}"
    10   fixes s
    10 
    11   assumes vt: "vt s"
    11 definition "hRAG (s::state) =  {(Cs cs, Th th) | th cs. holding s th cs}"
    12 
    12 
    13 context pip
    13 definition "tRAG s = wRAG s O hRAG s"
    14 begin
    14 
    15 
    15 lemma RAG_split: "RAG s = (wRAG s \<union> hRAG s)"
    16 interpretation rtree_RAG: rtree "RAG s"
    16   by (unfold s_RAG_abv wRAG_def hRAG_def s_waiting_abv 
       
    17              s_holding_abv cs_RAG_def, auto)
       
    18 
       
    19 lemma tRAG_alt_def: 
       
    20   "tRAG s = {(Th th1, Th th2) | th1 th2. 
       
    21                   \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
       
    22  by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
       
    23 
       
    24 lemma tRAG_mono:
       
    25   assumes "RAG s' \<subseteq> RAG s"
       
    26   shows "tRAG s' \<subseteq> tRAG s"
       
    27   using assms 
       
    28   by (unfold tRAG_alt_def, auto)
       
    29 
       
    30 lemma holding_next_thI:
       
    31   assumes "holding s th cs"
       
    32   and "length (wq s cs) > 1"
       
    33   obtains th' where "next_th s th cs th'"
       
    34 proof -
       
    35   from assms(1)[folded eq_holding, unfolded cs_holding_def]
       
    36   have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" .
       
    37   then obtain rest where h1: "wq s cs = th#rest" 
       
    38     by (cases "wq s cs", auto)
       
    39   with assms(2) have h2: "rest \<noteq> []" by auto
       
    40   let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)"
       
    41   have "next_th s th cs ?th'" using  h1(1) h2 
       
    42     by (unfold next_th_def, auto)
       
    43   from that[OF this] show ?thesis .
       
    44 qed
       
    45 
       
    46 lemma RAG_tRAG_transfer:
       
    47   assumes "vt s'"
       
    48   assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
       
    49   and "(Cs cs, Th th'') \<in> RAG s'"
       
    50   shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R")
       
    51 proof -
       
    52   interpret rtree: rtree "RAG s'"
       
    53   proof
       
    54   show "single_valued (RAG s')"
       
    55   apply (intro_locales)
       
    56     by (unfold single_valued_def, 
       
    57         auto intro:unique_RAG[OF assms(1)])
       
    58 
       
    59   show "acyclic (RAG s')"
       
    60      by (rule acyclic_RAG[OF assms(1)])
       
    61   qed
       
    62   { fix n1 n2
       
    63     assume "(n1, n2) \<in> ?L"
       
    64     from this[unfolded tRAG_alt_def]
       
    65     obtain th1 th2 cs' where 
       
    66       h: "n1 = Th th1" "n2 = Th th2" 
       
    67          "(Th th1, Cs cs') \<in> RAG s"
       
    68          "(Cs cs', Th th2) \<in> RAG s" by auto
       
    69     from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto
       
    70     from h(3) and assms(2) 
       
    71     have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> 
       
    72           (Th th1, Cs cs') \<in> RAG s'" by auto
       
    73     hence "(n1, n2) \<in> ?R"
       
    74     proof
       
    75       assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)"
       
    76       hence eq_th1: "th1 = th" by simp
       
    77       moreover have "th2 = th''"
       
    78       proof -
       
    79         from h1 have "cs' = cs" by simp
       
    80         from assms(3) cs_in[unfolded this] rtree.sgv
       
    81         show ?thesis
       
    82           by (unfold single_valued_def, auto)
       
    83       qed
       
    84       ultimately show ?thesis using h(1,2) by auto
       
    85     next
       
    86       assume "(Th th1, Cs cs') \<in> RAG s'"
       
    87       with cs_in have "(Th th1, Th th2) \<in> tRAG s'"
       
    88         by (unfold tRAG_alt_def, auto)
       
    89       from this[folded h(1, 2)] show ?thesis by auto
       
    90     qed
       
    91   } moreover {
       
    92     fix n1 n2
       
    93     assume "(n1, n2) \<in> ?R"
       
    94     hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto
       
    95     hence "(n1, n2) \<in> ?L" 
       
    96     proof
       
    97       assume "(n1, n2) \<in> tRAG s'"
       
    98       moreover have "... \<subseteq> ?L"
       
    99       proof(rule tRAG_mono)
       
   100         show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto)
       
   101       qed
       
   102       ultimately show ?thesis by auto
       
   103     next
       
   104       assume eq_n: "(n1, n2) = (Th th, Th th'')"
       
   105       from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto
       
   106       moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto
       
   107       ultimately show ?thesis 
       
   108         by (unfold eq_n tRAG_alt_def, auto)
       
   109     qed
       
   110   } ultimately show ?thesis by auto
       
   111 qed
       
   112 
       
   113 lemma readys_root:
       
   114   assumes "vt s"
       
   115   and "th \<in> readys s"
       
   116   shows "root (RAG s) (Th th)"
       
   117 proof -
       
   118   { fix x
       
   119     assume "x \<in> ancestors (RAG s) (Th th)"
       
   120     hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def)
       
   121     from tranclD[OF this]
       
   122     obtain z where "(Th th, z) \<in> RAG s" by auto
       
   123     with assms(2) have False
       
   124          apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
       
   125          by (fold wq_def, blast)
       
   126   } thus ?thesis by (unfold root_def, auto)
       
   127 qed
       
   128 
       
   129 lemma readys_in_no_subtree:
       
   130   assumes "vt s"
       
   131   and "th \<in> readys s"
       
   132   and "th' \<noteq> th"
       
   133   shows "Th th \<notin> subtree (RAG s) (Th th')" 
    17 proof
   134 proof
    18   show "single_valued (RAG s)"
   135    assume "Th th \<in> subtree (RAG s) (Th th')"
    19   apply (intro_locales)
   136    thus False
    20     by (unfold single_valued_def, auto intro: unique_RAG[OF vt])
   137    proof(cases rule:subtreeE)
    21 
   138       case 1
    22   show "acyclic (RAG s)"
   139       with assms show ?thesis by auto
    23      by (rule acyclic_RAG[OF vt])
   140    next
    24 qed
   141       case 2
    25 
   142       with readys_root[OF assms(1,2)]
    26 end
   143       show ?thesis by (auto simp:root_def)
    27 
   144    qed
    28 
   145 qed
       
   146 
       
   147 lemma image_id:
       
   148   assumes "\<And> x. x \<in> A \<Longrightarrow> f x = x"
       
   149   shows "f ` A = A"
       
   150   using assms by (auto simp:image_def)
    29 
   151 
    30 definition "the_preced s th = preced th s"
   152 definition "the_preced s th = preced th s"
    31 
   153 
    32 lemma cp_alt_def:
   154 lemma cp_alt_def:
    33   "cp s th =  
   155   "cp s th =  
    42     thus ?thesis by simp
   164     thus ?thesis by simp
    43   qed
   165   qed
    44   thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
   166   thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp)
    45 qed
   167 qed
    46 
   168 
       
   169 fun the_thread :: "node \<Rightarrow> thread" where
       
   170    "the_thread (Th th) = th"
       
   171 
       
   172 definition "cp_gen s x =
       
   173                   Max ((the_preced s \<circ> the_thread) ` subtree (tRAG s) x)"
       
   174 
       
   175 lemma cp_gen_alt_def:
       
   176   "cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))"
       
   177     by (auto simp:cp_gen_def)
       
   178 
       
   179 lemma tRAG_nodeE:
       
   180   assumes "(n1, n2) \<in> tRAG s"
       
   181   obtains th1 th2 where "n1 = Th th1" "n2 = Th th2"
       
   182   using assms
       
   183   by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def)
       
   184 
       
   185 lemma subtree_nodeE:
       
   186   assumes "n \<in> subtree (tRAG s) (Th th)"
       
   187   obtains th1 where "n = Th th1"
       
   188 proof -
       
   189   show ?thesis
       
   190   proof(rule subtreeE[OF assms])
       
   191     assume "n = Th th"
       
   192     from that[OF this] show ?thesis .
       
   193   next
       
   194     assume "Th th \<in> ancestors (tRAG s) n"
       
   195     hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
       
   196     hence "\<exists> th1. n = Th th1"
       
   197     proof(induct)
       
   198       case (base y)
       
   199       from tRAG_nodeE[OF this] show ?case by metis
       
   200     next
       
   201       case (step y z)
       
   202       thus ?case by auto
       
   203     qed
       
   204     with that show ?thesis by auto
       
   205   qed
       
   206 qed
       
   207 
       
   208 lemma threads_set_eq: 
       
   209    "the_thread ` (subtree (tRAG s) (Th th)) = 
       
   210                   {th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R")
       
   211 proof -
       
   212   { fix th'
       
   213     assume "th' \<in> ?L"
       
   214     then obtain n where h: "th' = the_thread n" "n \<in>  subtree (tRAG s) (Th th)" by auto
       
   215     from subtree_nodeE[OF this(2)]
       
   216     obtain th1 where "n = Th th1" by auto
       
   217     with h have "Th th' \<in>  subtree (tRAG s) (Th th)" by auto
       
   218     hence "Th th' \<in>  subtree (RAG s) (Th th)"
       
   219     proof(cases rule:subtreeE[consumes 1])
       
   220       case 1
       
   221       thus ?thesis by (auto simp:subtree_def)
       
   222     next
       
   223       case 2
       
   224       hence "(Th th', Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def)
       
   225       thus ?thesis
       
   226       proof(induct)
       
   227         case (step y z)
       
   228         from this(2)[unfolded tRAG_alt_def]
       
   229         obtain u where 
       
   230          h: "(y, u) \<in> RAG s" "(u, z) \<in> RAG s"
       
   231           by auto
       
   232         hence "y \<in> subtree (RAG s) z" by (auto simp:subtree_def)
       
   233         with step(3)
       
   234         show ?case by (auto simp:subtree_def)
       
   235       next
       
   236         case (base y)
       
   237         from this[unfolded tRAG_alt_def]
       
   238         show ?case by (auto simp:subtree_def)
       
   239       qed
       
   240     qed
       
   241     hence "th' \<in> ?R" by auto
       
   242   } moreover {
       
   243     fix th'
       
   244     assume "th' \<in> ?R"
       
   245     hence "(Th th', Th th) \<in> (RAG s)^*" by (auto simp:subtree_def)
       
   246     from star_rpath[OF this]
       
   247     obtain xs where "rpath (RAG s) (Th th') xs (Th th)" by auto
       
   248     hence "Th th' \<in> subtree (tRAG s) (Th th)"
       
   249     proof(induct xs arbitrary:th' th rule:length_induct)
       
   250       case (1 xs th' th)
       
   251       show ?case
       
   252       proof(cases xs)
       
   253         case Nil
       
   254           from rpath_nilE[OF 1(2)[unfolded this]]
       
   255           have "th' = th" by auto
       
   256           thus ?thesis by (auto simp:subtree_def)
       
   257       next
       
   258         case (Cons x1 xs1) note Cons1 = Cons
       
   259         show ?thesis
       
   260         proof(cases "xs1")
       
   261           case Nil
       
   262             from 1(2)[unfolded Cons[unfolded this]]
       
   263             have rp: "rpath (RAG s) (Th th') [x1] (Th th)" .
       
   264             hence "(Th th', x1) \<in> (RAG s)" by (cases, simp)
       
   265             then obtain cs where "x1 = Cs cs" 
       
   266               by (unfold s_RAG_def, auto)
       
   267               find_theorems rpath "_ = _@[_]"
       
   268             from rpath_nnl_lastE[OF rp[unfolded this]]
       
   269             show ?thesis by auto
       
   270         next
       
   271           case (Cons x2 xs2)
       
   272           from 1(2)[unfolded Cons1[unfolded this]]
       
   273           have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" .
       
   274           from rpath_edges_on[OF this]
       
   275           have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" .
       
   276           have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)"
       
   277             by (simp add: edges_on_unfold)
       
   278           with eds have rg1: "(Th th', x1) \<in> RAG s" by auto
       
   279           then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto)
       
   280           have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)"
       
   281             by (simp add: edges_on_unfold)
       
   282           from this eds
       
   283           have rg2: "(x1, x2) \<in> RAG s" by auto
       
   284           from this[unfolded eq_x1] 
       
   285           obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto)
       
   286           from rp have "rpath (RAG s) x2 xs2 (Th th)"
       
   287            by  (elim rpath_ConsE, simp)
       
   288           from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" .
       
   289           from 1(1)[rule_format, OF _ this, unfolded Cons1 Cons]
       
   290           have "Th th1 \<in> subtree (tRAG s) (Th th)" by simp
       
   291           moreover have "(Th th', Th th1) \<in> (tRAG s)^*"
       
   292           proof -
       
   293             from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2]
       
   294             show ?thesis by (unfold RAG_split tRAG_def wRAG_def hRAG_def, auto)
       
   295           qed
       
   296           ultimately show ?thesis by (auto simp:subtree_def)
       
   297         qed
       
   298       qed
       
   299     qed
       
   300     from imageI[OF this, of the_thread]
       
   301     have "th' \<in> ?L" by simp
       
   302   } ultimately show ?thesis by auto
       
   303 qed
       
   304                   
       
   305 lemma cp_alt_def1: 
       
   306   "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))"
       
   307 proof -
       
   308   have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) =
       
   309        ((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))"
       
   310        by auto
       
   311   thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto)
       
   312 qed
       
   313 
       
   314 lemma cp_gen_def_cond: 
       
   315   assumes "x = Th th"
       
   316   shows "cp s th = cp_gen s (Th th)"
       
   317 by (unfold cp_alt_def1 cp_gen_def, simp)
       
   318 
       
   319 lemma cp_gen_over_set:
       
   320   assumes "\<forall> x \<in> A. \<exists> th. x = Th th"
       
   321   shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A"
       
   322 proof(rule f_image_eq)
       
   323   fix a
       
   324   assume "a \<in> A"
       
   325   from assms[rule_format, OF this]
       
   326   obtain th where eq_a: "a = Th th" by auto
       
   327   show "cp_gen s a = (cp s \<circ> the_thread) a"
       
   328     by  (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp)
       
   329 qed
       
   330 
       
   331  
       
   332 
       
   333 locale valid_trace = 
       
   334   fixes s
       
   335   assumes vt : "vt s"
       
   336 
       
   337 context valid_trace
       
   338 begin
       
   339 
       
   340 lemma wf_RAG: "wf (RAG s)"
       
   341 proof(rule finite_acyclic_wf)
       
   342   from finite_RAG[OF vt] show "finite (RAG s)" .
       
   343 next
       
   344   from acyclic_RAG[OF vt] show "acyclic (RAG s)" .
       
   345 qed
       
   346 
       
   347 end
       
   348 
       
   349 context valid_trace
       
   350 begin
       
   351 
       
   352 lemma sgv_wRAG: "single_valued (wRAG s)"
       
   353   using waiting_unique[OF vt] 
       
   354   by (unfold single_valued_def wRAG_def, auto)
       
   355 
       
   356 lemma sgv_hRAG: "single_valued (hRAG s)"
       
   357   using holding_unique 
       
   358   by (unfold single_valued_def hRAG_def, auto)
       
   359 
       
   360 lemma sgv_tRAG: "single_valued (tRAG s)"
       
   361   by (unfold tRAG_def, rule single_valued_relcomp, 
       
   362               insert sgv_wRAG sgv_hRAG, auto)
       
   363 
       
   364 lemma acyclic_tRAG: "acyclic (tRAG s)"
       
   365 proof(unfold tRAG_def, rule acyclic_compose)
       
   366   show "acyclic (RAG s)" using acyclic_RAG[OF vt] .
       
   367 next
       
   368   show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
   369 next
       
   370   show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto
       
   371 qed
       
   372 
       
   373 lemma sgv_RAG: "single_valued (RAG s)"
       
   374   using unique_RAG[OF vt] by (auto simp:single_valued_def)
       
   375 
       
   376 lemma rtree_RAG: "rtree (RAG s)"
       
   377   using sgv_RAG acyclic_RAG[OF vt]
       
   378   by (unfold rtree_def rtree_axioms_def sgv_def, auto)
       
   379 
       
   380 end
       
   381 
       
   382 sublocale valid_trace < rtree_s: rtree "tRAG s"
       
   383 proof(unfold_locales)
       
   384   from sgv_tRAG show "single_valued (tRAG s)" .
       
   385 next
       
   386   from acyclic_tRAG show "acyclic (tRAG s)" .
       
   387 qed
       
   388 
       
   389 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s"
       
   390 proof -
       
   391   show "fsubtree (RAG s)"
       
   392   proof(intro_locales)
       
   393     show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG[OF vt]] .
       
   394   next
       
   395     show "fsubtree_axioms (RAG s)"
       
   396     proof(unfold fsubtree_axioms_def)
       
   397     find_theorems wf RAG
       
   398       from wf_RAG show "wf (RAG s)" .
       
   399     qed
       
   400   qed
       
   401 qed
       
   402 
       
   403 sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s"
       
   404 proof -
       
   405   have "fsubtree (tRAG s)"
       
   406   proof -
       
   407     have "fbranch (tRAG s)"
       
   408     proof(unfold tRAG_def, rule fbranch_compose)
       
   409         show "fbranch (wRAG s)"
       
   410         proof(rule finite_fbranchI)
       
   411            from finite_RAG[OF vt] show "finite (wRAG s)"
       
   412            by (unfold RAG_split, auto)
       
   413         qed
       
   414     next
       
   415         show "fbranch (hRAG s)"
       
   416         proof(rule finite_fbranchI)
       
   417            from finite_RAG[OF vt] 
       
   418            show "finite (hRAG s)" by (unfold RAG_split, auto)
       
   419         qed
       
   420     qed
       
   421     moreover have "wf (tRAG s)"
       
   422     proof(rule wf_subset)
       
   423       show "wf (RAG s O RAG s)" using wf_RAG
       
   424         by (fold wf_comp_self, simp)
       
   425     next
       
   426       show "tRAG s \<subseteq> (RAG s O RAG s)"
       
   427         by (unfold tRAG_alt_def, auto)
       
   428     qed
       
   429     ultimately show ?thesis
       
   430       by (unfold fsubtree_def fsubtree_axioms_def,auto)
       
   431   qed
       
   432   from this[folded tRAG_def] show "fsubtree (tRAG s)" .
       
   433 qed
       
   434 
       
   435 lemma Max_UNION: 
       
   436   assumes "finite A"
       
   437   and "A \<noteq> {}"
       
   438   and "\<forall> M \<in> f ` A. finite M"
       
   439   and "\<forall> M \<in> f ` A. M \<noteq> {}"
       
   440   shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R")
       
   441   using assms[simp]
       
   442 proof -
       
   443   have "?L = Max (\<Union>(f ` A))"
       
   444     by (fold Union_image_eq, simp)
       
   445   also have "... = ?R"
       
   446     by (subst Max_Union, simp+)
       
   447   finally show ?thesis .
       
   448 qed
       
   449 
       
   450 lemma max_Max_eq:
       
   451   assumes "finite A"
       
   452     and "A \<noteq> {}"
       
   453     and "x = y"
       
   454   shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R")
       
   455 proof -
       
   456   have "?R = Max (insert y A)" by simp
       
   457   also from assms have "... = ?L"
       
   458       by (subst Max.insert, simp+)
       
   459   finally show ?thesis by simp
       
   460 qed
       
   461 
       
   462 
       
   463 context valid_trace
       
   464 begin
       
   465 
       
   466 (* ddd *)
       
   467 lemma cp_gen_rec:
       
   468   assumes "x = Th th"
       
   469   shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)"
       
   470 proof(cases "children (tRAG s) x = {}")
       
   471   case True
       
   472   show ?thesis
       
   473     by (unfold True cp_gen_def subtree_children, simp add:assms)
       
   474 next
       
   475   case False
       
   476   hence [simp]: "children (tRAG s) x \<noteq> {}" by auto
       
   477   note fsbttRAGs.finite_subtree[simp]
       
   478   have [simp]: "finite (children (tRAG s) x)"
       
   479      by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], 
       
   480             rule children_subtree)
       
   481   { fix r x
       
   482     have "subtree r x \<noteq> {}" by (auto simp:subtree_def)
       
   483   } note this[simp]
       
   484   have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}"
       
   485   proof -
       
   486     from False obtain q where "q \<in> children (tRAG s) x" by blast
       
   487     moreover have "subtree (tRAG s) q \<noteq> {}" by simp
       
   488     ultimately show ?thesis by blast
       
   489   qed
       
   490   have h: "Max ((the_preced s \<circ> the_thread) `
       
   491                 ({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) =
       
   492         Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)"
       
   493                      (is "?L = ?R")
       
   494   proof -
       
   495     let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L
       
   496     let "Max (_ \<union> (?h ` ?B))" = ?R
       
   497     let ?L1 = "?f ` \<Union>(?g ` ?B)"
       
   498     have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)"
       
   499     proof -
       
   500       have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp
       
   501       also have "... =  (\<Union> x \<in> ?B. ?f ` (?g x))" by auto
       
   502       finally have "Max ?L1 = Max ..." by simp
       
   503       also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)"
       
   504         by (subst Max_UNION, simp+)
       
   505       also have "... = Max (cp_gen s ` children (tRAG s) x)"
       
   506           by (unfold image_comp cp_gen_alt_def, simp)
       
   507       finally show ?thesis .
       
   508     qed
       
   509     show ?thesis
       
   510     proof -
       
   511       have "?L = Max (?f ` ?A \<union> ?L1)" by simp
       
   512       also have "... = max (the_preced s (the_thread x)) (Max ?L1)"
       
   513             by (subst Max_Un, simp+)
       
   514       also have "... = max (?f x) (Max (?h ` ?B))"
       
   515         by (unfold eq_Max_L1, simp)
       
   516       also have "... =?R"
       
   517         by (rule max_Max_eq, (simp)+, unfold assms, simp)
       
   518       finally show ?thesis .
       
   519     qed
       
   520   qed  thus ?thesis 
       
   521           by (fold h subtree_children, unfold cp_gen_def, simp) 
       
   522 qed
       
   523 
       
   524 lemma cp_rec:
       
   525   "cp s th = Max ({the_preced s th} \<union> 
       
   526                      (cp s o the_thread) ` children (tRAG s) (Th th))"
       
   527 proof -
       
   528   have "Th th = Th th" by simp
       
   529   note h =  cp_gen_def_cond[OF this] cp_gen_rec[OF this]
       
   530   show ?thesis 
       
   531   proof -
       
   532     have "cp_gen s ` children (tRAG s) (Th th) = 
       
   533                 (cp s \<circ> the_thread) ` children (tRAG s) (Th th)"
       
   534     proof(rule cp_gen_over_set)
       
   535       show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th"
       
   536         by (unfold tRAG_alt_def, auto simp:children_def)
       
   537     qed
       
   538     thus ?thesis by (subst (1) h(1), unfold h(2), simp)
       
   539   qed
       
   540 qed
       
   541 
       
   542 end
       
   543 
       
   544 
    47 lemma eq_dependants: "dependants (wq s) = dependants s"
   545 lemma eq_dependants: "dependants (wq s) = dependants s"
    48   by (simp add: s_dependants_abv wq_def)
   546   by (simp add: s_dependants_abv wq_def)
    49   
   547  
       
   548 
    50 (* obvious lemma *)
   549 (* obvious lemma *)
    51 lemma not_thread_holdents:
   550 lemma not_thread_holdents:
    52   fixes th s
   551   fixes th s
    53   assumes vt: "vt s"
   552   assumes vt: "vt s"
    54   and not_in: "th \<notin> threads s" 
   553   and not_in: "th \<notin> threads s" 
   545       by (simp add: UN_exists)
  1044       by (simp add: UN_exists)
   546 
  1045 
   547     (* moving in Max *)
  1046     (* moving in Max *)
   548     also have "\<dots> = max (Max {preced th s})  
  1047     also have "\<dots> = max (Max {preced th s})  
   549       (Max ((\<lambda>th. Max (preceds s ({th} \<union> (dependants (wq s) th)))) ` children s th))" 
  1048       (Max ((\<lambda>th. Max (preceds s ({th} \<union> (dependants (wq s) th)))) ` children s th))" 
   550       by (subst Max_Union) (auto simp add: image_image)
  1049        by (subst Max_Union) (auto simp add: image_image) 
   551 
  1050 
   552     (* folding cp + moving out Max *)
  1051     (* folding cp + moving out Max *)
   553     also have "\<dots> = ?RHS" 
  1052     also have "\<dots> = ?RHS" 
   554       unfolding eq_cp by (simp add: Max_insert)
  1053       unfolding eq_cp by (simp add: Max_insert)
   555 
  1054 
   556     finally show "?LHS = ?RHS" .
  1055     finally show "?LHS = ?RHS" .
   557   qed
  1056   qed
   558 qed
  1057 qed
   559 
  1058 
   560 lemma next_waiting:
  1059 lemma next_th_holding:
   561   assumes vt: "vt s"
  1060   assumes vt: "vt s"
   562   and nxt: "next_th s th cs th'"
  1061   and nxt: "next_th s th cs th'"
   563   shows "waiting s th' cs"
  1062   shows "holding (wq s) th cs"
   564 proof -
  1063 proof -
   565   from assms show ?thesis
  1064   from nxt[unfolded next_th_def]
   566     apply (auto simp:next_th_def s_waiting_def[folded wq_def])
  1065   obtain rest where h: "wq s cs = th # rest"
   567   proof -
  1066                        "rest \<noteq> []" 
   568     fix rest
  1067                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
   569     assume ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest"
  1068   thus ?thesis
   570       and eq_wq: "wq s cs = th # rest"
  1069     by (unfold cs_holding_def, auto)
   571       and ne: "rest \<noteq> []"
  1070 qed
   572     have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
  1071 
   573     proof(rule someI2)
  1072 lemma next_th_waiting:
   574       from wq_distinct[OF vt, of cs] eq_wq
  1073   assumes vt: "vt s"
   575       show "distinct rest \<and> set rest = set rest" by auto
  1074   and nxt: "next_th s th cs th'"
   576     next
  1075   shows "waiting (wq s) th' cs"
   577       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
  1076 proof -
   578     qed
  1077   from nxt[unfolded next_th_def]
   579     with ni
  1078   obtain rest where h: "wq s cs = th # rest"
   580     have "hd (SOME q. distinct q \<and> set q = set rest) \<notin>  set (SOME q. distinct q \<and> set q = set rest)"
  1079                        "rest \<noteq> []" 
   581       by simp
  1080                        "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto
   582     moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
  1081   from wq_distinct[OF vt, of cs, unfolded h]
   583     proof(rule someI2)
  1082   have dst: "distinct (th # rest)" .
   584       from wq_distinct[OF vt, of cs] eq_wq
  1083   have in_rest: "th' \<in> set rest"
   585       show "distinct rest \<and> set rest = set rest" by auto
  1084   proof(unfold h, rule someI2)
   586     next
  1085     show "distinct rest \<and> set rest = set rest" using dst by auto
   587       from ne show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> x \<noteq> []" by auto
       
   588     qed
       
   589     ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto
       
   590   next
  1086   next
   591     fix rest
  1087     fix x assume "distinct x \<and> set x = set rest"
   592     assume eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest"
  1088     with h(2)
   593       and ne: "rest \<noteq> []"
  1089     show "hd x \<in> set (rest)" by (cases x, auto)
   594     have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []"
  1090   qed
   595     proof(rule someI2)
  1091   hence "th' \<in> set (wq s cs)" by (unfold h(1), auto)
   596       from wq_distinct[OF vt, of cs] eq_wq
  1092   moreover have "th' \<noteq> hd (wq s cs)"
   597       show "distinct rest \<and> set rest = set rest" by auto
  1093     by (unfold h(1), insert in_rest dst, auto)
   598     next
  1094   ultimately show ?thesis by (auto simp:cs_waiting_def)
   599       from ne show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> x \<noteq> []" by auto
  1095 qed
   600     qed
  1096 
   601     hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> set (SOME q. distinct q \<and> set q = set rest)"
  1097 lemma next_th_RAG:
   602       by auto
  1098   assumes vt: "vt s"
   603     moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest" 
  1099   and nxt: "next_th s th cs th'"
   604     proof(rule someI2)
  1100   shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s"
   605       from wq_distinct[OF vt, of cs] eq_wq
  1101   using assms next_th_holding next_th_waiting
   606       show "distinct rest \<and> set rest = set rest" by auto
  1102 by (unfold s_RAG_def, simp)
   607     next
       
   608       show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto
       
   609     qed
       
   610     ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<in> set rest" by simp
       
   611     with eq_wq and wq_distinct[OF vt, of cs]
       
   612     show False by auto
       
   613   qed
       
   614 qed
       
   615 
  1103 
   616 -- {* A useless definition *}
  1104 -- {* A useless definition *}
   617 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
  1105 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set"
   618 where "cps s = {(th, cp s th) | th . th \<in> threads s}"
  1106 where "cps s = {(th, cp s th) | th . th \<in> threads s}"
   619 
  1107 
   658   assumes vt_s: "vt s"
  1146   assumes vt_s: "vt s"
   659 
  1147 
   660 context step_set_cps 
  1148 context step_set_cps 
   661 begin
  1149 begin
   662 
  1150 
   663 interpretation rtree_RAGs: rtree "RAG s"
       
   664 proof
       
   665   show "single_valued (RAG s)"
       
   666   apply (intro_locales)
       
   667     by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s])
       
   668 
       
   669   show "acyclic (RAG s)"
       
   670      by (rule acyclic_RAG[OF vt_s])
       
   671 qed
       
   672 
       
   673 interpretation rtree_RAGs': rtree "RAG s'"
       
   674 proof
       
   675   show "single_valued (RAG s')"
       
   676   apply (intro_locales)
       
   677     by (unfold single_valued_def, 
       
   678         auto intro: unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]])
       
   679 
       
   680   show "acyclic (RAG s')"
       
   681      by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]])
       
   682 qed
       
   683 
       
   684 text {* (* ddd *)
  1151 text {* (* ddd *)
   685   The following lemma confirms that @{text "Set"}-operating only changes the precedence 
  1152   The following two lemmas confirm that @{text "Set"}-operating only changes the precedence 
   686   of initiating thread.
  1153   of the initiating thread.
   687 *}
  1154 *}
   688 
  1155 
   689 lemma eq_preced:
  1156 lemma eq_preced:
   690   fixes th'
  1157   fixes th'
   691   assumes "th' \<noteq> th"
  1158   assumes "th' \<noteq> th"
   693 proof -
  1160 proof -
   694   from assms show ?thesis 
  1161   from assms show ?thesis 
   695     by (unfold s_def, auto simp:preced_def)
  1162     by (unfold s_def, auto simp:preced_def)
   696 qed
  1163 qed
   697 
  1164 
   698 text {* (* ddd *)
  1165 lemma eq_the_preced: 
       
  1166   fixes th'
       
  1167   assumes "th' \<noteq> th"
       
  1168   shows "the_preced s th' = the_preced s' th'"
       
  1169   using assms
       
  1170   by (unfold the_preced_def, intro eq_preced, simp)
       
  1171 
       
  1172 text {*
   699   The following lemma assures that the resetting of priority does not change the RAG. 
  1173   The following lemma assures that the resetting of priority does not change the RAG. 
   700 *}
  1174 *}
   701 
  1175 
   702 lemma eq_dep: "RAG s = RAG s'"
  1176 lemma eq_dep: "RAG s = RAG s'"
   703   by (unfold s_def RAG_set_unchanged, auto)
  1177   by (unfold s_def RAG_set_unchanged, auto)
   704 
  1178 
   705 text {*
  1179 text {* (* ddd *)
   706   Th following lemma @{text "eq_cp_pre"} circumscribe a rough range of recalculation.
  1180   Th following lemma @{text "eq_cp_pre"} says the priority change of @{text "th"}
   707   It says, thread other than the initiating thread @{text "th"} does not need recalculation
  1181   only affects those threads, which as @{text "Th th"} in their sub-trees.
   708   unless it lies upstream of @{text "th"} in the RAG. 
  1182   
   709 
  1183   The proof of this lemma is simplified by using the alternative definition of @{text "cp"}. 
   710   The reason behind this lemma is that: 
       
   711   the change of precedence of one thread can only affect it's upstream threads, according to 
       
   712   lemma @{text "eq_preced"}. Since the only thread which might change precedence is
       
   713   @{text "th"}, so only @{text "th"} and its upstream threads need recalculation.
       
   714   (* ccc *)
       
   715 *}
  1184 *}
   716 
  1185 
   717 lemma eq_cp_pre:
  1186 lemma eq_cp_pre:
   718   fixes th' 
  1187   fixes th' 
   719   assumes neq_th: "th' \<noteq> th"
  1188   assumes nd: "Th th \<notin> subtree (RAG s') (Th th')"
   720   and nd: "th \<notin> dependants s th'"
       
   721   shows "cp s th' = cp s' th'"
  1189   shows "cp s th' = cp s' th'"
   722 proof -
  1190 proof -
   723   -- {* This is what we need to prove after expanding the definition of @{text "cp"} *}
  1191   -- {* After unfolding using the alternative definition, elements 
   724   have "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')) =
  1192         affecting the @{term "cp"}-value of threads become explicit. 
   725         Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependants (wq s') th'))"
  1193         We only need to prove the following: *}
   726    (is "Max (?f1 ` ({th'} \<union> ?A)) = Max (?f2 ` ({th'} \<union> ?B))") 
  1194   have "Max (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) =
       
  1195         Max (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})"
       
  1196         (is "Max (?f ` ?S1) = Max (?g ` ?S2)")
   727   proof -
  1197   proof -
   728       -- {* Since RAG is not changed by @{text "Set"}-operation, the dependants of 
  1198     -- {* The base sets are equal. *}
   729             any threads are not changed, this is one of key facts underpinning this 
  1199     have "?S1 = ?S2" using eq_dep by simp
   730             lemma *}
  1200     -- {* The function values on the base set are equal as well. *}
   731       have eq_ab: "?A = ?B" by (unfold cs_dependants_def, auto simp:eq_dep eq_RAG)
  1201     moreover have "\<forall> e \<in> ?S2. ?f e = ?g e"
   732       have "(?f1 ` ({th'} \<union> ?A)) =  (?f2 ` ({th'} \<union> ?B))"
  1202     proof
   733       proof(rule image_cong)
  1203       fix th1
   734         show "{th'} \<union> ?A = {th'} \<union> ?B" by (simp only:eq_ab)
  1204       assume "th1 \<in> ?S2"
   735       next  
  1205       with nd have "th1 \<noteq> th" by (auto)
   736         fix x
  1206       from eq_the_preced[OF this]
   737         assume x_in: "x \<in> {th'} \<union> ?B"
  1207       show "the_preced s th1 = the_preced s' th1" .
   738         show "?f1 x = ?f2 x"
  1208     qed
   739         proof(rule eq_preced) -- {* The other key fact underpinning this lemma is @{text "eq_preced"} *}
  1209     -- {* Therefore, the image of the functions are equal. *}
   740           from x_in[folded eq_ab, unfolded eq_dependants]
  1210     ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq)
   741           have "x \<in> {th'} \<union> dependants s th'" .
  1211     thus ?thesis by simp
   742           thus "x \<noteq> th"
  1212   qed
   743           proof
  1213   thus ?thesis by (simp add:cp_alt_def)
   744             assume "x \<in> {th'}" 
       
   745             with `th' \<noteq> th` show ?thesis by simp
       
   746           next
       
   747             assume "x \<in> dependants s th'"
       
   748             with `th \<notin> dependants s th'` show ?thesis by auto
       
   749           qed 
       
   750         qed 
       
   751       qed
       
   752       thus ?thesis by simp
       
   753   qed 
       
   754   thus ?thesis by (unfold cp_eq_cpreced cpreced_def)
       
   755 qed
  1214 qed
   756 
  1215 
   757 text {*
  1216 text {*
   758   The following lemma shows that no thread lies upstream of the initiating thread @{text "th"}. 
  1217   The following lemma shows that @{term "th"} is not in the 
   759   The reason for this is that only no-blocked thread can initiate 
  1218   sub-tree of any other thread. 
   760   a system call. Since thread @{text "th"} is non-blocked, it is not waiting for any 
       
   761   critical resource. Therefore, there is edge leading out of @{text "th"} in the RAG.
       
   762   Consequently, there is no node (neither resource nor thread) upstream of @{text "th"}.
       
   763 *}
  1219 *}
   764 lemma no_dependants:
  1220 lemma th_in_no_subtree:
   765   shows "th \<notin> dependants s th'"
  1221   assumes "th' \<noteq> th"
   766 proof
  1222   shows "Th th \<notin> subtree (RAG s') (Th th')"
   767   assume "th \<in> dependants s th'"
  1223 proof -
   768   from `th \<in> dependants s th'` have "(Th th, Th th') \<in> (RAG s')\<^sup>+"
  1224   have "th \<in> readys s'"
   769     by (unfold s_dependants_def, unfold eq_RAG, unfold eq_dep, auto)
       
   770   from tranclD[OF this]
       
   771   obtain z where "(Th th, z) \<in> RAG s'" by auto
       
   772   moreover have "th \<in> readys s'"
       
   773   proof -
  1225   proof -
   774     from step_back_step [OF vt_s[unfolded s_def]]
  1226     from step_back_step [OF vt_s[unfolded s_def]]
   775     have "step s' (Set th prio)" .
  1227     have "step s' (Set th prio)" .
   776     hence "th \<in> runing s'" by (cases, simp)
  1228     hence "th \<in> runing s'" by (cases, simp)
   777     thus ?thesis by (simp add:readys_def runing_def)
  1229     thus ?thesis by (simp add:readys_def runing_def)
   778   qed
  1230   qed
   779   ultimately show "False"
  1231   from readys_in_no_subtree[OF step_back_vt[OF vt_s[unfolded s_def]] this assms(1)]
   780     apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def)
  1232   show ?thesis by blast
   781     by (fold wq_def, blast)
  1233 qed
   782 qed
  1234 
   783 
       
   784 (* Result improved *)
       
   785 text {* 
  1235 text {* 
   786   A simple combination of @{text "eq_cp_pre"} and @{text "no_dependants"}
  1236   By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, 
   787   gives the main lemma of this locale, which shows that
  1237   it is obvious that the change of priority only affects the @{text "cp"}-value 
   788   only the initiating thread needs a recalculation of current precedence.
  1238   of the initiating thread @{text "th"}.
   789 *}
  1239 *}
   790 lemma eq_cp:
  1240 lemma eq_cp:
   791   fixes th' 
  1241   fixes th' 
   792   assumes "th' \<noteq> th"
  1242   assumes "th' \<noteq> th"
   793   shows "cp s th' = cp s' th'"
  1243   shows "cp s th' = cp s' th'"
   794   by (rule eq_cp_pre[OF assms no_dependants])
  1244   by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]])
   795 
  1245 
   796 end
  1246 end
   797 
  1247 
   798 text {*
  1248 text {*
   799   The following @{text "step_v_cps"} is the locale for @{text "V"}-operation.
  1249   The following @{text "step_v_cps"} is the locale for @{text "V"}-operation.
   804   -- {* @{text "cs"} is the critical resource release by the @{text "V"}-operation *}
  1254   -- {* @{text "cs"} is the critical resource release by the @{text "V"}-operation *}
   805   fixes s' th cs s    -- {* @{text "s'"} is the state before operation*}
  1255   fixes s' th cs s    -- {* @{text "s'"} is the state before operation*}
   806   defines s_def : "s \<equiv> (V th cs#s')" -- {* @{text "s"} is the state after operation*}
  1256   defines s_def : "s \<equiv> (V th cs#s')" -- {* @{text "s"} is the state after operation*}
   807   -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *}
  1257   -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *}
   808   assumes vt_s: "vt s"
  1258   assumes vt_s: "vt s"
       
  1259 
       
  1260 context step_v_cps
       
  1261 begin
       
  1262 
       
  1263 lemma rtree_RAGs: "rtree (RAG s)"
       
  1264 proof
       
  1265   show "single_valued (RAG s)"
       
  1266   apply (intro_locales)
       
  1267     by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s])
       
  1268 
       
  1269   show "acyclic (RAG s)"
       
  1270      by (rule acyclic_RAG[OF vt_s])
       
  1271 qed
       
  1272 
       
  1273 lemma rtree_RAGs': "rtree (RAG s')"
       
  1274 proof
       
  1275   show "single_valued (RAG s')"
       
  1276   apply (intro_locales)
       
  1277     by (unfold single_valued_def, 
       
  1278         auto intro:unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]])
       
  1279 
       
  1280   show "acyclic (RAG s')"
       
  1281      by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]])
       
  1282 qed
       
  1283 
       
  1284 lemmas vt_s' = step_back_vt[OF vt_s[unfolded s_def]]
       
  1285 
       
  1286 lemma ready_th_s': "th \<in> readys s'"
       
  1287   using step_back_step[OF vt_s[unfolded s_def]]
       
  1288   by (cases, simp add:runing_def)
       
  1289 
       
  1290 
       
  1291 lemma ancestors_th: "ancestors (RAG s') (Th th) = {}"
       
  1292 proof -
       
  1293   from readys_root[OF vt_s' ready_th_s']
       
  1294   show ?thesis
       
  1295   by (unfold root_def, simp)
       
  1296 qed
       
  1297 
       
  1298 lemma holding_th: "holding s' th cs"
       
  1299 proof -
       
  1300   from vt_s[unfolded s_def]
       
  1301   have " PIP s' (V th cs)" by (cases, simp)
       
  1302   thus ?thesis by (cases, auto)
       
  1303 qed
       
  1304 
       
  1305 lemma edge_of_th:
       
  1306     "(Cs cs, Th th) \<in> RAG s'" 
       
  1307 proof -
       
  1308  from holding_th
       
  1309  show ?thesis 
       
  1310     by (unfold s_RAG_def holding_eq, auto)
       
  1311 qed
       
  1312 
       
  1313 lemma ancestors_cs: 
       
  1314   "ancestors (RAG s') (Cs cs) = {Th th}"
       
  1315 proof -
       
  1316   find_theorems ancestors
       
  1317   have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th)  \<union>  {Th th}"
       
  1318   proof(rule RTree.rtree.ancestors_accum[OF rtree_RAGs'])
       
  1319     from vt_s[unfolded s_def]
       
  1320     have " PIP s' (V th cs)" by (cases, simp)
       
  1321     thus "(Cs cs, Th th) \<in> RAG s'" 
       
  1322     proof(cases)
       
  1323       assume "holding s' th cs"
       
  1324       from this[unfolded holding_eq]
       
  1325       show ?thesis by (unfold s_RAG_def, auto)
       
  1326     qed
       
  1327   qed
       
  1328   from this[unfolded ancestors_th] show ?thesis by simp
       
  1329 qed
       
  1330 
       
  1331 lemma preced_kept: "the_preced s = the_preced s'"
       
  1332   by (auto simp: s_def the_preced_def preced_def)
       
  1333 
       
  1334 end
       
  1335 
   809 
  1336 
   810 text {*
  1337 text {*
   811   The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, 
  1338   The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, 
   812   which represents the case when there is another thread @{text "th'"}
  1339   which represents the case when there is another thread @{text "th'"}
   813   to take over the critical resource released by the initiating thread @{text "th"}.
  1340   to take over the critical resource released by the initiating thread @{text "th"}.
   858     th6 ----|
  1385     th6 ----|
   859             |
  1386             |
   860     th7 ----|
  1387     th7 ----|
   861 *)
  1388 *)
   862 
  1389 
       
  1390 lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'"
       
  1391                 using next_th_RAG[OF vt_s' nt] .
       
  1392 
       
  1393 lemma ancestors_th': 
       
  1394   "ancestors (RAG s') (Th th') = {Th th, Cs cs}" 
       
  1395 proof -
       
  1396   have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}"
       
  1397   proof(rule  RTree.rtree.ancestors_accum[OF rtree_RAGs'])
       
  1398     from sub_RAGs' show "(Th th', Cs cs) \<in> RAG s'" by auto
       
  1399   qed
       
  1400   thus ?thesis using ancestors_th ancestors_cs by auto
       
  1401 qed
       
  1402 
   863 lemma RAG_s:
  1403 lemma RAG_s:
   864   "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union>
  1404   "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union>
   865                                          {(Cs cs, Th th')}"
  1405                                          {(Cs cs, Th th')}"
   866 proof -
  1406 proof -
   867   from step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
  1407   from step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
   868     and nt show ?thesis  by (auto intro:next_th_unique)
  1408     and nt show ?thesis  by (auto intro:next_th_unique)
   869 qed
  1409 qed
   870 
  1410 
   871 text {*
  1411 lemma subtree_kept:
   872   Lemma @{text "dependants_kept"} shows only @{text "th"} and @{text "th'"}
  1412   assumes "th1 \<notin> {th, th'}"
   873   have their dependants changed.
  1413   shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R")
   874 *}
  1414 proof -
   875 lemma dependants_kept:
  1415   let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})"
   876   fixes th''
  1416   let ?RAG'' = "?RAG' \<union> {(Cs cs, Th th')}"
   877   assumes neq1: "th'' \<noteq> th"
  1417   have "subtree ?RAG' (Th th1) = ?R" 
   878   and neq2: "th'' \<noteq> th'"
  1418   proof(rule subset_del_subtree_outside)
   879   shows "dependants (wq s) th'' = dependants (wq s') th''"
  1419     show "Range {(Cs cs, Th th), (Th th', Cs cs)} \<inter> subtree (RAG s') (Th th1) = {}"
   880 proof(auto) (* ccc *)
  1420     proof -
   881   fix x
  1421       have "(Th th) \<notin> subtree (RAG s') (Th th1)"
   882   assume "x \<in> dependants (wq s) th''"
  1422       proof(rule subtree_refute)
   883   hence dp: "(Th x, Th th'') \<in> (RAG s)^+"
  1423         show "Th th1 \<notin> ancestors (RAG s') (Th th)"
   884     by (auto simp:cs_dependants_def eq_RAG)
  1424           by (unfold ancestors_th, simp)
   885   { fix n
  1425       next
   886     have "(n, Th th'') \<in> (RAG s)^+ \<Longrightarrow>  (n, Th th'') \<in> (RAG s')^+"
  1426         from assms show "Th th1 \<noteq> Th th" by simp
   887     proof(induct rule:converse_trancl_induct)
  1427       qed
   888       fix y 
  1428       moreover have "(Cs cs) \<notin>  subtree (RAG s') (Th th1)"
   889       assume "(y, Th th'') \<in> RAG s"
  1429       proof(rule subtree_refute)
   890       with RAG_s neq1 neq2
  1430         show "Th th1 \<notin> ancestors (RAG s') (Cs cs)"
   891       have "(y, Th th'') \<in> RAG s'" by auto
  1431           by (unfold ancestors_cs, insert assms, auto)
   892       thus "(y, Th th'') \<in> (RAG s')\<^sup>+" by auto
  1432       qed simp
       
  1433       ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s') (Th th1) = {}" by auto
       
  1434       thus ?thesis by simp
       
  1435      qed
       
  1436   qed
       
  1437   moreover have "subtree ?RAG'' (Th th1) =  subtree ?RAG' (Th th1)"
       
  1438   proof(rule subtree_insert_next)
       
  1439     show "Th th' \<notin> subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)"
       
  1440     proof(rule subtree_refute)
       
  1441       show "Th th1 \<notin> ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')"
       
  1442             (is "_ \<notin> ?R")
       
  1443       proof -
       
  1444           have "?R \<subseteq> ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto)
       
  1445           moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp
       
  1446           ultimately show ?thesis by auto
       
  1447       qed
   893     next
  1448     next
   894       fix y z 
  1449       from assms show "Th th1 \<noteq> Th th'" by simp
   895       assume yz: "(y, z) \<in> RAG s"
  1450     qed
   896         and ztp: "(z, Th th'') \<in> (RAG s)\<^sup>+"
  1451   qed
   897         and ztp': "(z, Th th'') \<in> (RAG s')\<^sup>+"
  1452   ultimately show ?thesis by (unfold RAG_s, simp)
   898       have "y \<noteq> Cs cs \<and> y \<noteq> Th th'"
       
   899       proof
       
   900         show "y \<noteq> Cs cs"
       
   901         proof
       
   902           assume eq_y: "y = Cs cs"
       
   903           with yz have dp_yz: "(Cs cs, z) \<in> RAG s" by simp
       
   904           from RAG_s
       
   905           have cst': "(Cs cs, Th th') \<in> RAG s" by simp
       
   906           from unique_RAG[OF vt_s this dp_yz] 
       
   907           have eq_z: "z = Th th'" by simp
       
   908           with ztp have "(Th th', Th th'') \<in> (RAG s)^+" by simp
       
   909           from converse_tranclE[OF this]
       
   910           obtain cs' where dp'': "(Th th', Cs cs') \<in> RAG s"
       
   911             by (auto simp:s_RAG_def)
       
   912           with RAG_s have dp': "(Th th', Cs cs') \<in> RAG s'" by auto
       
   913           from dp'' eq_y yz eq_z have "(Cs cs, Cs cs') \<in> (RAG s)^+" by auto
       
   914           moreover have "cs' = cs"
       
   915           proof -
       
   916             from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt]
       
   917             have "(Th th', Cs cs) \<in> RAG s'"
       
   918               by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def)
       
   919             from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this dp']
       
   920             show ?thesis by simp
       
   921           qed
       
   922           ultimately have "(Cs cs, Cs cs) \<in> (RAG s)^+" by simp
       
   923           moreover note wf_trancl[OF wf_RAG[OF vt_s]]
       
   924           ultimately show False by auto
       
   925         qed
       
   926       next
       
   927         show "y \<noteq> Th th'"
       
   928         proof
       
   929           assume eq_y: "y = Th th'"
       
   930           with yz have dps: "(Th th', z) \<in> RAG s" by simp
       
   931           with RAG_s have dps': "(Th th', z) \<in> RAG s'" by auto
       
   932           have "z = Cs cs"
       
   933           proof -
       
   934             from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt]
       
   935             have "(Th th', Cs cs) \<in> RAG s'"
       
   936               by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def)
       
   937             from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] dps' this]
       
   938             show ?thesis .
       
   939           qed
       
   940           with dps RAG_s show False by auto
       
   941         qed
       
   942       qed
       
   943       with RAG_s yz have "(y, z) \<in> RAG s'" by auto
       
   944       with ztp'
       
   945       show "(y, Th th'') \<in> (RAG s')\<^sup>+" by auto
       
   946     qed    
       
   947   }
       
   948   from this[OF dp]
       
   949   show "x \<in> dependants (wq s') th''" 
       
   950     by (auto simp:cs_dependants_def eq_RAG)
       
   951 next
       
   952   fix x
       
   953   assume "x \<in> dependants (wq s') th''"
       
   954   hence dp: "(Th x, Th th'') \<in> (RAG s')^+"
       
   955     by (auto simp:cs_dependants_def eq_RAG)
       
   956   { fix n
       
   957     have "(n, Th th'') \<in> (RAG s')^+ \<Longrightarrow>  (n, Th th'') \<in> (RAG s)^+"
       
   958     proof(induct rule:converse_trancl_induct)
       
   959       fix y 
       
   960       assume "(y, Th th'') \<in> RAG s'"
       
   961       with RAG_s neq1 neq2
       
   962       have "(y, Th th'') \<in> RAG s" by auto
       
   963       thus "(y, Th th'') \<in> (RAG s)\<^sup>+" by auto
       
   964     next
       
   965       fix y z 
       
   966       assume yz: "(y, z) \<in> RAG s'"
       
   967         and ztp: "(z, Th th'') \<in> (RAG s')\<^sup>+"
       
   968         and ztp': "(z, Th th'') \<in> (RAG s)\<^sup>+"
       
   969       have "y \<noteq> Cs cs \<and> y \<noteq> Th th'"
       
   970       proof
       
   971         show "y \<noteq> Cs cs"
       
   972         proof
       
   973           assume eq_y: "y = Cs cs"
       
   974           with yz have dp_yz: "(Cs cs, z) \<in> RAG s'" by simp
       
   975           from this have eq_z: "z = Th th"
       
   976           proof -
       
   977             from step_back_step[OF vt_s[unfolded s_def]]
       
   978             have "(Cs cs, Th th) \<in> RAG s'"
       
   979               by(cases, auto simp: wq_def s_RAG_def cs_holding_def s_holding_def)
       
   980             from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this dp_yz]
       
   981             show ?thesis by simp
       
   982           qed
       
   983           from converse_tranclE[OF ztp]
       
   984           obtain u where "(z, u) \<in> RAG s'" by auto
       
   985           moreover 
       
   986           from step_back_step[OF vt_s[unfolded s_def]]
       
   987           have "th \<in> readys s'" by (cases, simp add:runing_def)
       
   988           moreover note eq_z
       
   989           ultimately show False 
       
   990             by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
       
   991         qed
       
   992       next
       
   993         show "y \<noteq> Th th'"
       
   994         proof
       
   995           assume eq_y: "y = Th th'"
       
   996           with yz have dps: "(Th th', z) \<in> RAG s'" by simp
       
   997           have "z = Cs cs"
       
   998           proof -
       
   999             from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt]
       
  1000             have "(Th th', Cs cs) \<in> RAG s'"
       
  1001               by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def)
       
  1002             from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] dps this]
       
  1003             show ?thesis .
       
  1004           qed
       
  1005           with ztp have cs_i: "(Cs cs, Th th'') \<in>  (RAG s')\<^sup>+" by simp
       
  1006           from step_back_step[OF vt_s[unfolded s_def]]
       
  1007           have cs_th: "(Cs cs, Th th) \<in> RAG s'"
       
  1008             by(cases, auto simp: s_RAG_def wq_def cs_holding_def s_holding_def)
       
  1009           have "(Cs cs, Th th'') \<notin>  RAG s'"
       
  1010           proof
       
  1011             assume "(Cs cs, Th th'') \<in> RAG s'"
       
  1012             from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this cs_th]
       
  1013             and neq1 show "False" by simp
       
  1014           qed
       
  1015           with converse_tranclE[OF cs_i]
       
  1016           obtain u where cu: "(Cs cs, u) \<in> RAG s'"  
       
  1017             and u_t: "(u, Th th'') \<in> (RAG s')\<^sup>+" by auto
       
  1018           have "u = Th th"
       
  1019           proof -
       
  1020             from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] cu cs_th]
       
  1021             show ?thesis .
       
  1022           qed
       
  1023           with u_t have "(Th th, Th th'') \<in> (RAG s')\<^sup>+" by simp
       
  1024           from converse_tranclE[OF this]
       
  1025           obtain v where "(Th th, v) \<in> (RAG s')" by auto
       
  1026           moreover from step_back_step[OF vt_s[unfolded s_def]]
       
  1027           have "th \<in> readys s'" by (cases, simp add:runing_def)
       
  1028           ultimately show False 
       
  1029             by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def)
       
  1030         qed
       
  1031       qed
       
  1032       with RAG_s yz have "(y, z) \<in> RAG s" by auto
       
  1033       with ztp'
       
  1034       show "(y, Th th'') \<in> (RAG s)\<^sup>+" by auto
       
  1035     qed    
       
  1036   }
       
  1037   from this[OF dp]
       
  1038   show "x \<in> dependants (wq s) th''"
       
  1039     by (auto simp:cs_dependants_def eq_RAG)
       
  1040 qed
  1453 qed
  1041 
  1454 
  1042 lemma cp_kept:
  1455 lemma cp_kept:
  1043   fixes th''
  1456   assumes "th1 \<notin> {th, th'}"
  1044   assumes neq1: "th'' \<noteq> th"
  1457   shows "cp s th1 = cp s' th1"
  1045   and neq2: "th'' \<noteq> th'"
  1458     by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp)
  1046   shows "cp s th'' = cp s' th''"
       
  1047 proof -
       
  1048   from dependants_kept[OF neq1 neq2]
       
  1049   have "dependants (wq s) th'' = dependants (wq s') th''" .
       
  1050   moreover {
       
  1051     fix th1
       
  1052     assume "th1 \<in> dependants (wq s) th''"
       
  1053     have "preced th1 s = preced th1 s'" 
       
  1054       by (unfold s_def, auto simp:preced_def)
       
  1055   }
       
  1056   moreover have "preced th'' s = preced th'' s'" 
       
  1057     by (unfold s_def, auto simp:preced_def)
       
  1058   ultimately have "((\<lambda>th. preced th s) ` ({th''} \<union> dependants (wq s) th'')) = 
       
  1059     ((\<lambda>th. preced th s') ` ({th''} \<union> dependants (wq s') th''))"
       
  1060     by (auto simp:image_def)
       
  1061   thus ?thesis
       
  1062     by (unfold cp_eq_cpreced cpreced_def, simp)
       
  1063 qed
       
  1064 
  1459 
  1065 end
  1460 end
  1066 
  1461 
  1067 locale step_v_cps_nnt = step_v_cps +
  1462 locale step_v_cps_nnt = step_v_cps +
  1068   assumes nnt: "\<And> th'. (\<not> next_th s' th cs th')"
  1463   assumes nnt: "\<And> th'. (\<not> next_th s' th cs th')"
  1069 
  1464 
  1070 context step_v_cps_nnt
  1465 context step_v_cps_nnt
  1071 begin
  1466 begin
  1072 
  1467 
  1073 lemma nw_cs: "(Th th1, Cs cs) \<notin> RAG s'"
       
  1074 proof
       
  1075   assume "(Th th1, Cs cs) \<in> RAG s'"
       
  1076   thus "False"
       
  1077     apply (auto simp:s_RAG_def cs_waiting_def)
       
  1078   proof -
       
  1079     assume h1: "th1 \<in> set (wq s' cs)"
       
  1080       and h2: "th1 \<noteq> hd (wq s' cs)"
       
  1081     from step_back_step[OF vt_s[unfolded s_def]]
       
  1082     show "False"
       
  1083     proof(cases)
       
  1084       assume "holding s' th cs" 
       
  1085       then obtain rest where
       
  1086         eq_wq: "wq s' cs = th#rest"
       
  1087         apply (unfold s_holding_def wq_def[symmetric])
       
  1088         by (case_tac "(wq s' cs)", auto)
       
  1089       with h1 h2 have ne: "rest \<noteq> []" by auto
       
  1090       with eq_wq
       
  1091       have "next_th s' th cs (hd (SOME q. distinct q \<and> set q = set rest))"
       
  1092         by(unfold next_th_def, rule_tac x = "rest" in exI, auto)
       
  1093       with nnt show ?thesis by auto
       
  1094     qed
       
  1095   qed
       
  1096 qed
       
  1097 
       
  1098 lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}"
  1468 lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}"
  1099 proof -
  1469 proof -
  1100   from nnt and  step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
  1470   from nnt and  step_RAG_v[OF vt_s[unfolded s_def], folded s_def]
  1101   show ?thesis by auto
  1471   show ?thesis by auto
  1102 qed
  1472 qed
  1103 
  1473 
  1104 lemma child_kept_left:
  1474 lemma subtree_kept:
  1105   assumes 
  1475   assumes "th1 \<noteq> th"
  1106   "(n1, n2) \<in> (child s')^+"
  1476   shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)"
  1107   shows "(n1, n2) \<in> (child s)^+"
  1477 proof(unfold RAG_s, rule subset_del_subtree_outside)
  1108 proof -
  1478   show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s') (Th th1) = {}"
  1109   from assms show ?thesis 
  1479   proof -
  1110   proof(induct rule: converse_trancl_induct)
  1480     have "(Th th) \<notin> subtree (RAG s') (Th th1)"
  1111     case (base y)
  1481     proof(rule subtree_refute)
  1112     from base obtain th1 cs1 th2
  1482       show "Th th1 \<notin> ancestors (RAG s') (Th th)"
  1113       where h1: "(Th th1, Cs cs1) \<in> RAG s'"
  1483           by (unfold ancestors_th, simp)
  1114       and h2: "(Cs cs1, Th th2) \<in> RAG s'"
  1484     next
  1115       and eq_y: "y = Th th1" and eq_n2: "n2 = Th th2"  by (auto simp:child_def)
  1485       from assms show "Th th1 \<noteq> Th th" by simp
  1116     have "cs1 \<noteq> cs"
  1486     qed
       
  1487     thus ?thesis by auto
       
  1488   qed
       
  1489 qed
       
  1490 
       
  1491 lemma cp_kept_1:
       
  1492   assumes "th1 \<noteq> th"
       
  1493   shows "cp s th1 = cp s' th1"
       
  1494     by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp)
       
  1495 
       
  1496 lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}"
       
  1497 proof -
       
  1498   { fix n
       
  1499     have "(Cs cs) \<notin> ancestors (RAG s') n"
  1117     proof
  1500     proof
  1118       assume eq_cs: "cs1 = cs"
  1501       assume "Cs cs \<in> ancestors (RAG s') n"
  1119       with h1 have "(Th th1, Cs cs1) \<in> RAG s'" by simp
  1502       hence "(n, Cs cs) \<in> (RAG s')^+" by (auto simp:ancestors_def)
  1120       with nw_cs eq_cs show False by auto
  1503       from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s'" by auto
  1121     qed
  1504       then obtain th' where "nn = Th th'"
  1122     with h1 h2 RAG_s have 
  1505         by (unfold s_RAG_def, auto)
  1123       h1': "(Th th1, Cs cs1) \<in> RAG s" and
  1506       from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s'" .
  1124       h2': "(Cs cs1, Th th2) \<in> RAG s" by auto
  1507       from this[unfolded s_RAG_def]
  1125     hence "(Th th1, Th th2) \<in> child s" by (auto simp:child_def)
  1508       have "waiting (wq s') th' cs" by auto
  1126     with eq_y eq_n2 have "(y, n2) \<in> child s" by simp
  1509       from this[unfolded cs_waiting_def]
  1127     thus ?case by auto
  1510       have "1 < length (wq s' cs)"
  1128   next
  1511           by (cases "wq s' cs", auto)
  1129     case (step y z)
  1512       from holding_next_thI[OF holding_th this]
  1130     have "(y, z) \<in> child s'" by fact
  1513       obtain th' where "next_th s' th cs th'" by auto
  1131     then obtain th1 cs1 th2
  1514       with nnt show False by auto
  1132       where h1: "(Th th1, Cs cs1) \<in> RAG s'"
  1515     qed
  1133       and h2: "(Cs cs1, Th th2) \<in> RAG s'"
  1516   } note h = this
  1134       and eq_y: "y = Th th1" and eq_z: "z = Th th2"  by (auto simp:child_def)
  1517   {  fix n
  1135     have "cs1 \<noteq> cs"
  1518      assume "n \<in> subtree (RAG s') (Cs cs)"
  1136     proof
  1519      hence "n = (Cs cs)"
  1137       assume eq_cs: "cs1 = cs"
  1520      by (elim subtreeE, insert h, auto)
  1138       with h1 have "(Th th1, Cs cs1) \<in> RAG s'" by simp
  1521   } moreover have "(Cs cs) \<in> subtree (RAG s') (Cs cs)"
  1139       with nw_cs eq_cs show False by auto
  1522       by (auto simp:subtree_def)
  1140     qed
  1523   ultimately show ?thesis by auto 
  1141     with h1 h2 RAG_s have 
  1524 qed
  1142       h1': "(Th th1, Cs cs1) \<in> RAG s" and
  1525 
  1143       h2': "(Cs cs1, Th th2) \<in> RAG s" by auto
  1526 lemma subtree_th: 
  1144     hence "(Th th1, Th th2) \<in> child s" by (auto simp:child_def)
  1527   "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}"
  1145     with eq_y eq_z have "(y, z) \<in> child s" by simp
  1528 proof(unfold RAG_s, fold subtree_cs, rule RTree.rtree.subtree_del_inside[OF rtree_RAGs'])
  1146     moreover have "(z, n2) \<in> (child s)^+" by fact
  1529   from edge_of_th
  1147     ultimately show ?case by auto
  1530   show "(Cs cs, Th th) \<in> edges_in (RAG s') (Th th)"
  1148   qed
  1531     by (unfold edges_in_def, auto simp:subtree_def)
  1149 qed
  1532 qed
  1150 
  1533 
  1151 lemma  child_kept_right:
  1534 lemma cp_kept_2: 
  1152   assumes
  1535   shows "cp s th = cp s' th" 
  1153   "(n1, n2) \<in> (child s)^+"
  1536  by (unfold cp_alt_def subtree_th preced_kept, auto)
  1154   shows "(n1, n2) \<in> (child s')^+"
       
  1155 proof -
       
  1156   from assms show ?thesis
       
  1157   proof(induct)
       
  1158     case (base y)
       
  1159     from base and RAG_s 
       
  1160     have "(n1, y) \<in> child s'"
       
  1161       by (auto simp:child_def)
       
  1162     thus ?case by auto
       
  1163   next
       
  1164     case (step y z)
       
  1165     have "(y, z) \<in> child s" by fact
       
  1166     with RAG_s have "(y, z) \<in> child s'"
       
  1167       by (auto simp:child_def)
       
  1168     moreover have "(n1, y) \<in> (child s')\<^sup>+" by fact
       
  1169     ultimately show ?case by auto
       
  1170   qed
       
  1171 qed
       
  1172 
       
  1173 lemma eq_child: "(child s)^+ = (child s')^+"
       
  1174   by (insert child_kept_left child_kept_right, auto)
       
  1175 
  1537 
  1176 lemma eq_cp:
  1538 lemma eq_cp:
  1177   fixes th' 
  1539   fixes th' 
  1178   shows "cp s th' = cp s' th'"
  1540   shows "cp s th' = cp s' th'"
  1179   apply (unfold cp_eq_cpreced cpreced_def)
  1541   using cp_kept_1 cp_kept_2
  1180 proof -
  1542   by (cases "th' = th", auto)
  1181   have eq_dp: "\<And> th. dependants (wq s) th = dependants (wq s') th"
  1543  
  1182     apply (unfold cs_dependants_def, unfold eq_RAG)
       
  1183   proof -
       
  1184     from eq_child
       
  1185     have "\<And>th. {th'. (Th th', Th th) \<in> (child s)\<^sup>+} = {th'. (Th th', Th th) \<in> (child s')\<^sup>+}"
       
  1186       by simp
       
  1187     with child_RAG_eq[OF vt_s] child_RAG_eq[OF step_back_vt[OF vt_s[unfolded s_def]]]
       
  1188     show "\<And>th. {th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} = {th'. (Th th', Th th) \<in> (RAG s')\<^sup>+}"
       
  1189       by simp
       
  1190   qed
       
  1191   moreover {
       
  1192     fix th1 
       
  1193     assume "th1 \<in> {th'} \<union> dependants (wq s') th'"
       
  1194     hence "th1 = th' \<or> th1 \<in> dependants (wq s') th'" by auto
       
  1195     hence "preced th1 s = preced th1 s'"
       
  1196     proof
       
  1197       assume "th1 = th'"
       
  1198       show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def)
       
  1199     next
       
  1200       assume "th1 \<in> dependants (wq s') th'"
       
  1201       show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def)
       
  1202     qed
       
  1203   } ultimately have "((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')) = 
       
  1204                      ((\<lambda>th. preced th s') ` ({th'} \<union> dependants (wq s') th'))" 
       
  1205     by (auto simp:image_def)
       
  1206   thus "Max ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')) =
       
  1207         Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependants (wq s') th'))" by simp
       
  1208 qed
       
  1209 
       
  1210 end
  1544 end
       
  1545 
       
  1546 find_theorems "_`_" "\<Union> _"
       
  1547 
       
  1548 find_theorems "Max" "\<Union> _"
       
  1549 
       
  1550 find_theorems wf RAG
       
  1551 
       
  1552 thm wf_def
       
  1553 
       
  1554 thm image_Union
  1211 
  1555 
  1212 locale step_P_cps =
  1556 locale step_P_cps =
  1213   fixes s' th cs s 
  1557   fixes s' th cs s 
  1214   defines s_def : "s \<equiv> (P th cs#s')"
  1558   defines s_def : "s \<equiv> (P th cs#s')"
  1215   assumes vt_s: "vt s"
  1559   assumes vt_s: "vt s"
  1216 
  1560 
       
  1561 sublocale step_P_cps < vat_s : valid_trace "s"
       
  1562 proof
       
  1563   from vt_s show "vt s" .
       
  1564 qed
       
  1565 
       
  1566 sublocale step_P_cps < vat_s' : valid_trace "s'"
       
  1567 proof
       
  1568   from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" .
       
  1569 qed
       
  1570 
       
  1571 
       
  1572 context step_P_cps
       
  1573 begin
       
  1574 
       
  1575 lemma rtree_RAGs: "rtree (RAG s)"
       
  1576 proof
       
  1577   show "single_valued (RAG s)"
       
  1578   apply (intro_locales)
       
  1579     by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s])
       
  1580 
       
  1581   show "acyclic (RAG s)"
       
  1582      by (rule acyclic_RAG[OF vt_s])
       
  1583 qed
       
  1584 
       
  1585 lemma rtree_RAGs': "rtree (RAG s')"
       
  1586 proof
       
  1587   show "single_valued (RAG s')"
       
  1588   apply (intro_locales)
       
  1589     by (unfold single_valued_def, 
       
  1590         auto intro:unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]])
       
  1591 
       
  1592   show "acyclic (RAG s')"
       
  1593      by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]])
       
  1594 qed
       
  1595 
       
  1596 lemma preced_kept: "the_preced s = the_preced s'"
       
  1597   by (auto simp: s_def the_preced_def preced_def)
       
  1598 
       
  1599 end
       
  1600 
  1217 locale step_P_cps_ne =step_P_cps +
  1601 locale step_P_cps_ne =step_P_cps +
       
  1602   fixes th'
  1218   assumes ne: "wq s' cs \<noteq> []"
  1603   assumes ne: "wq s' cs \<noteq> []"
       
  1604   defines th'_def: "th' \<equiv> hd (wq s' cs)"
  1219 
  1605 
  1220 locale step_P_cps_e =step_P_cps +
  1606 locale step_P_cps_e =step_P_cps +
  1221   assumes ee: "wq s' cs = []"
  1607   assumes ee: "wq s' cs = []"
  1222 
  1608 
  1223 context step_P_cps_e
  1609 context step_P_cps_e
  1351         Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependants (wq s') th'))" by simp
  1737         Max ((\<lambda>th. preced th s') ` ({th'} \<union> dependants (wq s') th'))" by simp
  1352 qed
  1738 qed
  1353 
  1739 
  1354 end
  1740 end
  1355 
  1741 
  1356 context step_P_cps_ne
  1742 lemma tRAG_ancestorsE:
       
  1743   assumes "x \<in> ancestors (tRAG s) u"
       
  1744   obtains th where "x = Th th"
       
  1745 proof -
       
  1746   from assms have "(u, x) \<in> (tRAG s)^+" 
       
  1747       by (unfold ancestors_def, auto)
       
  1748   from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto
       
  1749   then obtain th where "x = Th th"
       
  1750     by (unfold tRAG_alt_def, auto)
       
  1751   from that[OF this] show ?thesis .
       
  1752 qed
       
  1753 
       
  1754 
       
  1755 context step_P_cps_ne 
  1357 begin
  1756 begin
  1358 
  1757 
  1359 lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
  1758 lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
  1360 proof -
  1759 proof -
  1361   from step_RAG_p[OF vt_s[unfolded s_def]] and ne
  1760   from step_RAG_p[OF vt_s[unfolded s_def]] and ne
  1362   show ?thesis by (simp add:s_def)
  1761   show ?thesis by (simp add:s_def)
  1363 qed
  1762 qed
  1364 
  1763 
       
  1764 lemma cs_held: "(Cs cs, Th th') \<in> RAG s'"
       
  1765 proof -
       
  1766   have "(Cs cs, Th th') \<in> hRAG s'"
       
  1767   proof -
       
  1768     from ne
       
  1769     have " holding s' th' cs"
       
  1770       by (unfold th'_def holding_eq cs_holding_def, auto)
       
  1771     thus ?thesis 
       
  1772       by (unfold hRAG_def, auto)
       
  1773   qed
       
  1774   thus ?thesis by (unfold RAG_split, auto)
       
  1775 qed
       
  1776 
       
  1777 lemma tRAG_s: 
       
  1778   "tRAG s = tRAG s' \<union> {(Th th, Th th')}"
       
  1779   using RAG_tRAG_transfer[OF step_back_vt[OF vt_s[unfolded s_def]] RAG_s cs_held] .
       
  1780 
       
  1781 lemma cp_kept:
       
  1782   assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)"
       
  1783   shows "cp s th'' = cp s' th''"
       
  1784 proof -
       
  1785   have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')"
       
  1786   proof -
       
  1787     have "Th th' \<notin> subtree (tRAG s') (Th th'')"
       
  1788     proof
       
  1789       assume "Th th' \<in> subtree (tRAG s') (Th th'')"
       
  1790       thus False
       
  1791       proof(rule subtreeE)
       
  1792          assume "Th th' = Th th''"
       
  1793          from assms[unfolded tRAG_s ancestors_def, folded this]
       
  1794          show ?thesis by auto
       
  1795       next
       
  1796          assume "Th th'' \<in> ancestors (tRAG s') (Th th')"
       
  1797          moreover have "... \<subseteq> ancestors (tRAG s) (Th th')"
       
  1798          proof(rule ancestors_mono)
       
  1799             show "tRAG s' \<subseteq> tRAG s" by (unfold tRAG_s, auto)
       
  1800          qed 
       
  1801          ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th')" by auto
       
  1802          moreover have "Th th' \<in> ancestors (tRAG s) (Th th)"
       
  1803            by (unfold tRAG_s, auto simp:ancestors_def)
       
  1804          ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th)"
       
  1805                        by (auto simp:ancestors_def)
       
  1806          with assms show ?thesis by auto
       
  1807       qed
       
  1808     qed
       
  1809     from subtree_insert_next[OF this]
       
  1810     have "subtree (tRAG s' \<union> {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" .
       
  1811     from this[folded tRAG_s] show ?thesis .
       
  1812   qed
       
  1813   show ?thesis by (unfold cp_alt_def1 h preced_kept, simp)
       
  1814 qed
       
  1815 
       
  1816 lemma set_prop_split:
       
  1817   "A = {x. x \<in> A \<and> PP x} \<union> {x. x \<in> A \<and> \<not> PP x}"
       
  1818   by auto
       
  1819 
       
  1820 lemma f_image_union_eq:
       
  1821   assumes "f ` A = g ` A"
       
  1822   and "f ` B = g ` B"
       
  1823   shows "f ` (A \<union> B) = g ` (A \<union> B)"
       
  1824   using assms by auto
       
  1825 
       
  1826 (* ccc *)
       
  1827 
       
  1828 lemma cp_gen_update_stop:
       
  1829   assumes "u \<in> ancestors (tRAG s) (Th th)"
       
  1830   and "cp_gen s u = cp_gen s' u"
       
  1831   and "y \<in> ancestors (tRAG s) u"
       
  1832   shows "cp_gen s y = cp_gen s' y"
       
  1833   using assms(3)
       
  1834 proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf])
       
  1835   case (1 x)
       
  1836   show ?case (is "?L = ?R")
       
  1837   proof -
       
  1838     from tRAG_ancestorsE[OF 1(2)]
       
  1839     obtain th2 where eq_x: "x = Th th2" by blast
       
  1840     from vat_s.cp_gen_rec[OF this]
       
  1841     have "?L = 
       
  1842           Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" .
       
  1843     also have "... = 
       
  1844           Max ({the_preced s' th2} \<union> cp_gen s' ` RTree.children (tRAG s') x)"
       
  1845     proof -
       
  1846       from preced_kept have "the_preced s th2 = the_preced s' th2" by simp
       
  1847       moreover have "cp_gen s ` RTree.children (tRAG s) x =
       
  1848                      cp_gen s' ` RTree.children (tRAG s') x"
       
  1849       proof -
       
  1850         have "RTree.children (tRAG s) x =  RTree.children (tRAG s') x"
       
  1851         proof(unfold tRAG_s, rule children_union_kept)
       
  1852           have start: "(Th th, Th th') \<in> tRAG s"
       
  1853             by (unfold tRAG_s, auto)
       
  1854           note x_u = 1(2)
       
  1855           show "x \<notin> Range {(Th th, Th th')}"
       
  1856           proof
       
  1857             assume "x \<in> Range {(Th th, Th th')}"
       
  1858             hence eq_x: "x = Th th'" using RangeE by auto
       
  1859             show False
       
  1860             proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start])
       
  1861               case 1
       
  1862               from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG
       
  1863               show ?thesis by (auto simp:ancestors_def acyclic_def)
       
  1864             next
       
  1865               case 2
       
  1866               with x_u[unfolded eq_x]
       
  1867               have "(Th th', Th th') \<in> (tRAG s)^+" by (auto simp:ancestors_def)
       
  1868               with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def)
       
  1869             qed
       
  1870           qed
       
  1871         qed
       
  1872         moreover have "cp_gen s ` RTree.children (tRAG s) x =
       
  1873                        cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A")
       
  1874         proof(rule f_image_eq)
       
  1875           fix a
       
  1876           assume a_in: "a \<in> ?A"
       
  1877           from 1(2)
       
  1878           show "?f a = ?g a"
       
  1879           proof(cases rule:vat_s.rtree_s.ancestors_childrenE[case_names in_ch out_ch])
       
  1880              case in_ch
       
  1881              show ?thesis
       
  1882              proof(cases "a = u")
       
  1883                 case True
       
  1884                 from assms(2)[folded this] show ?thesis .
       
  1885              next
       
  1886                 case False
       
  1887                 have a_not_in: "a \<notin> ancestors (tRAG s) (Th th)"
       
  1888                 proof
       
  1889                   assume a_in': "a \<in> ancestors (tRAG s) (Th th)"
       
  1890                   have "a = u"
       
  1891                   proof(rule vat_s.rtree_s.ancestors_children_unique)
       
  1892                     from a_in' a_in show "a \<in> ancestors (tRAG s) (Th th) \<inter> 
       
  1893                                           RTree.children (tRAG s) x" by auto
       
  1894                   next 
       
  1895                     from assms(1) in_ch show "u \<in> ancestors (tRAG s) (Th th) \<inter> 
       
  1896                                       RTree.children (tRAG s) x" by auto
       
  1897                   qed
       
  1898                   with False show False by simp
       
  1899                 qed
       
  1900                 from a_in obtain th_a where eq_a: "a = Th th_a" 
       
  1901                     by (unfold RTree.children_def tRAG_alt_def, auto)
       
  1902                 from cp_kept[OF a_not_in[unfolded eq_a]]
       
  1903                 have "cp s th_a = cp s' th_a" .
       
  1904                 from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
       
  1905                 show ?thesis .
       
  1906              qed
       
  1907           next
       
  1908             case (out_ch z)
       
  1909             hence h: "z \<in> ancestors (tRAG s) u" "z \<in> RTree.children (tRAG s) x" by auto
       
  1910             show ?thesis
       
  1911             proof(cases "a = z")
       
  1912               case True
       
  1913               from h(2) have zx_in: "(z, x) \<in> (tRAG s)" by (auto simp:RTree.children_def)
       
  1914               from 1(1)[rule_format, OF this h(1)]
       
  1915               have eq_cp_gen: "cp_gen s z = cp_gen s' z" .
       
  1916               with True show ?thesis by metis
       
  1917             next
       
  1918               case False
       
  1919               from a_in obtain th_a where eq_a: "a = Th th_a"
       
  1920                 by (auto simp:RTree.children_def tRAG_alt_def)
       
  1921               have "a \<notin> ancestors (tRAG s) (Th th)" sorry
       
  1922               from cp_kept[OF this[unfolded eq_a]]
       
  1923               have "cp s th_a = cp s' th_a" .
       
  1924               from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a]
       
  1925               show ?thesis .
       
  1926             qed
       
  1927           qed
       
  1928         qed
       
  1929         ultimately show ?thesis by metis
       
  1930       qed
       
  1931       ultimately show ?thesis by simp
       
  1932     qed
       
  1933     also have "... = ?R"
       
  1934       by (fold vat_s'.cp_gen_rec[OF eq_x], simp)
       
  1935     finally show ?thesis .
       
  1936   qed
       
  1937 qed
       
  1938 
       
  1939 
       
  1940 
       
  1941 (* ccc *)
  1365 
  1942 
  1366 lemma eq_child_left:
  1943 lemma eq_child_left:
  1367   assumes nd: "(Th th, Th th') \<notin> (child s)^+"
  1944   assumes nd: "(Th th, Th th') \<notin> (child s)^+"
  1368   shows "(n1, Th th') \<in> (child s)^+ \<Longrightarrow> (n1, Th th') \<in> (child s')^+"
  1945   shows "(n1, Th th') \<in> (child s)^+ \<Longrightarrow> (n1, Th th') \<in> (child s')^+"
  1369 proof(induct rule:converse_trancl_induct)
  1946 proof(induct rule:converse_trancl_induct)