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