author | wenzelm |
Mon, 21 Aug 2023 13:01:45 +0200 | |
changeset 78553 | 66fc98b4557b |
parent 76217 | 8655344f1cf6 |
child 80917 | 2a77bc3b4eac |
permissions | -rw-r--r-- |
13356 | 1 |
(* Title: ZF/Trancl.thy |
1478 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 3 |
Copyright 1992 University of Cambridge |
4 |
*) |
|
5 |
||
60770 | 6 |
section\<open>Relations: Their General Properties and Transitive Closure\<close> |
13356 | 7 |
|
16417 | 8 |
theory Trancl imports Fixedpt Perm begin |
13239 | 9 |
|
24893 | 10 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
11 |
refl :: "[i,i]\<Rightarrow>o" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
12 |
"refl(A,r) \<equiv> (\<forall>x\<in>A. \<langle>x,x\<rangle> \<in> r)" |
13239 | 13 |
|
24893 | 14 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
15 |
irrefl :: "[i,i]\<Rightarrow>o" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
16 |
"irrefl(A,r) \<equiv> \<forall>x\<in>A. \<langle>x,x\<rangle> \<notin> r" |
13239 | 17 |
|
24893 | 18 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
19 |
sym :: "i\<Rightarrow>o" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
20 |
"sym(r) \<equiv> \<forall>x y. \<langle>x,y\<rangle>: r \<longrightarrow> \<langle>y,x\<rangle>: r" |
13239 | 21 |
|
24893 | 22 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
23 |
asym :: "i\<Rightarrow>o" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
24 |
"asym(r) \<equiv> \<forall>x y. \<langle>x,y\<rangle>:r \<longrightarrow> \<not> \<langle>y,x\<rangle>:r" |
13239 | 25 |
|
24893 | 26 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
27 |
antisym :: "i\<Rightarrow>o" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
28 |
"antisym(r) \<equiv> \<forall>x y.\<langle>x,y\<rangle>:r \<longrightarrow> \<langle>y,x\<rangle>:r \<longrightarrow> x=y" |
13239 | 29 |
|
24893 | 30 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
31 |
trans :: "i\<Rightarrow>o" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
32 |
"trans(r) \<equiv> \<forall>x y z. \<langle>x,y\<rangle>: r \<longrightarrow> \<langle>y,z\<rangle>: r \<longrightarrow> \<langle>x,z\<rangle>: r" |
13239 | 33 |
|
24893 | 34 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
35 |
trans_on :: "[i,i]\<Rightarrow>o" (\<open>trans[_]'(_')\<close>) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
36 |
"trans[A](r) \<equiv> \<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A. |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
37 |
\<langle>x,y\<rangle>: r \<longrightarrow> \<langle>y,z\<rangle>: r \<longrightarrow> \<langle>x,z\<rangle>: r" |
13239 | 38 |
|
24893 | 39 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
40 |
rtrancl :: "i\<Rightarrow>i" (\<open>(_^*)\<close> [100] 100) (*refl/transitive closure*) where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
41 |
"r^* \<equiv> lfp(field(r)*field(r), \<lambda>s. id(field(r)) \<union> (r O s))" |
13239 | 42 |
|
24893 | 43 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
44 |
trancl :: "i\<Rightarrow>i" (\<open>(_^+)\<close> [100] 100) (*transitive closure*) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
45 |
"r^+ \<equiv> r O r^*" |
13239 | 46 |
|
24893 | 47 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
48 |
equiv :: "[i,i]\<Rightarrow>o" where |
76214 | 49 |
"equiv(A,r) \<equiv> r \<subseteq> A*A \<and> refl(A,r) \<and> sym(r) \<and> trans(r)" |
14653 | 50 |
|
51 |
||
60770 | 52 |
subsection\<open>General properties of relations\<close> |
13239 | 53 |
|
60770 | 54 |
subsubsection\<open>irreflexivity\<close> |
13239 | 55 |
|
56 |
lemma irreflI: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
57 |
"\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> \<langle>x,x\<rangle> \<notin> r\<rbrakk> \<Longrightarrow> irrefl(A,r)" |
46820 | 58 |
by (simp add: irrefl_def) |
13239 | 59 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
60 |
lemma irreflE: "\<lbrakk>irrefl(A,r); x \<in> A\<rbrakk> \<Longrightarrow> \<langle>x,x\<rangle> \<notin> r" |
13240 | 61 |
by (simp add: irrefl_def) |
13239 | 62 |
|
60770 | 63 |
subsubsection\<open>symmetry\<close> |
13239 | 64 |
|
65 |
lemma symI: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
66 |
"\<lbrakk>\<And>x y.\<langle>x,y\<rangle>: r \<Longrightarrow> \<langle>y,x\<rangle>: r\<rbrakk> \<Longrightarrow> sym(r)" |
46820 | 67 |
by (unfold sym_def, blast) |
13239 | 68 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
69 |
lemma symE: "\<lbrakk>sym(r); \<langle>x,y\<rangle>: r\<rbrakk> \<Longrightarrow> \<langle>y,x\<rangle>: r" |
13240 | 70 |
by (unfold sym_def, blast) |
13239 | 71 |
|
60770 | 72 |
subsubsection\<open>antisymmetry\<close> |
13239 | 73 |
|
74 |
lemma antisymI: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
75 |
"\<lbrakk>\<And>x y.\<lbrakk>\<langle>x,y\<rangle>: r; \<langle>y,x\<rangle>: r\<rbrakk> \<Longrightarrow> x=y\<rbrakk> \<Longrightarrow> antisym(r)" |
46820 | 76 |
by (simp add: antisym_def, blast) |
13239 | 77 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
78 |
lemma antisymE: "\<lbrakk>antisym(r); \<langle>x,y\<rangle>: r; \<langle>y,x\<rangle>: r\<rbrakk> \<Longrightarrow> x=y" |
13240 | 79 |
by (simp add: antisym_def, blast) |
13239 | 80 |
|
60770 | 81 |
subsubsection\<open>transitivity\<close> |
13239 | 82 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
83 |
lemma transD: "\<lbrakk>trans(r); \<langle>a,b\<rangle>:r; \<langle>b,c\<rangle>:r\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle>:r" |
13240 | 84 |
by (unfold trans_def, blast) |
13239 | 85 |
|
46820 | 86 |
lemma trans_onD: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
87 |
"\<lbrakk>trans[A](r); \<langle>a,b\<rangle>:r; \<langle>b,c\<rangle>:r; a \<in> A; b \<in> A; c \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle>:r" |
13243 | 88 |
by (unfold trans_on_def, blast) |
89 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
90 |
lemma trans_imp_trans_on: "trans(r) \<Longrightarrow> trans[A](r)" |
13243 | 91 |
by (unfold trans_def trans_on_def, blast) |
13239 | 92 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
93 |
lemma trans_on_imp_trans: "\<lbrakk>trans[A](r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> trans(r)" |
13248 | 94 |
by (simp add: trans_on_def trans_def, blast) |
95 |
||
96 |
||
60770 | 97 |
subsection\<open>Transitive closure of a relation\<close> |
13239 | 98 |
|
99 |
lemma rtrancl_bnd_mono: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
100 |
"bnd_mono(field(r)*field(r), \<lambda>s. id(field(r)) \<union> (r O s))" |
13248 | 101 |
by (rule bnd_monoI, blast+) |
13239 | 102 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
103 |
lemma rtrancl_mono: "r<=s \<Longrightarrow> r^* \<subseteq> s^*" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
104 |
unfolding rtrancl_def |
13239 | 105 |
apply (rule lfp_mono) |
106 |
apply (rule rtrancl_bnd_mono)+ |
|
46820 | 107 |
apply blast |
13239 | 108 |
done |
109 |
||
46820 | 110 |
(* @{term"r^* = id(field(r)) \<union> ( r O r^* )"} *) |
13239 | 111 |
lemmas rtrancl_unfold = |
45602 | 112 |
rtrancl_bnd_mono [THEN rtrancl_def [THEN def_lfp_unfold]] |
13239 | 113 |
|
114 |
(** The relation rtrancl **) |
|
115 |
||
46820 | 116 |
(* @{term"r^* \<subseteq> field(r) * field(r)"} *) |
45602 | 117 |
lemmas rtrancl_type = rtrancl_def [THEN def_lfp_subset] |
13239 | 118 |
|
13248 | 119 |
lemma relation_rtrancl: "relation(r^*)" |
46820 | 120 |
apply (simp add: relation_def) |
121 |
apply (blast dest: rtrancl_type [THEN subsetD]) |
|
13248 | 122 |
done |
123 |
||
13239 | 124 |
(*Reflexivity of rtrancl*) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
125 |
lemma rtrancl_refl: "\<lbrakk>a \<in> field(r)\<rbrakk> \<Longrightarrow> \<langle>a,a\<rangle> \<in> r^*" |
13239 | 126 |
apply (rule rtrancl_unfold [THEN ssubst]) |
127 |
apply (erule idI [THEN UnI1]) |
|
128 |
done |
|
129 |
||
130 |
(*Closure under composition with r *) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
131 |
lemma rtrancl_into_rtrancl: "\<lbrakk>\<langle>a,b\<rangle> \<in> r^*; \<langle>b,c\<rangle> \<in> r\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle> \<in> r^*" |
13239 | 132 |
apply (rule rtrancl_unfold [THEN ssubst]) |
13269 | 133 |
apply (rule compI [THEN UnI2], assumption, assumption) |
13239 | 134 |
done |
135 |
||
136 |
(*rtrancl of r contains all pairs in r *) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
137 |
lemma r_into_rtrancl: "\<langle>a,b\<rangle> \<in> r \<Longrightarrow> \<langle>a,b\<rangle> \<in> r^*" |
13240 | 138 |
by (rule rtrancl_refl [THEN rtrancl_into_rtrancl], blast+) |
13239 | 139 |
|
140 |
(*The premise ensures that r consists entirely of pairs*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
141 |
lemma r_subset_rtrancl: "relation(r) \<Longrightarrow> r \<subseteq> r^*" |
13248 | 142 |
by (simp add: relation_def, blast intro: r_into_rtrancl) |
13239 | 143 |
|
144 |
lemma rtrancl_field: "field(r^*) = field(r)" |
|
13240 | 145 |
by (blast intro: r_into_rtrancl dest!: rtrancl_type [THEN subsetD]) |
13239 | 146 |
|
147 |
||
148 |
(** standard induction rule **) |
|
149 |
||
13534 | 150 |
lemma rtrancl_full_induct [case_names initial step, consumes 1]: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
151 |
"\<lbrakk>\<langle>a,b\<rangle> \<in> r^*; |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
152 |
\<And>x. x \<in> field(r) \<Longrightarrow> P(\<langle>x,x\<rangle>); |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
153 |
\<And>x y z.\<lbrakk>P(\<langle>x,y\<rangle>); \<langle>x,y\<rangle>: r^*; \<langle>y,z\<rangle>: r\<rbrakk> \<Longrightarrow> P(\<langle>x,z\<rangle>)\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
154 |
\<Longrightarrow> P(\<langle>a,b\<rangle>)" |
46820 | 155 |
by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast) |
13239 | 156 |
|
157 |
(*nice induction rule. |
|
46953 | 158 |
Tried adding the typing hypotheses y,z \<in> field(r), but these |
13239 | 159 |
caused expensive case splits!*) |
13534 | 160 |
lemma rtrancl_induct [case_names initial step, induct set: rtrancl]: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
161 |
"\<lbrakk>\<langle>a,b\<rangle> \<in> r^*; |
46820 | 162 |
P(a); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
163 |
\<And>y z.\<lbrakk>\<langle>a,y\<rangle> \<in> r^*; \<langle>y,z\<rangle> \<in> r; P(y)\<rbrakk> \<Longrightarrow> P(z) |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
164 |
\<rbrakk> \<Longrightarrow> P(b)" |
13239 | 165 |
(*by induction on this formula*) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
166 |
apply (subgoal_tac "\<forall>y. \<langle>a,b\<rangle> = \<langle>a,y\<rangle> \<longrightarrow> P (y) ") |
13239 | 167 |
(*now solve first subgoal: this formula is sufficient*) |
168 |
apply (erule spec [THEN mp], rule refl) |
|
169 |
(*now do the induction*) |
|
13240 | 170 |
apply (erule rtrancl_full_induct, blast+) |
13239 | 171 |
done |
172 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
173 |
(*transitivity of transitive closure\<And>-- by induction.*) |
13239 | 174 |
lemma trans_rtrancl: "trans(r^*)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
175 |
unfolding trans_def |
13239 | 176 |
apply (intro allI impI) |
13784 | 177 |
apply (erule_tac b = z in rtrancl_induct, assumption) |
46820 | 178 |
apply (blast intro: rtrancl_into_rtrancl) |
13239 | 179 |
done |
180 |
||
45602 | 181 |
lemmas rtrancl_trans = trans_rtrancl [THEN transD] |
13239 | 182 |
|
183 |
(*elimination of rtrancl -- by induction on a special formula*) |
|
184 |
lemma rtranclE: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
185 |
"\<lbrakk>\<langle>a,b\<rangle> \<in> r^*; (a=b) \<Longrightarrow> P; |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
186 |
\<And>y.\<lbrakk>\<langle>a,y\<rangle> \<in> r^*; \<langle>y,b\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
187 |
\<Longrightarrow> P" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
188 |
apply (subgoal_tac "a = b | (\<exists>y. \<langle>a,y\<rangle> \<in> r^* \<and> \<langle>y,b\<rangle> \<in> r) ") |
13239 | 189 |
(*see HOL/trancl*) |
46820 | 190 |
apply blast |
13240 | 191 |
apply (erule rtrancl_induct, blast+) |
13239 | 192 |
done |
193 |
||
194 |
||
195 |
(**** The relation trancl ****) |
|
196 |
||
197 |
(*Transitivity of r^+ is proved by transitivity of r^* *) |
|
198 |
lemma trans_trancl: "trans(r^+)" |
|
76217 | 199 |
unfolding trans_def trancl_def |
13239 | 200 |
apply (blast intro: rtrancl_into_rtrancl |
201 |
trans_rtrancl [THEN transD, THEN compI]) |
|
202 |
done |
|
203 |
||
13243 | 204 |
lemmas trans_on_trancl = trans_trancl [THEN trans_imp_trans_on] |
205 |
||
45602 | 206 |
lemmas trancl_trans = trans_trancl [THEN transD] |
13239 | 207 |
|
208 |
(** Conversions between trancl and rtrancl **) |
|
0 | 209 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
210 |
lemma trancl_into_rtrancl: "\<langle>a,b\<rangle> \<in> r^+ \<Longrightarrow> \<langle>a,b\<rangle> \<in> r^*" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
211 |
unfolding trancl_def |
13239 | 212 |
apply (blast intro: rtrancl_into_rtrancl) |
213 |
done |
|
214 |
||
215 |
(*r^+ contains all pairs in r *) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
216 |
lemma r_into_trancl: "\<langle>a,b\<rangle> \<in> r \<Longrightarrow> \<langle>a,b\<rangle> \<in> r^+" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
217 |
unfolding trancl_def |
13239 | 218 |
apply (blast intro!: rtrancl_refl) |
219 |
done |
|
220 |
||
221 |
(*The premise ensures that r consists entirely of pairs*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
222 |
lemma r_subset_trancl: "relation(r) \<Longrightarrow> r \<subseteq> r^+" |
13248 | 223 |
by (simp add: relation_def, blast intro: r_into_trancl) |
224 |
||
13239 | 225 |
|
226 |
(*intro rule by definition: from r^* and r *) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
227 |
lemma rtrancl_into_trancl1: "\<lbrakk>\<langle>a,b\<rangle> \<in> r^*; \<langle>b,c\<rangle> \<in> r\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle> \<in> r^+" |
13240 | 228 |
by (unfold trancl_def, blast) |
13239 | 229 |
|
230 |
(*intro rule from r and r^* *) |
|
231 |
lemma rtrancl_into_trancl2: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
232 |
"\<lbrakk>\<langle>a,b\<rangle> \<in> r; \<langle>b,c\<rangle> \<in> r^*\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle> \<in> r^+" |
13239 | 233 |
apply (erule rtrancl_induct) |
234 |
apply (erule r_into_trancl) |
|
46820 | 235 |
apply (blast intro: r_into_trancl trancl_trans) |
13239 | 236 |
done |
237 |
||
238 |
(*Nice induction rule for trancl*) |
|
13534 | 239 |
lemma trancl_induct [case_names initial step, induct set: trancl]: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
240 |
"\<lbrakk>\<langle>a,b\<rangle> \<in> r^+; |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
241 |
\<And>y. \<lbrakk>\<langle>a,y\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P(y); |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
242 |
\<And>y z.\<lbrakk>\<langle>a,y\<rangle> \<in> r^+; \<langle>y,z\<rangle> \<in> r; P(y)\<rbrakk> \<Longrightarrow> P(z) |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
243 |
\<rbrakk> \<Longrightarrow> P(b)" |
13239 | 244 |
apply (rule compEpair) |
245 |
apply (unfold trancl_def, assumption) |
|
246 |
(*by induction on this formula*) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
247 |
apply (subgoal_tac "\<forall>z. \<langle>y,z\<rangle> \<in> r \<longrightarrow> P (z) ") |
13239 | 248 |
(*now solve first subgoal: this formula is sufficient*) |
249 |
apply blast |
|
250 |
apply (erule rtrancl_induct) |
|
251 |
apply (blast intro: rtrancl_into_trancl1)+ |
|
252 |
done |
|
253 |
||
254 |
(*elimination of r^+ -- NOT an induction rule*) |
|
255 |
lemma tranclE: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
256 |
"\<lbrakk>\<langle>a,b\<rangle> \<in> r^+; |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
257 |
\<langle>a,b\<rangle> \<in> r \<Longrightarrow> P; |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
258 |
\<And>y.\<lbrakk>\<langle>a,y\<rangle> \<in> r^+; \<langle>y,b\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
259 |
\<rbrakk> \<Longrightarrow> P" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
260 |
apply (subgoal_tac "\<langle>a,b\<rangle> \<in> r | (\<exists>y. \<langle>a,y\<rangle> \<in> r^+ \<and> \<langle>y,b\<rangle> \<in> r) ") |
46820 | 261 |
apply blast |
13239 | 262 |
apply (rule compEpair) |
263 |
apply (unfold trancl_def, assumption) |
|
264 |
apply (erule rtranclE) |
|
265 |
apply (blast intro: rtrancl_into_trancl1)+ |
|
266 |
done |
|
267 |
||
46820 | 268 |
lemma trancl_type: "r^+ \<subseteq> field(r)*field(r)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
269 |
unfolding trancl_def |
13239 | 270 |
apply (blast elim: rtrancl_type [THEN subsetD, THEN SigmaE2]) |
271 |
done |
|
272 |
||
13248 | 273 |
lemma relation_trancl: "relation(r^+)" |
46820 | 274 |
apply (simp add: relation_def) |
275 |
apply (blast dest: trancl_type [THEN subsetD]) |
|
13248 | 276 |
done |
277 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
278 |
lemma trancl_subset_times: "r \<subseteq> A * A \<Longrightarrow> r^+ \<subseteq> A * A" |
13243 | 279 |
by (insert trancl_type [of r], blast) |
280 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
281 |
lemma trancl_mono: "r<=s \<Longrightarrow> r^+ \<subseteq> s^+" |
13239 | 282 |
by (unfold trancl_def, intro comp_mono rtrancl_mono) |
283 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
284 |
lemma trancl_eq_r: "\<lbrakk>relation(r); trans(r)\<rbrakk> \<Longrightarrow> r^+ = r" |
13248 | 285 |
apply (rule equalityI) |
46820 | 286 |
prefer 2 apply (erule r_subset_trancl, clarify) |
287 |
apply (frule trancl_type [THEN subsetD], clarify) |
|
13248 | 288 |
apply (erule trancl_induct, assumption) |
46820 | 289 |
apply (blast dest: transD) |
13248 | 290 |
done |
291 |
||
13239 | 292 |
|
293 |
(** Suggested by Sidi Ould Ehmety **) |
|
294 |
||
295 |
lemma rtrancl_idemp [simp]: "(r^*)^* = r^*" |
|
13240 | 296 |
apply (rule equalityI, auto) |
13239 | 297 |
prefer 2 |
298 |
apply (frule rtrancl_type [THEN subsetD]) |
|
46820 | 299 |
apply (blast intro: r_into_rtrancl ) |
60770 | 300 |
txt\<open>converse direction\<close> |
46820 | 301 |
apply (frule rtrancl_type [THEN subsetD], clarify) |
13239 | 302 |
apply (erule rtrancl_induct) |
303 |
apply (simp add: rtrancl_refl rtrancl_field) |
|
304 |
apply (blast intro: rtrancl_trans) |
|
305 |
done |
|
306 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
307 |
lemma rtrancl_subset: "\<lbrakk>R \<subseteq> S; S \<subseteq> R^*\<rbrakk> \<Longrightarrow> S^* = R^*" |
13239 | 308 |
apply (drule rtrancl_mono) |
13269 | 309 |
apply (drule rtrancl_mono, simp_all, blast) |
13239 | 310 |
done |
311 |
||
312 |
lemma rtrancl_Un_rtrancl: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
313 |
"\<lbrakk>relation(r); relation(s)\<rbrakk> \<Longrightarrow> (r^* \<union> s^*)^* = (r \<union> s)^*" |
13239 | 314 |
apply (rule rtrancl_subset) |
315 |
apply (blast dest: r_subset_rtrancl) |
|
316 |
apply (blast intro: rtrancl_mono [THEN subsetD]) |
|
317 |
done |
|
318 |
||
319 |
(*** "converse" laws by Sidi Ould Ehmety ***) |
|
320 |
||
321 |
(** rtrancl **) |
|
322 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
323 |
lemma rtrancl_converseD: "\<langle>x,y\<rangle>:converse(r)^* \<Longrightarrow> \<langle>x,y\<rangle>:converse(r^*)" |
13239 | 324 |
apply (rule converseI) |
325 |
apply (frule rtrancl_type [THEN subsetD]) |
|
326 |
apply (erule rtrancl_induct) |
|
327 |
apply (blast intro: rtrancl_refl) |
|
328 |
apply (blast intro: r_into_rtrancl rtrancl_trans) |
|
329 |
done |
|
330 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
331 |
lemma rtrancl_converseI: "\<langle>x,y\<rangle>:converse(r^*) \<Longrightarrow> \<langle>x,y\<rangle>:converse(r)^*" |
13239 | 332 |
apply (drule converseD) |
333 |
apply (frule rtrancl_type [THEN subsetD]) |
|
334 |
apply (erule rtrancl_induct) |
|
335 |
apply (blast intro: rtrancl_refl) |
|
336 |
apply (blast intro: r_into_rtrancl rtrancl_trans) |
|
337 |
done |
|
338 |
||
339 |
lemma rtrancl_converse: "converse(r)^* = converse(r^*)" |
|
340 |
apply (safe intro!: equalityI) |
|
341 |
apply (frule rtrancl_type [THEN subsetD]) |
|
342 |
apply (safe dest!: rtrancl_converseD intro!: rtrancl_converseI) |
|
343 |
done |
|
344 |
||
345 |
(** trancl **) |
|
346 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
347 |
lemma trancl_converseD: "\<langle>a, b\<rangle>:converse(r)^+ \<Longrightarrow> \<langle>a, b\<rangle>:converse(r^+)" |
13239 | 348 |
apply (erule trancl_induct) |
349 |
apply (auto intro: r_into_trancl trancl_trans) |
|
350 |
done |
|
351 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
352 |
lemma trancl_converseI: "\<langle>x,y\<rangle>:converse(r^+) \<Longrightarrow> \<langle>x,y\<rangle>:converse(r)^+" |
13239 | 353 |
apply (drule converseD) |
354 |
apply (erule trancl_induct) |
|
355 |
apply (auto intro: r_into_trancl trancl_trans) |
|
356 |
done |
|
357 |
||
358 |
lemma trancl_converse: "converse(r)^+ = converse(r^+)" |
|
359 |
apply (safe intro!: equalityI) |
|
360 |
apply (frule trancl_type [THEN subsetD]) |
|
361 |
apply (safe dest!: trancl_converseD intro!: trancl_converseI) |
|
362 |
done |
|
363 |
||
13534 | 364 |
lemma converse_trancl_induct [case_names initial step, consumes 1]: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
365 |
"\<lbrakk>\<langle>a, b\<rangle>:r^+; \<And>y. \<langle>y, b\<rangle> :r \<Longrightarrow> P(y); |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
366 |
\<And>y z. \<lbrakk>\<langle>y, z\<rangle> \<in> r; \<langle>z, b\<rangle> \<in> r^+; P(z)\<rbrakk> \<Longrightarrow> P(y)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
367 |
\<Longrightarrow> P(a)" |
13239 | 368 |
apply (drule converseI) |
369 |
apply (simp (no_asm_use) add: trancl_converse [symmetric]) |
|
370 |
apply (erule trancl_induct) |
|
371 |
apply (auto simp add: trancl_converse) |
|
372 |
done |
|
373 |
||
0 | 374 |
end |