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