|
1 theory RTree |
|
2 imports "~~/src/HOL/Library/Transitive_Closure_Table" Max |
|
3 begin |
|
4 |
|
5 section {* A theory of relational trees *} |
|
6 |
|
7 inductive_cases path_nilE [elim!]: "rtrancl_path r x [] y" |
|
8 inductive_cases path_consE [elim!]: "rtrancl_path r x (z#zs) y" |
|
9 |
|
10 subsection {* Definitions *} |
|
11 |
|
12 text {* |
|
13 In this theory, we are going to give a notion of of `Relational Graph` and |
|
14 its derived notion `Relational Tree`. Given a binary relation @{text "r"}, |
|
15 the `Relational Graph of @{text "r"}` is the graph, the edges of which |
|
16 are those in @{text "r"}. In this way, any binary relation can be viewed |
|
17 as a `Relational Graph`. Note, this notion of graph includes infinite graphs. |
|
18 |
|
19 A `Relation Graph` @{text "r"} is said to be a `Relational Tree` if it is both |
|
20 {\em single valued} and {\em acyclic}. |
|
21 *} |
|
22 |
|
23 text {* |
|
24 The following @{text "sgv"} specifies that relation @{text "r"} is {\em single valued}. |
|
25 *} |
|
26 locale sgv = |
|
27 fixes r |
|
28 assumes sgv: "single_valued r" |
|
29 |
|
30 text {* |
|
31 The following @{text "rtree"} specifies that @{text "r"} is a |
|
32 {\em Relational Tree}. |
|
33 *} |
|
34 locale rtree = sgv + |
|
35 assumes acl: "acyclic r" |
|
36 |
|
37 text {* |
|
38 The following two auxiliary functions @{text "rel_of"} and @{text "pred_of"} |
|
39 transfer between the predicate and set representation of binary relations. |
|
40 *} |
|
41 |
|
42 definition "rel_of r = {(x, y) | x y. r x y}" |
|
43 |
|
44 definition "pred_of r = (\<lambda> x y. (x, y) \<in> r)" |
|
45 |
|
46 text {* |
|
47 To reason about {\em Relational Graph}, a notion of path is |
|
48 needed, which is given by the following @{text "rpath"} (short |
|
49 for `relational path`). |
|
50 The path @{text "xs"} in proposition @{text "rpath r x xs y"} is |
|
51 a path leading from @{text "x"} to @{text "y"}, which serves as a |
|
52 witness of the fact @{text "(x, y) \<in> r^*"}. |
|
53 |
|
54 @{text "rpath"} |
|
55 is simply a wrapper of the @{text "rtrancl_path"} defined in the imported |
|
56 theory @{text "Transitive_Closure_Table"}, which defines |
|
57 a notion of path for the predicate form of binary relations. |
|
58 *} |
|
59 definition "rpath r x xs y = rtrancl_path (pred_of r) x xs y" |
|
60 |
|
61 text {* |
|
62 Given a path @{text "ps"}, @{text "edges_on ps"} is the |
|
63 set of edges along the path, which is defined as follows: |
|
64 *} |
|
65 |
|
66 definition "edges_on ps = {(a,b) | a b. \<exists> xs ys. ps = xs@[a,b]@ys}" |
|
67 |
|
68 text {* |
|
69 The following @{text "indep"} defines a notion of independence. |
|
70 Two nodes @{text "x"} and @{text "y"} are said to be independent |
|
71 (expressed as @{text "indep x y"}), if neither one is reachable |
|
72 from the other in relational graph @{text "r"}. |
|
73 *} |
|
74 definition "indep r x y = (((x, y) \<notin> r^*) \<and> ((y, x) \<notin> r^*))" |
|
75 |
|
76 text {* |
|
77 In relational tree @{text "r"}, the sub tree of node @{text "x"} is written |
|
78 @{text "subtree r x"}, which is defined to be the set of nodes (including itself) |
|
79 which can reach @{text "x"} by following some path in @{text "r"}: |
|
80 *} |
|
81 |
|
82 definition "subtree r x = {y . (y, x) \<in> r^*}" |
|
83 |
|
84 definition "ancestors r x = {y. (x, y) \<in> r^+}" |
|
85 |
|
86 definition "root r x = (ancestors r x = {})" |
|
87 |
|
88 text {* |
|
89 The following @{text "edge_in r x"} is the set of edges |
|
90 contained in the sub-tree of @{text "x"}, with @{text "r"} as the underlying graph. |
|
91 *} |
|
92 |
|
93 definition "edges_in r x = {(a, b) | a b. (a, b) \<in> r \<and> b \<in> subtree r x}" |
|
94 |
|
95 text {* |
|
96 The following lemma @{text "edges_in_meaning"} shows the intuitive meaning |
|
97 of `an edge @{text "(a, b)"} is in the sub-tree of @{text "x"}`, |
|
98 i.e., both @{text "a"} and @{text "b"} are in the sub-tree. |
|
99 *} |
|
100 lemma edges_in_meaning: |
|
101 "edges_in r x = {(a, b) | a b. (a, b) \<in> r \<and> a \<in> subtree r x \<and> b \<in> subtree r x}" |
|
102 proof - |
|
103 { fix a b |
|
104 assume h: "(a, b) \<in> r" "b \<in> subtree r x" |
|
105 moreover have "a \<in> subtree r x" |
|
106 proof - |
|
107 from h(2)[unfolded subtree_def] have "(b, x) \<in> r^*" by simp |
|
108 with h(1) have "(a, x) \<in> r^*" by auto |
|
109 thus ?thesis by (auto simp:subtree_def) |
|
110 qed |
|
111 ultimately have "((a, b) \<in> r \<and> a \<in> subtree r x \<and> b \<in> subtree r x)" |
|
112 by (auto) |
|
113 } thus ?thesis by (auto simp:edges_in_def) |
|
114 qed |
|
115 |
|
116 text {* |
|
117 The following lemma shows the meaning of @{term "edges_in"} from the other side, |
|
118 which says: for the edge @{text "(a,b)"} to be outside of the sub-tree of @{text "x"}, |
|
119 it is sufficient to show that @{text "b"} is. |
|
120 *} |
|
121 lemma edges_in_refutation: |
|
122 assumes "b \<notin> subtree r x" |
|
123 shows "(a, b) \<notin> edges_in r x" |
|
124 using assms by (unfold edges_in_def subtree_def, auto) |
|
125 |
|
126 definition "children r x = {y. (y, x) \<in> r}" |
|
127 |
|
128 locale fbranch = |
|
129 fixes r |
|
130 assumes fb: "\<forall> x \<in> Range r . finite (children r x)" |
|
131 begin |
|
132 |
|
133 lemma finite_children: "finite (children r x)" |
|
134 proof(cases "children r x = {}") |
|
135 case True |
|
136 thus ?thesis by auto |
|
137 next |
|
138 case False |
|
139 then obtain y where "(y, x) \<in> r" by (auto simp:children_def) |
|
140 hence "x \<in> Range r" by auto |
|
141 from fb[rule_format, OF this] |
|
142 show ?thesis . |
|
143 qed |
|
144 |
|
145 end |
|
146 |
|
147 locale fsubtree = fbranch + |
|
148 assumes wf: "wf r" |
|
149 |
|
150 (* ccc *) |
|
151 |
|
152 subsection {* Auxiliary lemmas *} |
|
153 |
|
154 lemma index_minimize: |
|
155 assumes "P (i::nat)" |
|
156 obtains j where "P j" and "\<forall> k < j. \<not> P k" |
|
157 proof - |
|
158 have "\<exists> j. P j \<and> (\<forall> k < j. \<not> P k)" |
|
159 using assms |
|
160 proof(induct i rule:less_induct) |
|
161 case (less t) |
|
162 show ?case |
|
163 proof(cases "\<forall> j < t. \<not> P j") |
|
164 case True |
|
165 with less (2) show ?thesis by blast |
|
166 next |
|
167 case False |
|
168 then obtain j where "j < t" "P j" by auto |
|
169 from less(1)[OF this] |
|
170 show ?thesis . |
|
171 qed |
|
172 qed |
|
173 with that show ?thesis by metis |
|
174 qed |
|
175 |
|
176 subsection {* Properties of Relational Graphs and Relational Trees *} |
|
177 |
|
178 subsubsection {* Properties of @{text "rel_of"} and @{text "pred_of"} *} |
|
179 |
|
180 text {* The following lemmas establish bijectivity of the two functions *} |
|
181 |
|
182 lemma pred_rel_eq: "pred_of (rel_of r) = r" by (auto simp:rel_of_def pred_of_def) |
|
183 |
|
184 lemma rel_pred_eq: "rel_of (pred_of r) = r" by (auto simp:rel_of_def pred_of_def) |
|
185 |
|
186 lemma rel_of_star: "rel_of (r^**) = (rel_of r)^*" |
|
187 by (unfold rel_of_def rtranclp_rtrancl_eq, auto) |
|
188 |
|
189 lemma pred_of_star: "pred_of (r^*) = (pred_of r)^**" |
|
190 proof - |
|
191 { fix x y |
|
192 have "pred_of (r^*) x y = (pred_of r)^** x y" |
|
193 by (unfold pred_of_def rtranclp_rtrancl_eq, auto) |
|
194 } thus ?thesis by auto |
|
195 qed |
|
196 |
|
197 lemma star_2_pstar: "(x, y) \<in> r^* = (pred_of (r^*)) x y" |
|
198 by (simp add: pred_of_def) |
|
199 |
|
200 subsubsection {* Properties of @{text "rpath"} *} |
|
201 |
|
202 text {* Induction rule for @{text "rpath"}: *} |
|
203 |
|
204 lemma rpath_induct [consumes 1, case_names rbase rstep, induct pred: rpath]: |
|
205 assumes "rpath r x1 x2 x3" |
|
206 and "\<And>x. P x [] x" |
|
207 and "\<And>x y ys z. (x, y) \<in> r \<Longrightarrow> rpath r y ys z \<Longrightarrow> P y ys z \<Longrightarrow> P x (y # ys) z" |
|
208 shows "P x1 x2 x3" |
|
209 using assms[unfolded rpath_def] |
|
210 by (induct, auto simp:pred_of_def rpath_def) |
|
211 |
|
212 lemma rpathE: |
|
213 assumes "rpath r x xs y" |
|
214 obtains (base) "y = x" "xs = []" |
|
215 | (step) z zs where "(x, z) \<in> r" "rpath r z zs y" "xs = z#zs" |
|
216 using assms |
|
217 by (induct, auto) |
|
218 |
|
219 text {* Introduction rule for empty path *} |
|
220 lemma rbaseI [intro!]: |
|
221 assumes "x = y" |
|
222 shows "rpath r x [] y" |
|
223 by (unfold rpath_def assms, |
|
224 rule Transitive_Closure_Table.rtrancl_path.base) |
|
225 |
|
226 text {* Introduction rule for non-empty path *} |
|
227 lemma rstepI [intro!]: |
|
228 assumes "(x, y) \<in> r" |
|
229 and "rpath r y ys z" |
|
230 shows "rpath r x (y#ys) z" |
|
231 proof(unfold rpath_def, rule Transitive_Closure_Table.rtrancl_path.step) |
|
232 from assms(1) show "pred_of r x y" by (auto simp:pred_of_def) |
|
233 next |
|
234 from assms(2) show "rtrancl_path (pred_of r) y ys z" |
|
235 by (auto simp:pred_of_def rpath_def) |
|
236 qed |
|
237 |
|
238 text {* Introduction rule for @{text "@"}-path *} |
|
239 lemma rpath_appendI [intro]: |
|
240 assumes "rpath r x xs a" and "rpath r a ys y" |
|
241 shows "rpath r x (xs @ ys) y" |
|
242 using assms |
|
243 by (unfold rpath_def, auto intro:rtrancl_path_trans) |
|
244 |
|
245 text {* Elimination rule for empty path *} |
|
246 |
|
247 lemma rpath_cases [cases pred:rpath]: |
|
248 assumes "rpath r a1 a2 a3" |
|
249 obtains (rbase) "a1 = a3" and "a2 = []" |
|
250 | (rstep) y :: "'a" and ys :: "'a list" |
|
251 where "(a1, y) \<in> r" and "a2 = y # ys" and "rpath r y ys a3" |
|
252 using assms [unfolded rpath_def] |
|
253 by (cases, auto simp:rpath_def pred_of_def) |
|
254 |
|
255 lemma rpath_nilE [elim!, cases pred:rpath]: |
|
256 assumes "rpath r x [] y" |
|
257 obtains "y = x" |
|
258 using assms[unfolded rpath_def] by auto |
|
259 |
|
260 -- {* This is a auxiliary lemmas used only in the proof of @{text "rpath_nnl_lastE"} *} |
|
261 lemma rpath_nnl_last: |
|
262 assumes "rtrancl_path r x xs y" |
|
263 and "xs \<noteq> []" |
|
264 obtains xs' where "xs = xs'@[y]" |
|
265 proof - |
|
266 from append_butlast_last_id[OF `xs \<noteq> []`, symmetric] |
|
267 obtain xs' y' where eq_xs: "xs = (xs' @ y' # [])" by simp |
|
268 with assms(1) |
|
269 have "rtrancl_path r x ... y" by simp |
|
270 hence "y = y'" by (rule rtrancl_path_appendE, auto) |
|
271 with eq_xs have "xs = xs'@[y]" by simp |
|
272 from that[OF this] show ?thesis . |
|
273 qed |
|
274 |
|
275 text {* |
|
276 Elimination rule for non-empty paths constructed with @{text "#"}. |
|
277 *} |
|
278 |
|
279 lemma rpath_ConsE [elim!, cases pred:rpath]: |
|
280 assumes "rpath r x (y # ys) x2" |
|
281 obtains (rstep) "(x, y) \<in> r" and "rpath r y ys x2" |
|
282 using assms[unfolded rpath_def] |
|
283 by (cases, auto simp:rpath_def pred_of_def) |
|
284 |
|
285 text {* |
|
286 Elimination rule for non-empty path, where the destination node |
|
287 @{text "y"} is shown to be at the end of the path. |
|
288 *} |
|
289 lemma rpath_nnl_lastE: |
|
290 assumes "rpath r x xs y" |
|
291 and "xs \<noteq> []" |
|
292 obtains xs' where "xs = xs'@[y]" |
|
293 using assms[unfolded rpath_def] |
|
294 by (rule rpath_nnl_last, auto) |
|
295 |
|
296 text {* Other elimination rules of @{text "rpath"} *} |
|
297 |
|
298 lemma rpath_appendE: |
|
299 assumes "rpath r x (xs @ [a] @ ys) y" |
|
300 obtains "rpath r x (xs @ [a]) a" and "rpath r a ys y" |
|
301 using rtrancl_path_appendE[OF assms[unfolded rpath_def, simplified], folded rpath_def] |
|
302 by auto |
|
303 |
|
304 lemma rpath_subE: |
|
305 assumes "rpath r x (xs @ [a] @ ys @ [b] @ zs) y" |
|
306 obtains "rpath r x (xs @ [a]) a" and "rpath r a (ys @ [b]) b" and "rpath r b zs y" |
|
307 using assms |
|
308 by (elim rpath_appendE, auto) |
|
309 |
|
310 text {* Every path has a unique end point. *} |
|
311 lemma rpath_dest_eq: |
|
312 assumes "rpath r x xs x1" |
|
313 and "rpath r x xs x2" |
|
314 shows "x1 = x2" |
|
315 using assms |
|
316 by (induct, auto) |
|
317 |
|
318 subsubsection {* Properites of @{text "edges_on"} *} |
|
319 |
|
320 lemma edges_on_unfold: |
|
321 "edges_on (a # b # xs) = {(a, b)} \<union> edges_on (b # xs)" (is "?L = ?R") |
|
322 proof - |
|
323 { fix c d |
|
324 assume "(c, d) \<in> ?L" |
|
325 then obtain l1 l2 where h: "(a # b # xs) = l1 @ [c, d] @ l2" |
|
326 by (auto simp:edges_on_def) |
|
327 have "(c, d) \<in> ?R" |
|
328 proof(cases "l1") |
|
329 case Nil |
|
330 with h have "(c, d) = (a, b)" by auto |
|
331 thus ?thesis by auto |
|
332 next |
|
333 case (Cons e es) |
|
334 from h[unfolded this] have "b#xs = es@[c, d]@l2" by auto |
|
335 thus ?thesis by (auto simp:edges_on_def) |
|
336 qed |
|
337 } moreover |
|
338 { fix c d |
|
339 assume "(c, d) \<in> ?R" |
|
340 moreover have "(a, b) \<in> ?L" |
|
341 proof - |
|
342 have "(a # b # xs) = []@[a,b]@xs" by simp |
|
343 hence "\<exists> l1 l2. (a # b # xs) = l1@[a,b]@l2" by auto |
|
344 thus ?thesis by (unfold edges_on_def, simp) |
|
345 qed |
|
346 moreover { |
|
347 assume "(c, d) \<in> edges_on (b#xs)" |
|
348 then obtain l1 l2 where "b#xs = l1@[c, d]@l2" by (unfold edges_on_def, auto) |
|
349 hence "a#b#xs = (a#l1)@[c,d]@l2" by simp |
|
350 hence "\<exists> l1 l2. (a # b # xs) = l1@[c,d]@l2" by metis |
|
351 hence "(c,d) \<in> ?L" by (unfold edges_on_def, simp) |
|
352 } |
|
353 ultimately have "(c, d) \<in> ?L" by auto |
|
354 } ultimately show ?thesis by auto |
|
355 qed |
|
356 |
|
357 lemma edges_on_len: |
|
358 assumes "(a,b) \<in> edges_on l" |
|
359 shows "length l \<ge> 2" |
|
360 using assms |
|
361 by (unfold edges_on_def, auto) |
|
362 |
|
363 text {* Elimination of @{text "edges_on"} for non-empty path *} |
|
364 |
|
365 lemma edges_on_consE [elim, cases set:edges_on]: |
|
366 assumes "(a,b) \<in> edges_on (x#xs)" |
|
367 obtains (head) xs' where "x = a" and "xs = b#xs'" |
|
368 | (tail) "(a,b) \<in> edges_on xs" |
|
369 proof - |
|
370 from assms obtain l1 l2 |
|
371 where h: "(x#xs) = l1 @ [a,b] @ l2" by (unfold edges_on_def, blast) |
|
372 have "(\<exists> xs'. x = a \<and> xs = b#xs') \<or> ((a,b) \<in> edges_on xs)" |
|
373 proof(cases "l1") |
|
374 case Nil with h |
|
375 show ?thesis by auto |
|
376 next |
|
377 case (Cons e el) |
|
378 from h[unfolded this] |
|
379 have "xs = el @ [a,b] @ l2" by auto |
|
380 thus ?thesis |
|
381 by (unfold edges_on_def, auto) |
|
382 qed |
|
383 thus ?thesis |
|
384 proof |
|
385 assume "(\<exists>xs'. x = a \<and> xs = b # xs')" |
|
386 then obtain xs' where "x = a" "xs = b#xs'" by blast |
|
387 from that(1)[OF this] show ?thesis . |
|
388 next |
|
389 assume "(a, b) \<in> edges_on xs" |
|
390 from that(2)[OF this] show ?thesis . |
|
391 qed |
|
392 qed |
|
393 |
|
394 text {* |
|
395 Every edges on the path is a graph edges: |
|
396 *} |
|
397 lemma rpath_edges_on: |
|
398 assumes "rpath r x xs y" |
|
399 shows "(edges_on (x#xs)) \<subseteq> r" |
|
400 using assms |
|
401 proof(induct arbitrary:y) |
|
402 case (rbase x) |
|
403 thus ?case by (unfold edges_on_def, auto) |
|
404 next |
|
405 case (rstep x y ys z) |
|
406 show ?case |
|
407 proof - |
|
408 { fix a b |
|
409 assume "(a, b) \<in> edges_on (x # y # ys)" |
|
410 hence "(a, b) \<in> r" by (cases, insert rstep, auto) |
|
411 } thus ?thesis by auto |
|
412 qed |
|
413 qed |
|
414 |
|
415 text {* @{text "edges_on"} is mono with respect to @{text "#"}-operation: *} |
|
416 lemma edges_on_Cons_mono: |
|
417 shows "edges_on xs \<subseteq> edges_on (x#xs)" |
|
418 proof - |
|
419 { fix a b |
|
420 assume "(a, b) \<in> edges_on xs" |
|
421 then obtain l1 l2 where "xs = l1 @ [a,b] @ l2" |
|
422 by (auto simp:edges_on_def) |
|
423 hence "x # xs = (x#l1) @ [a, b] @ l2" by auto |
|
424 hence "(a, b) \<in> edges_on (x#xs)" |
|
425 by (unfold edges_on_def, blast) |
|
426 } thus ?thesis by auto |
|
427 qed |
|
428 |
|
429 text {* |
|
430 The following rule @{text "rpath_transfer"} is used to show |
|
431 that one path is intact as long as all the edges on it are intact |
|
432 with the change of graph. |
|
433 |
|
434 If @{text "x#xs"} is path in graph @{text "r1"} and |
|
435 every edges along the path is also in @{text "r2"}, |
|
436 then @{text "x#xs"} is also a edge in graph @{text "r2"}: |
|
437 *} |
|
438 |
|
439 lemma rpath_transfer: |
|
440 assumes "rpath r1 x xs y" |
|
441 and "edges_on (x#xs) \<subseteq> r2" |
|
442 shows "rpath r2 x xs y" |
|
443 using assms |
|
444 proof(induct) |
|
445 case (rstep x y ys z) |
|
446 show ?case |
|
447 proof(rule rstepI) |
|
448 show "(x, y) \<in> r2" |
|
449 proof - |
|
450 have "(x, y) \<in> edges_on (x # y # ys)" |
|
451 by (unfold edges_on_def, auto) |
|
452 with rstep(4) show ?thesis by auto |
|
453 qed |
|
454 next |
|
455 show "rpath r2 y ys z" |
|
456 using rstep edges_on_Cons_mono[of "y#ys" "x"] by (auto) |
|
457 qed |
|
458 qed (unfold rpath_def, auto intro!:Transitive_Closure_Table.rtrancl_path.base) |
|
459 |
|
460 lemma edges_on_rpathI: |
|
461 assumes "edges_on (a#xs@[b]) \<subseteq> r" |
|
462 shows "rpath r a (xs@[b]) b" |
|
463 using assms |
|
464 proof(induct xs arbitrary: a b) |
|
465 case Nil |
|
466 moreover have "(a, b) \<in> edges_on (a # [] @ [b])" |
|
467 by (unfold edges_on_def, auto) |
|
468 ultimately have "(a, b) \<in> r" by auto |
|
469 thus ?case by auto |
|
470 next |
|
471 case (Cons x xs a b) |
|
472 from this(2) have "edges_on (x # xs @ [b]) \<subseteq> r" by (simp add:edges_on_unfold) |
|
473 from Cons(1)[OF this] have " rpath r x (xs @ [b]) b" . |
|
474 moreover from Cons(2) have "(a, x) \<in> r" by (auto simp:edges_on_unfold) |
|
475 ultimately show ?case by (auto) |
|
476 qed |
|
477 |
|
478 text {* |
|
479 The following lemma extracts the path from @{text "x"} to @{text "y"} |
|
480 from proposition @{text "(x, y) \<in> r^*"} |
|
481 *} |
|
482 lemma star_rpath: |
|
483 assumes "(x, y) \<in> r^*" |
|
484 obtains xs where "rpath r x xs y" |
|
485 proof - |
|
486 have "\<exists> xs. rpath r x xs y" |
|
487 proof(unfold rpath_def, rule iffD1[OF rtranclp_eq_rtrancl_path]) |
|
488 from assms |
|
489 show "(pred_of r)\<^sup>*\<^sup>* x y" |
|
490 apply (fold pred_of_star) |
|
491 by (auto simp:pred_of_def) |
|
492 qed |
|
493 from that and this show ?thesis by blast |
|
494 qed |
|
495 |
|
496 text {* |
|
497 The following lemma uses the path @{text "xs"} from @{text "x"} to @{text "y"} |
|
498 as a witness to show @{text "(x, y) \<in> r^*"}. |
|
499 *} |
|
500 lemma rpath_star: |
|
501 assumes "rpath r x xs y" |
|
502 shows "(x, y) \<in> r^*" |
|
503 proof - |
|
504 from iffD2[OF rtranclp_eq_rtrancl_path] and assms[unfolded rpath_def] |
|
505 have "(pred_of r)\<^sup>*\<^sup>* x y" by metis |
|
506 thus ?thesis by (simp add: pred_of_star star_2_pstar) |
|
507 qed |
|
508 |
|
509 lemma subtree_transfer: |
|
510 assumes "a \<in> subtree r1 a'" |
|
511 and "r1 \<subseteq> r2" |
|
512 shows "a \<in> subtree r2 a'" |
|
513 proof - |
|
514 from assms(1)[unfolded subtree_def] |
|
515 have "(a, a') \<in> r1^*" by auto |
|
516 from star_rpath[OF this] |
|
517 obtain xs where rp: "rpath r1 a xs a'" by blast |
|
518 hence "rpath r2 a xs a'" |
|
519 proof(rule rpath_transfer) |
|
520 from rpath_edges_on[OF rp] and assms(2) |
|
521 show "edges_on (a # xs) \<subseteq> r2" by simp |
|
522 qed |
|
523 from rpath_star[OF this] |
|
524 show ?thesis by (auto simp:subtree_def) |
|
525 qed |
|
526 |
|
527 lemma subtree_rev_transfer: |
|
528 assumes "a \<notin> subtree r2 a'" |
|
529 and "r1 \<subseteq> r2" |
|
530 shows "a \<notin> subtree r1 a'" |
|
531 using assms and subtree_transfer by metis |
|
532 |
|
533 text {* |
|
534 The following lemmas establishes a relation from paths in @{text "r"} |
|
535 to @{text "r^+"} relation. |
|
536 *} |
|
537 lemma rpath_plus: |
|
538 assumes "rpath r x xs y" |
|
539 and "xs \<noteq> []" |
|
540 shows "(x, y) \<in> r^+" |
|
541 proof - |
|
542 from assms(2) obtain e es where "xs = e#es" by (cases xs, auto) |
|
543 from assms(1)[unfolded this] |
|
544 show ?thesis |
|
545 proof(cases) |
|
546 case rstep |
|
547 show ?thesis |
|
548 proof - |
|
549 from rpath_star[OF rstep(2)] have "(e, y) \<in> r\<^sup>*" . |
|
550 with rstep(1) show "(x, y) \<in> r^+" by auto |
|
551 qed |
|
552 qed |
|
553 qed |
|
554 |
|
555 lemma plus_rpath: |
|
556 assumes "(x, y) \<in> r^+" |
|
557 obtains xs where "rpath r x xs y" and "xs \<noteq> []" |
|
558 proof - |
|
559 from assms |
|
560 show ?thesis |
|
561 proof(cases rule:converse_tranclE[consumes 1]) |
|
562 case 1 |
|
563 hence "rpath r x [y] y" by auto |
|
564 from that[OF this] show ?thesis by auto |
|
565 next |
|
566 case (2 z) |
|
567 from 2(2) have "(z, y) \<in> r^*" by auto |
|
568 from star_rpath[OF this] obtain xs where "rpath r z xs y" by auto |
|
569 from rstepI[OF 2(1) this] |
|
570 have "rpath r x (z # xs) y" . |
|
571 from that[OF this] show ?thesis by auto |
|
572 qed |
|
573 qed |
|
574 |
|
575 subsubsection {* Properties of @{text "subtree"} and @{term "ancestors"}*} |
|
576 |
|
577 lemma ancestors_subtreeI: |
|
578 assumes "b \<in> ancestors r a" |
|
579 shows "a \<in> subtree r b" |
|
580 using assms by (auto simp:subtree_def ancestors_def) |
|
581 |
|
582 lemma ancestors_Field: |
|
583 assumes "b \<in> ancestors r a" |
|
584 obtains "a \<in> Domain r" "b \<in> Range r" |
|
585 using assms |
|
586 apply (unfold ancestors_def, simp) |
|
587 by (metis Domain.DomainI Range.intros trancl_domain trancl_range) |
|
588 |
|
589 lemma subtreeE: |
|
590 assumes "a \<in> subtree r b" |
|
591 obtains "a = b" |
|
592 | "a \<noteq> b" and "b \<in> ancestors r a" |
|
593 proof - |
|
594 from assms have "(a, b) \<in> r^*" by (auto simp:subtree_def) |
|
595 from rtranclD[OF this] |
|
596 have " a = b \<or> a \<noteq> b \<and> (a, b) \<in> r\<^sup>+" . |
|
597 with that[unfolded ancestors_def] show ?thesis by auto |
|
598 qed |
|
599 |
|
600 lemma subtree_Field: |
|
601 assumes "a \<in> Field r" |
|
602 shows "subtree r a \<subseteq> Field r" |
|
603 by (metis Field_def UnI1 ancestors_Field assms subsetI subtreeE) |
|
604 |
|
605 lemma subtree_Field: |
|
606 "subtree r x \<subseteq> Field r \<union> {x}" |
|
607 proof |
|
608 fix y |
|
609 assume "y \<in> subtree r x" |
|
610 thus "y \<in> Field r \<union> {x}" |
|
611 proof(cases rule:subtreeE) |
|
612 case 1 |
|
613 thus ?thesis by auto |
|
614 next |
|
615 case 2 |
|
616 thus ?thesis apply (auto simp:ancestors_def) |
|
617 using Field_def tranclD by fastforce |
|
618 qed |
|
619 qed |
|
620 |
|
621 lemma subtree_ancestorsI: |
|
622 assumes "a \<in> subtree r b" |
|
623 and "a \<noteq> b" |
|
624 shows "b \<in> ancestors r a" |
|
625 using assms |
|
626 by (auto elim!:subtreeE) |
|
627 |
|
628 text {* |
|
629 @{text "subtree"} is mono with respect to the underlying graph. |
|
630 *} |
|
631 lemma subtree_mono: |
|
632 assumes "r1 \<subseteq> r2" |
|
633 shows "subtree r1 x \<subseteq> subtree r2 x" |
|
634 proof |
|
635 fix c |
|
636 assume "c \<in> subtree r1 x" |
|
637 hence "(c, x) \<in> r1^*" by (auto simp:subtree_def) |
|
638 from star_rpath[OF this] obtain xs |
|
639 where rp:"rpath r1 c xs x" by metis |
|
640 hence "rpath r2 c xs x" |
|
641 proof(rule rpath_transfer) |
|
642 from rpath_edges_on[OF rp] have "edges_on (c # xs) \<subseteq> r1" . |
|
643 with assms show "edges_on (c # xs) \<subseteq> r2" by auto |
|
644 qed |
|
645 thus "c \<in> subtree r2 x" |
|
646 by (rule rpath_star[elim_format], auto simp:subtree_def) |
|
647 qed |
|
648 |
|
649 text {* |
|
650 The following lemma characterizes the change of sub-tree of @{text "x"} |
|
651 with the removal of an outside edge @{text "(a,b)"}. |
|
652 |
|
653 Note that, according to lemma @{thm edges_in_refutation}, the assumption |
|
654 @{term "b \<notin> subtree r x"} amounts to saying @{text "(a, b)"} |
|
655 is outside the sub-tree of @{text "x"}. |
|
656 *} |
|
657 lemma subtree_del_outside: (* ddd *) |
|
658 assumes "b \<notin> subtree r x" |
|
659 shows "subtree (r - {(a, b)}) x = (subtree r x)" |
|
660 proof - |
|
661 { fix c |
|
662 assume "c \<in> (subtree r x)" |
|
663 hence "(c, x) \<in> r^*" by (auto simp:subtree_def) |
|
664 hence "c \<in> subtree (r - {(a, b)}) x" |
|
665 proof(rule star_rpath) |
|
666 fix xs |
|
667 assume rp: "rpath r c xs x" |
|
668 show ?thesis |
|
669 proof - |
|
670 from rp |
|
671 have "rpath (r - {(a, b)}) c xs x" |
|
672 proof(rule rpath_transfer) |
|
673 from rpath_edges_on[OF rp] have "edges_on (c # xs) \<subseteq> r" . |
|
674 moreover have "(a, b) \<notin> edges_on (c#xs)" |
|
675 proof |
|
676 assume "(a, b) \<in> edges_on (c # xs)" |
|
677 then obtain l1 l2 where h: "c#xs = l1@[a,b]@l2" by (auto simp:edges_on_def) |
|
678 hence "tl (c#xs) = tl (l1@[a,b]@l2)" by simp |
|
679 then obtain l1' where eq_xs_b: "xs = l1'@[b]@l2" by (cases l1, auto) |
|
680 from rp[unfolded this] |
|
681 show False |
|
682 proof(rule rpath_appendE) |
|
683 assume "rpath r b l2 x" |
|
684 thus ?thesis |
|
685 by(rule rpath_star[elim_format], insert assms(1), auto simp:subtree_def) |
|
686 qed |
|
687 qed |
|
688 ultimately show "edges_on (c # xs) \<subseteq> r - {(a,b)}" by auto |
|
689 qed |
|
690 thus ?thesis by (rule rpath_star[elim_format], auto simp:subtree_def) |
|
691 qed |
|
692 qed |
|
693 } moreover { |
|
694 fix c |
|
695 assume "c \<in> subtree (r - {(a, b)}) x" |
|
696 moreover have "... \<subseteq> (subtree r x)" by (rule subtree_mono, auto) |
|
697 ultimately have "c \<in> (subtree r x)" by auto |
|
698 } ultimately show ?thesis by auto |
|
699 qed |
|
700 |
|
701 (* ddd *) |
|
702 lemma subset_del_subtree_outside: (* ddd *) |
|
703 assumes "Range r' \<inter> subtree r x = {}" |
|
704 shows "subtree (r - r') x = (subtree r x)" |
|
705 proof - |
|
706 { fix c |
|
707 assume "c \<in> (subtree r x)" |
|
708 hence "(c, x) \<in> r^*" by (auto simp:subtree_def) |
|
709 hence "c \<in> subtree (r - r') x" |
|
710 proof(rule star_rpath) |
|
711 fix xs |
|
712 assume rp: "rpath r c xs x" |
|
713 show ?thesis |
|
714 proof - |
|
715 from rp |
|
716 have "rpath (r - r') c xs x" |
|
717 proof(rule rpath_transfer) |
|
718 from rpath_edges_on[OF rp] have "edges_on (c # xs) \<subseteq> r" . |
|
719 moreover { |
|
720 fix a b |
|
721 assume h: "(a, b) \<in> r'" |
|
722 have "(a, b) \<notin> edges_on (c#xs)" |
|
723 proof |
|
724 assume "(a, b) \<in> edges_on (c # xs)" |
|
725 then obtain l1 l2 where "c#xs = (l1@[a])@[b]@l2" by (auto simp:edges_on_def) |
|
726 hence "tl (c#xs) = tl (l1@[a,b]@l2)" by simp |
|
727 then obtain l1' where eq_xs_b: "xs = l1'@[b]@l2" by (cases l1, auto) |
|
728 from rp[unfolded this] |
|
729 show False |
|
730 proof(rule rpath_appendE) |
|
731 assume "rpath r b l2 x" |
|
732 from rpath_star[OF this] |
|
733 have "b \<in> subtree r x" by (auto simp:subtree_def) |
|
734 with assms (1) and h show ?thesis by (auto) |
|
735 qed |
|
736 qed |
|
737 } ultimately show "edges_on (c # xs) \<subseteq> r - r'" by auto |
|
738 qed |
|
739 thus ?thesis by (rule rpath_star[elim_format], auto simp:subtree_def) |
|
740 qed |
|
741 qed |
|
742 } moreover { |
|
743 fix c |
|
744 assume "c \<in> subtree (r - r') x" |
|
745 moreover have "... \<subseteq> (subtree r x)" by (rule subtree_mono, auto) |
|
746 ultimately have "c \<in> (subtree r x)" by auto |
|
747 } ultimately show ?thesis by auto |
|
748 qed |
|
749 |
|
750 lemma subtree_insert_ext: |
|
751 assumes "b \<in> subtree r x" |
|
752 shows "subtree (r \<union> {(a, b)}) x = (subtree r x) \<union> (subtree r a)" |
|
753 using assms by (auto simp:subtree_def rtrancl_insert) |
|
754 |
|
755 lemma subtree_insert_next: |
|
756 assumes "b \<notin> subtree r x" |
|
757 shows "subtree (r \<union> {(a, b)}) x = (subtree r x)" |
|
758 using assms |
|
759 by (auto simp:subtree_def rtrancl_insert) |
|
760 |
|
761 lemma set_add_rootI: |
|
762 assumes "root r a" |
|
763 and "a \<notin> Domain r1" |
|
764 shows "root (r \<union> r1) a" |
|
765 proof - |
|
766 let ?r = "r \<union> r1" |
|
767 { fix a' |
|
768 assume "a' \<in> ancestors ?r a" |
|
769 hence "(a, a') \<in> ?r^+" by (auto simp:ancestors_def) |
|
770 from tranclD[OF this] obtain z where "(a, z) \<in> ?r" by auto |
|
771 moreover have "(a, z) \<notin> r" |
|
772 proof |
|
773 assume "(a, z) \<in> r" |
|
774 with assms(1) show False |
|
775 by (auto simp:root_def ancestors_def) |
|
776 qed |
|
777 ultimately have "(a, z) \<in> r1" by auto |
|
778 with assms(2) |
|
779 have False by (auto) |
|
780 } thus ?thesis by (auto simp:root_def) |
|
781 qed |
|
782 |
|
783 lemma ancestors_mono: |
|
784 assumes "r1 \<subseteq> r2" |
|
785 shows "ancestors r1 x \<subseteq> ancestors r2 x" |
|
786 proof |
|
787 fix a |
|
788 assume "a \<in> ancestors r1 x" |
|
789 hence "(x, a) \<in> r1^+" by (auto simp:ancestors_def) |
|
790 from plus_rpath[OF this] obtain xs where |
|
791 h: "rpath r1 x xs a" "xs \<noteq> []" . |
|
792 have "rpath r2 x xs a" |
|
793 proof(rule rpath_transfer[OF h(1)]) |
|
794 from rpath_edges_on[OF h(1)] and assms |
|
795 show "edges_on (x # xs) \<subseteq> r2" by auto |
|
796 qed |
|
797 from rpath_plus[OF this h(2)] |
|
798 show "a \<in> ancestors r2 x" by (auto simp:ancestors_def) |
|
799 qed |
|
800 |
|
801 lemma subtree_refute: |
|
802 assumes "x \<notin> ancestors r y" |
|
803 and "x \<noteq> y" |
|
804 shows "y \<notin> subtree r x" |
|
805 proof |
|
806 assume "y \<in> subtree r x" |
|
807 thus False |
|
808 by(elim subtreeE, insert assms, auto) |
|
809 qed |
|
810 |
|
811 subsubsection {* Properties about relational trees *} |
|
812 |
|
813 context rtree |
|
814 begin |
|
815 |
|
816 lemma ancestors_headE: |
|
817 assumes "c \<in> ancestors r a" |
|
818 assumes "(a, b) \<in> r" |
|
819 obtains "b = c" |
|
820 | "c \<in> ancestors r b" |
|
821 proof - |
|
822 from assms(1) |
|
823 have "(a, c) \<in> r^+" by (auto simp:ancestors_def) |
|
824 hence "b = c \<or> c \<in> ancestors r b" |
|
825 proof(cases rule:converse_tranclE[consumes 1]) |
|
826 case 1 |
|
827 with assms(2) and sgv have "b = c" by (auto simp:single_valued_def) |
|
828 thus ?thesis by auto |
|
829 next |
|
830 case (2 y) |
|
831 from 2(1) and assms(2) and sgv have "y = b" by (auto simp:single_valued_def) |
|
832 from 2(2)[unfolded this] have "c \<in> ancestors r b" by (auto simp:ancestors_def) |
|
833 thus ?thesis by auto |
|
834 qed |
|
835 with that show ?thesis by metis |
|
836 qed |
|
837 |
|
838 lemma ancestors_accum: |
|
839 assumes "(a, b) \<in> r" |
|
840 shows "ancestors r a = ancestors r b \<union> {b}" |
|
841 proof - |
|
842 { fix c |
|
843 assume "c \<in> ancestors r a" |
|
844 hence "(a, c) \<in> r^+" by (auto simp:ancestors_def) |
|
845 hence "c \<in> ancestors r b \<union> {b}" |
|
846 proof(cases rule:converse_tranclE[consumes 1]) |
|
847 case 1 |
|
848 with sgv assms have "c = b" by (unfold single_valued_def, auto) |
|
849 thus ?thesis by auto |
|
850 next |
|
851 case (2 c') |
|
852 with sgv assms have "c' = b" by (unfold single_valued_def, auto) |
|
853 from 2(2)[unfolded this] |
|
854 show ?thesis by (auto simp:ancestors_def) |
|
855 qed |
|
856 } moreover { |
|
857 fix c |
|
858 assume "c \<in> ancestors r b \<union> {b}" |
|
859 hence "c = b \<or> c \<in> ancestors r b" by auto |
|
860 hence "c \<in> ancestors r a" |
|
861 proof |
|
862 assume "c = b" |
|
863 from assms[folded this] |
|
864 show ?thesis by (auto simp:ancestors_def) |
|
865 next |
|
866 assume "c \<in> ancestors r b" |
|
867 with assms show ?thesis by (auto simp:ancestors_def) |
|
868 qed |
|
869 } ultimately show ?thesis by auto |
|
870 qed |
|
871 |
|
872 lemma rootI: |
|
873 assumes h: "\<And> x'. x' \<noteq> x \<Longrightarrow> x \<notin> subtree r' x'" |
|
874 and "r' \<subseteq> r" |
|
875 shows "root r' x" |
|
876 proof - |
|
877 from acyclic_subset[OF acl assms(2)] |
|
878 have acl': "acyclic r'" . |
|
879 { fix x' |
|
880 assume "x' \<in> ancestors r' x" |
|
881 hence h1: "(x, x') \<in> r'^+" by (auto simp:ancestors_def) |
|
882 have "x' \<noteq> x" |
|
883 proof |
|
884 assume eq_x: "x' = x" |
|
885 from h1[unfolded this] and acl' |
|
886 show False by (auto simp:acyclic_def) |
|
887 qed |
|
888 moreover from h1 have "x \<in> subtree r' x'" by (auto simp:subtree_def) |
|
889 ultimately have False using h by auto |
|
890 } thus ?thesis by (auto simp:root_def) |
|
891 qed |
|
892 |
|
893 lemma rpath_overlap_oneside: (* ddd *) |
|
894 assumes "rpath r x xs1 x1" |
|
895 and "rpath r x xs2 x2" |
|
896 and "length xs1 \<le> length xs2" |
|
897 obtains xs3 where "xs2 = xs1 @ xs3" |
|
898 proof(cases "xs1 = []") |
|
899 case True |
|
900 with that show ?thesis by auto |
|
901 next |
|
902 case False |
|
903 have "\<forall> i \<le> length xs1. take i xs1 = take i xs2" |
|
904 proof - |
|
905 { assume "\<not> (\<forall> i \<le> length xs1. take i xs1 = take i xs2)" |
|
906 then obtain i where "i \<le> length xs1 \<and> take i xs1 \<noteq> take i xs2" by auto |
|
907 from this(1) have "False" |
|
908 proof(rule index_minimize) |
|
909 fix j |
|
910 assume h1: "j \<le> length xs1 \<and> take j xs1 \<noteq> take j xs2" |
|
911 and h2: " \<forall>k<j. \<not> (k \<le> length xs1 \<and> take k xs1 \<noteq> take k xs2)" |
|
912 -- {* @{text "j - 1"} is the branch point between @{text "xs1"} and @{text "xs2"} *} |
|
913 let ?idx = "j - 1" |
|
914 -- {* A number of inequalities concerning @{text "j - 1"} are derived first *} |
|
915 have lt_i: "?idx < length xs1" using False h1 |
|
916 by (metis Suc_diff_1 le_neq_implies_less length_greater_0_conv lessI less_imp_diff_less) |
|
917 have lt_i': "?idx < length xs2" using lt_i and assms(3) by auto |
|
918 have lt_j: "?idx < j" using h1 by (cases j, auto) |
|
919 -- {* From thesis inequalities, a number of equations concerning @{text "xs1"} |
|
920 and @{text "xs2"} are derived *} |
|
921 have eq_take: "take ?idx xs1 = take ?idx xs2" |
|
922 using h2[rule_format, OF lt_j] and h1 by auto |
|
923 have eq_xs1: " xs1 = take ?idx xs1 @ xs1 ! (?idx) # drop (Suc (?idx)) xs1" |
|
924 using id_take_nth_drop[OF lt_i] . |
|
925 have eq_xs2: "xs2 = take ?idx xs2 @ xs2 ! (?idx) # drop (Suc (?idx)) xs2" |
|
926 using id_take_nth_drop[OF lt_i'] . |
|
927 -- {* The branch point along the path is finally pinpointed *} |
|
928 have neq_idx: "xs1!?idx \<noteq> xs2!?idx" |
|
929 proof - |
|
930 have "take j xs1 = take ?idx xs1 @ [xs1 ! ?idx]" |
|
931 using eq_xs1 Suc_diff_1 lt_i lt_j take_Suc_conv_app_nth by fastforce |
|
932 moreover have eq_tk2: "take j xs2 = take ?idx xs2 @ [xs2 ! ?idx]" |
|
933 using Suc_diff_1 lt_i' lt_j take_Suc_conv_app_nth by fastforce |
|
934 ultimately show ?thesis using eq_take h1 by auto |
|
935 qed |
|
936 show ?thesis |
|
937 proof(cases " take (j - 1) xs1 = []") |
|
938 case True |
|
939 have "(x, xs1!?idx) \<in> r" |
|
940 proof - |
|
941 from eq_xs1[unfolded True, simplified, symmetric] assms(1) |
|
942 have "rpath r x ( xs1 ! ?idx # drop (Suc ?idx) xs1) x1" by simp |
|
943 from this[unfolded rpath_def] |
|
944 show ?thesis by (auto simp:pred_of_def) |
|
945 qed |
|
946 moreover have "(x, xs2!?idx) \<in> r" |
|
947 proof - |
|
948 from eq_xs2[folded eq_take, unfolded True, simplified, symmetric] assms(2) |
|
949 have "rpath r x ( xs2 ! ?idx # drop (Suc ?idx) xs2) x2" by simp |
|
950 from this[unfolded rpath_def] |
|
951 show ?thesis by (auto simp:pred_of_def) |
|
952 qed |
|
953 ultimately show ?thesis using neq_idx sgv[unfolded single_valued_def] by metis |
|
954 next |
|
955 case False |
|
956 then obtain e es where eq_es: "take ?idx xs1 = es@[e]" |
|
957 using rev_exhaust by blast |
|
958 have "(e, xs1!?idx) \<in> r" |
|
959 proof - |
|
960 from eq_xs1[unfolded eq_es] |
|
961 have "xs1 = es@[e, xs1!?idx]@drop (Suc ?idx) xs1" by simp |
|
962 hence "(e, xs1!?idx) \<in> edges_on xs1" by (simp add:edges_on_def, metis) |
|
963 with rpath_edges_on[OF assms(1)] edges_on_Cons_mono[of xs1 x] |
|
964 show ?thesis by auto |
|
965 qed moreover have "(e, xs2!?idx) \<in> r" |
|
966 proof - |
|
967 from eq_xs2[folded eq_take, unfolded eq_es] |
|
968 have "xs2 = es@[e, xs2!?idx]@drop (Suc ?idx) xs2" by simp |
|
969 hence "(e, xs2!?idx) \<in> edges_on xs2" by (simp add:edges_on_def, metis) |
|
970 with rpath_edges_on[OF assms(2)] edges_on_Cons_mono[of xs2 x] |
|
971 show ?thesis by auto |
|
972 qed |
|
973 ultimately show ?thesis |
|
974 using sgv[unfolded single_valued_def] neq_idx by metis |
|
975 qed |
|
976 qed |
|
977 } thus ?thesis by auto |
|
978 qed |
|
979 from this[rule_format, of "length xs1"] |
|
980 have "take (length xs1) xs1 = take (length xs1) xs2" by simp |
|
981 moreover have "xs2 = take (length xs1) xs2 @ drop (length xs1) xs2" by simp |
|
982 ultimately have "xs2 = xs1 @ drop (length xs1) xs2" by auto |
|
983 from that[OF this] show ?thesis . |
|
984 qed |
|
985 |
|
986 lemma rpath_overlap [consumes 2, cases pred:rpath]: |
|
987 assumes "rpath r x xs1 x1" |
|
988 and "rpath r x xs2 x2" |
|
989 obtains (less_1) xs3 where "xs2 = xs1 @ xs3" |
|
990 | (less_2) xs3 where "xs1 = xs2 @ xs3" |
|
991 proof - |
|
992 have "length xs1 \<le> length xs2 \<or> length xs2 \<le> length xs1" by auto |
|
993 with assms rpath_overlap_oneside that show ?thesis by metis |
|
994 qed |
|
995 |
|
996 text {* |
|
997 As a corollary of @{thm "rpath_overlap_oneside"}, |
|
998 the following two lemmas gives one important property of relation tree, |
|
999 i.e. there is at most one path between any two nodes. |
|
1000 Similar to the proof of @{thm rpath_overlap}, we starts with |
|
1001 the one side version first. |
|
1002 *} |
|
1003 |
|
1004 lemma rpath_unique_oneside: |
|
1005 assumes "rpath r x xs1 y" |
|
1006 and "rpath r x xs2 y" |
|
1007 and "length xs1 \<le> length xs2" |
|
1008 shows "xs1 = xs2" |
|
1009 proof - |
|
1010 from rpath_overlap_oneside[OF assms] |
|
1011 obtain xs3 where less_1: "xs2 = xs1 @ xs3" by blast |
|
1012 show ?thesis |
|
1013 proof(cases "xs3 = []") |
|
1014 case True |
|
1015 from less_1[unfolded this] show ?thesis by simp |
|
1016 next |
|
1017 case False |
|
1018 note FalseH = this |
|
1019 show ?thesis |
|
1020 proof(cases "xs1 = []") |
|
1021 case True |
|
1022 have "(x, x) \<in> r^+" |
|
1023 proof(rule rpath_plus) |
|
1024 from assms(1)[unfolded True] |
|
1025 have "y = x" by (cases rule:rpath_nilE, simp) |
|
1026 from assms(2)[unfolded this] show "rpath r x xs2 x" . |
|
1027 next |
|
1028 from less_1 and False show "xs2 \<noteq> []" by simp |
|
1029 qed |
|
1030 with acl show ?thesis by (unfold acyclic_def, auto) |
|
1031 next |
|
1032 case False |
|
1033 then obtain e es where eq_xs1: "xs1 = es@[e]" using rev_exhaust by auto |
|
1034 from assms(2)[unfolded less_1 this] |
|
1035 have "rpath r x (es @ [e] @ xs3) y" by simp |
|
1036 thus ?thesis |
|
1037 proof(cases rule:rpath_appendE) |
|
1038 case 1 |
|
1039 from rpath_dest_eq [OF 1(1)[folded eq_xs1] assms(1)] |
|
1040 have "e = y" . |
|
1041 from rpath_plus [OF 1(2)[unfolded this] FalseH] |
|
1042 have "(y, y) \<in> r^+" . |
|
1043 with acl show ?thesis by (unfold acyclic_def, auto) |
|
1044 qed |
|
1045 qed |
|
1046 qed |
|
1047 qed |
|
1048 |
|
1049 text {* |
|
1050 The following is the full version of path uniqueness. |
|
1051 *} |
|
1052 lemma rpath_unique: |
|
1053 assumes "rpath r x xs1 y" |
|
1054 and "rpath r x xs2 y" |
|
1055 shows "xs1 = xs2" |
|
1056 proof(cases "length xs1 \<le> length xs2") |
|
1057 case True |
|
1058 from rpath_unique_oneside[OF assms this] show ?thesis . |
|
1059 next |
|
1060 case False |
|
1061 hence "length xs2 \<le> length xs1" by simp |
|
1062 from rpath_unique_oneside[OF assms(2,1) this] |
|
1063 show ?thesis by simp |
|
1064 qed |
|
1065 |
|
1066 text {* |
|
1067 The following lemma shows that the `independence` relation is symmetric. |
|
1068 It is an obvious auxiliary lemma which will be used later. |
|
1069 *} |
|
1070 lemma sym_indep: "indep r x y \<Longrightarrow> indep r y x" |
|
1071 by (unfold indep_def, auto) |
|
1072 |
|
1073 text {* |
|
1074 This is another `obvious` lemma about trees, which says trees rooted at |
|
1075 independent nodes are disjoint. |
|
1076 *} |
|
1077 lemma subtree_disjoint: |
|
1078 assumes "indep r x y" |
|
1079 shows "subtree r x \<inter> subtree r y = {}" |
|
1080 proof - |
|
1081 { fix z x y xs1 xs2 xs3 |
|
1082 assume ind: "indep r x y" |
|
1083 and rp1: "rpath r z xs1 x" |
|
1084 and rp2: "rpath r z xs2 y" |
|
1085 and h: "xs2 = xs1 @ xs3" |
|
1086 have False |
|
1087 proof(cases "xs1 = []") |
|
1088 case True |
|
1089 from rp1[unfolded this] have "x = z" by auto |
|
1090 from rp2[folded this] rpath_star ind[unfolded indep_def] |
|
1091 show ?thesis by metis |
|
1092 next |
|
1093 case False |
|
1094 then obtain e es where eq_xs1: "xs1 = es@[e]" using rev_exhaust by blast |
|
1095 from rp2[unfolded h this] |
|
1096 have "rpath r z (es @ [e] @ xs3) y" by simp |
|
1097 thus ?thesis |
|
1098 proof(cases rule:rpath_appendE) |
|
1099 case 1 |
|
1100 have "e = x" using 1(1)[folded eq_xs1] rp1 rpath_dest_eq by metis |
|
1101 from rpath_star[OF 1(2)[unfolded this]] ind[unfolded indep_def] |
|
1102 show ?thesis by auto |
|
1103 qed |
|
1104 qed |
|
1105 } note my_rule = this |
|
1106 { fix z |
|
1107 assume h: "z \<in> subtree r x" "z \<in> subtree r y" |
|
1108 from h(1) have "(z, x) \<in> r^*" by (unfold subtree_def, auto) |
|
1109 then obtain xs1 where rp1: "rpath r z xs1 x" using star_rpath by metis |
|
1110 from h(2) have "(z, y) \<in> r^*" by (unfold subtree_def, auto) |
|
1111 then obtain xs2 where rp2: "rpath r z xs2 y" using star_rpath by metis |
|
1112 from rp1 rp2 |
|
1113 have False |
|
1114 by (cases, insert my_rule[OF sym_indep[OF assms(1)] rp2 rp1] |
|
1115 my_rule[OF assms(1) rp1 rp2], auto) |
|
1116 } thus ?thesis by auto |
|
1117 qed |
|
1118 |
|
1119 text {* |
|
1120 The following lemma @{text "subtree_del"} characterizes the change of sub-tree of |
|
1121 @{text "x"} with the removal of an inside edge @{text "(a, b)"}. |
|
1122 Note that, the case for the removal of an outside edge has already been dealt with |
|
1123 in lemma @{text "subtree_del_outside"}). |
|
1124 |
|
1125 This lemma is underpinned by the following two `obvious` facts: |
|
1126 \begin{enumearte} |
|
1127 \item |
|
1128 In graph @{text "r"}, for an inside edge @{text "(a,b) \<in> edges_in r x"}, |
|
1129 every node @{text "c"} in the sub-tree of @{text "a"} has a path |
|
1130 which goes first from @{text "c"} to @{text "a"}, then through edge @{text "(a, b)"}, and |
|
1131 finally reaches @{text "x"}. By the uniqueness of path in a tree, |
|
1132 all paths from sub-tree of @{text "a"} to @{text "x"} are such constructed, therefore |
|
1133 must go through @{text "(a, b)"}. The consequence is: with the removal of @{text "(a,b)"}, |
|
1134 all such paths will be broken. |
|
1135 |
|
1136 \item |
|
1137 On the other hand, all paths not originate from within the sub-tree of @{text "a"} |
|
1138 will not be affected by the removal of edge @{text "(a, b)"}. |
|
1139 The reason is simple: if the path is affected by the removal, it must |
|
1140 contain @{text "(a, b)"}, then it must originate from within the sub-tree of @{text "a"}. |
|
1141 \end{enumearte} |
|
1142 *} |
|
1143 |
|
1144 lemma subtree_del_inside: (* ddd *) |
|
1145 assumes "(a,b) \<in> edges_in r x" |
|
1146 shows "subtree (r - {(a, b)}) x = (subtree r x) - subtree r a" |
|
1147 proof - |
|
1148 from assms have asm: "b \<in> subtree r x" "(a, b) \<in> r" by (auto simp:edges_in_def) |
|
1149 -- {* The proof follows a common pattern to prove the equality of sets. *} |
|
1150 { -- {* The `left to right` direction. |
|
1151 *} |
|
1152 fix c |
|
1153 -- {* Assuming @{text "c"} is inside the sub-tree of @{text "x"} in the reduced graph *} |
|
1154 assume h: "c \<in> subtree (r - {(a, b)}) x" |
|
1155 -- {* We are going to show that @{text "c"} can not be in the sub-tree of @{text "a"} in |
|
1156 the original graph. *} |
|
1157 -- {* In other words, all nodes inside the sub-tree of @{text "a"} in the original |
|
1158 graph will be removed from the sub-tree of @{text "x"} in the reduced graph. *} |
|
1159 -- {* The reason, as analyzed before, is that all paths from within the |
|
1160 sub-tree of @{text "a"} are broken with the removal of edge @{text "(a,b)"}. |
|
1161 *} |
|
1162 have "c \<in> (subtree r x) - subtree r a" |
|
1163 proof - |
|
1164 let ?r' = "r - {(a, b)}" -- {* The reduced graph is abbreviated as @{text "?r'"} *} |
|
1165 from h have "(c, x) \<in> ?r'^*" by (auto simp:subtree_def) |
|
1166 -- {* Extract from the reduced graph the path @{text "xs"} from @{text "c"} to @{text "x"}. *} |
|
1167 then obtain xs where rp0: "rpath ?r' c xs x" by (rule star_rpath, auto) |
|
1168 -- {* It is easy to show @{text "xs"} is also a path in the original graph *} |
|
1169 hence rp1: "rpath r c xs x" |
|
1170 proof(rule rpath_transfer) |
|
1171 from rpath_edges_on[OF rp0] |
|
1172 show "edges_on (c # xs) \<subseteq> r" by auto |
|
1173 qed |
|
1174 -- {* @{text "xs"} is used as the witness to show that @{text "c"} |
|
1175 in the sub-tree of @{text "x"} in the original graph. *} |
|
1176 hence "c \<in> subtree r x" |
|
1177 by (rule rpath_star[elim_format], auto simp:subtree_def) |
|
1178 -- {* The next step is to show that @{text "c"} can not be in the sub-tree of @{text "a"} |
|
1179 in the original graph. *} |
|
1180 -- {* We need to use the fact that all paths originate from within sub-tree of @{text "a"} |
|
1181 are broken. *} |
|
1182 moreover have "c \<notin> subtree r a" |
|
1183 proof |
|
1184 -- {* Proof by contradiction, suppose otherwise *} |
|
1185 assume otherwise: "c \<in> subtree r a" |
|
1186 -- {* Then there is a path in original graph leading from @{text "c"} to @{text "a"} *} |
|
1187 obtain xs1 where rp_c: "rpath r c xs1 a" |
|
1188 proof - |
|
1189 from otherwise have "(c, a) \<in> r^*" by (auto simp:subtree_def) |
|
1190 thus ?thesis by (rule star_rpath, auto intro!:that) |
|
1191 qed |
|
1192 -- {* Starting from this path, we are going to construct a fictional |
|
1193 path from @{text "c"} to @{text "x"}, which, as explained before, |
|
1194 is broken, so that contradiction can be derived. *} |
|
1195 -- {* First, there is a path from @{text "b"} to @{text "x"} *} |
|
1196 obtain ys where rp_b: "rpath r b ys x" |
|
1197 proof - |
|
1198 from asm have "(b, x) \<in> r^*" by (auto simp:subtree_def) |
|
1199 thus ?thesis by (rule star_rpath, auto intro!:that) |
|
1200 qed |
|
1201 -- {* The paths @{text "xs1"} and @{text "ys"} can be |
|
1202 tied together using @{text "(a,b)"} to form a path |
|
1203 from @{text "c"} to @{text "x"}: *} |
|
1204 have "rpath r c (xs1 @ b # ys) x" |
|
1205 proof - |
|
1206 from rstepI[OF asm(2) rp_b] have "rpath r a (b # ys) x" . |
|
1207 from rpath_appendI[OF rp_c this] |
|
1208 show ?thesis . |
|
1209 qed |
|
1210 -- {* By the uniqueness of path between two nodes of a tree, we have: *} |
|
1211 from rpath_unique[OF rp1 this] have eq_xs: "xs = xs1 @ b # ys" . |
|
1212 -- {* Contradiction can be derived from from this fictional path . *} |
|
1213 show False |
|
1214 proof - |
|
1215 -- {* It can be shown that @{term "(a,b)"} is on this fictional path. *} |
|
1216 have "(a, b) \<in> edges_on (c#xs)" |
|
1217 proof(cases "xs1 = []") |
|
1218 case True |
|
1219 from rp_c[unfolded this] have "rpath r c [] a" . |
|
1220 hence eq_c: "c = a" by (rule rpath_nilE, simp) |
|
1221 hence "c#xs = a#xs" by simp |
|
1222 from this and eq_xs have "c#xs = a # xs1 @ b # ys" by simp |
|
1223 from this[unfolded True] have "c#xs = []@[a,b]@ys" by simp |
|
1224 thus ?thesis by (auto simp:edges_on_def) |
|
1225 next |
|
1226 case False |
|
1227 from rpath_nnl_lastE[OF rp_c this] |
|
1228 obtain xs' where "xs1 = xs'@[a]" by auto |
|
1229 from eq_xs[unfolded this] have "c#xs = (c#xs')@[a,b]@ys" by simp |
|
1230 thus ?thesis by (unfold edges_on_def, blast) |
|
1231 qed |
|
1232 -- {* It can also be shown that @{term "(a,b)"} is not on this fictional path. *} |
|
1233 moreover have "(a, b) \<notin> edges_on (c#xs)" |
|
1234 using rpath_edges_on[OF rp0] by auto |
|
1235 -- {* Contradiction is thus derived. *} |
|
1236 ultimately show False by auto |
|
1237 qed |
|
1238 qed |
|
1239 ultimately show ?thesis by auto |
|
1240 qed |
|
1241 } moreover { |
|
1242 -- {* The `right to left` direction. |
|
1243 *} |
|
1244 fix c |
|
1245 -- {* Assuming that @{text "c"} is in the sub-tree of @{text "x"}, but |
|
1246 outside of the sub-tree of @{text "a"} in the original graph, *} |
|
1247 assume h: "c \<in> (subtree r x) - subtree r a" |
|
1248 -- {* we need to show that in the reduced graph, @{text "c"} is still in |
|
1249 the sub-tree of @{text "x"}. *} |
|
1250 have "c \<in> subtree (r - {(a, b)}) x" |
|
1251 proof - |
|
1252 -- {* The proof goes by showing that the path from @{text "c"} to @{text "x"} |
|
1253 in the original graph is not affected by the removal of @{text "(a,b)"}. |
|
1254 *} |
|
1255 from h have "(c, x) \<in> r^*" by (unfold subtree_def, auto) |
|
1256 -- {* Extract the path @{text "xs"} from @{text "c"} to @{text "x"} in the original graph. *} |
|
1257 from star_rpath[OF this] obtain xs where rp: "rpath r c xs x" by auto |
|
1258 -- {* Show that it is also a path in the reduced graph. *} |
|
1259 hence "rpath (r - {(a, b)}) c xs x" |
|
1260 -- {* The proof goes by using rule @{thm rpath_transfer} *} |
|
1261 proof(rule rpath_transfer) |
|
1262 -- {* We need to show all edges on the path are still in the reduced graph. *} |
|
1263 show "edges_on (c # xs) \<subseteq> r - {(a, b)}" |
|
1264 proof - |
|
1265 -- {* It is easy to show that all the edges are in the original graph. *} |
|
1266 from rpath_edges_on [OF rp] have " edges_on (c # xs) \<subseteq> r" . |
|
1267 -- {* The essential part is to show that @{text "(a, b)"} is not on the path. *} |
|
1268 moreover have "(a,b) \<notin> edges_on (c#xs)" |
|
1269 proof |
|
1270 -- {* Proof by contradiction, suppose otherwise: *} |
|
1271 assume otherwise: "(a, b) \<in> edges_on (c#xs)" |
|
1272 -- {* Then @{text "(a, b)"} is in the middle of the path. |
|
1273 with @{text "l1"} and @{text "l2"} be the nodes in |
|
1274 the front and rear respectively. *} |
|
1275 then obtain l1 l2 where eq_xs: |
|
1276 "c#xs = l1 @ [a, b] @ l2" by (unfold edges_on_def, blast) |
|
1277 -- {* From this, it can be shown that @{text "c"} is |
|
1278 in the sub-tree of @{text "a"} *} |
|
1279 have "c \<in> subtree r a" |
|
1280 proof(cases "l1 = []") |
|
1281 case True |
|
1282 -- {* If @{text "l1"} is null, it can be derived that @{text "c = a"}. *} |
|
1283 with eq_xs have "c = a" by auto |
|
1284 -- {* So, @{text "c"} is obviously in the sub-tree of @{text "a"}. *} |
|
1285 thus ?thesis by (unfold subtree_def, auto) |
|
1286 next |
|
1287 case False |
|
1288 -- {* When @{text "l1"} is not null, it must have a tail @{text "es"}: *} |
|
1289 then obtain e es where "l1 = e#es" by (cases l1, auto) |
|
1290 -- {* The relation of this tail with @{text "xs"} is derived: *} |
|
1291 with eq_xs have "xs = es@[a,b]@l2" by auto |
|
1292 -- {* From this, a path from @{text "c"} to @{text "a"} is made visible: *} |
|
1293 from rp[unfolded this] have "rpath r c (es @ [a] @ (b#l2)) x" by simp |
|
1294 thus ?thesis |
|
1295 proof(cases rule:rpath_appendE) |
|
1296 -- {* The path from @{text "c"} to @{text "a"} is extraced |
|
1297 using @{thm "rpath_appendE"}: *} |
|
1298 case 1 |
|
1299 from rpath_star[OF this(1)] |
|
1300 -- {* The extracted path servers as a witness that @{text "c"} is |
|
1301 in the sub-tree of @{text "a"}: *} |
|
1302 show ?thesis by (simp add:subtree_def) |
|
1303 qed |
|
1304 qed with h show False by auto |
|
1305 qed ultimately show ?thesis by auto |
|
1306 qed |
|
1307 qed |
|
1308 -- {* From , it is shown that @{text "c"} is in the sub-tree of @{text "x"} |
|
1309 inthe reduced graph. *} |
|
1310 from rpath_star[OF this] show ?thesis by (auto simp:subtree_def) |
|
1311 qed |
|
1312 } |
|
1313 -- {* The equality of sets is derived from the two directions just proved. *} |
|
1314 ultimately show ?thesis by auto |
|
1315 qed |
|
1316 |
|
1317 lemma set_del_rootI: |
|
1318 assumes "r1 \<subseteq> r" |
|
1319 and "a \<in> Domain r1" |
|
1320 shows "root (r - r1) a" |
|
1321 proof - |
|
1322 let ?r = "r - r1" |
|
1323 { fix a' |
|
1324 assume neq: "a' \<noteq> a" |
|
1325 have "a \<notin> subtree ?r a'" |
|
1326 proof |
|
1327 assume "a \<in> subtree ?r a'" |
|
1328 hence "(a, a') \<in> ?r^*" by (auto simp:subtree_def) |
|
1329 from star_rpath[OF this] obtain xs |
|
1330 where rp: "rpath ?r a xs a'" by auto |
|
1331 from rpathE[OF this] and neq |
|
1332 obtain z zs where h: "(a, z) \<in> ?r" "rpath ?r z zs a'" "xs = z#zs" by auto |
|
1333 from assms(2) obtain z' where z'_in: "(a, z') \<in> r1" by (auto simp:DomainE) |
|
1334 with assms(1) have "(a, z') \<in> r" by auto |
|
1335 moreover from h(1) have "(a, z) \<in> r" by simp |
|
1336 ultimately have "z' = z" using sgv by (auto simp:single_valued_def) |
|
1337 from z'_in[unfolded this] and h(1) show False by auto |
|
1338 qed |
|
1339 } thus ?thesis by (intro rootI, auto) |
|
1340 qed |
|
1341 |
|
1342 lemma edge_del_no_rootI: |
|
1343 assumes "(a, b) \<in> r" |
|
1344 shows "root (r - {(a, b)}) a" |
|
1345 by (rule set_del_rootI, insert assms, auto) |
|
1346 |
|
1347 lemma ancestors_children_unique: |
|
1348 assumes "z1 \<in> ancestors r x \<inter> children r y" |
|
1349 and "z2 \<in> ancestors r x \<inter> children r y" |
|
1350 shows "z1 = z2" |
|
1351 proof - |
|
1352 from assms have h: |
|
1353 "(x, z1) \<in> r^+" "(z1, y) \<in> r" |
|
1354 "(x, z2) \<in> r^+" "(z2, y) \<in> r" |
|
1355 by (auto simp:ancestors_def children_def) |
|
1356 |
|
1357 -- {* From this, a path containing @{text "z1"} is obtained. *} |
|
1358 from plus_rpath[OF h(1)] obtain xs1 |
|
1359 where h1: "rpath r x xs1 z1" "xs1 \<noteq> []" by auto |
|
1360 from rpath_nnl_lastE[OF this] obtain xs1' where eq_xs1: "xs1 = xs1' @ [z1]" |
|
1361 by auto |
|
1362 from h(2) have h2: "rpath r z1 [y] y" by auto |
|
1363 from rpath_appendI[OF h1(1) h2, unfolded eq_xs1] |
|
1364 have rp1: "rpath r x (xs1' @ [z1, y]) y" by simp |
|
1365 |
|
1366 -- {* Then, another path containing @{text "z2"} is obtained. *} |
|
1367 from plus_rpath[OF h(3)] obtain xs2 |
|
1368 where h3: "rpath r x xs2 z2" "xs2 \<noteq> []" by auto |
|
1369 from rpath_nnl_lastE[OF this] obtain xs2' where eq_xs2: "xs2 = xs2' @ [z2]" |
|
1370 by auto |
|
1371 from h(4) have h4: "rpath r z2 [y] y" by auto |
|
1372 from rpath_appendI[OF h3(1) h4, unfolded eq_xs2] |
|
1373 have "rpath r x (xs2' @ [z2, y]) y" by simp |
|
1374 |
|
1375 -- {* Finally @{text "z1 = z2"} is proved by uniqueness of path. *} |
|
1376 from rpath_unique[OF rp1 this] |
|
1377 have "xs1' @ [z1, y] = xs2' @ [z2, y]" . |
|
1378 thus ?thesis by auto |
|
1379 qed |
|
1380 |
|
1381 lemma ancestors_childrenE: |
|
1382 assumes "y \<in> ancestors r x" |
|
1383 obtains "x \<in> children r y" |
|
1384 | z where "z \<in> ancestors r x \<inter> children r y" |
|
1385 proof - |
|
1386 from assms(1) have "(x, y) \<in> r^+" by (auto simp:ancestors_def) |
|
1387 from tranclD2[OF this] obtain z where |
|
1388 h: "(x, z) \<in> r\<^sup>*" "(z, y) \<in> r" by auto |
|
1389 from h(1) |
|
1390 show ?thesis |
|
1391 proof(cases rule:rtranclE) |
|
1392 case base |
|
1393 from h(2)[folded this] have "x \<in> children r y" |
|
1394 by (auto simp:children_def) |
|
1395 thus ?thesis by (intro that, auto) |
|
1396 next |
|
1397 case (step u) |
|
1398 hence "z \<in> ancestors r x" by (auto simp:ancestors_def) |
|
1399 moreover from h(2) have "z \<in> children r y" |
|
1400 by (auto simp:children_def) |
|
1401 ultimately show ?thesis by (intro that, auto) |
|
1402 qed |
|
1403 qed |
|
1404 |
|
1405 |
|
1406 end (* of rtree *) |
|
1407 |
|
1408 lemma subtree_children: |
|
1409 "subtree r x = {x} \<union> (\<Union> (subtree r ` (children r x)))" (is "?L = ?R") |
|
1410 proof - |
|
1411 { fix z |
|
1412 assume "z \<in> ?L" |
|
1413 hence "z \<in> ?R" |
|
1414 proof(cases rule:subtreeE[consumes 1]) |
|
1415 case 2 |
|
1416 hence "(z, x) \<in> r^+" by (auto simp:ancestors_def) |
|
1417 thus ?thesis |
|
1418 proof(rule tranclE) |
|
1419 assume "(z, x) \<in> r" |
|
1420 hence "z \<in> children r x" by (unfold children_def, auto) |
|
1421 moreover have "z \<in> subtree r z" by (auto simp:subtree_def) |
|
1422 ultimately show ?thesis by auto |
|
1423 next |
|
1424 fix c |
|
1425 assume h: "(z, c) \<in> r\<^sup>+" "(c, x) \<in> r" |
|
1426 hence "c \<in> children r x" by (auto simp:children_def) |
|
1427 moreover from h have "z \<in> subtree r c" by (auto simp:subtree_def) |
|
1428 ultimately show ?thesis by auto |
|
1429 qed |
|
1430 qed auto |
|
1431 } moreover { |
|
1432 fix z |
|
1433 assume h: "z \<in> ?R" |
|
1434 have "x \<in> subtree r x" by (auto simp:subtree_def) |
|
1435 moreover { |
|
1436 assume "z \<in> \<Union>(subtree r ` children r x)" |
|
1437 then obtain y where "(y, x) \<in> r" "(z, y) \<in> r^*" |
|
1438 by (auto simp:subtree_def children_def) |
|
1439 hence "(z, x) \<in> r^*" by auto |
|
1440 hence "z \<in> ?L" by (auto simp:subtree_def) |
|
1441 } ultimately have "z \<in> ?L" using h by auto |
|
1442 } ultimately show ?thesis by auto |
|
1443 qed |
|
1444 |
|
1445 context fsubtree |
|
1446 begin |
|
1447 |
|
1448 lemma finite_subtree: |
|
1449 shows "finite (subtree r x)" |
|
1450 proof(induct rule:wf_induct[OF wf]) |
|
1451 case (1 x) |
|
1452 have "finite (\<Union>(subtree r ` children r x))" |
|
1453 proof(rule finite_Union) |
|
1454 show "finite (subtree r ` children r x)" |
|
1455 proof(cases "children r x = {}") |
|
1456 case True |
|
1457 thus ?thesis by auto |
|
1458 next |
|
1459 case False |
|
1460 hence "x \<in> Range r" by (auto simp:children_def) |
|
1461 from fb[rule_format, OF this] |
|
1462 have "finite (children r x)" . |
|
1463 thus ?thesis by (rule finite_imageI) |
|
1464 qed |
|
1465 next |
|
1466 fix M |
|
1467 assume "M \<in> subtree r ` children r x" |
|
1468 then obtain y where h: "y \<in> children r x" "M = subtree r y" by auto |
|
1469 hence "(y, x) \<in> r" by (auto simp:children_def) |
|
1470 from 1[rule_format, OF this, folded h(2)] |
|
1471 show "finite M" . |
|
1472 qed |
|
1473 thus ?case |
|
1474 by (unfold subtree_children finite_Un, auto) |
|
1475 qed |
|
1476 |
|
1477 end |
|
1478 |
|
1479 definition "pairself f = (\<lambda>(a, b). (f a, f b))" |
|
1480 |
|
1481 definition "rel_map f r = (pairself f ` r)" |
|
1482 |
|
1483 lemma rel_mapE: |
|
1484 assumes "(a, b) \<in> rel_map f r" |
|
1485 obtains c d |
|
1486 where "(c, d) \<in> r" "(a, b) = (f c, f d)" |
|
1487 using assms |
|
1488 by (unfold rel_map_def pairself_def, auto) |
|
1489 |
|
1490 lemma rel_mapI: |
|
1491 assumes "(a, b) \<in> r" |
|
1492 and "c = f a" |
|
1493 and "d = f b" |
|
1494 shows "(c, d) \<in> rel_map f r" |
|
1495 using assms |
|
1496 by (unfold rel_map_def pairself_def, auto) |
|
1497 |
|
1498 lemma map_appendE: |
|
1499 assumes "map f zs = xs @ ys" |
|
1500 obtains xs' ys' |
|
1501 where "zs = xs' @ ys'" "xs = map f xs'" "ys = map f ys'" |
|
1502 proof - |
|
1503 have "\<exists> xs' ys'. zs = xs' @ ys' \<and> xs = map f xs' \<and> ys = map f ys'" |
|
1504 using assms |
|
1505 proof(induct xs arbitrary:zs ys) |
|
1506 case (Nil zs ys) |
|
1507 thus ?case by auto |
|
1508 next |
|
1509 case (Cons x xs zs ys) |
|
1510 note h = this |
|
1511 show ?case |
|
1512 proof(cases zs) |
|
1513 case (Cons e es) |
|
1514 with h have eq_x: "map f es = xs @ ys" "x = f e" by auto |
|
1515 from h(1)[OF this(1)] |
|
1516 obtain xs' ys' where "es = xs' @ ys'" "xs = map f xs'" "ys = map f ys'" |
|
1517 by blast |
|
1518 with Cons eq_x |
|
1519 have "zs = (e#xs') @ ys' \<and> x # xs = map f (e#xs') \<and> ys = map f ys'" by auto |
|
1520 thus ?thesis by metis |
|
1521 qed (insert h, auto) |
|
1522 qed |
|
1523 thus ?thesis by (auto intro!:that) |
|
1524 qed |
|
1525 |
|
1526 lemma rel_map_mono: |
|
1527 assumes "r1 \<subseteq> r2" |
|
1528 shows "rel_map f r1 \<subseteq> rel_map f r2" |
|
1529 using assms |
|
1530 by (auto simp:rel_map_def pairself_def) |
|
1531 |
|
1532 lemma rel_map_compose [simp]: |
|
1533 shows "rel_map f1 (rel_map f2 r) = rel_map (f1 o f2) r" |
|
1534 by (auto simp:rel_map_def pairself_def) |
|
1535 |
|
1536 lemma edges_on_map: "edges_on (map f xs) = rel_map f (edges_on xs)" |
|
1537 proof - |
|
1538 { fix a b |
|
1539 assume "(a, b) \<in> edges_on (map f xs)" |
|
1540 then obtain l1 l2 where eq_map: "map f xs = l1 @ [a, b] @ l2" |
|
1541 by (unfold edges_on_def, auto) |
|
1542 hence "(a, b) \<in> rel_map f (edges_on xs)" |
|
1543 by (auto elim!:map_appendE intro!:rel_mapI simp:edges_on_def) |
|
1544 } moreover { |
|
1545 fix a b |
|
1546 assume "(a, b) \<in> rel_map f (edges_on xs)" |
|
1547 then obtain c d where |
|
1548 h: "(c, d) \<in> edges_on xs" "(a, b) = (f c, f d)" |
|
1549 by (elim rel_mapE, auto) |
|
1550 then obtain l1 l2 where |
|
1551 eq_xs: "xs = l1 @ [c, d] @ l2" |
|
1552 by (auto simp:edges_on_def) |
|
1553 hence eq_map: "map f xs = map f l1 @ [f c, f d] @ map f l2" by auto |
|
1554 have "(a, b) \<in> edges_on (map f xs)" |
|
1555 proof - |
|
1556 from h(2) have "[f c, f d] = [a, b]" by simp |
|
1557 from eq_map[unfolded this] show ?thesis by (auto simp:edges_on_def) |
|
1558 qed |
|
1559 } ultimately show ?thesis by auto |
|
1560 qed |
|
1561 |
|
1562 lemma image_id: |
|
1563 assumes "\<And> x. x \<in> A \<Longrightarrow> f x = x" |
|
1564 shows "f ` A = A" |
|
1565 using assms by (auto simp:image_def) |
|
1566 |
|
1567 lemma rel_map_inv_id: |
|
1568 assumes "inj_on f ((Domain r) \<union> (Range r))" |
|
1569 shows "(rel_map (inv_into ((Domain r) \<union> (Range r)) f \<circ> f) r) = r" |
|
1570 proof - |
|
1571 let ?f = "(inv_into (Domain r \<union> Range r) f \<circ> f)" |
|
1572 { |
|
1573 fix a b |
|
1574 assume h0: "(a, b) \<in> r" |
|
1575 have "pairself ?f (a, b) = (a, b)" |
|
1576 proof - |
|
1577 from assms h0 have "?f a = a" by (auto intro:inv_into_f_f) |
|
1578 moreover have "?f b = b" |
|
1579 by (insert h0, simp, intro inv_into_f_f[OF assms], auto intro!:RangeI) |
|
1580 ultimately show ?thesis by (auto simp:pairself_def) |
|
1581 qed |
|
1582 } thus ?thesis by (unfold rel_map_def, intro image_id, case_tac x, auto) |
|
1583 qed |
|
1584 |
|
1585 lemma rel_map_acyclic: |
|
1586 assumes "acyclic r" |
|
1587 and "inj_on f ((Domain r) \<union> (Range r))" |
|
1588 shows "acyclic (rel_map f r)" |
|
1589 proof - |
|
1590 let ?D = "Domain r \<union> Range r" |
|
1591 { fix a |
|
1592 assume "(a, a) \<in> (rel_map f r)^+" |
|
1593 from plus_rpath[OF this] |
|
1594 obtain xs where rp: "rpath (rel_map f r) a xs a" "xs \<noteq> []" by auto |
|
1595 from rpath_nnl_lastE[OF this] obtain xs' where eq_xs: "xs = xs'@[a]" by auto |
|
1596 from rpath_edges_on[OF rp(1)] |
|
1597 have h: "edges_on (a # xs) \<subseteq> rel_map f r" . |
|
1598 from edges_on_map[of "inv_into ?D f" "a#xs"] |
|
1599 have "edges_on (map (inv_into ?D f) (a # xs)) = rel_map (inv_into ?D f) (edges_on (a # xs))" . |
|
1600 with rel_map_mono[OF h, of "inv_into ?D f"] |
|
1601 have "edges_on (map (inv_into ?D f) (a # xs)) \<subseteq> rel_map ((inv_into ?D f) o f) r" by simp |
|
1602 from this[unfolded eq_xs] |
|
1603 have subr: "edges_on (map (inv_into ?D f) (a # xs' @ [a])) \<subseteq> rel_map (inv_into ?D f \<circ> f) r" . |
|
1604 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]" |
|
1605 by simp |
|
1606 from edges_on_rpathI[OF subr[unfolded this]] |
|
1607 have "rpath (rel_map (inv_into ?D f \<circ> f) r) |
|
1608 (inv_into ?D f a) (map (inv_into ?D f) xs' @ [inv_into ?D f a]) (inv_into ?D f a)" . |
|
1609 hence "(inv_into ?D f a, inv_into ?D f a) \<in> (rel_map (inv_into ?D f \<circ> f) r)^+" |
|
1610 by (rule rpath_plus, simp) |
|
1611 moreover have "(rel_map (inv_into ?D f \<circ> f) r) = r" by (rule rel_map_inv_id[OF assms(2)]) |
|
1612 moreover note assms(1) |
|
1613 ultimately have False by (unfold acyclic_def, auto) |
|
1614 } thus ?thesis by (auto simp:acyclic_def) |
|
1615 qed |
|
1616 |
|
1617 lemma relpow_mult: |
|
1618 "((r::'a rel) ^^ m) ^^ n = r ^^ (m*n)" |
|
1619 proof(induct n arbitrary:m) |
|
1620 case (Suc k m) |
|
1621 thus ?case |
|
1622 proof - |
|
1623 have h: "(m * k + m) = (m + m * k)" by auto |
|
1624 show ?thesis |
|
1625 apply (simp add:Suc relpow_add[symmetric]) |
|
1626 by (unfold h, simp) |
|
1627 qed |
|
1628 qed simp |
|
1629 |
|
1630 lemma compose_relpow_2: |
|
1631 assumes "r1 \<subseteq> r" |
|
1632 and "r2 \<subseteq> r" |
|
1633 shows "r1 O r2 \<subseteq> r ^^ (2::nat)" |
|
1634 proof - |
|
1635 { fix a b |
|
1636 assume "(a, b) \<in> r1 O r2" |
|
1637 then obtain e where "(a, e) \<in> r1" "(e, b) \<in> r2" |
|
1638 by auto |
|
1639 with assms have "(a, e) \<in> r" "(e, b) \<in> r" by auto |
|
1640 hence "(a, b) \<in> r ^^ (Suc (Suc 0))" by auto |
|
1641 } thus ?thesis by (auto simp:numeral_2_eq_2) |
|
1642 qed |
|
1643 |
|
1644 lemma acyclic_compose: |
|
1645 assumes "acyclic r" |
|
1646 and "r1 \<subseteq> r" |
|
1647 and "r2 \<subseteq> r" |
|
1648 shows "acyclic (r1 O r2)" |
|
1649 proof - |
|
1650 { fix a |
|
1651 assume "(a, a) \<in> (r1 O r2)^+" |
|
1652 from trancl_mono[OF this compose_relpow_2[OF assms(2, 3)]] |
|
1653 have "(a, a) \<in> (r ^^ 2) ^+" . |
|
1654 from trancl_power[THEN iffD1, OF this] |
|
1655 obtain n where h: "(a, a) \<in> (r ^^ 2) ^^ n" "n > 0" by blast |
|
1656 from this(1)[unfolded relpow_mult] have h2: "(a, a) \<in> r ^^ (2 * n)" . |
|
1657 have "(a, a) \<in> r^+" |
|
1658 proof(cases rule:trancl_power[THEN iffD2]) |
|
1659 from h(2) h2 show "\<exists>n>0. (a, a) \<in> r ^^ n" |
|
1660 by (rule_tac x = "2*n" in exI, auto) |
|
1661 qed |
|
1662 with assms have "False" by (auto simp:acyclic_def) |
|
1663 } thus ?thesis by (auto simp:acyclic_def) |
|
1664 qed |
|
1665 |
|
1666 lemma children_compose_unfold: |
|
1667 "children (r1 O r2) x = \<Union> (children r1 ` (children r2 x))" |
|
1668 by (auto simp:children_def) |
|
1669 |
|
1670 lemma fbranch_compose: |
|
1671 assumes "fbranch r1" |
|
1672 and "fbranch r2" |
|
1673 shows "fbranch (r1 O r2)" |
|
1674 proof - |
|
1675 { fix x |
|
1676 assume "x\<in>Range (r1 O r2)" |
|
1677 then obtain y z where h: "(y, z) \<in> r1" "(z, x) \<in> r2" by auto |
|
1678 have "finite (children (r1 O r2) x)" |
|
1679 proof(unfold children_compose_unfold, rule finite_Union) |
|
1680 show "finite (children r1 ` children r2 x)" |
|
1681 proof(rule finite_imageI) |
|
1682 from h(2) have "x \<in> Range r2" by auto |
|
1683 from assms(2)[unfolded fbranch_def, rule_format, OF this] |
|
1684 show "finite (children r2 x)" . |
|
1685 qed |
|
1686 next |
|
1687 fix M |
|
1688 assume "M \<in> children r1 ` children r2 x" |
|
1689 then obtain y where h1: "y \<in> children r2 x" "M = children r1 y" by auto |
|
1690 show "finite M" |
|
1691 proof(cases "children r1 y = {}") |
|
1692 case True |
|
1693 with h1(2) show ?thesis by auto |
|
1694 next |
|
1695 case False |
|
1696 hence "y \<in> Range r1" by (unfold children_def, auto) |
|
1697 from assms(1)[unfolded fbranch_def, rule_format, OF this, folded h1(2)] |
|
1698 show ?thesis . |
|
1699 qed |
|
1700 qed |
|
1701 } thus ?thesis by (unfold fbranch_def, auto) |
|
1702 qed |
|
1703 |
|
1704 lemma finite_fbranchI: |
|
1705 assumes "finite r" |
|
1706 shows "fbranch r" |
|
1707 proof - |
|
1708 { fix x |
|
1709 assume "x \<in>Range r" |
|
1710 have "finite (children r x)" |
|
1711 proof - |
|
1712 have "{y. (y, x) \<in> r} \<subseteq> Domain r" by (auto) |
|
1713 from rev_finite_subset[OF finite_Domain[OF assms] this] |
|
1714 have "finite {y. (y, x) \<in> r}" . |
|
1715 thus ?thesis by (unfold children_def, simp) |
|
1716 qed |
|
1717 } thus ?thesis by (auto simp:fbranch_def) |
|
1718 qed |
|
1719 |
|
1720 lemma subset_fbranchI: |
|
1721 assumes "fbranch r1" |
|
1722 and "r2 \<subseteq> r1" |
|
1723 shows "fbranch r2" |
|
1724 proof - |
|
1725 { fix x |
|
1726 assume "x \<in>Range r2" |
|
1727 with assms(2) have "x \<in> Range r1" by auto |
|
1728 from assms(1)[unfolded fbranch_def, rule_format, OF this] |
|
1729 have "finite (children r1 x)" . |
|
1730 hence "finite (children r2 x)" |
|
1731 proof(rule rev_finite_subset) |
|
1732 from assms(2) |
|
1733 show "children r2 x \<subseteq> children r1 x" by (auto simp:children_def) |
|
1734 qed |
|
1735 } thus ?thesis by (auto simp:fbranch_def) |
|
1736 qed |
|
1737 |
|
1738 lemma children_subtree: |
|
1739 shows "children r x \<subseteq> subtree r x" |
|
1740 by (auto simp:children_def subtree_def) |
|
1741 |
|
1742 lemma children_union_kept: |
|
1743 assumes "x \<notin> Range r'" |
|
1744 shows "children (r \<union> r') x = children r x" |
|
1745 using assms |
|
1746 by (auto simp:children_def) |
|
1747 |
|
1748 end |