7 |
7 |
8 theory Trancl |
8 theory Trancl |
9 imports CCL |
9 imports CCL |
10 begin |
10 begin |
11 |
11 |
12 definition trans :: "i set => o" (*transitivity predicate*) |
12 definition trans :: "i set \<Rightarrow> o" (*transitivity predicate*) |
13 where "trans(r) == (ALL x y z. <x,y>:r --> <y,z>:r --> <x,z>:r)" |
13 where "trans(r) == (ALL x y z. <x,y>:r \<longrightarrow> <y,z>:r \<longrightarrow> <x,z>:r)" |
14 |
14 |
15 definition id :: "i set" (*the identity relation*) |
15 definition id :: "i set" (*the identity relation*) |
16 where "id == {p. EX x. p = <x,x>}" |
16 where "id == {p. EX x. p = <x,x>}" |
17 |
17 |
18 definition relcomp :: "[i set,i set] => i set" (infixr "O" 60) (*composition of relations*) |
18 definition relcomp :: "[i set,i set] \<Rightarrow> i set" (infixr "O" 60) (*composition of relations*) |
19 where "r O s == {xz. EX x y z. xz = <x,z> & <x,y>:s & <y,z>:r}" |
19 where "r O s == {xz. EX x y z. xz = <x,z> \<and> <x,y>:s \<and> <y,z>:r}" |
20 |
20 |
21 definition rtrancl :: "i set => i set" ("(_^*)" [100] 100) |
21 definition rtrancl :: "i set \<Rightarrow> i set" ("(_^*)" [100] 100) |
22 where "r^* == lfp(%s. id Un (r O s))" |
22 where "r^* == lfp(\<lambda>s. id Un (r O s))" |
23 |
23 |
24 definition trancl :: "i set => i set" ("(_^+)" [100] 100) |
24 definition trancl :: "i set \<Rightarrow> i set" ("(_^+)" [100] 100) |
25 where "r^+ == r O rtrancl(r)" |
25 where "r^+ == r O rtrancl(r)" |
26 |
26 |
27 |
27 |
28 subsection {* Natural deduction for @{text "trans(r)"} *} |
28 subsection {* Natural deduction for @{text "trans(r)"} *} |
29 |
29 |
30 lemma transI: |
30 lemma transI: "(\<And>x y z. \<lbrakk><x,y>:r; <y,z>:r\<rbrakk> \<Longrightarrow> <x,z>:r) \<Longrightarrow> trans(r)" |
31 "(!! x y z. [| <x,y>:r; <y,z>:r |] ==> <x,z>:r) ==> trans(r)" |
|
32 unfolding trans_def by blast |
31 unfolding trans_def by blast |
33 |
32 |
34 lemma transD: "[| trans(r); <a,b>:r; <b,c>:r |] ==> <a,c>:r" |
33 lemma transD: "\<lbrakk>trans(r); <a,b>:r; <b,c>:r\<rbrakk> \<Longrightarrow> <a,c>:r" |
35 unfolding trans_def by blast |
34 unfolding trans_def by blast |
36 |
35 |
37 |
36 |
38 subsection {* Identity relation *} |
37 subsection {* Identity relation *} |
39 |
38 |
42 apply (rule CollectI) |
41 apply (rule CollectI) |
43 apply (rule exI) |
42 apply (rule exI) |
44 apply (rule refl) |
43 apply (rule refl) |
45 done |
44 done |
46 |
45 |
47 lemma idE: |
46 lemma idE: "\<lbrakk>p: id; \<And>x. p = <x,x> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
48 "[| p: id; !!x.[| p = <x,x> |] ==> P |] ==> P" |
|
49 apply (unfold id_def) |
47 apply (unfold id_def) |
50 apply (erule CollectE) |
48 apply (erule CollectE) |
51 apply blast |
49 apply blast |
52 done |
50 done |
53 |
51 |
54 |
52 |
55 subsection {* Composition of two relations *} |
53 subsection {* Composition of two relations *} |
56 |
54 |
57 lemma compI: "[| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s" |
55 lemma compI: "\<lbrakk><a,b>:s; <b,c>:r\<rbrakk> \<Longrightarrow> <a,c> : r O s" |
58 unfolding relcomp_def by blast |
56 unfolding relcomp_def by blast |
59 |
57 |
60 (*proof requires higher-level assumptions or a delaying of hyp_subst_tac*) |
58 (*proof requires higher-level assumptions or a delaying of hyp_subst_tac*) |
61 lemma compE: |
59 lemma compE: "\<lbrakk>xz : r O s; \<And>x y z. \<lbrakk>xz = <x,z>; <x,y>:s; <y,z>:r\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
62 "[| xz : r O s; |
|
63 !!x y z. [| xz = <x,z>; <x,y>:s; <y,z>:r |] ==> P |
|
64 |] ==> P" |
|
65 unfolding relcomp_def by blast |
60 unfolding relcomp_def by blast |
66 |
61 |
67 lemma compEpair: |
62 lemma compEpair: "\<lbrakk><a,c> : r O s; \<And>y. \<lbrakk><a,y>:s; <y,c>:r\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
68 "[| <a,c> : r O s; |
|
69 !!y. [| <a,y>:s; <y,c>:r |] ==> P |
|
70 |] ==> P" |
|
71 apply (erule compE) |
63 apply (erule compE) |
72 apply (simp add: pair_inject) |
64 apply (simp add: pair_inject) |
73 done |
65 done |
74 |
66 |
75 lemmas [intro] = compI idI |
67 lemmas [intro] = compI idI |
76 and [elim] = compE idE |
68 and [elim] = compE idE |
77 and [elim!] = pair_inject |
69 and [elim!] = pair_inject |
78 |
70 |
79 lemma comp_mono: "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)" |
71 lemma comp_mono: "\<lbrakk>r'<=r; s'<=s\<rbrakk> \<Longrightarrow> (r' O s') <= (r O s)" |
80 by blast |
72 by blast |
81 |
73 |
82 |
74 |
83 subsection {* The relation rtrancl *} |
75 subsection {* The relation rtrancl *} |
84 |
76 |
85 lemma rtrancl_fun_mono: "mono(%s. id Un (r O s))" |
77 lemma rtrancl_fun_mono: "mono(\<lambda>s. id Un (r O s))" |
86 apply (rule monoI) |
78 apply (rule monoI) |
87 apply (rule monoI subset_refl comp_mono Un_mono)+ |
79 apply (rule monoI subset_refl comp_mono Un_mono)+ |
88 apply assumption |
80 apply assumption |
89 done |
81 done |
90 |
82 |
96 apply (subst rtrancl_unfold) |
88 apply (subst rtrancl_unfold) |
97 apply blast |
89 apply blast |
98 done |
90 done |
99 |
91 |
100 (*Closure under composition with r*) |
92 (*Closure under composition with r*) |
101 lemma rtrancl_into_rtrancl: "[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^*" |
93 lemma rtrancl_into_rtrancl: "\<lbrakk><a,b> : r^*; <b,c> : r\<rbrakk> \<Longrightarrow> <a,c> : r^*" |
102 apply (subst rtrancl_unfold) |
94 apply (subst rtrancl_unfold) |
103 apply blast |
95 apply blast |
104 done |
96 done |
105 |
97 |
106 (*rtrancl of r contains r*) |
98 (*rtrancl of r contains r*) |
107 lemma r_into_rtrancl: "[| <a,b> : r |] ==> <a,b> : r^*" |
99 lemma r_into_rtrancl: "<a,b> : r \<Longrightarrow> <a,b> : r^*" |
108 apply (rule rtrancl_refl [THEN rtrancl_into_rtrancl]) |
100 apply (rule rtrancl_refl [THEN rtrancl_into_rtrancl]) |
109 apply assumption |
101 apply assumption |
110 done |
102 done |
111 |
103 |
112 |
104 |
113 subsection {* standard induction rule *} |
105 subsection {* standard induction rule *} |
114 |
106 |
115 lemma rtrancl_full_induct: |
107 lemma rtrancl_full_induct: |
116 "[| <a,b> : r^*; |
108 "\<lbrakk><a,b> : r^*; |
117 !!x. P(<x,x>); |
109 \<And>x. P(<x,x>); |
118 !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |] ==> P(<x,z>) |] |
110 \<And>x y z. \<lbrakk>P(<x,y>); <x,y>: r^*; <y,z>: r\<rbrakk> \<Longrightarrow> P(<x,z>)\<rbrakk> |
119 ==> P(<a,b>)" |
111 \<Longrightarrow> P(<a,b>)" |
120 apply (erule def_induct [OF rtrancl_def]) |
112 apply (erule def_induct [OF rtrancl_def]) |
121 apply (rule rtrancl_fun_mono) |
113 apply (rule rtrancl_fun_mono) |
122 apply blast |
114 apply blast |
123 done |
115 done |
124 |
116 |
125 (*nice induction rule*) |
117 (*nice induction rule*) |
126 lemma rtrancl_induct: |
118 lemma rtrancl_induct: |
127 "[| <a,b> : r^*; |
119 "\<lbrakk><a,b> : r^*; |
128 P(a); |
120 P(a); |
129 !!y z.[| <a,y> : r^*; <y,z> : r; P(y) |] ==> P(z) |] |
121 \<And>y z. \<lbrakk><a,y> : r^*; <y,z> : r; P(y)\<rbrakk> \<Longrightarrow> P(z) \<rbrakk> |
130 ==> P(b)" |
122 \<Longrightarrow> P(b)" |
131 (*by induction on this formula*) |
123 (*by induction on this formula*) |
132 apply (subgoal_tac "ALL y. <a,b> = <a,y> --> P(y)") |
124 apply (subgoal_tac "ALL y. <a,b> = <a,y> \<longrightarrow> P(y)") |
133 (*now solve first subgoal: this formula is sufficient*) |
125 (*now solve first subgoal: this formula is sufficient*) |
134 apply blast |
126 apply blast |
135 (*now do the induction*) |
127 (*now do the induction*) |
136 apply (erule rtrancl_full_induct) |
128 apply (erule rtrancl_full_induct) |
137 apply blast |
129 apply blast |
145 apply (fast elim: rtrancl_into_rtrancl)+ |
137 apply (fast elim: rtrancl_into_rtrancl)+ |
146 done |
138 done |
147 |
139 |
148 (*elimination of rtrancl -- by induction on a special formula*) |
140 (*elimination of rtrancl -- by induction on a special formula*) |
149 lemma rtranclE: |
141 lemma rtranclE: |
150 "[| <a,b> : r^*; (a = b) ==> P; |
142 "\<lbrakk><a,b> : r^*; a = b \<Longrightarrow> P; \<And>y. \<lbrakk><a,y> : r^*; <y,b> : r\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
151 !!y.[| <a,y> : r^*; <y,b> : r |] ==> P |] |
143 apply (subgoal_tac "a = b | (EX y. <a,y> : r^* \<and> <y,b> : r)") |
152 ==> P" |
|
153 apply (subgoal_tac "a = b | (EX y. <a,y> : r^* & <y,b> : r)") |
|
154 prefer 2 |
144 prefer 2 |
155 apply (erule rtrancl_induct) |
145 apply (erule rtrancl_induct) |
156 apply blast |
146 apply blast |
157 apply blast |
147 apply blast |
158 apply blast |
148 apply blast |
161 |
151 |
162 subsection {* The relation trancl *} |
152 subsection {* The relation trancl *} |
163 |
153 |
164 subsubsection {* Conversions between trancl and rtrancl *} |
154 subsubsection {* Conversions between trancl and rtrancl *} |
165 |
155 |
166 lemma trancl_into_rtrancl: "[| <a,b> : r^+ |] ==> <a,b> : r^*" |
156 lemma trancl_into_rtrancl: "<a,b> : r^+ \<Longrightarrow> <a,b> : r^*" |
167 apply (unfold trancl_def) |
157 apply (unfold trancl_def) |
168 apply (erule compEpair) |
158 apply (erule compEpair) |
169 apply (erule rtrancl_into_rtrancl) |
159 apply (erule rtrancl_into_rtrancl) |
170 apply assumption |
160 apply assumption |
171 done |
161 done |
172 |
162 |
173 (*r^+ contains r*) |
163 (*r^+ contains r*) |
174 lemma r_into_trancl: "[| <a,b> : r |] ==> <a,b> : r^+" |
164 lemma r_into_trancl: "<a,b> : r \<Longrightarrow> <a,b> : r^+" |
175 unfolding trancl_def by (blast intro: rtrancl_refl) |
165 unfolding trancl_def by (blast intro: rtrancl_refl) |
176 |
166 |
177 (*intro rule by definition: from rtrancl and r*) |
167 (*intro rule by definition: from rtrancl and r*) |
178 lemma rtrancl_into_trancl1: "[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^+" |
168 lemma rtrancl_into_trancl1: "\<lbrakk><a,b> : r^*; <b,c> : r\<rbrakk> \<Longrightarrow> <a,c> : r^+" |
179 unfolding trancl_def by blast |
169 unfolding trancl_def by blast |
180 |
170 |
181 (*intro rule from r and rtrancl*) |
171 (*intro rule from r and rtrancl*) |
182 lemma rtrancl_into_trancl2: "[| <a,b> : r; <b,c> : r^* |] ==> <a,c> : r^+" |
172 lemma rtrancl_into_trancl2: "\<lbrakk><a,b> : r; <b,c> : r^*\<rbrakk> \<Longrightarrow> <a,c> : r^+" |
183 apply (erule rtranclE) |
173 apply (erule rtranclE) |
184 apply (erule subst) |
174 apply (erule subst) |
185 apply (erule r_into_trancl) |
175 apply (erule r_into_trancl) |
186 apply (rule trans_rtrancl [THEN transD, THEN rtrancl_into_trancl1]) |
176 apply (rule trans_rtrancl [THEN transD, THEN rtrancl_into_trancl1]) |
187 apply (assumption | rule r_into_rtrancl)+ |
177 apply (assumption | rule r_into_rtrancl)+ |
188 done |
178 done |
189 |
179 |
190 (*elimination of r^+ -- NOT an induction rule*) |
180 (*elimination of r^+ -- NOT an induction rule*) |
191 lemma tranclE: |
181 lemma tranclE: |
192 "[| <a,b> : r^+; |
182 "\<lbrakk><a,b> : r^+; |
193 <a,b> : r ==> P; |
183 <a,b> : r \<Longrightarrow> P; |
194 !!y.[| <a,y> : r^+; <y,b> : r |] ==> P |
184 \<And>y. \<lbrakk><a,y> : r^+; <y,b> : r\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
195 |] ==> P" |
185 apply (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+ \<and> <y,b> : r)") |
196 apply (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+ & <y,b> : r)") |
|
197 apply blast |
186 apply blast |
198 apply (unfold trancl_def) |
187 apply (unfold trancl_def) |
199 apply (erule compEpair) |
188 apply (erule compEpair) |
200 apply (erule rtranclE) |
189 apply (erule rtranclE) |
201 apply blast |
190 apply blast |