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