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