# HG changeset patch # User zhangx # Date 1450187146 -28800 # Node ID ad57323fd4d651dcbc674a48e4b0f76ce4150c03 # Parent f1b39d77db00ec34d139eb7323ff35000481c48e Extended RTree.thy diff -r f1b39d77db00 -r ad57323fd4d6 CpsG.thy --- a/CpsG.thy Thu Dec 03 14:34:29 2015 +0800 +++ b/CpsG.thy Tue Dec 15 21:45:46 2015 +0800 @@ -6,26 +6,148 @@ imports PrioG Max RTree begin -locale pip = - fixes s - assumes vt: "vt s" +definition "wRAG (s::state) = {(Th th, Cs cs) | th cs. waiting s th cs}" + +definition "hRAG (s::state) = {(Cs cs, Th th) | th cs. holding s th cs}" + +definition "tRAG s = wRAG s O hRAG s" -context pip -begin +lemma RAG_split: "RAG s = (wRAG s \ hRAG s)" + by (unfold s_RAG_abv wRAG_def hRAG_def s_waiting_abv + s_holding_abv cs_RAG_def, auto) + +lemma tRAG_alt_def: + "tRAG s = {(Th th1, Th th2) | th1 th2. + \ cs. (Th th1, Cs cs) \ RAG s \ (Cs cs, Th th2) \ RAG s}" + by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def) -interpretation rtree_RAG: rtree "RAG s" -proof - show "single_valued (RAG s)" - apply (intro_locales) - by (unfold single_valued_def, auto intro: unique_RAG[OF vt]) +lemma tRAG_mono: + assumes "RAG s' \ RAG s" + shows "tRAG s' \ tRAG s" + using assms + by (unfold tRAG_alt_def, auto) - show "acyclic (RAG s)" - by (rule acyclic_RAG[OF vt]) +lemma holding_next_thI: + assumes "holding s th cs" + and "length (wq s cs) > 1" + obtains th' where "next_th s th cs th'" +proof - + from assms(1)[folded eq_holding, unfolded cs_holding_def] + have " th \ set (wq s cs) \ th = hd (wq s cs)" . + then obtain rest where h1: "wq s cs = th#rest" + by (cases "wq s cs", auto) + with assms(2) have h2: "rest \ []" by auto + let ?th' = "hd (SOME q. distinct q \ set q = set rest)" + have "next_th s th cs ?th'" using h1(1) h2 + by (unfold next_th_def, auto) + from that[OF this] show ?thesis . qed -end +lemma RAG_tRAG_transfer: + assumes "vt s'" + assumes "RAG s = RAG s' \ {(Th th, Cs cs)}" + and "(Cs cs, Th th'') \ RAG s'" + shows "tRAG s = tRAG s' \ {(Th th, Th th'')}" (is "?L = ?R") +proof - + interpret rtree: rtree "RAG s'" + proof + show "single_valued (RAG s')" + apply (intro_locales) + by (unfold single_valued_def, + auto intro:unique_RAG[OF assms(1)]) + show "acyclic (RAG s')" + by (rule acyclic_RAG[OF assms(1)]) + qed + { fix n1 n2 + assume "(n1, n2) \ ?L" + from this[unfolded tRAG_alt_def] + obtain th1 th2 cs' where + h: "n1 = Th th1" "n2 = Th th2" + "(Th th1, Cs cs') \ RAG s" + "(Cs cs', Th th2) \ RAG s" by auto + from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \ RAG s'" by auto + from h(3) and assms(2) + have "(Th th1, Cs cs') = (Th th, Cs cs) \ + (Th th1, Cs cs') \ RAG s'" by auto + hence "(n1, n2) \ ?R" + proof + assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)" + hence eq_th1: "th1 = th" by simp + moreover have "th2 = th''" + proof - + from h1 have "cs' = cs" by simp + from assms(3) cs_in[unfolded this] rtree.sgv + show ?thesis + by (unfold single_valued_def, auto) + qed + ultimately show ?thesis using h(1,2) by auto + next + assume "(Th th1, Cs cs') \ RAG s'" + with cs_in have "(Th th1, Th th2) \ tRAG s'" + by (unfold tRAG_alt_def, auto) + from this[folded h(1, 2)] show ?thesis by auto + qed + } moreover { + fix n1 n2 + assume "(n1, n2) \ ?R" + hence "(n1, n2) \tRAG s' \ (n1, n2) = (Th th, Th th'')" by auto + hence "(n1, n2) \ ?L" + proof + assume "(n1, n2) \ tRAG s'" + moreover have "... \ ?L" + proof(rule tRAG_mono) + show "RAG s' \ RAG s" by (unfold assms(2), auto) + qed + ultimately show ?thesis by auto + next + assume eq_n: "(n1, n2) = (Th th, Th th'')" + from assms(2, 3) have "(Cs cs, Th th'') \ RAG s" by auto + moreover have "(Th th, Cs cs) \ RAG s" using assms(2) by auto + ultimately show ?thesis + by (unfold eq_n tRAG_alt_def, auto) + qed + } ultimately show ?thesis by auto +qed +lemma readys_root: + assumes "vt s" + and "th \ readys s" + shows "root (RAG s) (Th th)" +proof - + { fix x + assume "x \ ancestors (RAG s) (Th th)" + hence h: "(Th th, x) \ (RAG s)^+" by (auto simp:ancestors_def) + from tranclD[OF this] + obtain z where "(Th th, z) \ RAG s" by auto + with assms(2) have False + apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def) + by (fold wq_def, blast) + } thus ?thesis by (unfold root_def, auto) +qed + +lemma readys_in_no_subtree: + assumes "vt s" + and "th \ readys s" + and "th' \ th" + shows "Th th \ subtree (RAG s) (Th th')" +proof + assume "Th th \ subtree (RAG s) (Th th')" + thus False + proof(cases rule:subtreeE) + case 1 + with assms show ?thesis by auto + next + case 2 + with readys_root[OF assms(1,2)] + show ?thesis by (auto simp:root_def) + qed +qed + +lemma image_id: + assumes "\ x. x \ A \ f x = x" + shows "f ` A = A" + using assms by (auto simp:image_def) definition "the_preced s th = preced th s" @@ -44,9 +166,386 @@ thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp) qed +fun the_thread :: "node \ thread" where + "the_thread (Th th) = th" + +definition "cp_gen s x = + Max ((the_preced s \ the_thread) ` subtree (tRAG s) x)" + +lemma cp_gen_alt_def: + "cp_gen s = (Max \ (\x. (the_preced s \ the_thread) ` subtree (tRAG s) x))" + by (auto simp:cp_gen_def) + +lemma tRAG_nodeE: + assumes "(n1, n2) \ tRAG s" + obtains th1 th2 where "n1 = Th th1" "n2 = Th th2" + using assms + by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def) + +lemma subtree_nodeE: + assumes "n \ subtree (tRAG s) (Th th)" + obtains th1 where "n = Th th1" +proof - + show ?thesis + proof(rule subtreeE[OF assms]) + assume "n = Th th" + from that[OF this] show ?thesis . + next + assume "Th th \ ancestors (tRAG s) n" + hence "(n, Th th) \ (tRAG s)^+" by (auto simp:ancestors_def) + hence "\ th1. n = Th th1" + proof(induct) + case (base y) + from tRAG_nodeE[OF this] show ?case by metis + next + case (step y z) + thus ?case by auto + qed + with that show ?thesis by auto + qed +qed + +lemma threads_set_eq: + "the_thread ` (subtree (tRAG s) (Th th)) = + {th'. Th th' \ (subtree (RAG s) (Th th))}" (is "?L = ?R") +proof - + { fix th' + assume "th' \ ?L" + then obtain n where h: "th' = the_thread n" "n \ subtree (tRAG s) (Th th)" by auto + from subtree_nodeE[OF this(2)] + obtain th1 where "n = Th th1" by auto + with h have "Th th' \ subtree (tRAG s) (Th th)" by auto + hence "Th th' \ subtree (RAG s) (Th th)" + proof(cases rule:subtreeE[consumes 1]) + case 1 + thus ?thesis by (auto simp:subtree_def) + next + case 2 + hence "(Th th', Th th) \ (tRAG s)^+" by (auto simp:ancestors_def) + thus ?thesis + proof(induct) + case (step y z) + from this(2)[unfolded tRAG_alt_def] + obtain u where + h: "(y, u) \ RAG s" "(u, z) \ RAG s" + by auto + hence "y \ subtree (RAG s) z" by (auto simp:subtree_def) + with step(3) + show ?case by (auto simp:subtree_def) + next + case (base y) + from this[unfolded tRAG_alt_def] + show ?case by (auto simp:subtree_def) + qed + qed + hence "th' \ ?R" by auto + } moreover { + fix th' + assume "th' \ ?R" + hence "(Th th', Th th) \ (RAG s)^*" by (auto simp:subtree_def) + from star_rpath[OF this] + obtain xs where "rpath (RAG s) (Th th') xs (Th th)" by auto + hence "Th th' \ subtree (tRAG s) (Th th)" + proof(induct xs arbitrary:th' th rule:length_induct) + case (1 xs th' th) + show ?case + proof(cases xs) + case Nil + from rpath_nilE[OF 1(2)[unfolded this]] + have "th' = th" by auto + thus ?thesis by (auto simp:subtree_def) + next + case (Cons x1 xs1) note Cons1 = Cons + show ?thesis + proof(cases "xs1") + case Nil + from 1(2)[unfolded Cons[unfolded this]] + have rp: "rpath (RAG s) (Th th') [x1] (Th th)" . + hence "(Th th', x1) \ (RAG s)" by (cases, simp) + then obtain cs where "x1 = Cs cs" + by (unfold s_RAG_def, auto) + find_theorems rpath "_ = _@[_]" + from rpath_nnl_lastE[OF rp[unfolded this]] + show ?thesis by auto + next + case (Cons x2 xs2) + from 1(2)[unfolded Cons1[unfolded this]] + have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" . + from rpath_edges_on[OF this] + have eds: "edges_on (Th th' # x1 # x2 # xs2) \ RAG s" . + have "(Th th', x1) \ edges_on (Th th' # x1 # x2 # xs2)" + by (simp add: edges_on_unfold) + with eds have rg1: "(Th th', x1) \ RAG s" by auto + then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto) + have "(x1, x2) \ edges_on (Th th' # x1 # x2 # xs2)" + by (simp add: edges_on_unfold) + from this eds + have rg2: "(x1, x2) \ RAG s" by auto + from this[unfolded eq_x1] + obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto) + from rp have "rpath (RAG s) x2 xs2 (Th th)" + by (elim rpath_ConsE, simp) + from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" . + from 1(1)[rule_format, OF _ this, unfolded Cons1 Cons] + have "Th th1 \ subtree (tRAG s) (Th th)" by simp + moreover have "(Th th', Th th1) \ (tRAG s)^*" + proof - + from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2] + show ?thesis by (unfold RAG_split tRAG_def wRAG_def hRAG_def, auto) + qed + ultimately show ?thesis by (auto simp:subtree_def) + qed + qed + qed + from imageI[OF this, of the_thread] + have "th' \ ?L" by simp + } ultimately show ?thesis by auto +qed + +lemma cp_alt_def1: + "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))" +proof - + have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) = + ((the_preced s \ the_thread) ` subtree (tRAG s) (Th th))" + by auto + thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto) +qed + +lemma cp_gen_def_cond: + assumes "x = Th th" + shows "cp s th = cp_gen s (Th th)" +by (unfold cp_alt_def1 cp_gen_def, simp) + +lemma cp_gen_over_set: + assumes "\ x \ A. \ th. x = Th th" + shows "cp_gen s ` A = (cp s \ the_thread) ` A" +proof(rule f_image_eq) + fix a + assume "a \ A" + from assms[rule_format, OF this] + obtain th where eq_a: "a = Th th" by auto + show "cp_gen s a = (cp s \ the_thread) a" + by (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp) +qed + + + +locale valid_trace = + fixes s + assumes vt : "vt s" + +context valid_trace +begin + +lemma wf_RAG: "wf (RAG s)" +proof(rule finite_acyclic_wf) + from finite_RAG[OF vt] show "finite (RAG s)" . +next + from acyclic_RAG[OF vt] show "acyclic (RAG s)" . +qed + +end + +context valid_trace +begin + +lemma sgv_wRAG: "single_valued (wRAG s)" + using waiting_unique[OF vt] + by (unfold single_valued_def wRAG_def, auto) + +lemma sgv_hRAG: "single_valued (hRAG s)" + using holding_unique + by (unfold single_valued_def hRAG_def, auto) + +lemma sgv_tRAG: "single_valued (tRAG s)" + by (unfold tRAG_def, rule single_valued_relcomp, + insert sgv_wRAG sgv_hRAG, auto) + +lemma acyclic_tRAG: "acyclic (tRAG s)" +proof(unfold tRAG_def, rule acyclic_compose) + show "acyclic (RAG s)" using acyclic_RAG[OF vt] . +next + show "wRAG s \ RAG s" unfolding RAG_split by auto +next + show "hRAG s \ RAG s" unfolding RAG_split by auto +qed + +lemma sgv_RAG: "single_valued (RAG s)" + using unique_RAG[OF vt] by (auto simp:single_valued_def) + +lemma rtree_RAG: "rtree (RAG s)" + using sgv_RAG acyclic_RAG[OF vt] + by (unfold rtree_def rtree_axioms_def sgv_def, auto) + +end + +sublocale valid_trace < rtree_s: rtree "tRAG s" +proof(unfold_locales) + from sgv_tRAG show "single_valued (tRAG s)" . +next + from acyclic_tRAG show "acyclic (tRAG s)" . +qed + +sublocale valid_trace < fsbtRAGs : fsubtree "RAG s" +proof - + show "fsubtree (RAG s)" + proof(intro_locales) + show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG[OF vt]] . + next + show "fsubtree_axioms (RAG s)" + proof(unfold fsubtree_axioms_def) + find_theorems wf RAG + from wf_RAG show "wf (RAG s)" . + qed + qed +qed + +sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s" +proof - + have "fsubtree (tRAG s)" + proof - + have "fbranch (tRAG s)" + proof(unfold tRAG_def, rule fbranch_compose) + show "fbranch (wRAG s)" + proof(rule finite_fbranchI) + from finite_RAG[OF vt] show "finite (wRAG s)" + by (unfold RAG_split, auto) + qed + next + show "fbranch (hRAG s)" + proof(rule finite_fbranchI) + from finite_RAG[OF vt] + show "finite (hRAG s)" by (unfold RAG_split, auto) + qed + qed + moreover have "wf (tRAG s)" + proof(rule wf_subset) + show "wf (RAG s O RAG s)" using wf_RAG + by (fold wf_comp_self, simp) + next + show "tRAG s \ (RAG s O RAG s)" + by (unfold tRAG_alt_def, auto) + qed + ultimately show ?thesis + by (unfold fsubtree_def fsubtree_axioms_def,auto) + qed + from this[folded tRAG_def] show "fsubtree (tRAG s)" . +qed + +lemma Max_UNION: + assumes "finite A" + and "A \ {}" + and "\ M \ f ` A. finite M" + and "\ M \ f ` A. M \ {}" + shows "Max (\x\ A. f x) = Max (Max ` f ` A)" (is "?L = ?R") + using assms[simp] +proof - + have "?L = Max (\(f ` A))" + by (fold Union_image_eq, simp) + also have "... = ?R" + by (subst Max_Union, simp+) + finally show ?thesis . +qed + +lemma max_Max_eq: + assumes "finite A" + and "A \ {}" + and "x = y" + shows "max x (Max A) = Max ({y} \ A)" (is "?L = ?R") +proof - + have "?R = Max (insert y A)" by simp + also from assms have "... = ?L" + by (subst Max.insert, simp+) + finally show ?thesis by simp +qed + + +context valid_trace +begin + +(* ddd *) +lemma cp_gen_rec: + assumes "x = Th th" + shows "cp_gen s x = Max ({the_preced s th} \ (cp_gen s) ` children (tRAG s) x)" +proof(cases "children (tRAG s) x = {}") + case True + show ?thesis + by (unfold True cp_gen_def subtree_children, simp add:assms) +next + case False + hence [simp]: "children (tRAG s) x \ {}" by auto + note fsbttRAGs.finite_subtree[simp] + have [simp]: "finite (children (tRAG s) x)" + by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], + rule children_subtree) + { fix r x + have "subtree r x \ {}" by (auto simp:subtree_def) + } note this[simp] + have [simp]: "\x\children (tRAG s) x. subtree (tRAG s) x \ {}" + proof - + from False obtain q where "q \ children (tRAG s) x" by blast + moreover have "subtree (tRAG s) q \ {}" by simp + ultimately show ?thesis by blast + qed + have h: "Max ((the_preced s \ the_thread) ` + ({x} \ \(subtree (tRAG s) ` children (tRAG s) x))) = + Max ({the_preced s th} \ cp_gen s ` children (tRAG s) x)" + (is "?L = ?R") + proof - + let "Max (?f ` (?A \ \ (?g ` ?B)))" = ?L + let "Max (_ \ (?h ` ?B))" = ?R + let ?L1 = "?f ` \(?g ` ?B)" + have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)" + proof - + have "?L1 = ?f ` (\ x \ ?B.(?g x))" by simp + also have "... = (\ x \ ?B. ?f ` (?g x))" by auto + finally have "Max ?L1 = Max ..." by simp + also have "... = Max (Max ` (\x. ?f ` subtree (tRAG s) x) ` ?B)" + by (subst Max_UNION, simp+) + also have "... = Max (cp_gen s ` children (tRAG s) x)" + by (unfold image_comp cp_gen_alt_def, simp) + finally show ?thesis . + qed + show ?thesis + proof - + have "?L = Max (?f ` ?A \ ?L1)" by simp + also have "... = max (the_preced s (the_thread x)) (Max ?L1)" + by (subst Max_Un, simp+) + also have "... = max (?f x) (Max (?h ` ?B))" + by (unfold eq_Max_L1, simp) + also have "... =?R" + by (rule max_Max_eq, (simp)+, unfold assms, simp) + finally show ?thesis . + qed + qed thus ?thesis + by (fold h subtree_children, unfold cp_gen_def, simp) +qed + +lemma cp_rec: + "cp s th = Max ({the_preced s th} \ + (cp s o the_thread) ` children (tRAG s) (Th th))" +proof - + have "Th th = Th th" by simp + note h = cp_gen_def_cond[OF this] cp_gen_rec[OF this] + show ?thesis + proof - + have "cp_gen s ` children (tRAG s) (Th th) = + (cp s \ the_thread) ` children (tRAG s) (Th th)" + proof(rule cp_gen_over_set) + show " \x\children (tRAG s) (Th th). \th. x = Th th" + by (unfold tRAG_alt_def, auto simp:children_def) + qed + thus ?thesis by (subst (1) h(1), unfold h(2), simp) + qed +qed + +end + + lemma eq_dependants: "dependants (wq s) = dependants s" by (simp add: s_dependants_abv wq_def) - + + (* obvious lemma *) lemma not_thread_holdents: fixes th s @@ -547,7 +1046,7 @@ (* moving in Max *) also have "\ = max (Max {preced th s}) (Max ((\th. Max (preceds s ({th} \ (dependants (wq s) th)))) ` children s th))" - by (subst Max_Union) (auto simp add: image_image) + by (subst Max_Union) (auto simp add: image_image) (* folding cp + moving out Max *) also have "\ = ?RHS" @@ -557,62 +1056,51 @@ qed qed -lemma next_waiting: +lemma next_th_holding: + assumes vt: "vt s" + and nxt: "next_th s th cs th'" + shows "holding (wq s) th cs" +proof - + from nxt[unfolded next_th_def] + obtain rest where h: "wq s cs = th # rest" + "rest \ []" + "th' = hd (SOME q. distinct q \ set q = set rest)" by auto + thus ?thesis + by (unfold cs_holding_def, auto) +qed + +lemma next_th_waiting: assumes vt: "vt s" and nxt: "next_th s th cs th'" - shows "waiting s th' cs" + shows "waiting (wq s) th' cs" proof - - from assms show ?thesis - apply (auto simp:next_th_def s_waiting_def[folded wq_def]) - proof - - fix rest - assume ni: "hd (SOME q. distinct q \ set q = set rest) \ set rest" - and eq_wq: "wq s cs = th # rest" - and ne: "rest \ []" - have "set (SOME q. distinct q \ set q = set rest) = set rest" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - show "\x. distinct x \ set x = set rest \ set x = set rest" by auto - qed - with ni - have "hd (SOME q. distinct q \ set q = set rest) \ set (SOME q. distinct q \ set q = set rest)" - by simp - moreover have "(SOME q. distinct q \ set q = set rest) \ []" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - from ne show "\x. distinct x \ set x = set rest \ x \ []" by auto - qed - ultimately show "hd (SOME q. distinct q \ set q = set rest) = th" by auto + from nxt[unfolded next_th_def] + obtain rest where h: "wq s cs = th # rest" + "rest \ []" + "th' = hd (SOME q. distinct q \ set q = set rest)" by auto + from wq_distinct[OF vt, of cs, unfolded h] + have dst: "distinct (th # rest)" . + have in_rest: "th' \ set rest" + proof(unfold h, rule someI2) + show "distinct rest \ set rest = set rest" using dst by auto next - fix rest - assume eq_wq: "wq s cs = hd (SOME q. distinct q \ set q = set rest) # rest" - and ne: "rest \ []" - have "(SOME q. distinct q \ set q = set rest) \ []" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - from ne show "\x. distinct x \ set x = set rest \ x \ []" by auto - qed - hence "hd (SOME q. distinct q \ set q = set rest) \ set (SOME q. distinct q \ set q = set rest)" - by auto - moreover have "set (SOME q. distinct q \ set q = set rest) = set rest" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - show "\x. distinct x \ set x = set rest \ set x = set rest" by auto - qed - ultimately have "hd (SOME q. distinct q \ set q = set rest) \ set rest" by simp - with eq_wq and wq_distinct[OF vt, of cs] - show False by auto + fix x assume "distinct x \ set x = set rest" + with h(2) + show "hd x \ set (rest)" by (cases x, auto) qed + hence "th' \ set (wq s cs)" by (unfold h(1), auto) + moreover have "th' \ hd (wq s cs)" + by (unfold h(1), insert in_rest dst, auto) + ultimately show ?thesis by (auto simp:cs_waiting_def) qed +lemma next_th_RAG: + assumes vt: "vt s" + and nxt: "next_th s th cs th'" + shows "{(Cs cs, Th th), (Th th', Cs cs)} \ RAG s" + using assms next_th_holding next_th_waiting +by (unfold s_RAG_def, simp) + -- {* A useless definition *} definition cps:: "state \ (thread \ precedence) set" where "cps s = {(th, cp s th) | th . th \ threads s}" @@ -660,30 +1148,9 @@ context step_set_cps begin -interpretation rtree_RAGs: rtree "RAG s" -proof - show "single_valued (RAG s)" - apply (intro_locales) - by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s]) - - show "acyclic (RAG s)" - by (rule acyclic_RAG[OF vt_s]) -qed - -interpretation rtree_RAGs': rtree "RAG s'" -proof - show "single_valued (RAG s')" - apply (intro_locales) - by (unfold single_valued_def, - auto intro: unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) - - show "acyclic (RAG s')" - by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) -qed - text {* (* ddd *) - The following lemma confirms that @{text "Set"}-operating only changes the precedence - of initiating thread. + The following two lemmas confirm that @{text "Set"}-operating only changes the precedence + of the initiating thread. *} lemma eq_preced: @@ -695,103 +1162,86 @@ by (unfold s_def, auto simp:preced_def) qed -text {* (* ddd *) +lemma eq_the_preced: + fixes th' + assumes "th' \ th" + shows "the_preced s th' = the_preced s' th'" + using assms + by (unfold the_preced_def, intro eq_preced, simp) + +text {* The following lemma assures that the resetting of priority does not change the RAG. *} lemma eq_dep: "RAG s = RAG s'" by (unfold s_def RAG_set_unchanged, auto) -text {* - Th following lemma @{text "eq_cp_pre"} circumscribe a rough range of recalculation. - It says, thread other than the initiating thread @{text "th"} does not need recalculation - unless it lies upstream of @{text "th"} in the RAG. - - The reason behind this lemma is that: - the change of precedence of one thread can only affect it's upstream threads, according to - lemma @{text "eq_preced"}. Since the only thread which might change precedence is - @{text "th"}, so only @{text "th"} and its upstream threads need recalculation. - (* ccc *) +text {* (* ddd *) + Th following lemma @{text "eq_cp_pre"} says the priority change of @{text "th"} + only affects those threads, which as @{text "Th th"} in their sub-trees. + + The proof of this lemma is simplified by using the alternative definition of @{text "cp"}. *} lemma eq_cp_pre: fixes th' - assumes neq_th: "th' \ th" - and nd: "th \ dependants s th'" + assumes nd: "Th th \ subtree (RAG s') (Th th')" shows "cp s th' = cp s' th'" proof - - -- {* This is what we need to prove after expanding the definition of @{text "cp"} *} - have "Max ((\th. preced th s) ` ({th'} \ dependants (wq s) th')) = - Max ((\th. preced th s') ` ({th'} \ dependants (wq s') th'))" - (is "Max (?f1 ` ({th'} \ ?A)) = Max (?f2 ` ({th'} \ ?B))") + -- {* After unfolding using the alternative definition, elements + affecting the @{term "cp"}-value of threads become explicit. + We only need to prove the following: *} + have "Max (the_preced s ` {th'a. Th th'a \ subtree (RAG s) (Th th')}) = + Max (the_preced s' ` {th'a. Th th'a \ subtree (RAG s') (Th th')})" + (is "Max (?f ` ?S1) = Max (?g ` ?S2)") proof - - -- {* Since RAG is not changed by @{text "Set"}-operation, the dependants of - any threads are not changed, this is one of key facts underpinning this - lemma *} - have eq_ab: "?A = ?B" by (unfold cs_dependants_def, auto simp:eq_dep eq_RAG) - have "(?f1 ` ({th'} \ ?A)) = (?f2 ` ({th'} \ ?B))" - proof(rule image_cong) - show "{th'} \ ?A = {th'} \ ?B" by (simp only:eq_ab) - next - fix x - assume x_in: "x \ {th'} \ ?B" - show "?f1 x = ?f2 x" - proof(rule eq_preced) -- {* The other key fact underpinning this lemma is @{text "eq_preced"} *} - from x_in[folded eq_ab, unfolded eq_dependants] - have "x \ {th'} \ dependants s th'" . - thus "x \ th" - proof - assume "x \ {th'}" - with `th' \ th` show ?thesis by simp - next - assume "x \ dependants s th'" - with `th \ dependants s th'` show ?thesis by auto - qed - qed - qed - thus ?thesis by simp - qed - thus ?thesis by (unfold cp_eq_cpreced cpreced_def) + -- {* The base sets are equal. *} + have "?S1 = ?S2" using eq_dep by simp + -- {* The function values on the base set are equal as well. *} + moreover have "\ e \ ?S2. ?f e = ?g e" + proof + fix th1 + assume "th1 \ ?S2" + with nd have "th1 \ th" by (auto) + from eq_the_preced[OF this] + show "the_preced s th1 = the_preced s' th1" . + qed + -- {* Therefore, the image of the functions are equal. *} + ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq) + thus ?thesis by simp + qed + thus ?thesis by (simp add:cp_alt_def) qed text {* - The following lemma shows that no thread lies upstream of the initiating thread @{text "th"}. - The reason for this is that only no-blocked thread can initiate - a system call. Since thread @{text "th"} is non-blocked, it is not waiting for any - critical resource. Therefore, there is edge leading out of @{text "th"} in the RAG. - Consequently, there is no node (neither resource nor thread) upstream of @{text "th"}. + The following lemma shows that @{term "th"} is not in the + sub-tree of any other thread. *} -lemma no_dependants: - shows "th \ dependants s th'" -proof - assume "th \ dependants s th'" - from `th \ dependants s th'` have "(Th th, Th th') \ (RAG s')\<^sup>+" - by (unfold s_dependants_def, unfold eq_RAG, unfold eq_dep, auto) - from tranclD[OF this] - obtain z where "(Th th, z) \ RAG s'" by auto - moreover have "th \ readys s'" +lemma th_in_no_subtree: + assumes "th' \ th" + shows "Th th \ subtree (RAG s') (Th th')" +proof - + have "th \ readys s'" proof - from step_back_step [OF vt_s[unfolded s_def]] have "step s' (Set th prio)" . hence "th \ runing s'" by (cases, simp) thus ?thesis by (simp add:readys_def runing_def) qed - ultimately show "False" - apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def) - by (fold wq_def, blast) + from readys_in_no_subtree[OF step_back_vt[OF vt_s[unfolded s_def]] this assms(1)] + show ?thesis by blast qed -(* Result improved *) text {* - A simple combination of @{text "eq_cp_pre"} and @{text "no_dependants"} - gives the main lemma of this locale, which shows that - only the initiating thread needs a recalculation of current precedence. + By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, + it is obvious that the change of priority only affects the @{text "cp"}-value + of the initiating thread @{text "th"}. *} lemma eq_cp: fixes th' assumes "th' \ th" shows "cp s th' = cp s' th'" - by (rule eq_cp_pre[OF assms no_dependants]) + by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) end @@ -807,6 +1257,83 @@ -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *} assumes vt_s: "vt s" +context step_v_cps +begin + +lemma rtree_RAGs: "rtree (RAG s)" +proof + show "single_valued (RAG s)" + apply (intro_locales) + by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s]) + + show "acyclic (RAG s)" + by (rule acyclic_RAG[OF vt_s]) +qed + +lemma rtree_RAGs': "rtree (RAG s')" +proof + show "single_valued (RAG s')" + apply (intro_locales) + by (unfold single_valued_def, + auto intro:unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) + + show "acyclic (RAG s')" + by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) +qed + +lemmas vt_s' = step_back_vt[OF vt_s[unfolded s_def]] + +lemma ready_th_s': "th \ readys s'" + using step_back_step[OF vt_s[unfolded s_def]] + by (cases, simp add:runing_def) + + +lemma ancestors_th: "ancestors (RAG s') (Th th) = {}" +proof - + from readys_root[OF vt_s' ready_th_s'] + show ?thesis + by (unfold root_def, simp) +qed + +lemma holding_th: "holding s' th cs" +proof - + from vt_s[unfolded s_def] + have " PIP s' (V th cs)" by (cases, simp) + thus ?thesis by (cases, auto) +qed + +lemma edge_of_th: + "(Cs cs, Th th) \ RAG s'" +proof - + from holding_th + show ?thesis + by (unfold s_RAG_def holding_eq, auto) +qed + +lemma ancestors_cs: + "ancestors (RAG s') (Cs cs) = {Th th}" +proof - + find_theorems ancestors + have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th) \ {Th th}" + proof(rule RTree.rtree.ancestors_accum[OF rtree_RAGs']) + from vt_s[unfolded s_def] + have " PIP s' (V th cs)" by (cases, simp) + thus "(Cs cs, Th th) \ RAG s'" + proof(cases) + assume "holding s' th cs" + from this[unfolded holding_eq] + show ?thesis by (unfold s_RAG_def, auto) + qed + qed + from this[unfolded ancestors_th] show ?thesis by simp +qed + +lemma preced_kept: "the_preced s = the_preced s'" + by (auto simp: s_def the_preced_def preced_def) + +end + + text {* The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, which represents the case when there is another thread @{text "th'"} @@ -860,6 +1387,19 @@ th7 ----| *) +lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \ RAG s'" + using next_th_RAG[OF vt_s' nt] . + +lemma ancestors_th': + "ancestors (RAG s') (Th th') = {Th th, Cs cs}" +proof - + have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \ {Cs cs}" + proof(rule RTree.rtree.ancestors_accum[OF rtree_RAGs']) + from sub_RAGs' show "(Th th', Cs cs) \ RAG s'" by auto + qed + thus ?thesis using ancestors_th ancestors_cs by auto +qed + lemma RAG_s: "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \ {(Cs cs, Th th')}" @@ -868,199 +1408,54 @@ and nt show ?thesis by (auto intro:next_th_unique) qed -text {* - Lemma @{text "dependants_kept"} shows only @{text "th"} and @{text "th'"} - have their dependants changed. -*} -lemma dependants_kept: - fixes th'' - assumes neq1: "th'' \ th" - and neq2: "th'' \ th'" - shows "dependants (wq s) th'' = dependants (wq s') th''" -proof(auto) (* ccc *) - fix x - assume "x \ dependants (wq s) th''" - hence dp: "(Th x, Th th'') \ (RAG s)^+" - by (auto simp:cs_dependants_def eq_RAG) - { fix n - have "(n, Th th'') \ (RAG s)^+ \ (n, Th th'') \ (RAG s')^+" - proof(induct rule:converse_trancl_induct) - fix y - assume "(y, Th th'') \ RAG s" - with RAG_s neq1 neq2 - have "(y, Th th'') \ RAG s'" by auto - thus "(y, Th th'') \ (RAG s')\<^sup>+" by auto - next - fix y z - assume yz: "(y, z) \ RAG s" - and ztp: "(z, Th th'') \ (RAG s)\<^sup>+" - and ztp': "(z, Th th'') \ (RAG s')\<^sup>+" - have "y \ Cs cs \ y \ Th th'" - proof - show "y \ Cs cs" - proof - assume eq_y: "y = Cs cs" - with yz have dp_yz: "(Cs cs, z) \ RAG s" by simp - from RAG_s - have cst': "(Cs cs, Th th') \ RAG s" by simp - from unique_RAG[OF vt_s this dp_yz] - have eq_z: "z = Th th'" by simp - with ztp have "(Th th', Th th'') \ (RAG s)^+" by simp - from converse_tranclE[OF this] - obtain cs' where dp'': "(Th th', Cs cs') \ RAG s" - by (auto simp:s_RAG_def) - with RAG_s have dp': "(Th th', Cs cs') \ RAG s'" by auto - from dp'' eq_y yz eq_z have "(Cs cs, Cs cs') \ (RAG s)^+" by auto - moreover have "cs' = cs" - proof - - from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt] - have "(Th th', Cs cs) \ RAG s'" - by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this dp'] - show ?thesis by simp - qed - ultimately have "(Cs cs, Cs cs) \ (RAG s)^+" by simp - moreover note wf_trancl[OF wf_RAG[OF vt_s]] - ultimately show False by auto - qed +lemma subtree_kept: + assumes "th1 \ {th, th'}" + shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R") +proof - + let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})" + let ?RAG'' = "?RAG' \ {(Cs cs, Th th')}" + have "subtree ?RAG' (Th th1) = ?R" + proof(rule subset_del_subtree_outside) + show "Range {(Cs cs, Th th), (Th th', Cs cs)} \ subtree (RAG s') (Th th1) = {}" + proof - + have "(Th th) \ subtree (RAG s') (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s') (Th th)" + by (unfold ancestors_th, simp) next - show "y \ Th th'" - proof - assume eq_y: "y = Th th'" - with yz have dps: "(Th th', z) \ RAG s" by simp - with RAG_s have dps': "(Th th', z) \ RAG s'" by auto - have "z = Cs cs" - proof - - from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt] - have "(Th th', Cs cs) \ RAG s'" - by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] dps' this] - show ?thesis . - qed - with dps RAG_s show False by auto - qed + from assms show "Th th1 \ Th th" by simp qed - with RAG_s yz have "(y, z) \ RAG s'" by auto - with ztp' - show "(y, Th th'') \ (RAG s')\<^sup>+" by auto - qed - } - from this[OF dp] - show "x \ dependants (wq s') th''" - by (auto simp:cs_dependants_def eq_RAG) -next - fix x - assume "x \ dependants (wq s') th''" - hence dp: "(Th x, Th th'') \ (RAG s')^+" - by (auto simp:cs_dependants_def eq_RAG) - { fix n - have "(n, Th th'') \ (RAG s')^+ \ (n, Th th'') \ (RAG s)^+" - proof(induct rule:converse_trancl_induct) - fix y - assume "(y, Th th'') \ RAG s'" - with RAG_s neq1 neq2 - have "(y, Th th'') \ RAG s" by auto - thus "(y, Th th'') \ (RAG s)\<^sup>+" by auto + moreover have "(Cs cs) \ subtree (RAG s') (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s') (Cs cs)" + by (unfold ancestors_cs, insert assms, auto) + qed simp + ultimately have "{Th th, Cs cs} \ subtree (RAG s') (Th th1) = {}" by auto + thus ?thesis by simp + qed + qed + moreover have "subtree ?RAG'' (Th th1) = subtree ?RAG' (Th th1)" + proof(rule subtree_insert_next) + show "Th th' \ subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')" + (is "_ \ ?R") + proof - + have "?R \ ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto) + moreover have "Th th1 \ ..." using ancestors_th' assms by simp + ultimately show ?thesis by auto + qed next - fix y z - assume yz: "(y, z) \ RAG s'" - and ztp: "(z, Th th'') \ (RAG s')\<^sup>+" - and ztp': "(z, Th th'') \ (RAG s)\<^sup>+" - have "y \ Cs cs \ y \ Th th'" - proof - show "y \ Cs cs" - proof - assume eq_y: "y = Cs cs" - with yz have dp_yz: "(Cs cs, z) \ RAG s'" by simp - from this have eq_z: "z = Th th" - proof - - from step_back_step[OF vt_s[unfolded s_def]] - have "(Cs cs, Th th) \ RAG s'" - by(cases, auto simp: wq_def s_RAG_def cs_holding_def s_holding_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this dp_yz] - show ?thesis by simp - qed - from converse_tranclE[OF ztp] - obtain u where "(z, u) \ RAG s'" by auto - moreover - from step_back_step[OF vt_s[unfolded s_def]] - have "th \ readys s'" by (cases, simp add:runing_def) - moreover note eq_z - ultimately show False - by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def) - qed - next - show "y \ Th th'" - proof - assume eq_y: "y = Th th'" - with yz have dps: "(Th th', z) \ RAG s'" by simp - have "z = Cs cs" - proof - - from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt] - have "(Th th', Cs cs) \ RAG s'" - by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] dps this] - show ?thesis . - qed - with ztp have cs_i: "(Cs cs, Th th'') \ (RAG s')\<^sup>+" by simp - from step_back_step[OF vt_s[unfolded s_def]] - have cs_th: "(Cs cs, Th th) \ RAG s'" - by(cases, auto simp: s_RAG_def wq_def cs_holding_def s_holding_def) - have "(Cs cs, Th th'') \ RAG s'" - proof - assume "(Cs cs, Th th'') \ RAG s'" - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this cs_th] - and neq1 show "False" by simp - qed - with converse_tranclE[OF cs_i] - obtain u where cu: "(Cs cs, u) \ RAG s'" - and u_t: "(u, Th th'') \ (RAG s')\<^sup>+" by auto - have "u = Th th" - proof - - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] cu cs_th] - show ?thesis . - qed - with u_t have "(Th th, Th th'') \ (RAG s')\<^sup>+" by simp - from converse_tranclE[OF this] - obtain v where "(Th th, v) \ (RAG s')" by auto - moreover from step_back_step[OF vt_s[unfolded s_def]] - have "th \ readys s'" by (cases, simp add:runing_def) - ultimately show False - by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def) - qed - qed - with RAG_s yz have "(y, z) \ RAG s" by auto - with ztp' - show "(y, Th th'') \ (RAG s)\<^sup>+" by auto - qed - } - from this[OF dp] - show "x \ dependants (wq s) th''" - by (auto simp:cs_dependants_def eq_RAG) + from assms show "Th th1 \ Th th'" by simp + qed + qed + ultimately show ?thesis by (unfold RAG_s, simp) qed lemma cp_kept: - fixes th'' - assumes neq1: "th'' \ th" - and neq2: "th'' \ th'" - shows "cp s th'' = cp s' th''" -proof - - from dependants_kept[OF neq1 neq2] - have "dependants (wq s) th'' = dependants (wq s') th''" . - moreover { - fix th1 - assume "th1 \ dependants (wq s) th''" - have "preced th1 s = preced th1 s'" - by (unfold s_def, auto simp:preced_def) - } - moreover have "preced th'' s = preced th'' s'" - by (unfold s_def, auto simp:preced_def) - ultimately have "((\th. preced th s) ` ({th''} \ dependants (wq s) th'')) = - ((\th. preced th s') ` ({th''} \ dependants (wq s') th''))" - by (auto simp:image_def) - thus ?thesis - by (unfold cp_eq_cpreced cpreced_def, simp) -qed + assumes "th1 \ {th, th'}" + shows "cp s th1 = cp s' th1" + by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) end @@ -1070,152 +1465,143 @@ context step_v_cps_nnt begin -lemma nw_cs: "(Th th1, Cs cs) \ RAG s'" -proof - assume "(Th th1, Cs cs) \ RAG s'" - thus "False" - apply (auto simp:s_RAG_def cs_waiting_def) - proof - - assume h1: "th1 \ set (wq s' cs)" - and h2: "th1 \ hd (wq s' cs)" - from step_back_step[OF vt_s[unfolded s_def]] - show "False" - proof(cases) - assume "holding s' th cs" - then obtain rest where - eq_wq: "wq s' cs = th#rest" - apply (unfold s_holding_def wq_def[symmetric]) - by (case_tac "(wq s' cs)", auto) - with h1 h2 have ne: "rest \ []" by auto - with eq_wq - have "next_th s' th cs (hd (SOME q. distinct q \ set q = set rest))" - by(unfold next_th_def, rule_tac x = "rest" in exI, auto) - with nnt show ?thesis by auto - qed - qed -qed - lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}" proof - from nnt and step_RAG_v[OF vt_s[unfolded s_def], folded s_def] show ?thesis by auto qed -lemma child_kept_left: - assumes - "(n1, n2) \ (child s')^+" - shows "(n1, n2) \ (child s)^+" -proof - - from assms show ?thesis - proof(induct rule: converse_trancl_induct) - case (base y) - from base obtain th1 cs1 th2 - where h1: "(Th th1, Cs cs1) \ RAG s'" - and h2: "(Cs cs1, Th th2) \ RAG s'" - and eq_y: "y = Th th1" and eq_n2: "n2 = Th th2" by (auto simp:child_def) - have "cs1 \ cs" - proof - assume eq_cs: "cs1 = cs" - with h1 have "(Th th1, Cs cs1) \ RAG s'" by simp - with nw_cs eq_cs show False by auto +lemma subtree_kept: + assumes "th1 \ th" + shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" +proof(unfold RAG_s, rule subset_del_subtree_outside) + show "Range {(Cs cs, Th th)} \ subtree (RAG s') (Th th1) = {}" + proof - + have "(Th th) \ subtree (RAG s') (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s') (Th th)" + by (unfold ancestors_th, simp) + next + from assms show "Th th1 \ Th th" by simp qed - with h1 h2 RAG_s have - h1': "(Th th1, Cs cs1) \ RAG s" and - h2': "(Cs cs1, Th th2) \ RAG s" by auto - hence "(Th th1, Th th2) \ child s" by (auto simp:child_def) - with eq_y eq_n2 have "(y, n2) \ child s" by simp - thus ?case by auto - next - case (step y z) - have "(y, z) \ child s'" by fact - then obtain th1 cs1 th2 - where h1: "(Th th1, Cs cs1) \ RAG s'" - and h2: "(Cs cs1, Th th2) \ RAG s'" - and eq_y: "y = Th th1" and eq_z: "z = Th th2" by (auto simp:child_def) - have "cs1 \ cs" - proof - assume eq_cs: "cs1 = cs" - with h1 have "(Th th1, Cs cs1) \ RAG s'" by simp - with nw_cs eq_cs show False by auto - qed - with h1 h2 RAG_s have - h1': "(Th th1, Cs cs1) \ RAG s" and - h2': "(Cs cs1, Th th2) \ RAG s" by auto - hence "(Th th1, Th th2) \ child s" by (auto simp:child_def) - with eq_y eq_z have "(y, z) \ child s" by simp - moreover have "(z, n2) \ (child s)^+" by fact - ultimately show ?case by auto + thus ?thesis by auto qed qed -lemma child_kept_right: - assumes - "(n1, n2) \ (child s)^+" - shows "(n1, n2) \ (child s')^+" +lemma cp_kept_1: + assumes "th1 \ th" + shows "cp s th1 = cp s' th1" + by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) + +lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}" proof - - from assms show ?thesis - proof(induct) - case (base y) - from base and RAG_s - have "(n1, y) \ child s'" - by (auto simp:child_def) - thus ?case by auto - next - case (step y z) - have "(y, z) \ child s" by fact - with RAG_s have "(y, z) \ child s'" - by (auto simp:child_def) - moreover have "(n1, y) \ (child s')\<^sup>+" by fact - ultimately show ?case by auto - qed + { fix n + have "(Cs cs) \ ancestors (RAG s') n" + proof + assume "Cs cs \ ancestors (RAG s') n" + hence "(n, Cs cs) \ (RAG s')^+" by (auto simp:ancestors_def) + from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \ RAG s'" by auto + then obtain th' where "nn = Th th'" + by (unfold s_RAG_def, auto) + from h[unfolded this] have "(Th th', Cs cs) \ RAG s'" . + from this[unfolded s_RAG_def] + have "waiting (wq s') th' cs" by auto + from this[unfolded cs_waiting_def] + have "1 < length (wq s' cs)" + by (cases "wq s' cs", auto) + from holding_next_thI[OF holding_th this] + obtain th' where "next_th s' th cs th'" by auto + with nnt show False by auto + qed + } note h = this + { fix n + assume "n \ subtree (RAG s') (Cs cs)" + hence "n = (Cs cs)" + by (elim subtreeE, insert h, auto) + } moreover have "(Cs cs) \ subtree (RAG s') (Cs cs)" + by (auto simp:subtree_def) + ultimately show ?thesis by auto qed -lemma eq_child: "(child s)^+ = (child s')^+" - by (insert child_kept_left child_kept_right, auto) +lemma subtree_th: + "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}" +proof(unfold RAG_s, fold subtree_cs, rule RTree.rtree.subtree_del_inside[OF rtree_RAGs']) + from edge_of_th + show "(Cs cs, Th th) \ edges_in (RAG s') (Th th)" + by (unfold edges_in_def, auto simp:subtree_def) +qed + +lemma cp_kept_2: + shows "cp s th = cp s' th" + by (unfold cp_alt_def subtree_th preced_kept, auto) lemma eq_cp: fixes th' shows "cp s th' = cp s' th'" - apply (unfold cp_eq_cpreced cpreced_def) -proof - - have eq_dp: "\ th. dependants (wq s) th = dependants (wq s') th" - apply (unfold cs_dependants_def, unfold eq_RAG) - proof - - from eq_child - have "\th. {th'. (Th th', Th th) \ (child s)\<^sup>+} = {th'. (Th th', Th th) \ (child s')\<^sup>+}" - by simp - with child_RAG_eq[OF vt_s] child_RAG_eq[OF step_back_vt[OF vt_s[unfolded s_def]]] - show "\th. {th'. (Th th', Th th) \ (RAG s)\<^sup>+} = {th'. (Th th', Th th) \ (RAG s')\<^sup>+}" - by simp - qed - moreover { - fix th1 - assume "th1 \ {th'} \ dependants (wq s') th'" - hence "th1 = th' \ th1 \ dependants (wq s') th'" by auto - hence "preced th1 s = preced th1 s'" - proof - assume "th1 = th'" - show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def) - next - assume "th1 \ dependants (wq s') th'" - show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def) - qed - } ultimately have "((\th. preced th s) ` ({th'} \ dependants (wq s) th')) = - ((\th. preced th s') ` ({th'} \ dependants (wq s') th'))" - by (auto simp:image_def) - thus "Max ((\th. preced th s) ` ({th'} \ dependants (wq s) th')) = - Max ((\th. preced th s') ` ({th'} \ dependants (wq s') th'))" by simp -qed + using cp_kept_1 cp_kept_2 + by (cases "th' = th", auto) + +end + +find_theorems "_`_" "\ _" -end +find_theorems "Max" "\ _" + +find_theorems wf RAG + +thm wf_def + +thm image_Union locale step_P_cps = fixes s' th cs s defines s_def : "s \ (P th cs#s')" assumes vt_s: "vt s" +sublocale step_P_cps < vat_s : valid_trace "s" +proof + from vt_s show "vt s" . +qed + +sublocale step_P_cps < vat_s' : valid_trace "s'" +proof + from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . +qed + + +context step_P_cps +begin + +lemma rtree_RAGs: "rtree (RAG s)" +proof + show "single_valued (RAG s)" + apply (intro_locales) + by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s]) + + show "acyclic (RAG s)" + by (rule acyclic_RAG[OF vt_s]) +qed + +lemma rtree_RAGs': "rtree (RAG s')" +proof + show "single_valued (RAG s')" + apply (intro_locales) + by (unfold single_valued_def, + auto intro:unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) + + show "acyclic (RAG s')" + by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) +qed + +lemma preced_kept: "the_preced s = the_preced s'" + by (auto simp: s_def the_preced_def preced_def) + +end + locale step_P_cps_ne =step_P_cps + + fixes th' assumes ne: "wq s' cs \ []" + defines th'_def: "th' \ hd (wq s' cs)" locale step_P_cps_e =step_P_cps + assumes ee: "wq s' cs = []" @@ -1353,7 +1739,20 @@ end -context step_P_cps_ne +lemma tRAG_ancestorsE: + assumes "x \ ancestors (tRAG s) u" + obtains th where "x = Th th" +proof - + from assms have "(u, x) \ (tRAG s)^+" + by (unfold ancestors_def, auto) + from tranclE[OF this] obtain c where "(c, x) \ tRAG s" by auto + then obtain th where "x = Th th" + by (unfold tRAG_alt_def, auto) + from that[OF this] show ?thesis . +qed + + +context step_P_cps_ne begin lemma RAG_s: "RAG s = RAG s' \ {(Th th, Cs cs)}" @@ -1362,6 +1761,184 @@ show ?thesis by (simp add:s_def) qed +lemma cs_held: "(Cs cs, Th th') \ RAG s'" +proof - + have "(Cs cs, Th th') \ hRAG s'" + proof - + from ne + have " holding s' th' cs" + by (unfold th'_def holding_eq cs_holding_def, auto) + thus ?thesis + by (unfold hRAG_def, auto) + qed + thus ?thesis by (unfold RAG_split, auto) +qed + +lemma tRAG_s: + "tRAG s = tRAG s' \ {(Th th, Th th')}" + using RAG_tRAG_transfer[OF step_back_vt[OF vt_s[unfolded s_def]] RAG_s cs_held] . + +lemma cp_kept: + assumes "Th th'' \ ancestors (tRAG s) (Th th)" + shows "cp s th'' = cp s' th''" +proof - + have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')" + proof - + have "Th th' \ subtree (tRAG s') (Th th'')" + proof + assume "Th th' \ subtree (tRAG s') (Th th'')" + thus False + proof(rule subtreeE) + assume "Th th' = Th th''" + from assms[unfolded tRAG_s ancestors_def, folded this] + show ?thesis by auto + next + assume "Th th'' \ ancestors (tRAG s') (Th th')" + moreover have "... \ ancestors (tRAG s) (Th th')" + proof(rule ancestors_mono) + show "tRAG s' \ tRAG s" by (unfold tRAG_s, auto) + qed + ultimately have "Th th'' \ ancestors (tRAG s) (Th th')" by auto + moreover have "Th th' \ ancestors (tRAG s) (Th th)" + by (unfold tRAG_s, auto simp:ancestors_def) + ultimately have "Th th'' \ ancestors (tRAG s) (Th th)" + by (auto simp:ancestors_def) + with assms show ?thesis by auto + qed + qed + from subtree_insert_next[OF this] + have "subtree (tRAG s' \ {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" . + from this[folded tRAG_s] show ?thesis . + qed + show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) +qed + +lemma set_prop_split: + "A = {x. x \ A \ PP x} \ {x. x \ A \ \ PP x}" + by auto + +lemma f_image_union_eq: + assumes "f ` A = g ` A" + and "f ` B = g ` B" + shows "f ` (A \ B) = g ` (A \ B)" + using assms by auto + +(* ccc *) + +lemma cp_gen_update_stop: + assumes "u \ ancestors (tRAG s) (Th th)" + and "cp_gen s u = cp_gen s' u" + and "y \ ancestors (tRAG s) u" + shows "cp_gen s y = cp_gen s' y" + using assms(3) +proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf]) + case (1 x) + show ?case (is "?L = ?R") + proof - + from tRAG_ancestorsE[OF 1(2)] + obtain th2 where eq_x: "x = Th th2" by blast + from vat_s.cp_gen_rec[OF this] + have "?L = + Max ({the_preced s th2} \ cp_gen s ` RTree.children (tRAG s) x)" . + also have "... = + Max ({the_preced s' th2} \ cp_gen s' ` RTree.children (tRAG s') x)" + proof - + from preced_kept have "the_preced s th2 = the_preced s' th2" by simp + moreover have "cp_gen s ` RTree.children (tRAG s) x = + cp_gen s' ` RTree.children (tRAG s') x" + proof - + have "RTree.children (tRAG s) x = RTree.children (tRAG s') x" + proof(unfold tRAG_s, rule children_union_kept) + have start: "(Th th, Th th') \ tRAG s" + by (unfold tRAG_s, auto) + note x_u = 1(2) + show "x \ Range {(Th th, Th th')}" + proof + assume "x \ Range {(Th th, Th th')}" + hence eq_x: "x = Th th'" using RangeE by auto + show False + proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start]) + case 1 + from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG + show ?thesis by (auto simp:ancestors_def acyclic_def) + next + case 2 + with x_u[unfolded eq_x] + have "(Th th', Th th') \ (tRAG s)^+" by (auto simp:ancestors_def) + with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) + qed + qed + qed + moreover have "cp_gen s ` RTree.children (tRAG s) x = + cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A") + proof(rule f_image_eq) + fix a + assume a_in: "a \ ?A" + from 1(2) + show "?f a = ?g a" + proof(cases rule:vat_s.rtree_s.ancestors_childrenE[case_names in_ch out_ch]) + case in_ch + show ?thesis + proof(cases "a = u") + case True + from assms(2)[folded this] show ?thesis . + next + case False + have a_not_in: "a \ ancestors (tRAG s) (Th th)" + proof + assume a_in': "a \ ancestors (tRAG s) (Th th)" + have "a = u" + proof(rule vat_s.rtree_s.ancestors_children_unique) + from a_in' a_in show "a \ ancestors (tRAG s) (Th th) \ + RTree.children (tRAG s) x" by auto + next + from assms(1) in_ch show "u \ ancestors (tRAG s) (Th th) \ + RTree.children (tRAG s) x" by auto + qed + with False show False by simp + qed + from a_in obtain th_a where eq_a: "a = Th th_a" + by (unfold RTree.children_def tRAG_alt_def, auto) + from cp_kept[OF a_not_in[unfolded eq_a]] + have "cp s th_a = cp s' th_a" . + from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a] + show ?thesis . + qed + next + case (out_ch z) + hence h: "z \ ancestors (tRAG s) u" "z \ RTree.children (tRAG s) x" by auto + show ?thesis + proof(cases "a = z") + case True + from h(2) have zx_in: "(z, x) \ (tRAG s)" by (auto simp:RTree.children_def) + from 1(1)[rule_format, OF this h(1)] + have eq_cp_gen: "cp_gen s z = cp_gen s' z" . + with True show ?thesis by metis + next + case False + from a_in obtain th_a where eq_a: "a = Th th_a" + by (auto simp:RTree.children_def tRAG_alt_def) + have "a \ ancestors (tRAG s) (Th th)" sorry + from cp_kept[OF this[unfolded eq_a]] + have "cp s th_a = cp s' th_a" . + from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a] + show ?thesis . + qed + qed + qed + ultimately show ?thesis by metis + qed + ultimately show ?thesis by simp + qed + also have "... = ?R" + by (fold vat_s'.cp_gen_rec[OF eq_x], simp) + finally show ?thesis . + qed +qed + + + +(* ccc *) lemma eq_child_left: assumes nd: "(Th th, Th th') \ (child s)^+" diff -r f1b39d77db00 -r ad57323fd4d6 CpsG.thy~ --- a/CpsG.thy~ Thu Dec 03 14:34:29 2015 +0800 +++ b/CpsG.thy~ Tue Dec 15 21:45:46 2015 +0800 @@ -6,27 +6,148 @@ imports PrioG Max RTree begin -locale pip = - fixes s - assumes vt: "vt s" +definition "wRAG (s::state) = {(Th th, Cs cs) | th cs. waiting s th cs}" + +definition "hRAG (s::state) = {(Cs cs, Th th) | th cs. holding s th cs}" + +definition "tRAG s = wRAG s O hRAG s" -context pip -begin +lemma RAG_split: "RAG s = (wRAG s \ hRAG s)" + by (unfold s_RAG_abv wRAG_def hRAG_def s_waiting_abv + s_holding_abv cs_RAG_def, auto) + +lemma tRAG_alt_def: + "tRAG s = {(Th th1, Th th2) | th1 th2. + \ cs. (Th th1, Cs cs) \ RAG s \ (Cs cs, Th th2) \ RAG s}" + by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def) -interpretation rtree_RAG: rtree "RAG s" -proof - show "single_valued (RAG s)" - by (unfold single_valued_def, auto intro: unique_RAG[OF vt]) +lemma tRAG_mono: + assumes "RAG s' \ RAG s" + shows "tRAG s' \ tRAG s" + using assms + by (unfold tRAG_alt_def, auto) - show "acyclic (RAG s)" - by (rule acyclic_RAG[OF vt]) +lemma holding_next_thI: + assumes "holding s th cs" + and "length (wq s cs) > 1" + obtains th' where "next_th s th cs th'" +proof - + from assms(1)[folded eq_holding, unfolded cs_holding_def] + have " th \ set (wq s cs) \ th = hd (wq s cs)" . + then obtain rest where h1: "wq s cs = th#rest" + by (cases "wq s cs", auto) + with assms(2) have h2: "rest \ []" by auto + let ?th' = "hd (SOME q. distinct q \ set q = set rest)" + have "next_th s th cs ?th'" using h1(1) h2 + by (unfold next_th_def, auto) + from that[OF this] show ?thesis . qed -thm rtree_RAG.rpath_overlap_oneside +lemma RAG_tRAG_transfer: + assumes "vt s'" + assumes "RAG s = RAG s' \ {(Th th, Cs cs)}" + and "(Cs cs, Th th'') \ RAG s'" + shows "tRAG s = tRAG s' \ {(Th th, Th th'')}" (is "?L = ?R") +proof - + interpret rtree: rtree "RAG s'" + proof + show "single_valued (RAG s')" + apply (intro_locales) + by (unfold single_valued_def, + auto intro:unique_RAG[OF assms(1)]) -end + show "acyclic (RAG s')" + by (rule acyclic_RAG[OF assms(1)]) + qed + { fix n1 n2 + assume "(n1, n2) \ ?L" + from this[unfolded tRAG_alt_def] + obtain th1 th2 cs' where + h: "n1 = Th th1" "n2 = Th th2" + "(Th th1, Cs cs') \ RAG s" + "(Cs cs', Th th2) \ RAG s" by auto + from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \ RAG s'" by auto + from h(3) and assms(2) + have "(Th th1, Cs cs') = (Th th, Cs cs) \ + (Th th1, Cs cs') \ RAG s'" by auto + hence "(n1, n2) \ ?R" + proof + assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)" + hence eq_th1: "th1 = th" by simp + moreover have "th2 = th''" + proof - + from h1 have "cs' = cs" by simp + from assms(3) cs_in[unfolded this] rtree.sgv + show ?thesis + by (unfold single_valued_def, auto) + qed + ultimately show ?thesis using h(1,2) by auto + next + assume "(Th th1, Cs cs') \ RAG s'" + with cs_in have "(Th th1, Th th2) \ tRAG s'" + by (unfold tRAG_alt_def, auto) + from this[folded h(1, 2)] show ?thesis by auto + qed + } moreover { + fix n1 n2 + assume "(n1, n2) \ ?R" + hence "(n1, n2) \tRAG s' \ (n1, n2) = (Th th, Th th'')" by auto + hence "(n1, n2) \ ?L" + proof + assume "(n1, n2) \ tRAG s'" + moreover have "... \ ?L" + proof(rule tRAG_mono) + show "RAG s' \ RAG s" by (unfold assms(2), auto) + qed + ultimately show ?thesis by auto + next + assume eq_n: "(n1, n2) = (Th th, Th th'')" + from assms(2, 3) have "(Cs cs, Th th'') \ RAG s" by auto + moreover have "(Th th, Cs cs) \ RAG s" using assms(2) by auto + ultimately show ?thesis + by (unfold eq_n tRAG_alt_def, auto) + qed + } ultimately show ?thesis by auto +qed +lemma readys_root: + assumes "vt s" + and "th \ readys s" + shows "root (RAG s) (Th th)" +proof - + { fix x + assume "x \ ancestors (RAG s) (Th th)" + hence h: "(Th th, x) \ (RAG s)^+" by (auto simp:ancestors_def) + from tranclD[OF this] + obtain z where "(Th th, z) \ RAG s" by auto + with assms(2) have False + apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def) + by (fold wq_def, blast) + } thus ?thesis by (unfold root_def, auto) +qed +lemma readys_in_no_subtree: + assumes "vt s" + and "th \ readys s" + and "th' \ th" + shows "Th th \ subtree (RAG s) (Th th')" +proof + assume "Th th \ subtree (RAG s) (Th th')" + thus False + proof(cases rule:subtreeE) + case 1 + with assms show ?thesis by auto + next + case 2 + with readys_root[OF assms(1,2)] + show ?thesis by (auto simp:root_def) + qed +qed + +lemma image_id: + assumes "\ x. x \ A \ f x = x" + shows "f ` A = A" + using assms by (auto simp:image_def) definition "the_preced s th = preced th s" @@ -45,9 +166,386 @@ thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp) qed +fun the_thread :: "node \ thread" where + "the_thread (Th th) = th" + +definition "cp_gen s x = + Max ((the_preced s \ the_thread) ` subtree (tRAG s) x)" + +lemma cp_gen_alt_def: + "cp_gen s = (Max \ (\x. (the_preced s \ the_thread) ` subtree (tRAG s) x))" + by (auto simp:cp_gen_def) + +lemma tRAG_nodeE: + assumes "(n1, n2) \ tRAG s" + obtains th1 th2 where "n1 = Th th1" "n2 = Th th2" + using assms + by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def) + +lemma subtree_nodeE: + assumes "n \ subtree (tRAG s) (Th th)" + obtains th1 where "n = Th th1" +proof - + show ?thesis + proof(rule subtreeE[OF assms]) + assume "n = Th th" + from that[OF this] show ?thesis . + next + assume "Th th \ ancestors (tRAG s) n" + hence "(n, Th th) \ (tRAG s)^+" by (auto simp:ancestors_def) + hence "\ th1. n = Th th1" + proof(induct) + case (base y) + from tRAG_nodeE[OF this] show ?case by metis + next + case (step y z) + thus ?case by auto + qed + with that show ?thesis by auto + qed +qed + +lemma threads_set_eq: + "the_thread ` (subtree (tRAG s) (Th th)) = + {th'. Th th' \ (subtree (RAG s) (Th th))}" (is "?L = ?R") +proof - + { fix th' + assume "th' \ ?L" + then obtain n where h: "th' = the_thread n" "n \ subtree (tRAG s) (Th th)" by auto + from subtree_nodeE[OF this(2)] + obtain th1 where "n = Th th1" by auto + with h have "Th th' \ subtree (tRAG s) (Th th)" by auto + hence "Th th' \ subtree (RAG s) (Th th)" + proof(cases rule:subtreeE[consumes 1]) + case 1 + thus ?thesis by (auto simp:subtree_def) + next + case 2 + hence "(Th th', Th th) \ (tRAG s)^+" by (auto simp:ancestors_def) + thus ?thesis + proof(induct) + case (step y z) + from this(2)[unfolded tRAG_alt_def] + obtain u where + h: "(y, u) \ RAG s" "(u, z) \ RAG s" + by auto + hence "y \ subtree (RAG s) z" by (auto simp:subtree_def) + with step(3) + show ?case by (auto simp:subtree_def) + next + case (base y) + from this[unfolded tRAG_alt_def] + show ?case by (auto simp:subtree_def) + qed + qed + hence "th' \ ?R" by auto + } moreover { + fix th' + assume "th' \ ?R" + hence "(Th th', Th th) \ (RAG s)^*" by (auto simp:subtree_def) + from star_rpath[OF this] + obtain xs where "rpath (RAG s) (Th th') xs (Th th)" by auto + hence "Th th' \ subtree (tRAG s) (Th th)" + proof(induct xs arbitrary:th' th rule:length_induct) + case (1 xs th' th) + show ?case + proof(cases xs) + case Nil + from rpath_nilE[OF 1(2)[unfolded this]] + have "th' = th" by auto + thus ?thesis by (auto simp:subtree_def) + next + case (Cons x1 xs1) note Cons1 = Cons + show ?thesis + proof(cases "xs1") + case Nil + from 1(2)[unfolded Cons[unfolded this]] + have rp: "rpath (RAG s) (Th th') [x1] (Th th)" . + hence "(Th th', x1) \ (RAG s)" by (cases, simp) + then obtain cs where "x1 = Cs cs" + by (unfold s_RAG_def, auto) + find_theorems rpath "_ = _@[_]" + from rpath_nnl_lastE[OF rp[unfolded this]] + show ?thesis by auto + next + case (Cons x2 xs2) + from 1(2)[unfolded Cons1[unfolded this]] + have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" . + from rpath_edges_on[OF this] + have eds: "edges_on (Th th' # x1 # x2 # xs2) \ RAG s" . + have "(Th th', x1) \ edges_on (Th th' # x1 # x2 # xs2)" + by (simp add: edges_on_unfold) + with eds have rg1: "(Th th', x1) \ RAG s" by auto + then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto) + have "(x1, x2) \ edges_on (Th th' # x1 # x2 # xs2)" + by (simp add: edges_on_unfold) + from this eds + have rg2: "(x1, x2) \ RAG s" by auto + from this[unfolded eq_x1] + obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto) + from rp have "rpath (RAG s) x2 xs2 (Th th)" + by (elim rpath_ConsE, simp) + from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" . + from 1(1)[rule_format, OF _ this, unfolded Cons1 Cons] + have "Th th1 \ subtree (tRAG s) (Th th)" by simp + moreover have "(Th th', Th th1) \ (tRAG s)^*" + proof - + from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2] + show ?thesis by (unfold RAG_split tRAG_def wRAG_def hRAG_def, auto) + qed + ultimately show ?thesis by (auto simp:subtree_def) + qed + qed + qed + from imageI[OF this, of the_thread] + have "th' \ ?L" by simp + } ultimately show ?thesis by auto +qed + +lemma cp_alt_def1: + "cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))" +proof - + have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) = + ((the_preced s \ the_thread) ` subtree (tRAG s) (Th th))" + by auto + thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto) +qed + +lemma cp_gen_def_cond: + assumes "x = Th th" + shows "cp s th = cp_gen s (Th th)" +by (unfold cp_alt_def1 cp_gen_def, simp) + +lemma cp_gen_over_set: + assumes "\ x \ A. \ th. x = Th th" + shows "cp_gen s ` A = (cp s \ the_thread) ` A" +proof(rule f_image_eq) + fix a + assume "a \ A" + from assms[rule_format, OF this] + obtain th where eq_a: "a = Th th" by auto + show "cp_gen s a = (cp s \ the_thread) a" + by (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp) +qed + + + +locale valid_trace = + fixes s + assumes vt : "vt s" + +context valid_trace +begin + +lemma wf_RAG: "wf (RAG s)" +proof(rule finite_acyclic_wf) + from finite_RAG[OF vt] show "finite (RAG s)" . +next + from acyclic_RAG[OF vt] show "acyclic (RAG s)" . +qed + +end + +context valid_trace +begin + +lemma sgv_wRAG: "single_valued (wRAG s)" + using waiting_unique[OF vt] + by (unfold single_valued_def wRAG_def, auto) + +lemma sgv_hRAG: "single_valued (hRAG s)" + using holding_unique + by (unfold single_valued_def hRAG_def, auto) + +lemma sgv_tRAG: "single_valued (tRAG s)" + by (unfold tRAG_def, rule single_valued_relcomp, + insert sgv_wRAG sgv_hRAG, auto) + +lemma acyclic_tRAG: "acyclic (tRAG s)" +proof(unfold tRAG_def, rule acyclic_compose) + show "acyclic (RAG s)" using acyclic_RAG[OF vt] . +next + show "wRAG s \ RAG s" unfolding RAG_split by auto +next + show "hRAG s \ RAG s" unfolding RAG_split by auto +qed + +lemma sgv_RAG: "single_valued (RAG s)" + using unique_RAG[OF vt] by (auto simp:single_valued_def) + +lemma rtree_RAG: "rtree (RAG s)" + using sgv_RAG acyclic_RAG[OF vt] + by (unfold rtree_def rtree_axioms_def sgv_def, auto) + +end + +sublocale valid_trace < rtree_s: rtree "tRAG s" +proof(unfold_locales) + from sgv_tRAG show "single_valued (tRAG s)" . +next + from acyclic_tRAG show "acyclic (tRAG s)" . +qed + +sublocale valid_trace < fsbtRAGs : fsubtree "RAG s" +proof - + show "fsubtree (RAG s)" + proof(intro_locales) + show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG[OF vt]] . + next + show "fsubtree_axioms (RAG s)" + proof(unfold fsubtree_axioms_def) + find_theorems wf RAG + from wf_RAG show "wf (RAG s)" . + qed + qed +qed + +sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s" +proof - + have "fsubtree (tRAG s)" + proof - + have "fbranch (tRAG s)" + proof(unfold tRAG_def, rule fbranch_compose) + show "fbranch (wRAG s)" + proof(rule finite_fbranchI) + from finite_RAG[OF vt] show "finite (wRAG s)" + by (unfold RAG_split, auto) + qed + next + show "fbranch (hRAG s)" + proof(rule finite_fbranchI) + from finite_RAG[OF vt] + show "finite (hRAG s)" by (unfold RAG_split, auto) + qed + qed + moreover have "wf (tRAG s)" + proof(rule wf_subset) + show "wf (RAG s O RAG s)" using wf_RAG + by (fold wf_comp_self, simp) + next + show "tRAG s \ (RAG s O RAG s)" + by (unfold tRAG_alt_def, auto) + qed + ultimately show ?thesis + by (unfold fsubtree_def fsubtree_axioms_def,auto) + qed + from this[folded tRAG_def] show "fsubtree (tRAG s)" . +qed + +lemma Max_UNION: + assumes "finite A" + and "A \ {}" + and "\ M \ f ` A. finite M" + and "\ M \ f ` A. M \ {}" + shows "Max (\x\ A. f x) = Max (Max ` f ` A)" (is "?L = ?R") + using assms[simp] +proof - + have "?L = Max (\(f ` A))" + by (fold Union_image_eq, simp) + also have "... = ?R" + by (subst Max_Union, simp+) + finally show ?thesis . +qed + +lemma max_Max_eq: + assumes "finite A" + and "A \ {}" + and "x = y" + shows "max x (Max A) = Max ({y} \ A)" (is "?L = ?R") +proof - + have "?R = Max (insert y A)" by simp + also from assms have "... = ?L" + by (subst Max.insert, simp+) + finally show ?thesis by simp +qed + + +context valid_trace +begin + +(* ddd *) +lemma cp_gen_rec: + assumes "x = Th th" + shows "cp_gen s x = Max ({the_preced s th} \ (cp_gen s) ` children (tRAG s) x)" +proof(cases "children (tRAG s) x = {}") + case True + show ?thesis + by (unfold True cp_gen_def subtree_children, simp add:assms) +next + case False + hence [simp]: "children (tRAG s) x \ {}" by auto + note fsbttRAGs.finite_subtree[simp] + have [simp]: "finite (children (tRAG s) x)" + by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], + rule children_subtree) + { fix r x + have "subtree r x \ {}" by (auto simp:subtree_def) + } note this[simp] + have [simp]: "\x\children (tRAG s) x. subtree (tRAG s) x \ {}" + proof - + from False obtain q where "q \ children (tRAG s) x" by blast + moreover have "subtree (tRAG s) q \ {}" by simp + ultimately show ?thesis by blast + qed + have h: "Max ((the_preced s \ the_thread) ` + ({x} \ \(subtree (tRAG s) ` children (tRAG s) x))) = + Max ({the_preced s th} \ cp_gen s ` children (tRAG s) x)" + (is "?L = ?R") + proof - + let "Max (?f ` (?A \ \ (?g ` ?B)))" = ?L + let "Max (_ \ (?h ` ?B))" = ?R + let ?L1 = "?f ` \(?g ` ?B)" + have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)" + proof - + have "?L1 = ?f ` (\ x \ ?B.(?g x))" by simp + also have "... = (\ x \ ?B. ?f ` (?g x))" by auto + finally have "Max ?L1 = Max ..." by simp + also have "... = Max (Max ` (\x. ?f ` subtree (tRAG s) x) ` ?B)" + by (subst Max_UNION, simp+) + also have "... = Max (cp_gen s ` children (tRAG s) x)" + by (unfold image_comp cp_gen_alt_def, simp) + finally show ?thesis . + qed + show ?thesis + proof - + have "?L = Max (?f ` ?A \ ?L1)" by simp + also have "... = max (the_preced s (the_thread x)) (Max ?L1)" + by (subst Max_Un, simp+) + also have "... = max (?f x) (Max (?h ` ?B))" + by (unfold eq_Max_L1, simp) + also have "... =?R" + by (rule max_Max_eq, (simp)+, unfold assms, simp) + finally show ?thesis . + qed + qed thus ?thesis + by (fold h subtree_children, unfold cp_gen_def, simp) +qed + +lemma cp_rec: + "cp s th = Max ({the_preced s th} \ + (cp s o the_thread) ` children (tRAG s) (Th th))" +proof - + have "Th th = Th th" by simp + note h = cp_gen_def_cond[OF this] cp_gen_rec[OF this] + show ?thesis + proof - + have "cp_gen s ` children (tRAG s) (Th th) = + (cp s \ the_thread) ` children (tRAG s) (Th th)" + proof(rule cp_gen_over_set) + show " \x\children (tRAG s) (Th th). \th. x = Th th" + by (unfold tRAG_alt_def, auto simp:children_def) + qed + thus ?thesis by (subst (1) h(1), unfold h(2), simp) + qed +qed + +end + + lemma eq_dependants: "dependants (wq s) = dependants s" by (simp add: s_dependants_abv wq_def) - + + (* obvious lemma *) lemma not_thread_holdents: fixes th s @@ -548,7 +1046,7 @@ (* moving in Max *) also have "\ = max (Max {preced th s}) (Max ((\th. Max (preceds s ({th} \ (dependants (wq s) th)))) ` children s th))" - by (subst Max_Union) (auto simp add: image_image) + by (subst Max_Union) (auto simp add: image_image) (* folding cp + moving out Max *) also have "\ = ?RHS" @@ -558,62 +1056,51 @@ qed qed -lemma next_waiting: +lemma next_th_holding: + assumes vt: "vt s" + and nxt: "next_th s th cs th'" + shows "holding (wq s) th cs" +proof - + from nxt[unfolded next_th_def] + obtain rest where h: "wq s cs = th # rest" + "rest \ []" + "th' = hd (SOME q. distinct q \ set q = set rest)" by auto + thus ?thesis + by (unfold cs_holding_def, auto) +qed + +lemma next_th_waiting: assumes vt: "vt s" and nxt: "next_th s th cs th'" - shows "waiting s th' cs" + shows "waiting (wq s) th' cs" proof - - from assms show ?thesis - apply (auto simp:next_th_def s_waiting_def[folded wq_def]) - proof - - fix rest - assume ni: "hd (SOME q. distinct q \ set q = set rest) \ set rest" - and eq_wq: "wq s cs = th # rest" - and ne: "rest \ []" - have "set (SOME q. distinct q \ set q = set rest) = set rest" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - show "\x. distinct x \ set x = set rest \ set x = set rest" by auto - qed - with ni - have "hd (SOME q. distinct q \ set q = set rest) \ set (SOME q. distinct q \ set q = set rest)" - by simp - moreover have "(SOME q. distinct q \ set q = set rest) \ []" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - from ne show "\x. distinct x \ set x = set rest \ x \ []" by auto - qed - ultimately show "hd (SOME q. distinct q \ set q = set rest) = th" by auto + from nxt[unfolded next_th_def] + obtain rest where h: "wq s cs = th # rest" + "rest \ []" + "th' = hd (SOME q. distinct q \ set q = set rest)" by auto + from wq_distinct[OF vt, of cs, unfolded h] + have dst: "distinct (th # rest)" . + have in_rest: "th' \ set rest" + proof(unfold h, rule someI2) + show "distinct rest \ set rest = set rest" using dst by auto next - fix rest - assume eq_wq: "wq s cs = hd (SOME q. distinct q \ set q = set rest) # rest" - and ne: "rest \ []" - have "(SOME q. distinct q \ set q = set rest) \ []" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - from ne show "\x. distinct x \ set x = set rest \ x \ []" by auto - qed - hence "hd (SOME q. distinct q \ set q = set rest) \ set (SOME q. distinct q \ set q = set rest)" - by auto - moreover have "set (SOME q. distinct q \ set q = set rest) = set rest" - proof(rule someI2) - from wq_distinct[OF vt, of cs] eq_wq - show "distinct rest \ set rest = set rest" by auto - next - show "\x. distinct x \ set x = set rest \ set x = set rest" by auto - qed - ultimately have "hd (SOME q. distinct q \ set q = set rest) \ set rest" by simp - with eq_wq and wq_distinct[OF vt, of cs] - show False by auto + fix x assume "distinct x \ set x = set rest" + with h(2) + show "hd x \ set (rest)" by (cases x, auto) qed + hence "th' \ set (wq s cs)" by (unfold h(1), auto) + moreover have "th' \ hd (wq s cs)" + by (unfold h(1), insert in_rest dst, auto) + ultimately show ?thesis by (auto simp:cs_waiting_def) qed +lemma next_th_RAG: + assumes vt: "vt s" + and nxt: "next_th s th cs th'" + shows "{(Cs cs, Th th), (Th th', Cs cs)} \ RAG s" + using assms next_th_holding next_th_waiting +by (unfold s_RAG_def, simp) + -- {* A useless definition *} definition cps:: "state \ (thread \ precedence) set" where "cps s = {(th, cp s th) | th . th \ threads s}" @@ -661,16 +1148,9 @@ context step_set_cps begin -interpretation h: pip "s" - by (unfold pip_def, insert vt_s, simp) - -find_theorems - -(* *) - text {* (* ddd *) - The following lemma confirms that @{text "Set"}-operating only changes the precedence - of initiating thread. + The following two lemmas confirm that @{text "Set"}-operating only changes the precedence + of the initiating thread. *} lemma eq_preced: @@ -682,103 +1162,86 @@ by (unfold s_def, auto simp:preced_def) qed -text {* (* ddd *) +lemma eq_the_preced: + fixes th' + assumes "th' \ th" + shows "the_preced s th' = the_preced s' th'" + using assms + by (unfold the_preced_def, intro eq_preced, simp) + +text {* The following lemma assures that the resetting of priority does not change the RAG. *} lemma eq_dep: "RAG s = RAG s'" by (unfold s_def RAG_set_unchanged, auto) -text {* - Th following lemma @{text "eq_cp_pre"} circumscribe a rough range of recalculation. - It says, thread other than the initiating thread @{text "th"} does not need recalculation - unless it lies upstream of @{text "th"} in the RAG. - - The reason behind this lemma is that: - the change of precedence of one thread can only affect it's upstream threads, according to - lemma @{text "eq_preced"}. Since the only thread which might change precedence is - @{text "th"}, so only @{text "th"} and its upstream threads need recalculation. - (* ccc *) +text {* (* ddd *) + Th following lemma @{text "eq_cp_pre"} says the priority change of @{text "th"} + only affects those threads, which as @{text "Th th"} in their sub-trees. + + The proof of this lemma is simplified by using the alternative definition of @{text "cp"}. *} lemma eq_cp_pre: fixes th' - assumes neq_th: "th' \ th" - and nd: "th \ dependants s th'" + assumes nd: "Th th \ subtree (RAG s') (Th th')" shows "cp s th' = cp s' th'" proof - - -- {* This is what we need to prove after expanding the definition of @{text "cp"} *} - have "Max ((\th. preced th s) ` ({th'} \ dependants (wq s) th')) = - Max ((\th. preced th s') ` ({th'} \ dependants (wq s') th'))" - (is "Max (?f1 ` ({th'} \ ?A)) = Max (?f2 ` ({th'} \ ?B))") + -- {* After unfolding using the alternative definition, elements + affecting the @{term "cp"}-value of threads become explicit. + We only need to prove the following: *} + have "Max (the_preced s ` {th'a. Th th'a \ subtree (RAG s) (Th th')}) = + Max (the_preced s' ` {th'a. Th th'a \ subtree (RAG s') (Th th')})" + (is "Max (?f ` ?S1) = Max (?g ` ?S2)") proof - - -- {* Since RAG is not changed by @{text "Set"}-operation, the dependants of - any threads are not changed, this is one of key facts underpinning this - lemma *} - have eq_ab: "?A = ?B" by (unfold cs_dependants_def, auto simp:eq_dep eq_RAG) - have "(?f1 ` ({th'} \ ?A)) = (?f2 ` ({th'} \ ?B))" - proof(rule image_cong) - show "{th'} \ ?A = {th'} \ ?B" by (simp only:eq_ab) - next - fix x - assume x_in: "x \ {th'} \ ?B" - show "?f1 x = ?f2 x" - proof(rule eq_preced) -- {* The other key fact underpinning this lemma is @{text "eq_preced"} *} - from x_in[folded eq_ab, unfolded eq_dependants] - have "x \ {th'} \ dependants s th'" . - thus "x \ th" - proof - assume "x \ {th'}" - with `th' \ th` show ?thesis by simp - next - assume "x \ dependants s th'" - with `th \ dependants s th'` show ?thesis by auto - qed - qed - qed - thus ?thesis by simp - qed - thus ?thesis by (unfold cp_eq_cpreced cpreced_def) + -- {* The base sets are equal. *} + have "?S1 = ?S2" using eq_dep by simp + -- {* The function values on the base set are equal as well. *} + moreover have "\ e \ ?S2. ?f e = ?g e" + proof + fix th1 + assume "th1 \ ?S2" + with nd have "th1 \ th" by (auto) + from eq_the_preced[OF this] + show "the_preced s th1 = the_preced s' th1" . + qed + -- {* Therefore, the image of the functions are equal. *} + ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq) + thus ?thesis by simp + qed + thus ?thesis by (simp add:cp_alt_def) qed text {* - The following lemma shows that no thread lies upstream of the initiating thread @{text "th"}. - The reason for this is that only no-blocked thread can initiate - a system call. Since thread @{text "th"} is non-blocked, it is not waiting for any - critical resource. Therefore, there is edge leading out of @{text "th"} in the RAG. - Consequently, there is no node (neither resource nor thread) upstream of @{text "th"}. + The following lemma shows that @{term "th"} is not in the + sub-tree of any other thread. *} -lemma no_dependants: - shows "th \ dependants s th'" -proof - assume "th \ dependants s th'" - from `th \ dependants s th'` have "(Th th, Th th') \ (RAG s')\<^sup>+" - by (unfold s_dependants_def, unfold eq_RAG, unfold eq_dep, auto) - from tranclD[OF this] - obtain z where "(Th th, z) \ RAG s'" by auto - moreover have "th \ readys s'" +lemma th_in_no_subtree: + assumes "th' \ th" + shows "Th th \ subtree (RAG s') (Th th')" +proof - + have "th \ readys s'" proof - from step_back_step [OF vt_s[unfolded s_def]] have "step s' (Set th prio)" . hence "th \ runing s'" by (cases, simp) thus ?thesis by (simp add:readys_def runing_def) qed - ultimately show "False" - apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def) - by (fold wq_def, blast) + from readys_in_no_subtree[OF step_back_vt[OF vt_s[unfolded s_def]] this assms(1)] + show ?thesis by blast qed -(* Result improved *) text {* - A simple combination of @{text "eq_cp_pre"} and @{text "no_dependants"} - gives the main lemma of this locale, which shows that - only the initiating thread needs a recalculation of current precedence. + By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, + it is obvious that the change of priority only affects the @{text "cp"}-value + of the initiating thread @{text "th"}. *} lemma eq_cp: fixes th' assumes "th' \ th" shows "cp s th' = cp s' th'" - by (rule eq_cp_pre[OF assms no_dependants]) + by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) end @@ -794,6 +1257,83 @@ -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *} assumes vt_s: "vt s" +context step_v_cps +begin + +lemma rtree_RAGs: "rtree (RAG s)" +proof + show "single_valued (RAG s)" + apply (intro_locales) + by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s]) + + show "acyclic (RAG s)" + by (rule acyclic_RAG[OF vt_s]) +qed + +lemma rtree_RAGs': "rtree (RAG s')" +proof + show "single_valued (RAG s')" + apply (intro_locales) + by (unfold single_valued_def, + auto intro:unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) + + show "acyclic (RAG s')" + by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) +qed + +lemmas vt_s' = step_back_vt[OF vt_s[unfolded s_def]] + +lemma ready_th_s': "th \ readys s'" + using step_back_step[OF vt_s[unfolded s_def]] + by (cases, simp add:runing_def) + + +lemma ancestors_th: "ancestors (RAG s') (Th th) = {}" +proof - + from readys_root[OF vt_s' ready_th_s'] + show ?thesis + by (unfold root_def, simp) +qed + +lemma holding_th: "holding s' th cs" +proof - + from vt_s[unfolded s_def] + have " PIP s' (V th cs)" by (cases, simp) + thus ?thesis by (cases, auto) +qed + +lemma edge_of_th: + "(Cs cs, Th th) \ RAG s'" +proof - + from holding_th + show ?thesis + by (unfold s_RAG_def holding_eq, auto) +qed + +lemma ancestors_cs: + "ancestors (RAG s') (Cs cs) = {Th th}" +proof - + find_theorems ancestors + have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th) \ {Th th}" + proof(rule RTree.rtree.ancestors_accum[OF rtree_RAGs']) + from vt_s[unfolded s_def] + have " PIP s' (V th cs)" by (cases, simp) + thus "(Cs cs, Th th) \ RAG s'" + proof(cases) + assume "holding s' th cs" + from this[unfolded holding_eq] + show ?thesis by (unfold s_RAG_def, auto) + qed + qed + from this[unfolded ancestors_th] show ?thesis by simp +qed + +lemma preced_kept: "the_preced s = the_preced s'" + by (auto simp: s_def the_preced_def preced_def) + +end + + text {* The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, which represents the case when there is another thread @{text "th'"} @@ -847,6 +1387,19 @@ th7 ----| *) +lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \ RAG s'" + using next_th_RAG[OF vt_s' nt] . + +lemma ancestors_th': + "ancestors (RAG s') (Th th') = {Th th, Cs cs}" +proof - + have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \ {Cs cs}" + proof(rule RTree.rtree.ancestors_accum[OF rtree_RAGs']) + from sub_RAGs' show "(Th th', Cs cs) \ RAG s'" by auto + qed + thus ?thesis using ancestors_th ancestors_cs by auto +qed + lemma RAG_s: "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \ {(Cs cs, Th th')}" @@ -855,199 +1408,54 @@ and nt show ?thesis by (auto intro:next_th_unique) qed -text {* - Lemma @{text "dependants_kept"} shows only @{text "th"} and @{text "th'"} - have their dependants changed. -*} -lemma dependants_kept: - fixes th'' - assumes neq1: "th'' \ th" - and neq2: "th'' \ th'" - shows "dependants (wq s) th'' = dependants (wq s') th''" -proof(auto) (* ccc *) - fix x - assume "x \ dependants (wq s) th''" - hence dp: "(Th x, Th th'') \ (RAG s)^+" - by (auto simp:cs_dependants_def eq_RAG) - { fix n - have "(n, Th th'') \ (RAG s)^+ \ (n, Th th'') \ (RAG s')^+" - proof(induct rule:converse_trancl_induct) - fix y - assume "(y, Th th'') \ RAG s" - with RAG_s neq1 neq2 - have "(y, Th th'') \ RAG s'" by auto - thus "(y, Th th'') \ (RAG s')\<^sup>+" by auto - next - fix y z - assume yz: "(y, z) \ RAG s" - and ztp: "(z, Th th'') \ (RAG s)\<^sup>+" - and ztp': "(z, Th th'') \ (RAG s')\<^sup>+" - have "y \ Cs cs \ y \ Th th'" - proof - show "y \ Cs cs" - proof - assume eq_y: "y = Cs cs" - with yz have dp_yz: "(Cs cs, z) \ RAG s" by simp - from RAG_s - have cst': "(Cs cs, Th th') \ RAG s" by simp - from unique_RAG[OF vt_s this dp_yz] - have eq_z: "z = Th th'" by simp - with ztp have "(Th th', Th th'') \ (RAG s)^+" by simp - from converse_tranclE[OF this] - obtain cs' where dp'': "(Th th', Cs cs') \ RAG s" - by (auto simp:s_RAG_def) - with RAG_s have dp': "(Th th', Cs cs') \ RAG s'" by auto - from dp'' eq_y yz eq_z have "(Cs cs, Cs cs') \ (RAG s)^+" by auto - moreover have "cs' = cs" - proof - - from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt] - have "(Th th', Cs cs) \ RAG s'" - by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this dp'] - show ?thesis by simp - qed - ultimately have "(Cs cs, Cs cs) \ (RAG s)^+" by simp - moreover note wf_trancl[OF wf_RAG[OF vt_s]] - ultimately show False by auto - qed +lemma subtree_kept: + assumes "th1 \ {th, th'}" + shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R") +proof - + let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})" + let ?RAG'' = "?RAG' \ {(Cs cs, Th th')}" + have "subtree ?RAG' (Th th1) = ?R" + proof(rule subset_del_subtree_outside) + show "Range {(Cs cs, Th th), (Th th', Cs cs)} \ subtree (RAG s') (Th th1) = {}" + proof - + have "(Th th) \ subtree (RAG s') (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s') (Th th)" + by (unfold ancestors_th, simp) next - show "y \ Th th'" - proof - assume eq_y: "y = Th th'" - with yz have dps: "(Th th', z) \ RAG s" by simp - with RAG_s have dps': "(Th th', z) \ RAG s'" by auto - have "z = Cs cs" - proof - - from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt] - have "(Th th', Cs cs) \ RAG s'" - by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] dps' this] - show ?thesis . - qed - with dps RAG_s show False by auto - qed + from assms show "Th th1 \ Th th" by simp qed - with RAG_s yz have "(y, z) \ RAG s'" by auto - with ztp' - show "(y, Th th'') \ (RAG s')\<^sup>+" by auto - qed - } - from this[OF dp] - show "x \ dependants (wq s') th''" - by (auto simp:cs_dependants_def eq_RAG) -next - fix x - assume "x \ dependants (wq s') th''" - hence dp: "(Th x, Th th'') \ (RAG s')^+" - by (auto simp:cs_dependants_def eq_RAG) - { fix n - have "(n, Th th'') \ (RAG s')^+ \ (n, Th th'') \ (RAG s)^+" - proof(induct rule:converse_trancl_induct) - fix y - assume "(y, Th th'') \ RAG s'" - with RAG_s neq1 neq2 - have "(y, Th th'') \ RAG s" by auto - thus "(y, Th th'') \ (RAG s)\<^sup>+" by auto + moreover have "(Cs cs) \ subtree (RAG s') (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s') (Cs cs)" + by (unfold ancestors_cs, insert assms, auto) + qed simp + ultimately have "{Th th, Cs cs} \ subtree (RAG s') (Th th1) = {}" by auto + thus ?thesis by simp + qed + qed + moreover have "subtree ?RAG'' (Th th1) = subtree ?RAG' (Th th1)" + proof(rule subtree_insert_next) + show "Th th' \ subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')" + (is "_ \ ?R") + proof - + have "?R \ ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto) + moreover have "Th th1 \ ..." using ancestors_th' assms by simp + ultimately show ?thesis by auto + qed next - fix y z - assume yz: "(y, z) \ RAG s'" - and ztp: "(z, Th th'') \ (RAG s')\<^sup>+" - and ztp': "(z, Th th'') \ (RAG s)\<^sup>+" - have "y \ Cs cs \ y \ Th th'" - proof - show "y \ Cs cs" - proof - assume eq_y: "y = Cs cs" - with yz have dp_yz: "(Cs cs, z) \ RAG s'" by simp - from this have eq_z: "z = Th th" - proof - - from step_back_step[OF vt_s[unfolded s_def]] - have "(Cs cs, Th th) \ RAG s'" - by(cases, auto simp: wq_def s_RAG_def cs_holding_def s_holding_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this dp_yz] - show ?thesis by simp - qed - from converse_tranclE[OF ztp] - obtain u where "(z, u) \ RAG s'" by auto - moreover - from step_back_step[OF vt_s[unfolded s_def]] - have "th \ readys s'" by (cases, simp add:runing_def) - moreover note eq_z - ultimately show False - by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def) - qed - next - show "y \ Th th'" - proof - assume eq_y: "y = Th th'" - with yz have dps: "(Th th', z) \ RAG s'" by simp - have "z = Cs cs" - proof - - from next_waiting[OF step_back_vt[OF vt_s[unfolded s_def]] nt] - have "(Th th', Cs cs) \ RAG s'" - by (auto simp:s_waiting_def wq_def s_RAG_def cs_waiting_def) - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] dps this] - show ?thesis . - qed - with ztp have cs_i: "(Cs cs, Th th'') \ (RAG s')\<^sup>+" by simp - from step_back_step[OF vt_s[unfolded s_def]] - have cs_th: "(Cs cs, Th th) \ RAG s'" - by(cases, auto simp: s_RAG_def wq_def cs_holding_def s_holding_def) - have "(Cs cs, Th th'') \ RAG s'" - proof - assume "(Cs cs, Th th'') \ RAG s'" - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] this cs_th] - and neq1 show "False" by simp - qed - with converse_tranclE[OF cs_i] - obtain u where cu: "(Cs cs, u) \ RAG s'" - and u_t: "(u, Th th'') \ (RAG s')\<^sup>+" by auto - have "u = Th th" - proof - - from unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]] cu cs_th] - show ?thesis . - qed - with u_t have "(Th th, Th th'') \ (RAG s')\<^sup>+" by simp - from converse_tranclE[OF this] - obtain v where "(Th th, v) \ (RAG s')" by auto - moreover from step_back_step[OF vt_s[unfolded s_def]] - have "th \ readys s'" by (cases, simp add:runing_def) - ultimately show False - by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def) - qed - qed - with RAG_s yz have "(y, z) \ RAG s" by auto - with ztp' - show "(y, Th th'') \ (RAG s)\<^sup>+" by auto - qed - } - from this[OF dp] - show "x \ dependants (wq s) th''" - by (auto simp:cs_dependants_def eq_RAG) + from assms show "Th th1 \ Th th'" by simp + qed + qed + ultimately show ?thesis by (unfold RAG_s, simp) qed lemma cp_kept: - fixes th'' - assumes neq1: "th'' \ th" - and neq2: "th'' \ th'" - shows "cp s th'' = cp s' th''" -proof - - from dependants_kept[OF neq1 neq2] - have "dependants (wq s) th'' = dependants (wq s') th''" . - moreover { - fix th1 - assume "th1 \ dependants (wq s) th''" - have "preced th1 s = preced th1 s'" - by (unfold s_def, auto simp:preced_def) - } - moreover have "preced th'' s = preced th'' s'" - by (unfold s_def, auto simp:preced_def) - ultimately have "((\th. preced th s) ` ({th''} \ dependants (wq s) th'')) = - ((\th. preced th s') ` ({th''} \ dependants (wq s') th''))" - by (auto simp:image_def) - thus ?thesis - by (unfold cp_eq_cpreced cpreced_def, simp) -qed + assumes "th1 \ {th, th'}" + shows "cp s th1 = cp s' th1" + by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) end @@ -1057,152 +1465,143 @@ context step_v_cps_nnt begin -lemma nw_cs: "(Th th1, Cs cs) \ RAG s'" -proof - assume "(Th th1, Cs cs) \ RAG s'" - thus "False" - apply (auto simp:s_RAG_def cs_waiting_def) - proof - - assume h1: "th1 \ set (wq s' cs)" - and h2: "th1 \ hd (wq s' cs)" - from step_back_step[OF vt_s[unfolded s_def]] - show "False" - proof(cases) - assume "holding s' th cs" - then obtain rest where - eq_wq: "wq s' cs = th#rest" - apply (unfold s_holding_def wq_def[symmetric]) - by (case_tac "(wq s' cs)", auto) - with h1 h2 have ne: "rest \ []" by auto - with eq_wq - have "next_th s' th cs (hd (SOME q. distinct q \ set q = set rest))" - by(unfold next_th_def, rule_tac x = "rest" in exI, auto) - with nnt show ?thesis by auto - qed - qed -qed - lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}" proof - from nnt and step_RAG_v[OF vt_s[unfolded s_def], folded s_def] show ?thesis by auto qed -lemma child_kept_left: - assumes - "(n1, n2) \ (child s')^+" - shows "(n1, n2) \ (child s)^+" -proof - - from assms show ?thesis - proof(induct rule: converse_trancl_induct) - case (base y) - from base obtain th1 cs1 th2 - where h1: "(Th th1, Cs cs1) \ RAG s'" - and h2: "(Cs cs1, Th th2) \ RAG s'" - and eq_y: "y = Th th1" and eq_n2: "n2 = Th th2" by (auto simp:child_def) - have "cs1 \ cs" - proof - assume eq_cs: "cs1 = cs" - with h1 have "(Th th1, Cs cs1) \ RAG s'" by simp - with nw_cs eq_cs show False by auto +lemma subtree_kept: + assumes "th1 \ th" + shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" +proof(unfold RAG_s, rule subset_del_subtree_outside) + show "Range {(Cs cs, Th th)} \ subtree (RAG s') (Th th1) = {}" + proof - + have "(Th th) \ subtree (RAG s') (Th th1)" + proof(rule subtree_refute) + show "Th th1 \ ancestors (RAG s') (Th th)" + by (unfold ancestors_th, simp) + next + from assms show "Th th1 \ Th th" by simp qed - with h1 h2 RAG_s have - h1': "(Th th1, Cs cs1) \ RAG s" and - h2': "(Cs cs1, Th th2) \ RAG s" by auto - hence "(Th th1, Th th2) \ child s" by (auto simp:child_def) - with eq_y eq_n2 have "(y, n2) \ child s" by simp - thus ?case by auto - next - case (step y z) - have "(y, z) \ child s'" by fact - then obtain th1 cs1 th2 - where h1: "(Th th1, Cs cs1) \ RAG s'" - and h2: "(Cs cs1, Th th2) \ RAG s'" - and eq_y: "y = Th th1" and eq_z: "z = Th th2" by (auto simp:child_def) - have "cs1 \ cs" - proof - assume eq_cs: "cs1 = cs" - with h1 have "(Th th1, Cs cs1) \ RAG s'" by simp - with nw_cs eq_cs show False by auto - qed - with h1 h2 RAG_s have - h1': "(Th th1, Cs cs1) \ RAG s" and - h2': "(Cs cs1, Th th2) \ RAG s" by auto - hence "(Th th1, Th th2) \ child s" by (auto simp:child_def) - with eq_y eq_z have "(y, z) \ child s" by simp - moreover have "(z, n2) \ (child s)^+" by fact - ultimately show ?case by auto + thus ?thesis by auto qed qed -lemma child_kept_right: - assumes - "(n1, n2) \ (child s)^+" - shows "(n1, n2) \ (child s')^+" +lemma cp_kept_1: + assumes "th1 \ th" + shows "cp s th1 = cp s' th1" + by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) + +lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}" proof - - from assms show ?thesis - proof(induct) - case (base y) - from base and RAG_s - have "(n1, y) \ child s'" - by (auto simp:child_def) - thus ?case by auto - next - case (step y z) - have "(y, z) \ child s" by fact - with RAG_s have "(y, z) \ child s'" - by (auto simp:child_def) - moreover have "(n1, y) \ (child s')\<^sup>+" by fact - ultimately show ?case by auto - qed + { fix n + have "(Cs cs) \ ancestors (RAG s') n" + proof + assume "Cs cs \ ancestors (RAG s') n" + hence "(n, Cs cs) \ (RAG s')^+" by (auto simp:ancestors_def) + from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \ RAG s'" by auto + then obtain th' where "nn = Th th'" + by (unfold s_RAG_def, auto) + from h[unfolded this] have "(Th th', Cs cs) \ RAG s'" . + from this[unfolded s_RAG_def] + have "waiting (wq s') th' cs" by auto + from this[unfolded cs_waiting_def] + have "1 < length (wq s' cs)" + by (cases "wq s' cs", auto) + from holding_next_thI[OF holding_th this] + obtain th' where "next_th s' th cs th'" by auto + with nnt show False by auto + qed + } note h = this + { fix n + assume "n \ subtree (RAG s') (Cs cs)" + hence "n = (Cs cs)" + by (elim subtreeE, insert h, auto) + } moreover have "(Cs cs) \ subtree (RAG s') (Cs cs)" + by (auto simp:subtree_def) + ultimately show ?thesis by auto qed -lemma eq_child: "(child s)^+ = (child s')^+" - by (insert child_kept_left child_kept_right, auto) +lemma subtree_th: + "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}" +proof(unfold RAG_s, fold subtree_cs, rule RTree.rtree.subtree_del_inside[OF rtree_RAGs']) + from edge_of_th + show "(Cs cs, Th th) \ edges_in (RAG s') (Th th)" + by (unfold edges_in_def, auto simp:subtree_def) +qed + +lemma cp_kept_2: + shows "cp s th = cp s' th" + by (unfold cp_alt_def subtree_th preced_kept, auto) lemma eq_cp: fixes th' shows "cp s th' = cp s' th'" - apply (unfold cp_eq_cpreced cpreced_def) -proof - - have eq_dp: "\ th. dependants (wq s) th = dependants (wq s') th" - apply (unfold cs_dependants_def, unfold eq_RAG) - proof - - from eq_child - have "\th. {th'. (Th th', Th th) \ (child s)\<^sup>+} = {th'. (Th th', Th th) \ (child s')\<^sup>+}" - by simp - with child_RAG_eq[OF vt_s] child_RAG_eq[OF step_back_vt[OF vt_s[unfolded s_def]]] - show "\th. {th'. (Th th', Th th) \ (RAG s)\<^sup>+} = {th'. (Th th', Th th) \ (RAG s')\<^sup>+}" - by simp - qed - moreover { - fix th1 - assume "th1 \ {th'} \ dependants (wq s') th'" - hence "th1 = th' \ th1 \ dependants (wq s') th'" by auto - hence "preced th1 s = preced th1 s'" - proof - assume "th1 = th'" - show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def) - next - assume "th1 \ dependants (wq s') th'" - show "preced th1 s = preced th1 s'" by (simp add:s_def preced_def) - qed - } ultimately have "((\th. preced th s) ` ({th'} \ dependants (wq s) th')) = - ((\th. preced th s') ` ({th'} \ dependants (wq s') th'))" - by (auto simp:image_def) - thus "Max ((\th. preced th s) ` ({th'} \ dependants (wq s) th')) = - Max ((\th. preced th s') ` ({th'} \ dependants (wq s') th'))" by simp -qed + using cp_kept_1 cp_kept_2 + by (cases "th' = th", auto) + +end + +find_theorems "_`_" "\ _" -end +find_theorems "Max" "\ _" + +find_theorems wf RAG + +thm wf_def + +thm image_Union locale step_P_cps = fixes s' th cs s defines s_def : "s \ (P th cs#s')" assumes vt_s: "vt s" +sublocale step_P_cps < vat_s : valid_trace "s" +proof + from vt_s show "vt s" . +qed + +sublocale step_P_cps < vat_s' : valid_trace "s'" +proof + from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . +qed + + +context step_P_cps +begin + +lemma rtree_RAGs: "rtree (RAG s)" +proof + show "single_valued (RAG s)" + apply (intro_locales) + by (unfold single_valued_def, auto intro: unique_RAG[OF vt_s]) + + show "acyclic (RAG s)" + by (rule acyclic_RAG[OF vt_s]) +qed + +lemma rtree_RAGs': "rtree (RAG s')" +proof + show "single_valued (RAG s')" + apply (intro_locales) + by (unfold single_valued_def, + auto intro:unique_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) + + show "acyclic (RAG s')" + by (rule acyclic_RAG[OF step_back_vt[OF vt_s[unfolded s_def]]]) +qed + +lemma preced_kept: "the_preced s = the_preced s'" + by (auto simp: s_def the_preced_def preced_def) + +end + locale step_P_cps_ne =step_P_cps + + fixes th' assumes ne: "wq s' cs \ []" + defines th'_def: "th' \ hd (wq s' cs)" locale step_P_cps_e =step_P_cps + assumes ee: "wq s' cs = []" @@ -1340,7 +1739,20 @@ end -context step_P_cps_ne +lemma tRAG_ancestorsE: + assumes "x \ ancestors (tRAG s) u" + obtains th where "x = Th th" +proof - + from assms have "(u, x) \ (tRAG s)^+" + by (unfold ancestors_def, auto) + from tranclE[OF this] obtain c where "(c, x) \ tRAG s" by auto + then obtain th where "x = Th th" + by (unfold tRAG_alt_def, auto) + from that[OF this] show ?thesis . +qed + + +context step_P_cps_ne begin lemma RAG_s: "RAG s = RAG s' \ {(Th th, Cs cs)}" @@ -1349,6 +1761,178 @@ show ?thesis by (simp add:s_def) qed +lemma cs_held: "(Cs cs, Th th') \ RAG s'" +proof - + have "(Cs cs, Th th') \ hRAG s'" + proof - + from ne + have " holding s' th' cs" + by (unfold th'_def holding_eq cs_holding_def, auto) + thus ?thesis + by (unfold hRAG_def, auto) + qed + thus ?thesis by (unfold RAG_split, auto) +qed + +lemma tRAG_s: + "tRAG s = tRAG s' \ {(Th th, Th th')}" + using RAG_tRAG_transfer[OF step_back_vt[OF vt_s[unfolded s_def]] RAG_s cs_held] . + +lemma cp_kept: + assumes "Th th'' \ ancestors (tRAG s) (Th th)" + shows "cp s th'' = cp s' th''" +proof - + have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')" + proof - + have "Th th' \ subtree (tRAG s') (Th th'')" + proof + assume "Th th' \ subtree (tRAG s') (Th th'')" + thus False + proof(rule subtreeE) + assume "Th th' = Th th''" + from assms[unfolded tRAG_s ancestors_def, folded this] + show ?thesis by auto + next + assume "Th th'' \ ancestors (tRAG s') (Th th')" + moreover have "... \ ancestors (tRAG s) (Th th')" + proof(rule ancestors_mono) + show "tRAG s' \ tRAG s" by (unfold tRAG_s, auto) + qed + ultimately have "Th th'' \ ancestors (tRAG s) (Th th')" by auto + moreover have "Th th' \ ancestors (tRAG s) (Th th)" + by (unfold tRAG_s, auto simp:ancestors_def) + ultimately have "Th th'' \ ancestors (tRAG s) (Th th)" + by (auto simp:ancestors_def) + with assms show ?thesis by auto + qed + qed + from subtree_insert_next[OF this] + have "subtree (tRAG s' \ {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" . + from this[folded tRAG_s] show ?thesis . + qed + show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) +qed + +lemma set_prop_split: + "A = {x. x \ A \ PP x} \ {x. x \ A \ \ PP x}" + by auto + +lemma f_image_union_eq: + assumes "f ` A = g ` A" + and "f ` B = g ` B" + shows "f ` (A \ B) = g ` (A \ B)" + using assms by auto + +(* ccc *) + +lemma cp_gen_update_stop: + assumes "u \ ancestors (tRAG s) (Th th)" + and "cp_gen s u = cp_gen s' u" + and "y \ ancestors (tRAG s) u" + shows "cp_gen s y = cp_gen s' y" + using assms(3) +proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf]) + case (1 x) + show ?case (is "?L = ?R") + proof - + from tRAG_ancestorsE[OF 1(2)] + obtain th2 where eq_x: "x = Th th2" by blast + from vat_s.cp_gen_rec[OF this] + have "?L = + Max ({the_preced s th2} \ cp_gen s ` RTree.children (tRAG s) x)" . + also have "... = + Max ({the_preced s' th2} \ cp_gen s' ` RTree.children (tRAG s') x)" + proof - + have "the_preced s th2 = the_preced s' th2" sorry + moreover have "cp_gen s ` RTree.children (tRAG s) x = + cp_gen s' ` RTree.children (tRAG s') x" + proof - + have "RTree.children (tRAG s) x = RTree.children (tRAG s') x" + proof(unfold tRAG_s, rule children_union_kept) + have start: "(Th th, Th th') \ tRAG s" + by (unfold tRAG_s, auto) + note x_u = 1(2) + show "x \ Range {(Th th, Th th')}" + proof + assume "x \ Range {(Th th, Th th')}" + hence eq_x: "x = Th th'" using RangeE by auto + show False + proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start]) + case 1 + from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG + show ?thesis by (auto simp:ancestors_def acyclic_def) + next + case 2 + with x_u[unfolded eq_x] + have "(Th th', Th th') \ (tRAG s)^+" by (auto simp:ancestors_def) + with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) + qed + qed + qed + moreover have "cp_gen s ` RTree.children (tRAG s) x = + cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A") + proof - + have h: "?A = ({y. y \ ?A \ y \ ancestors (tRAG s) u} \ + {y. y \ ?A \ \ (y \ ancestors (tRAG s) u)})" (is "_ = ?B \ ?C") + by (rule set_prop_split) + show ?thesis + proof(subst h, subst (3) h, rule f_image_union_eq) + show "?f ` ?B = ?g ` ?B" + proof(rule f_image_eq) + fix a + assume "a \ ?B" + hence h: "a \ RTree.children (tRAG s) x" "a \ ancestors (tRAG s) u" by auto + show "?f a = ?g a" + proof(rule 1(1)[rule_format]) + from h(2) show "a \ ancestors (tRAG s) u" . + next + from h(1) show "(a, x) \ tRAG s" + unfolding RTree.children_def by auto + qed + qed + next + show "?f ` ?C = ?g ` ?C" + proof(rule f_image_eq) + fix a + assume "a \ ?C" + hence h: "a \ RTree.children (tRAG s) x" "a \ ancestors (tRAG s) u" by auto + from h(1) obtain th3 where eq_a: "a = Th th3" + by (unfold RTree.children_def tRAG_alt_def, auto) + thm cp_gen_def_cond[OF eq_a, folded eq_a] + show "?f a = ?g a" + proof(fold cp_gen_def_cond[OF eq_a, folded eq_a], rule cp_kept) + show "Th th3 \ ancestors (tRAG s) (Th th)" + proof + assume "Th th3 \ ancestors (tRAG s) (Th th)" + from this[folded eq_a] have "(Th th, a) \ (tRAG s)^+" by (auto simp:ancestors_def) + from plus_rpath[OF this] obtain xs1 + where h_a: "rpath (tRAG s) (Th th) xs1 a" "xs1 \ []" by auto + from assms(1) have "(Th th, u) \ (tRAG s)^+" by (auto simp:ancestors_def) + from plus_rpath[OF this] obtain xs2 + where h_u: "rpath (tRAG s) (Th th) xs2 u" "xs2 \ []" by auto + show False + proof(cases rule:vat_s.rtree_s.rpath_overlap[OF h_a(1) h_u(1)]) + case (1 xs3) + (* ccc *) + qed + qed + qed + qed + qed + qed + ultimately show ?thesis by metis + qed + ultimately show ?thesis by simp + qed + also have "... = ?R" + by (fold vat_s'.cp_gen_rec[OF eq_x], simp) + finally show ?thesis . + qed +qed + + + +(* ccc *) lemma eq_child_left: assumes nd: "(Th th, Th th') \ (child s)^+" diff -r f1b39d77db00 -r ad57323fd4d6 PrioG.thy --- a/PrioG.thy Thu Dec 03 14:34:29 2015 +0800 +++ b/PrioG.thy Tue Dec 15 21:45:46 2015 +0800 @@ -2192,11 +2192,9 @@ hence eq_max: "Max ((\th. preced th s) ` ({th1} \ dependants (wq s) th1)) = Max ((\th. preced th s) ` ({th2} \ dependants (wq s) th2))" (is "Max (?f ` ?A) = Max (?f ` ?B)") - thm cp_def image_Collect unfolding cp_eq_cpreced unfolding cpreced_def . obtain th1' where th1_in: "th1' \ ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)" - thm Max_in proof - have h1: "finite (?f ` ?A)" proof - @@ -2231,7 +2229,6 @@ have "?A \ {}" by simp thus ?thesis by simp qed - thm Max_in from Max_in [OF h1 h2] have "Max (?f ` ?A) \ (?f ` ?A)" . thus ?thesis diff -r f1b39d77db00 -r ad57323fd4d6 PrioG.thy~ --- a/PrioG.thy~ Thu Dec 03 14:34:29 2015 +0800 +++ b/PrioG.thy~ Tue Dec 15 21:45:46 2015 +0800 @@ -210,7 +210,7 @@ lemma \thread \ set (wq_fun cs1 s); thread \ set (wq_fun cs2 s)\ \ cs1 = cs2" *) -text {* (* ??? *) +text {* (* ddd *) The nature of the work is like this: since it starts from a very simple and basic model, even intuitively very `basic` and `obvious` properties need to derived from scratch. For instance, the fact @@ -842,7 +842,7 @@ qed qed -text {* (* ??? *) +text {* (* ddd *) The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed with the happening of @{text "V"}-events: *} @@ -1544,7 +1544,7 @@ ultimately show ?thesis by (simp add:cntCS_def) qed -text {* (* ??? *) \noindent +text {* (* ddd *) \noindent The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} of one particular thread. *} diff -r f1b39d77db00 -r ad57323fd4d6 RTree.thy --- a/RTree.thy Thu Dec 03 14:34:29 2015 +0800 +++ b/RTree.thy Tue Dec 15 21:45:46 2015 +0800 @@ -1,5 +1,5 @@ theory RTree -imports "~~/src/HOL/Library/Transitive_Closure_Table" +imports "~~/src/HOL/Library/Transitive_Closure_Table" Max begin section {* A theory of relational trees *} @@ -10,7 +10,7 @@ subsection {* Definitions *} text {* - In this theory, we are giving to give a notion of of `Relational Graph` and + In this theory, we are going to give a notion of of `Relational Graph` and its derived notion `Relational Tree`. Given a binary relation @{text "r"}, the `Relational Graph of @{text "r"}` is the graph, the edges of which are those in @{text "r"}. In this way, any binary relation can be viewed @@ -81,6 +81,10 @@ definition "subtree r x = {y . (y, x) \ r^*}" +definition "ancestors r x = {y. (x, y) \ r^+}" + +definition "root r x = (ancestors r x = {})" + text {* The following @{text "edge_in r x"} is the set of edges contained in the sub-tree of @{text "x"}, with @{text "r"} as the underlying graph. @@ -110,15 +114,26 @@ qed text {* - The following lemma shows the means of @{term "edges_in"} from the other side, - which says to for the edge @{text "(a,b)"} to be outside of the sub-tree of @{text "x"}, - it is sufficient if @{text "b"} is. + The following lemma shows the meaning of @{term "edges_in"} from the other side, + which says: for the edge @{text "(a,b)"} to be outside of the sub-tree of @{text "x"}, + it is sufficient to show that @{text "b"} is. *} lemma edges_in_refutation: assumes "b \ subtree r x" shows "(a, b) \ edges_in r x" using assms by (unfold edges_in_def subtree_def, auto) +definition "children r x = {y. (y, x) \ r}" + +locale fbranch = + fixes r + assumes fb: "\ x \ Range r . finite (children r x)" + +locale fsubtree = fbranch + + assumes wf: "wf r" + +(* ccc *) + subsection {* Auxiliary lemmas *} lemma index_minimize: @@ -171,8 +186,6 @@ text {* Induction rule for @{text "rpath"}: *} -print_statement rtrancl_path.induct - lemma rpath_induct [consumes 1, case_names rbase rstep, induct pred: rpath]: assumes "rpath r x1 x2 x3" and "\x. P x [] x" @@ -181,6 +194,13 @@ using assms[unfolded rpath_def] by (induct, auto simp:pred_of_def rpath_def) +lemma rpathE: + assumes "rpath r x xs y" + obtains (base) "y = x" "xs = []" + | (step) z zs where "(x, z) \ r" "rpath r z zs y" "xs = z#zs" + using assms + by (induct, auto) + text {* Introduction rule for empty path *} lemma rbaseI [intro!]: assumes "x = y" @@ -282,6 +302,43 @@ subsubsection {* Properites of @{text "edges_on"} *} +lemma edges_on_unfold: + "edges_on (a # b # xs) = {(a, b)} \ edges_on (b # xs)" (is "?L = ?R") +proof - + { fix c d + assume "(c, d) \ ?L" + then obtain l1 l2 where h: "(a # b # xs) = l1 @ [c, d] @ l2" + by (auto simp:edges_on_def) + have "(c, d) \ ?R" + proof(cases "l1") + case Nil + with h have "(c, d) = (a, b)" by auto + thus ?thesis by auto + next + case (Cons e es) + from h[unfolded this] have "b#xs = es@[c, d]@l2" by auto + thus ?thesis by (auto simp:edges_on_def) + qed + } moreover + { fix c d + assume "(c, d) \ ?R" + moreover have "(a, b) \ ?L" + proof - + have "(a # b # xs) = []@[a,b]@xs" by simp + hence "\ l1 l2. (a # b # xs) = l1@[a,b]@l2" by auto + thus ?thesis by (unfold edges_on_def, simp) + qed + moreover { + assume "(c, d) \ edges_on (b#xs)" + then obtain l1 l2 where "b#xs = l1@[c, d]@l2" by (unfold edges_on_def, auto) + hence "a#b#xs = (a#l1)@[c,d]@l2" by simp + hence "\ l1 l2. (a # b # xs) = l1@[c,d]@l2" by metis + hence "(c,d) \ ?L" by (unfold edges_on_def, simp) + } + ultimately have "(c, d) \ ?L" by auto + } ultimately show ?thesis by auto +qed + lemma edges_on_len: assumes "(a,b) \ edges_on l" shows "length l \ 2" @@ -289,6 +346,7 @@ by (unfold edges_on_def, auto) text {* Elimination of @{text "edges_on"} for non-empty path *} + lemma edges_on_consE [elim, cases set:edges_on]: assumes "(a,b) \ edges_on (x#xs)" obtains (head) xs' where "x = a" and "xs = b#xs'" @@ -384,6 +442,23 @@ qed qed (unfold rpath_def, auto intro!:Transitive_Closure_Table.rtrancl_path.base) +lemma edges_on_rpathI: + assumes "edges_on (a#xs@[b]) \ r" + shows "rpath r a (xs@[b]) b" + using assms +proof(induct xs arbitrary: a b) + case Nil + moreover have "(a, b) \ edges_on (a # [] @ [b])" + by (unfold edges_on_def, auto) + ultimately have "(a, b) \ r" by auto + thus ?case by auto +next + case (Cons x xs a b) + from this(2) have "edges_on (x # xs @ [b]) \ r" by (simp add:edges_on_unfold) + from Cons(1)[OF this] have " rpath r x (xs @ [b]) b" . + moreover from Cons(2) have "(a, x) \ r" by (auto simp:edges_on_unfold) + ultimately show ?case by (auto) +qed text {* The following lemma extracts the path from @{text "x"} to @{text "y"} @@ -416,8 +491,32 @@ thus ?thesis by (simp add: pred_of_star star_2_pstar) qed +lemma subtree_transfer: + assumes "a \ subtree r1 a'" + and "r1 \ r2" + shows "a \ subtree r2 a'" +proof - + from assms(1)[unfolded subtree_def] + have "(a, a') \ r1^*" by auto + from star_rpath[OF this] + obtain xs where rp: "rpath r1 a xs a'" by blast + hence "rpath r2 a xs a'" + proof(rule rpath_transfer) + from rpath_edges_on[OF rp] and assms(2) + show "edges_on (a # xs) \ r2" by simp + qed + from rpath_star[OF this] + show ?thesis by (auto simp:subtree_def) +qed + +lemma subtree_rev_transfer: + assumes "a \ subtree r2 a'" + and "r1 \ r2" + shows "a \ subtree r1 a'" + using assms and subtree_transfer by metis + text {* - The following lemmas establishes a relation from pathes in @{text "r"} + The following lemmas establishes a relation from paths in @{text "r"} to @{text "r^+"} relation. *} lemma rpath_plus: @@ -438,7 +537,50 @@ qed qed -subsubsection {* Properties of @{text "subtree"} *} +lemma plus_rpath: + assumes "(x, y) \ r^+" + obtains xs where "rpath r x xs y" and "xs \ []" +proof - + from assms + show ?thesis + proof(cases rule:converse_tranclE[consumes 1]) + case 1 + hence "rpath r x [y] y" by auto + from that[OF this] show ?thesis by auto + next + case (2 z) + from 2(2) have "(z, y) \ r^*" by auto + from star_rpath[OF this] obtain xs where "rpath r z xs y" by auto + from rstepI[OF 2(1) this] + have "rpath r x (z # xs) y" . + from that[OF this] show ?thesis by auto + qed +qed + +subsubsection {* Properties of @{text "subtree"} and @{term "ancestors"}*} + +lemma ancestors_subtreeI: + assumes "b \ ancestors r a" + shows "a \ subtree r b" + using assms by (auto simp:subtree_def ancestors_def) + +lemma subtreeE: + assumes "a \ subtree r b" + obtains "a = b" + | "a \ b" and "b \ ancestors r a" +proof - + from assms have "(a, b) \ r^*" by (auto simp:subtree_def) + from rtranclD[OF this] + have " a = b \ a \ b \ (a, b) \ r\<^sup>+" . + with that[unfolded ancestors_def] show ?thesis by auto +qed + +lemma subtree_ancestorsI: + assumes "a \ subtree r b" + and "a \ b" + shows "b \ ancestors r a" + using assms + by (auto elim!:subtreeE) text {* @{text "subtree"} is mono with respect to the underlying graph. @@ -513,6 +655,55 @@ } ultimately show ?thesis by auto qed +(* ddd *) +lemma subset_del_subtree_outside: (* ddd *) + assumes "Range r' \ subtree r x = {}" + shows "subtree (r - r') x = (subtree r x)" +proof - + { fix c + assume "c \ (subtree r x)" + hence "(c, x) \ r^*" by (auto simp:subtree_def) + hence "c \ subtree (r - r') x" + proof(rule star_rpath) + fix xs + assume rp: "rpath r c xs x" + show ?thesis + proof - + from rp + have "rpath (r - r') c xs x" + proof(rule rpath_transfer) + from rpath_edges_on[OF rp] have "edges_on (c # xs) \ r" . + moreover { + fix a b + assume h: "(a, b) \ r'" + have "(a, b) \ edges_on (c#xs)" + proof + assume "(a, b) \ edges_on (c # xs)" + then obtain l1 l2 where "c#xs = (l1@[a])@[b]@l2" by (auto simp:edges_on_def) + hence "tl (c#xs) = tl (l1@[a,b]@l2)" by simp + then obtain l1' where eq_xs_b: "xs = l1'@[b]@l2" by (cases l1, auto) + from rp[unfolded this] + show False + proof(rule rpath_appendE) + assume "rpath r b l2 x" + from rpath_star[OF this] + have "b \ subtree r x" by (auto simp:subtree_def) + with assms (1) and h show ?thesis by (auto) + qed + qed + } ultimately show "edges_on (c # xs) \ r - r'" by auto + qed + thus ?thesis by (rule rpath_star[elim_format], auto simp:subtree_def) + qed + qed + } moreover { + fix c + assume "c \ subtree (r - r') x" + moreover have "... \ (subtree r x)" by (rule subtree_mono, auto) + ultimately have "c \ (subtree r x)" by auto + } ultimately show ?thesis by auto +qed + lemma subtree_insert_ext: assumes "b \ subtree r x" shows "subtree (r \ {(a, b)}) x = (subtree r x) \ (subtree r a)" @@ -524,11 +715,138 @@ using assms by (auto simp:subtree_def rtrancl_insert) +lemma set_add_rootI: + assumes "root r a" + and "a \ Domain r1" + shows "root (r \ r1) a" +proof - + let ?r = "r \ r1" + { fix a' + assume "a' \ ancestors ?r a" + hence "(a, a') \ ?r^+" by (auto simp:ancestors_def) + from tranclD[OF this] obtain z where "(a, z) \ ?r" by auto + moreover have "(a, z) \ r" + proof + assume "(a, z) \ r" + with assms(1) show False + by (auto simp:root_def ancestors_def) + qed + ultimately have "(a, z) \ r1" by auto + with assms(2) + have False by (auto) + } thus ?thesis by (auto simp:root_def) +qed + +lemma ancestors_mono: + assumes "r1 \ r2" + shows "ancestors r1 x \ ancestors r2 x" +proof + fix a + assume "a \ ancestors r1 x" + hence "(x, a) \ r1^+" by (auto simp:ancestors_def) + from plus_rpath[OF this] obtain xs where + h: "rpath r1 x xs a" "xs \ []" . + have "rpath r2 x xs a" + proof(rule rpath_transfer[OF h(1)]) + from rpath_edges_on[OF h(1)] and assms + show "edges_on (x # xs) \ r2" by auto + qed + from rpath_plus[OF this h(2)] + show "a \ ancestors r2 x" by (auto simp:ancestors_def) +qed + +lemma subtree_refute: + assumes "x \ ancestors r y" + and "x \ y" + shows "y \ subtree r x" +proof + assume "y \ subtree r x" + thus False + by(elim subtreeE, insert assms, auto) +qed + subsubsection {* Properties about relational trees *} context rtree begin +lemma ancestors_headE: + assumes "c \ ancestors r a" + assumes "(a, b) \ r" + obtains "b = c" + | "c \ ancestors r b" +proof - + from assms(1) + have "(a, c) \ r^+" by (auto simp:ancestors_def) + hence "b = c \ c \ ancestors r b" + proof(cases rule:converse_tranclE[consumes 1]) + case 1 + with assms(2) and sgv have "b = c" by (auto simp:single_valued_def) + thus ?thesis by auto + next + case (2 y) + from 2(1) and assms(2) and sgv have "y = b" by (auto simp:single_valued_def) + from 2(2)[unfolded this] have "c \ ancestors r b" by (auto simp:ancestors_def) + thus ?thesis by auto + qed + with that show ?thesis by metis +qed + +lemma ancestors_accum: + assumes "(a, b) \ r" + shows "ancestors r a = ancestors r b \ {b}" +proof - + { fix c + assume "c \ ancestors r a" + hence "(a, c) \ r^+" by (auto simp:ancestors_def) + hence "c \ ancestors r b \ {b}" + proof(cases rule:converse_tranclE[consumes 1]) + case 1 + with sgv assms have "c = b" by (unfold single_valued_def, auto) + thus ?thesis by auto + next + case (2 c') + with sgv assms have "c' = b" by (unfold single_valued_def, auto) + from 2(2)[unfolded this] + show ?thesis by (auto simp:ancestors_def) + qed + } moreover { + fix c + assume "c \ ancestors r b \ {b}" + hence "c = b \ c \ ancestors r b" by auto + hence "c \ ancestors r a" + proof + assume "c = b" + from assms[folded this] + show ?thesis by (auto simp:ancestors_def) + next + assume "c \ ancestors r b" + with assms show ?thesis by (auto simp:ancestors_def) + qed + } ultimately show ?thesis by auto +qed + +lemma rootI: + assumes h: "\ x'. x' \ x \ x \ subtree r' x'" + and "r' \ r" + shows "root r' x" +proof - + from acyclic_subset[OF acl assms(2)] + have acl': "acyclic r'" . + { fix x' + assume "x' \ ancestors r' x" + hence h1: "(x, x') \ r'^+" by (auto simp:ancestors_def) + have "x' \ x" + proof + assume eq_x: "x' = x" + from h1[unfolded this] and acl' + show False by (auto simp:acyclic_def) + qed + moreover from h1 have "x \ subtree r' x'" by (auto simp:subtree_def) + ultimately have False using h by auto + } thus ?thesis by (auto simp:root_def) +qed + lemma rpath_overlap_oneside: (* ddd *) assumes "rpath r x xs1 x1" and "rpath r x xs2 x2" @@ -953,6 +1271,435 @@ ultimately show ?thesis by auto qed +lemma set_del_rootI: + assumes "r1 \ r" + and "a \ Domain r1" + shows "root (r - r1) a" +proof - + let ?r = "r - r1" + { fix a' + assume neq: "a' \ a" + have "a \ subtree ?r a'" + proof + assume "a \ subtree ?r a'" + hence "(a, a') \ ?r^*" by (auto simp:subtree_def) + from star_rpath[OF this] obtain xs + where rp: "rpath ?r a xs a'" by auto + from rpathE[OF this] and neq + obtain z zs where h: "(a, z) \ ?r" "rpath ?r z zs a'" "xs = z#zs" by auto + from assms(2) obtain z' where z'_in: "(a, z') \ r1" by (auto simp:DomainE) + with assms(1) have "(a, z') \ r" by auto + moreover from h(1) have "(a, z) \ r" by simp + ultimately have "z' = z" using sgv by (auto simp:single_valued_def) + from z'_in[unfolded this] and h(1) show False by auto + qed + } thus ?thesis by (intro rootI, auto) +qed + +lemma edge_del_no_rootI: + assumes "(a, b) \ r" + shows "root (r - {(a, b)}) a" + by (rule set_del_rootI, insert assms, auto) + +lemma ancestors_children_unique: + assumes "z1 \ ancestors r x \ children r y" + and "z2 \ ancestors r x \ children r y" + shows "z1 = z2" +proof - + from assms have h: + "(x, z1) \ r^+" "(z1, y) \ r" + "(x, z2) \ r^+" "(z2, y) \ r" + by (auto simp:ancestors_def children_def) + + -- {* From this, a path containing @{text "z1"} is obtained. *} + from plus_rpath[OF h(1)] obtain xs1 + where h1: "rpath r x xs1 z1" "xs1 \ []" by auto + from rpath_nnl_lastE[OF this] obtain xs1' where eq_xs1: "xs1 = xs1' @ [z1]" + by auto + from h(2) have h2: "rpath r z1 [y] y" by auto + from rpath_appendI[OF h1(1) h2, unfolded eq_xs1] + have rp1: "rpath r x (xs1' @ [z1, y]) y" by simp + + -- {* Then, another path containing @{text "z2"} is obtained. *} + from plus_rpath[OF h(3)] obtain xs2 + where h3: "rpath r x xs2 z2" "xs2 \ []" by auto + from rpath_nnl_lastE[OF this] obtain xs2' where eq_xs2: "xs2 = xs2' @ [z2]" + by auto + from h(4) have h4: "rpath r z2 [y] y" by auto + from rpath_appendI[OF h3(1) h4, unfolded eq_xs2] + have "rpath r x (xs2' @ [z2, y]) y" by simp + + -- {* Finally @{text "z1 = z2"} is proved by uniqueness of path. *} + from rpath_unique[OF rp1 this] + have "xs1' @ [z1, y] = xs2' @ [z2, y]" . + thus ?thesis by auto +qed + +lemma ancestors_childrenE: + assumes "y \ ancestors r x" + obtains "x \ children r y" + | z where "z \ ancestors r x \ children r y" +proof - + from assms(1) have "(x, y) \ r^+" by (auto simp:ancestors_def) + from tranclD2[OF this] obtain z where + h: "(x, z) \ r\<^sup>*" "(z, y) \ r" by auto + from h(1) + show ?thesis + proof(cases rule:rtranclE) + case base + from h(2)[folded this] have "x \ children r y" + by (auto simp:children_def) + thus ?thesis by (intro that, auto) + next + case (step u) + hence "z \ ancestors r x" by (auto simp:ancestors_def) + moreover from h(2) have "z \ children r y" + by (auto simp:children_def) + ultimately show ?thesis by (intro that, auto) + qed +qed + + +end (* of rtree *) + +lemma subtree_children: + "subtree r x = {x} \ (\ (subtree r ` (children r x)))" (is "?L = ?R") +proof - + { fix z + assume "z \ ?L" + hence "z \ ?R" + proof(cases rule:subtreeE[consumes 1]) + case 2 + hence "(z, x) \ r^+" by (auto simp:ancestors_def) + thus ?thesis + proof(rule tranclE) + assume "(z, x) \ r" + hence "z \ children r x" by (unfold children_def, auto) + moreover have "z \ subtree r z" by (auto simp:subtree_def) + ultimately show ?thesis by auto + next + fix c + assume h: "(z, c) \ r\<^sup>+" "(c, x) \ r" + hence "c \ children r x" by (auto simp:children_def) + moreover from h have "z \ subtree r c" by (auto simp:subtree_def) + ultimately show ?thesis by auto + qed + qed auto + } moreover { + fix z + assume h: "z \ ?R" + have "x \ subtree r x" by (auto simp:subtree_def) + moreover { + assume "z \ \(subtree r ` children r x)" + then obtain y where "(y, x) \ r" "(z, y) \ r^*" + by (auto simp:subtree_def children_def) + hence "(z, x) \ r^*" by auto + hence "z \ ?L" by (auto simp:subtree_def) + } ultimately have "z \ ?L" using h by auto + } ultimately show ?thesis by auto +qed + +context fsubtree +begin + +lemma finite_subtree: + shows "finite (subtree r x)" +proof(induct rule:wf_induct[OF wf]) + case (1 x) + have "finite (\(subtree r ` children r x))" + proof(rule finite_Union) + show "finite (subtree r ` children r x)" + proof(cases "children r x = {}") + case True + thus ?thesis by auto + next + case False + hence "x \ Range r" by (auto simp:children_def) + from fb[rule_format, OF this] + have "finite (children r x)" . + thus ?thesis by (rule finite_imageI) + qed + next + fix M + assume "M \ subtree r ` children r x" + then obtain y where h: "y \ children r x" "M = subtree r y" by auto + hence "(y, x) \ r" by (auto simp:children_def) + from 1[rule_format, OF this, folded h(2)] + show "finite M" . + qed + thus ?case + by (unfold subtree_children finite_Un, auto) +qed + end +definition "pairself f = (\(a, b). (f a, f b))" + +definition "rel_map f r = (pairself f ` r)" + +lemma rel_mapE: + assumes "(a, b) \ rel_map f r" + obtains c d + where "(c, d) \ r" "(a, b) = (f c, f d)" + using assms + by (unfold rel_map_def pairself_def, auto) + +lemma rel_mapI: + assumes "(a, b) \ r" + and "c = f a" + and "d = f b" + shows "(c, d) \ rel_map f r" + using assms + by (unfold rel_map_def pairself_def, auto) + +lemma map_appendE: + assumes "map f zs = xs @ ys" + obtains xs' ys' + where "zs = xs' @ ys'" "xs = map f xs'" "ys = map f ys'" +proof - + have "\ xs' ys'. zs = xs' @ ys' \ xs = map f xs' \ ys = map f ys'" + using assms + proof(induct xs arbitrary:zs ys) + case (Nil zs ys) + thus ?case by auto + next + case (Cons x xs zs ys) + note h = this + show ?case + proof(cases zs) + case (Cons e es) + with h have eq_x: "map f es = xs @ ys" "x = f e" by auto + from h(1)[OF this(1)] + obtain xs' ys' where "es = xs' @ ys'" "xs = map f xs'" "ys = map f ys'" + by blast + with Cons eq_x + have "zs = (e#xs') @ ys' \ x # xs = map f (e#xs') \ ys = map f ys'" by auto + thus ?thesis by metis + qed (insert h, auto) + qed + thus ?thesis by (auto intro!:that) +qed + +lemma rel_map_mono: + assumes "r1 \ r2" + shows "rel_map f r1 \ rel_map f r2" + using assms + by (auto simp:rel_map_def pairself_def) + +lemma rel_map_compose [simp]: + shows "rel_map f1 (rel_map f2 r) = rel_map (f1 o f2) r" + by (auto simp:rel_map_def pairself_def) + +lemma edges_on_map: "edges_on (map f xs) = rel_map f (edges_on xs)" +proof - + { fix a b + assume "(a, b) \ edges_on (map f xs)" + then obtain l1 l2 where eq_map: "map f xs = l1 @ [a, b] @ l2" + by (unfold edges_on_def, auto) + hence "(a, b) \ rel_map f (edges_on xs)" + by (auto elim!:map_appendE intro!:rel_mapI simp:edges_on_def) + } moreover { + fix a b + assume "(a, b) \ rel_map f (edges_on xs)" + then obtain c d where + h: "(c, d) \ edges_on xs" "(a, b) = (f c, f d)" + by (elim rel_mapE, auto) + then obtain l1 l2 where + eq_xs: "xs = l1 @ [c, d] @ l2" + by (auto simp:edges_on_def) + hence eq_map: "map f xs = map f l1 @ [f c, f d] @ map f l2" by auto + have "(a, b) \ edges_on (map f xs)" + proof - + from h(2) have "[f c, f d] = [a, b]" by simp + from eq_map[unfolded this] show ?thesis by (auto simp:edges_on_def) + qed + } ultimately show ?thesis by auto +qed + +lemma image_id: + assumes "\ x. x \ A \ f x = x" + shows "f ` A = A" + using assms by (auto simp:image_def) + +lemma rel_map_inv_id: + assumes "inj_on f ((Domain r) \ (Range r))" + shows "(rel_map (inv_into ((Domain r) \ (Range r)) f \ f) r) = r" +proof - + let ?f = "(inv_into (Domain r \ Range r) f \ f)" + { + fix a b + assume h0: "(a, b) \ r" + have "pairself ?f (a, b) = (a, b)" + proof - + from assms h0 have "?f a = a" by (auto intro:inv_into_f_f) + moreover have "?f b = b" + by (insert h0, simp, intro inv_into_f_f[OF assms], auto intro!:RangeI) + ultimately show ?thesis by (auto simp:pairself_def) + qed + } thus ?thesis by (unfold rel_map_def, intro image_id, case_tac x, auto) +qed + +lemma rel_map_acyclic: + assumes "acyclic r" + and "inj_on f ((Domain r) \ (Range r))" + shows "acyclic (rel_map f r)" +proof - + let ?D = "Domain r \ Range r" + { fix a + assume "(a, a) \ (rel_map f r)^+" + from plus_rpath[OF this] + obtain xs where rp: "rpath (rel_map f r) a xs a" "xs \ []" by auto + from rpath_nnl_lastE[OF this] obtain xs' where eq_xs: "xs = xs'@[a]" by auto + from rpath_edges_on[OF rp(1)] + have h: "edges_on (a # xs) \ rel_map f r" . + from edges_on_map[of "inv_into ?D f" "a#xs"] + have "edges_on (map (inv_into ?D f) (a # xs)) = rel_map (inv_into ?D f) (edges_on (a # xs))" . + with rel_map_mono[OF h, of "inv_into ?D f"] + have "edges_on (map (inv_into ?D f) (a # xs)) \ rel_map ((inv_into ?D f) o f) r" by simp + from this[unfolded eq_xs] + have subr: "edges_on (map (inv_into ?D f) (a # xs' @ [a])) \ rel_map (inv_into ?D f \ f) r" . + have "(map (inv_into ?D f) (a # xs' @ [a])) = (inv_into ?D f a) # map (inv_into ?D f) xs' @ [inv_into ?D f a]" + by simp + from edges_on_rpathI[OF subr[unfolded this]] + have "rpath (rel_map (inv_into ?D f \ f) r) + (inv_into ?D f a) (map (inv_into ?D f) xs' @ [inv_into ?D f a]) (inv_into ?D f a)" . + hence "(inv_into ?D f a, inv_into ?D f a) \ (rel_map (inv_into ?D f \ f) r)^+" + by (rule rpath_plus, simp) + moreover have "(rel_map (inv_into ?D f \ f) r) = r" by (rule rel_map_inv_id[OF assms(2)]) + moreover note assms(1) + ultimately have False by (unfold acyclic_def, auto) + } thus ?thesis by (auto simp:acyclic_def) +qed + +lemma relpow_mult: + "((r::'a rel) ^^ m) ^^ n = r ^^ (m*n)" +proof(induct n arbitrary:m) + case (Suc k m) + thus ?case + proof - + have h: "(m * k + m) = (m + m * k)" by auto + show ?thesis + apply (simp add:Suc relpow_add[symmetric]) + by (unfold h, simp) + qed +qed simp + +lemma compose_relpow_2: + assumes "r1 \ r" + and "r2 \ r" + shows "r1 O r2 \ r ^^ (2::nat)" +proof - + { fix a b + assume "(a, b) \ r1 O r2" + then obtain e where "(a, e) \ r1" "(e, b) \ r2" + by auto + with assms have "(a, e) \ r" "(e, b) \ r" by auto + hence "(a, b) \ r ^^ (Suc (Suc 0))" by auto + } thus ?thesis by (auto simp:numeral_2_eq_2) +qed + +lemma acyclic_compose: + assumes "acyclic r" + and "r1 \ r" + and "r2 \ r" + shows "acyclic (r1 O r2)" +proof - + { fix a + assume "(a, a) \ (r1 O r2)^+" + from trancl_mono[OF this compose_relpow_2[OF assms(2, 3)]] + have "(a, a) \ (r ^^ 2) ^+" . + from trancl_power[THEN iffD1, OF this] + obtain n where h: "(a, a) \ (r ^^ 2) ^^ n" "n > 0" by blast + from this(1)[unfolded relpow_mult] have h2: "(a, a) \ r ^^ (2 * n)" . + have "(a, a) \ r^+" + proof(cases rule:trancl_power[THEN iffD2]) + from h(2) h2 show "\n>0. (a, a) \ r ^^ n" + by (rule_tac x = "2*n" in exI, auto) + qed + with assms have "False" by (auto simp:acyclic_def) + } thus ?thesis by (auto simp:acyclic_def) +qed + +lemma children_compose_unfold: + "children (r1 O r2) x = \ (children r1 ` (children r2 x))" + by (auto simp:children_def) + +lemma fbranch_compose: + assumes "fbranch r1" + and "fbranch r2" + shows "fbranch (r1 O r2)" +proof - + { fix x + assume "x\Range (r1 O r2)" + then obtain y z where h: "(y, z) \ r1" "(z, x) \ r2" by auto + have "finite (children (r1 O r2) x)" + proof(unfold children_compose_unfold, rule finite_Union) + show "finite (children r1 ` children r2 x)" + proof(rule finite_imageI) + from h(2) have "x \ Range r2" by auto + from assms(2)[unfolded fbranch_def, rule_format, OF this] + show "finite (children r2 x)" . + qed + next + fix M + assume "M \ children r1 ` children r2 x" + then obtain y where h1: "y \ children r2 x" "M = children r1 y" by auto + show "finite M" + proof(cases "children r1 y = {}") + case True + with h1(2) show ?thesis by auto + next + case False + hence "y \ Range r1" by (unfold children_def, auto) + from assms(1)[unfolded fbranch_def, rule_format, OF this, folded h1(2)] + show ?thesis . + qed + qed + } thus ?thesis by (unfold fbranch_def, auto) +qed + +lemma finite_fbranchI: + assumes "finite r" + shows "fbranch r" +proof - + { fix x + assume "x \Range r" + have "finite (children r x)" + proof - + have "{y. (y, x) \ r} \ Domain r" by (auto) + from rev_finite_subset[OF finite_Domain[OF assms] this] + have "finite {y. (y, x) \ r}" . + thus ?thesis by (unfold children_def, simp) + qed + } thus ?thesis by (auto simp:fbranch_def) +qed + +lemma subset_fbranchI: + assumes "fbranch r1" + and "r2 \ r1" + shows "fbranch r2" +proof - + { fix x + assume "x \Range r2" + with assms(2) have "x \ Range r1" by auto + from assms(1)[unfolded fbranch_def, rule_format, OF this] + have "finite (children r1 x)" . + hence "finite (children r2 x)" + proof(rule rev_finite_subset) + from assms(2) + show "children r2 x \ children r1 x" by (auto simp:children_def) + qed + } thus ?thesis by (auto simp:fbranch_def) +qed + +lemma children_subtree: + shows "children r x \ subtree r x" + by (auto simp:children_def subtree_def) + +lemma children_union_kept: + assumes "x \ Range r'" + shows "children (r \ r') x = children r x" + using assms + by (auto simp:children_def) + end \ No newline at end of file