| author | urbanc | 
| Thu, 10 Feb 2011 13:10:16 +0000 | |
| changeset 95 | 9540c2f2ea77 | 
| parent 94 | 5b12cd0a3b3c | 
| child 96 | 3b9deda4f459 | 
| permissions | -rw-r--r-- | 
| 42 | 1 | theory Myhill_1 | 
| 91 | 2 | imports Main Folds While_Combinator | 
| 42 | 3 | begin | 
| 4 | ||
| 5 | section {* Preliminary definitions *}
 | |
| 6 | ||
| 43 | 7 | types lang = "string set" | 
| 8 | ||
| 86 | 9 | |
| 70 | 10 | text {*  Sequential composition of two languages *}
 | 
| 43 | 11 | |
| 60 | 12 | definition | 
| 13 | Seq :: "lang \<Rightarrow> lang \<Rightarrow> lang" (infixr ";;" 100) | |
| 42 | 14 | where | 
| 54 | 15 |   "A ;; B = {s\<^isub>1 @ s\<^isub>2 | s\<^isub>1 s\<^isub>2. s\<^isub>1 \<in> A \<and> s\<^isub>2 \<in> B}"
 | 
| 42 | 16 | |
| 70 | 17 | |
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 18 | text {* Some properties of operator @{text ";;"}. *}
 | 
| 50 | 19 | |
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 20 | lemma seq_add_left: | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 21 | assumes a: "A = B" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 22 | shows "C ;; A = C ;; B" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 23 | using a by simp | 
| 42 | 24 | |
| 50 | 25 | lemma seq_union_distrib_right: | 
| 26 | shows "(A \<union> B) ;; C = (A ;; C) \<union> (B ;; C)" | |
| 27 | unfolding Seq_def by auto | |
| 28 | ||
| 29 | lemma seq_union_distrib_left: | |
| 30 | shows "C ;; (A \<union> B) = (C ;; A) \<union> (C ;; B)" | |
| 31 | unfolding Seq_def by auto | |
| 42 | 32 | |
| 33 | lemma seq_intro: | |
| 70 | 34 | assumes a: "x \<in> A" "y \<in> B" | 
| 35 | shows "x @ y \<in> A ;; B " | |
| 36 | using a by (auto simp: Seq_def) | |
| 42 | 37 | |
| 38 | lemma seq_assoc: | |
| 50 | 39 | shows "(A ;; B) ;; C = A ;; (B ;; C)" | 
| 40 | unfolding Seq_def | |
| 41 | apply(auto) | |
| 42 | apply(blast) | |
| 42 | 43 | by (metis append_assoc) | 
| 44 | ||
| 50 | 45 | lemma seq_empty [simp]: | 
| 46 |   shows "A ;; {[]} = A"
 | |
| 47 |   and   "{[]} ;; A = A"
 | |
| 48 | by (simp_all add: Seq_def) | |
| 49 | ||
| 70 | 50 | |
| 51 | text {* Power and Star of a language *}
 | |
| 52 | ||
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 53 | fun | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 54 | pow :: "lang \<Rightarrow> nat \<Rightarrow> lang" (infixl "\<up>" 100) | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 55 | where | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 56 |   "A \<up> 0 = {[]}"
 | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 57 | | "A \<up> (Suc n) = A ;; (A \<up> n)" | 
| 50 | 58 | |
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 59 | definition | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 60 |   Star :: "lang \<Rightarrow> lang" ("_\<star>" [101] 102)
 | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 61 | where | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 62 | "A\<star> \<equiv> (\<Union>n. A \<up> n)" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 63 | |
| 70 | 64 | |
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 65 | lemma star_start[intro]: | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 66 | shows "[] \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 67 | proof - | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 68 | have "[] \<in> A \<up> 0" by auto | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 69 | then show "[] \<in> A\<star>" unfolding Star_def by blast | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 70 | qed | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 71 | |
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 72 | lemma star_step [intro]: | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 73 | assumes a: "s1 \<in> A" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 74 | and b: "s2 \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 75 | shows "s1 @ s2 \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 76 | proof - | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 77 | from b obtain n where "s2 \<in> A \<up> n" unfolding Star_def by auto | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 78 | then have "s1 @ s2 \<in> A \<up> (Suc n)" using a by (auto simp add: Seq_def) | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 79 | then show "s1 @ s2 \<in> A\<star>" unfolding Star_def by blast | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 80 | qed | 
| 42 | 81 | |
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 82 | lemma star_induct[consumes 1, case_names start step]: | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 83 | assumes a: "x \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 84 | and b: "P []" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 85 | and c: "\<And>s1 s2. \<lbrakk>s1 \<in> A; s2 \<in> A\<star>; P s2\<rbrakk> \<Longrightarrow> P (s1 @ s2)" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 86 | shows "P x" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 87 | proof - | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 88 | from a obtain n where "x \<in> A \<up> n" unfolding Star_def by auto | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 89 | then show "P x" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 90 | by (induct n arbitrary: x) | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 91 | (auto intro!: b c simp add: Seq_def Star_def) | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 92 | qed | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 93 | |
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 94 | lemma star_intro1: | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 95 | assumes a: "x \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 96 | and b: "y \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 97 | shows "x @ y \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 98 | using a b | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 99 | by (induct rule: star_induct) (auto) | 
| 42 | 100 | |
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 101 | lemma star_intro2: | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 102 | assumes a: "y \<in> A" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 103 | shows "y \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 104 | proof - | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 105 | from a have "y @ [] \<in> A\<star>" by blast | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 106 | then show "y \<in> A\<star>" by simp | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 107 | qed | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 108 | |
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 109 | lemma star_intro3: | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 110 | assumes a: "x \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 111 | and b: "y \<in> A" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 112 | shows "x @ y \<in> A\<star>" | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 113 | using a b by (blast intro: star_intro1 star_intro2) | 
| 42 | 114 | |
| 71 | 115 | lemma star_cases: | 
| 116 |   shows "A\<star> =  {[]} \<union> A ;; A\<star>"
 | |
| 117 | proof | |
| 118 |   { fix x
 | |
| 119 |     have "x \<in> A\<star> \<Longrightarrow> x \<in> {[]} \<union> A ;; A\<star>"
 | |
| 120 | unfolding Seq_def | |
| 121 | by (induct rule: star_induct) (auto) | |
| 122 | } | |
| 123 |   then show "A\<star> \<subseteq> {[]} \<union> A ;; A\<star>" by auto
 | |
| 124 | next | |
| 125 |   show "{[]} \<union> A ;; A\<star> \<subseteq> A\<star>"
 | |
| 126 | unfolding Seq_def by auto | |
| 127 | qed | |
| 128 | ||
| 42 | 129 | lemma star_decom: | 
| 71 | 130 | assumes a: "x \<in> A\<star>" "x \<noteq> []" | 
| 131 | shows "\<exists>a b. x = a @ b \<and> a \<noteq> [] \<and> a \<in> A \<and> b \<in> A\<star>" | |
| 132 | using a | |
| 86 | 133 | by (induct rule: star_induct) (blast)+ | 
| 42 | 134 | |
| 50 | 135 | lemma | 
| 136 | shows seq_Union_left: "B ;; (\<Union>n. A \<up> n) = (\<Union>n. B ;; (A \<up> n))" | |
| 137 | and seq_Union_right: "(\<Union>n. A \<up> n) ;; B = (\<Union>n. (A \<up> n) ;; B)" | |
| 138 | unfolding Seq_def by auto | |
| 139 | ||
| 140 | lemma seq_pow_comm: | |
| 141 | shows "A ;; (A \<up> n) = (A \<up> n) ;; A" | |
| 142 | by (induct n) (simp_all add: seq_assoc[symmetric]) | |
| 143 | ||
| 144 | lemma seq_star_comm: | |
| 145 | shows "A ;; A\<star> = A\<star> ;; A" | |
| 86 | 146 | unfolding Star_def seq_Union_left | 
| 147 | unfolding seq_pow_comm seq_Union_right | |
| 50 | 148 | by simp | 
| 149 | ||
| 86 | 150 | |
| 50 | 151 | text {* Two lemmas about the length of strings in @{text "A \<up> n"} *}
 | 
| 152 | ||
| 153 | lemma pow_length: | |
| 154 | assumes a: "[] \<notin> A" | |
| 155 | and b: "s \<in> A \<up> Suc n" | |
| 156 | shows "n < length s" | |
| 157 | using b | |
| 158 | proof (induct n arbitrary: s) | |
| 159 | case 0 | |
| 160 | have "s \<in> A \<up> Suc 0" by fact | |
| 161 | with a have "s \<noteq> []" by auto | |
| 162 | then show "0 < length s" by auto | |
| 163 | next | |
| 164 | case (Suc n) | |
| 165 | have ih: "\<And>s. s \<in> A \<up> Suc n \<Longrightarrow> n < length s" by fact | |
| 166 | have "s \<in> A \<up> Suc (Suc n)" by fact | |
| 167 | then obtain s1 s2 where eq: "s = s1 @ s2" and *: "s1 \<in> A" and **: "s2 \<in> A \<up> Suc n" | |
| 168 | by (auto simp add: Seq_def) | |
| 169 | from ih ** have "n < length s2" by simp | |
| 170 | moreover have "0 < length s1" using * a by auto | |
| 171 | ultimately show "Suc n < length s" unfolding eq | |
| 172 | by (simp only: length_append) | |
| 173 | qed | |
| 174 | ||
| 175 | lemma seq_pow_length: | |
| 176 | assumes a: "[] \<notin> A" | |
| 177 | and b: "s \<in> B ;; (A \<up> Suc n)" | |
| 178 | shows "n < length s" | |
| 179 | proof - | |
| 180 | from b obtain s1 s2 where eq: "s = s1 @ s2" and *: "s2 \<in> A \<up> Suc n" | |
| 181 | unfolding Seq_def by auto | |
| 182 | from * have " n < length s2" by (rule pow_length[OF a]) | |
| 183 | then show "n < length s" using eq by simp | |
| 184 | qed | |
| 185 | ||
| 186 | ||
| 86 | 187 | |
| 188 | section {* A modified version of Arden's lemma *}
 | |
| 50 | 189 | |
| 70 | 190 | |
| 191 | text {*  A helper lemma for Arden *}
 | |
| 50 | 192 | |
| 86 | 193 | lemma arden_helper: | 
| 50 | 194 | assumes eq: "X = X ;; A \<union> B" | 
| 195 |   shows "X = X ;; (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))"
 | |
| 196 | proof (induct n) | |
| 197 | case 0 | |
| 198 |   show "X = X ;; (A \<up> Suc 0) \<union> (\<Union>(m::nat)\<in>{0..0}. B ;; (A \<up> m))"
 | |
| 199 | using eq by simp | |
| 200 | next | |
| 201 | case (Suc n) | |
| 202 |   have ih: "X = X ;; (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))" by fact
 | |
| 203 |   also have "\<dots> = (X ;; A \<union> B) ;; (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))" using eq by simp
 | |
| 204 |   also have "\<dots> = X ;; (A \<up> Suc (Suc n)) \<union> (B ;; (A \<up> Suc n)) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))"
 | |
| 205 | by (simp add: seq_union_distrib_right seq_assoc) | |
| 206 |   also have "\<dots> = X ;; (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B ;; (A \<up> m))"
 | |
| 207 | by (auto simp add: le_Suc_eq) | |
| 208 |   finally show "X = X ;; (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B ;; (A \<up> m))" .
 | |
| 209 | qed | |
| 210 | ||
| 86 | 211 | theorem arden: | 
| 50 | 212 | assumes nemp: "[] \<notin> A" | 
| 213 | shows "X = X ;; A \<union> B \<longleftrightarrow> X = B ;; A\<star>" | |
| 214 | proof | |
| 215 | assume eq: "X = B ;; A\<star>" | |
| 216 |   have "A\<star> = {[]} \<union> A\<star> ;; A" 
 | |
| 217 | unfolding seq_star_comm[symmetric] | |
| 71 | 218 | by (rule star_cases) | 
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 219 |   then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)"
 | 
| 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 220 | by (rule seq_add_left) | 
| 50 | 221 | also have "\<dots> = B \<union> B ;; (A\<star> ;; A)" | 
| 222 | unfolding seq_union_distrib_left by simp | |
| 223 | also have "\<dots> = B \<union> (B ;; A\<star>) ;; A" | |
| 224 | by (simp only: seq_assoc) | |
| 225 | finally show "X = X ;; A \<union> B" | |
| 226 | using eq by blast | |
| 227 | next | |
| 228 | assume eq: "X = X ;; A \<union> B" | |
| 229 |   { fix n::nat
 | |
| 86 | 230 | have "B ;; (A \<up> n) \<subseteq> X" using arden_helper[OF eq, of "n"] by auto } | 
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 231 | then have "B ;; A\<star> \<subseteq> X" | 
| 86 | 232 | unfolding Seq_def Star_def UNION_def by auto | 
| 50 | 233 | moreover | 
| 234 |   { fix s::string
 | |
| 235 | obtain k where "k = length s" by auto | |
| 236 | then have not_in: "s \<notin> X ;; (A \<up> Suc k)" | |
| 237 | using seq_pow_length[OF nemp] by blast | |
| 238 | assume "s \<in> X" | |
| 239 |     then have "s \<in> X ;; (A \<up> Suc k) \<union> (\<Union>m\<in>{0..k}. B ;; (A \<up> m))"
 | |
| 86 | 240 | using arden_helper[OF eq, of "k"] by auto | 
| 50 | 241 |     then have "s \<in> (\<Union>m\<in>{0..k}. B ;; (A \<up> m))" using not_in by auto
 | 
| 242 | moreover | |
| 243 |     have "(\<Union>m\<in>{0..k}. B ;; (A \<up> m)) \<subseteq> (\<Union>n. B ;; (A \<up> n))" by auto
 | |
| 244 | ultimately | |
| 56 
b3898315e687
removed the inductive definition of Star and replaced it by a definition in terms of pow
 urbanc parents: 
54diff
changeset | 245 | have "s \<in> B ;; A\<star>" | 
| 86 | 246 | unfolding seq_Union_left Star_def by auto } | 
| 50 | 247 | then have "X \<subseteq> B ;; A\<star>" by auto | 
| 248 | ultimately | |
| 249 | show "X = B ;; A\<star>" by simp | |
| 250 | qed | |
| 251 | ||
| 94 | 252 | (* | 
| 253 | corollary arden_eq: | |
| 254 | assumes nemp: "[] \<notin> A" | |
| 255 | shows "(X ;; A \<union> B) = (B ;; A\<star>)" | |
| 256 | proof - | |
| 257 |   { assume "X = X ;; A \<union> B"
 | |
| 258 | then have "X = B ;; A\<star>" | |
| 259 | then have ?thesis | |
| 260 | thm arden[THEN iffD1] | |
| 261 | apply(rule set_eqI) | |
| 262 | thm arden[THEN iffD1] | |
| 263 | apply(rule iffI) | |
| 264 | apply(rule trans) | |
| 265 | apply(rule arden[THEN iffD2, symmetric]) | |
| 266 | apply(rule arden[OF iffD1, symmetric]) | |
| 267 | thm trans | |
| 268 | proof - | |
| 269 |   { assume "X = X ;; A \<union> B"
 | |
| 270 | then have "X = B ;; A\<star>" using arden[OF nemp] by blast | |
| 271 | moreover | |
| 272 | ||
| 273 | ||
| 274 | using arden[of "A" "X" "B", OF nemp] | |
| 275 | apply(erule_tac iffE) | |
| 276 | apply(blast) | |
| 277 | *) | |
| 278 | ||
| 42 | 279 | |
| 70 | 280 | section {* Regular Expressions *}
 | 
| 48 | 281 | |
| 282 | datatype rexp = | |
| 283 | NULL | |
| 284 | | EMPTY | |
| 285 | | CHAR char | |
| 286 | | SEQ rexp rexp | |
| 287 | | ALT rexp rexp | |
| 288 | | STAR rexp | |
| 289 | ||
| 290 | ||
| 291 | text {* 
 | |
| 86 | 292 |   The function @{text L} is overloaded, with the idea that @{text "L x"} 
 | 
| 293 |   evaluates to the language represented by the object @{text x}.
 | |
| 48 | 294 | *} | 
| 295 | ||
| 70 | 296 | consts L:: "'a \<Rightarrow> lang" | 
| 48 | 297 | |
| 70 | 298 | overloading L_rexp \<equiv> "L:: rexp \<Rightarrow> lang" | 
| 48 | 299 | begin | 
| 300 | fun | |
| 88 | 301 | L_rexp :: "rexp \<Rightarrow> lang" | 
| 48 | 302 | where | 
| 303 |     "L_rexp (NULL) = {}"
 | |
| 304 |   | "L_rexp (EMPTY) = {[]}"
 | |
| 305 |   | "L_rexp (CHAR c) = {[c]}"
 | |
| 306 | | "L_rexp (SEQ r1 r2) = (L_rexp r1) ;; (L_rexp r2)" | |
| 307 | | "L_rexp (ALT r1 r2) = (L_rexp r1) \<union> (L_rexp r2)" | |
| 308 | | "L_rexp (STAR r) = (L_rexp r)\<star>" | |
| 309 | end | |
| 310 | ||
| 88 | 311 | |
| 86 | 312 | text {* ALT-combination of a set or regulare expressions *}
 | 
| 50 | 313 | |
| 76 | 314 | abbreviation | 
| 315 |   Setalt  ("\<Uplus>_" [1000] 999) 
 | |
| 316 | where | |
| 94 | 317 | "\<Uplus>A \<equiv> folds ALT NULL A" | 
| 76 | 318 | |
| 50 | 319 | text {* 
 | 
| 86 | 320 |   For finite sets, @{term Setalt} is preserved under @{term L}.
 | 
| 79 | 321 | *} | 
| 70 | 322 | |
| 50 | 323 | lemma folds_alt_simp [simp]: | 
| 88 | 324 | fixes rs::"rexp set" | 
| 70 | 325 | assumes a: "finite rs" | 
| 76 | 326 | shows "L (\<Uplus>rs) = \<Union> (L ` rs)" | 
| 94 | 327 | unfolding folds_def | 
| 75 | 328 | apply(rule set_eqI) | 
| 70 | 329 | apply(rule someI2_ex) | 
| 330 | apply(rule_tac finite_imp_fold_graph[OF a]) | |
| 331 | apply(erule fold_graph.induct) | |
| 332 | apply(auto) | |
| 333 | done | |
| 50 | 334 | |
| 70 | 335 | |
| 86 | 336 | |
| 337 | section {* Direction @{text "finite partition \<Rightarrow> regular language"} *}
 | |
| 338 | ||
| 339 | ||
| 70 | 340 | text {* Just a technical lemma for collections and pairs *}
 | 
| 341 | ||
| 75 | 342 | lemma Pair_Collect[simp]: | 
| 48 | 343 |   shows "(x, y) \<in> {(x, y). P x y} \<longleftrightarrow> P x y"
 | 
| 344 | by simp | |
| 345 | ||
| 86 | 346 | text {* Myhill-Nerode relation *}
 | 
| 347 | ||
| 48 | 348 | definition | 
| 71 | 349 |   str_eq_rel :: "lang \<Rightarrow> (string \<times> string) set" ("\<approx>_" [100] 100)
 | 
| 48 | 350 | where | 
| 70 | 351 |   "\<approx>A \<equiv> {(x, y).  (\<forall>z. x @ z \<in> A \<longleftrightarrow> y @ z \<in> A)}"
 | 
| 48 | 352 | |
| 353 | text {* 
 | |
| 86 | 354 |   Among the equivalence clases of @{text "\<approx>A"}, the set @{text "finals A"} 
 | 
| 355 |   singles out those which contains the strings from @{text A}.
 | |
| 48 | 356 | *} | 
| 357 | ||
| 358 | definition | |
| 71 | 359 | finals :: "lang \<Rightarrow> lang set" | 
| 360 | where | |
| 361 |   "finals A \<equiv> {\<approx>A `` {x} | x . x \<in> A}"
 | |
| 48 | 362 | |
| 70 | 363 | |
| 48 | 364 | lemma lang_is_union_of_finals: | 
| 70 | 365 | shows "A = \<Union> finals A" | 
| 366 | unfolding finals_def | |
| 367 | unfolding Image_def | |
| 368 | unfolding str_eq_rel_def | |
| 369 | apply(auto) | |
| 370 | apply(drule_tac x = "[]" in spec) | |
| 371 | apply(auto) | |
| 372 | done | |
| 373 | ||
| 79 | 374 | lemma finals_in_partitions: | 
| 375 | shows "finals A \<subseteq> (UNIV // \<approx>A)" | |
| 94 | 376 | unfolding finals_def quotient_def | 
| 76 | 377 | by auto | 
| 378 | ||
| 48 | 379 | |
| 86 | 380 | section {* Equational systems *}
 | 
| 42 | 381 | |
| 89 | 382 | |
| 383 | text {* The two kinds of terms in the rhs of equations. *}
 | |
| 384 | ||
| 42 | 385 | datatype rhs_item = | 
| 86 | 386 | Lam "rexp" (* Lambda-marker *) | 
| 70 | 387 | | Trn "lang" "rexp" (* Transition *) | 
| 388 | ||
| 42 | 389 | |
| 86 | 390 | overloading L_rhs_item \<equiv> "L:: rhs_item \<Rightarrow> lang" | 
| 42 | 391 | begin | 
| 86 | 392 | fun L_rhs_item:: "rhs_item \<Rightarrow> lang" | 
| 42 | 393 | where | 
| 86 | 394 | "L_rhs_item (Lam r) = L r" | 
| 395 | | "L_rhs_item (Trn X r) = X ;; L r" | |
| 42 | 396 | end | 
| 397 | ||
| 70 | 398 | overloading L_rhs \<equiv> "L:: rhs_item set \<Rightarrow> lang" | 
| 42 | 399 | begin | 
| 70 | 400 | fun L_rhs:: "rhs_item set \<Rightarrow> lang" | 
| 401 | where | |
| 402 | "L_rhs rhs = \<Union> (L ` rhs)" | |
| 42 | 403 | end | 
| 404 | ||
| 86 | 405 | text {* Transitions between equivalence classes *}
 | 
| 71 | 406 | |
| 407 | definition | |
| 92 | 408 |   transition :: "lang \<Rightarrow> char \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
 | 
| 71 | 409 | where | 
| 92 | 410 |   "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y ;; {[c]} \<subseteq> X"
 | 
| 42 | 411 | |
| 86 | 412 | text {* Initial equational system *}
 | 
| 413 | ||
| 42 | 414 | definition | 
| 415 | "init_rhs CS X \<equiv> | |
| 416 | if ([] \<in> X) then | |
| 92 | 417 |           {Lam EMPTY} \<union> {Trn Y (CHAR c) | Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}
 | 
| 42 | 418 | else | 
| 92 | 419 |           {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}"
 | 
| 42 | 420 | |
| 86 | 421 | definition | 
| 422 |   "eqs CS \<equiv> {(X, init_rhs CS X) | X.  X \<in> CS}"
 | |
| 75 | 423 | |
| 424 | ||
| 425 | ||
| 86 | 426 | section {* Arden Operation on equations *}
 | 
| 42 | 427 | |
| 428 | text {*
 | |
| 86 | 429 |   The function @{text "attach_rexp r item"} SEQ-composes @{text r} to the
 | 
| 430 | right of every rhs-item. | |
| 75 | 431 | *} | 
| 42 | 432 | |
| 70 | 433 | fun | 
| 92 | 434 | append_rexp :: "rexp \<Rightarrow> rhs_item \<Rightarrow> rhs_item" | 
| 42 | 435 | where | 
| 92 | 436 | "append_rexp r (Lam rexp) = Lam (SEQ rexp r)" | 
| 437 | | "append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)" | |
| 42 | 438 | |
| 439 | ||
| 440 | definition | |
| 92 | 441 | "append_rhs_rexp rhs rexp \<equiv> (append_rexp rexp) ` rhs" | 
| 42 | 442 | |
| 86 | 443 | definition | 
| 94 | 444 | "Arden X rhs \<equiv> | 
| 445 |      append_rhs_rexp (rhs - {Trn X r | r. Trn X r \<in> rhs}) (STAR (\<Uplus> {r. Trn X r \<in> rhs}))"
 | |
| 86 | 446 | |
| 447 | ||
| 448 | section {* Substitution Operation on equations *}
 | |
| 449 | ||
| 450 | text {* 
 | |
| 95 | 451 |   Suppose and equation @{text "X = xrhs"}, @{text "Subst"} substitutes 
 | 
| 86 | 452 |   all occurences of @{text "X"} in @{text "rhs"} by @{text "xrhs"}.
 | 
| 71 | 453 | *} | 
| 454 | ||
| 42 | 455 | definition | 
| 94 | 456 | "Subst rhs X xrhs \<equiv> | 
| 457 |         (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> (append_rhs_rexp xrhs (\<Uplus> {r. Trn X r \<in> rhs}))"
 | |
| 42 | 458 | |
| 459 | text {*
 | |
| 86 | 460 |   @{text "eqs_subst ES X xrhs"} substitutes @{text xrhs} into every 
 | 
| 461 |   equation of the equational system @{text ES}.
 | |
| 462 | *} | |
| 42 | 463 | |
| 464 | definition | |
| 94 | 465 |   "Subst_all ES X xrhs \<equiv> {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}"
 | 
| 86 | 466 | |
| 467 | ||
| 91 | 468 | section {* While-combinator *}
 | 
| 42 | 469 | |
| 470 | text {*
 | |
| 91 | 471 |   The following term @{text "remove ES Y yrhs"} removes the equation
 | 
| 472 |   @{text "Y = yrhs"} from equational system @{text "ES"} by replacing
 | |
| 473 |   all occurences of @{text "Y"} by its definition (using @{text "eqs_subst"}).
 | |
| 474 |   The @{text "Y"}-definition is made non-recursive using Arden's transformation
 | |
| 475 |   @{text "arden_variate Y yrhs"}.
 | |
| 476 | *} | |
| 477 | ||
| 478 | definition | |
| 95 | 479 | "Remove ES Y yrhs \<equiv> | 
| 94 | 480 |       Subst_all  (ES - {(Y, yrhs)}) Y (Arden Y yrhs)"
 | 
| 91 | 481 | |
| 482 | text {*
 | |
| 483 |   The following term @{text "iterm X ES"} represents one iteration in the while loop.
 | |
| 484 |   It arbitrarily chooses a @{text "Y"} different from @{text "X"} to remove.
 | |
| 71 | 485 | *} | 
| 42 | 486 | |
| 91 | 487 | definition | 
| 488 | "iter X ES \<equiv> (let (Y, yrhs) = SOME (Y, yrhs). (Y, yrhs) \<in> ES \<and> (X \<noteq> Y) | |
| 95 | 489 | in Remove ES Y yrhs)" | 
| 42 | 490 | |
| 491 | text {*
 | |
| 91 | 492 |   The following term @{text "reduce X ES"} repeatedly removes characteriztion equations
 | 
| 493 |   for unknowns other than @{text "X"} until one is left.
 | |
| 42 | 494 | *} | 
| 495 | ||
| 91 | 496 | definition | 
| 497 | "reduce X ES \<equiv> while (\<lambda> ES. card ES \<noteq> 1) (iter X) ES" | |
| 498 | ||
| 499 | text {*
 | |
| 500 |   Since the @{text "while"} combinator from HOL library is used to implement @{text "reduce X ES"},
 | |
| 501 |   the induction principle @{thm [source] while_rule} is used to proved the desired properties
 | |
| 502 |   of @{text "reduce X ES"}. For this purpose, an invariant predicate @{text "invariant"} is defined
 | |
| 503 | in terms of a series of auxilliary predicates: | |
| 504 | *} | |
| 86 | 505 | |
| 506 | section {* Invariants *}
 | |
| 507 | ||
| 508 | text {* Every variable is defined at most onece in @{text ES}. *}
 | |
| 75 | 509 | |
| 42 | 510 | definition | 
| 511 | "distinct_equas ES \<equiv> | |
| 86 | 512 | \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'" | 
| 70 | 513 | |
| 42 | 514 | text {* 
 | 
| 86 | 515 |   Every equation in @{text ES} (represented by @{text "(X, rhs)"}) 
 | 
| 516 |   is valid, i.e. @{text "(X = L rhs)"}.
 | |
| 517 | *} | |
| 518 | ||
| 42 | 519 | definition | 
| 520 | "valid_eqns ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> (X = L rhs)" | |
| 521 | ||
| 522 | text {*
 | |
| 86 | 523 |   @{text "rhs_nonempty rhs"} requires regular expressions occuring in 
 | 
| 524 |   transitional items of @{text "rhs"} do not contain empty string. This is 
 | |
| 525 |   necessary for the application of Arden's transformation to @{text "rhs"}.
 | |
| 526 | *} | |
| 70 | 527 | |
| 42 | 528 | definition | 
| 529 | "rhs_nonempty rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)" | |
| 530 | ||
| 531 | text {*
 | |
| 86 | 532 |   The following @{text "ardenable ES"} requires that Arden's transformation 
 | 
| 533 |   is applicable to every equation of equational system @{text "ES"}.
 | |
| 534 | *} | |
| 70 | 535 | |
| 42 | 536 | definition | 
| 537 | "ardenable ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> rhs_nonempty rhs" | |
| 538 | ||
| 86 | 539 | text {* 
 | 
| 540 |   @{text "finite_rhs ES"} requires every equation in @{text "rhs"} 
 | |
| 541 | be finite. | |
| 542 | *} | |
| 42 | 543 | definition | 
| 544 | "finite_rhs ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> finite rhs" | |
| 545 | ||
| 546 | text {*
 | |
| 86 | 547 |   @{text "classes_of rhs"} returns all variables (or equivalent classes)
 | 
| 42 | 548 |   occuring in @{text "rhs"}.
 | 
| 549 | *} | |
| 86 | 550 | |
| 42 | 551 | definition | 
| 552 |   "classes_of rhs \<equiv> {X. \<exists> r. Trn X r \<in> rhs}"
 | |
| 553 | ||
| 554 | text {*
 | |
| 86 | 555 |   @{text "lefts_of ES"} returns all variables defined by an 
 | 
| 556 |   equational system @{text "ES"}.
 | |
| 557 | *} | |
| 42 | 558 | definition | 
| 559 |   "lefts_of ES \<equiv> {Y | Y yrhs. (Y, yrhs) \<in> ES}"
 | |
| 560 | ||
| 561 | text {*
 | |
| 86 | 562 |   The following @{text "self_contained ES"} requires that every variable occuring 
 | 
| 563 |   on the right hand side of equations is already defined by some equation in @{text "ES"}.
 | |
| 564 | *} | |
| 42 | 565 | definition | 
| 566 | "self_contained ES \<equiv> \<forall> (X, xrhs) \<in> ES. classes_of xrhs \<subseteq> lefts_of ES" | |
| 567 | ||
| 568 | ||
| 569 | text {*
 | |
| 86 | 570 |   The invariant @{text "invariant(ES)"} is a conjunction of all the previously defined constaints.
 | 
| 42 | 571 | *} | 
| 572 | definition | |
| 86 | 573 | "invariant ES \<equiv> valid_eqns ES \<and> finite ES \<and> distinct_equas ES \<and> ardenable ES \<and> | 
| 87 | 574 | finite_rhs ES \<and> self_contained ES" | 
| 42 | 575 | |
| 576 | subsection {* The proof of this direction *}
 | |
| 577 | ||
| 578 | subsubsection {* Basic properties *}
 | |
| 579 | ||
| 580 | text {*
 | |
| 581 | The following are some basic properties of the above definitions. | |
| 582 | *} | |
| 583 | ||
| 584 | lemma L_rhs_union_distrib: | |
| 70 | 585 | fixes A B::"rhs_item set" | 
| 586 | shows "L A \<union> L B = L (A \<union> B)" | |
| 42 | 587 | by simp | 
| 588 | ||
| 79 | 589 | lemma finite_Trn: | 
| 590 | assumes fin: "finite rhs" | |
| 591 |   shows "finite {r. Trn Y r \<in> rhs}"
 | |
| 592 | proof - | |
| 593 |   have "finite {Trn Y r | Y r. Trn Y r \<in> rhs}"
 | |
| 594 | by (rule rev_finite_subset[OF fin]) (auto) | |
| 81 | 595 |   then have "finite ((\<lambda>(Y, r). Trn Y r) ` {(Y, r) | Y r. Trn Y r \<in> rhs})"
 | 
| 596 | by (simp add: image_Collect) | |
| 597 |   then have "finite {(Y, r) | Y r. Trn Y r \<in> rhs}"
 | |
| 598 | by (erule_tac finite_imageD) (simp add: inj_on_def) | |
| 79 | 599 |   then show "finite {r. Trn Y r \<in> rhs}"
 | 
| 81 | 600 | by (erule_tac f="snd" in finite_surj) (auto simp add: image_def) | 
| 79 | 601 | qed | 
| 602 | ||
| 603 | lemma finite_Lam: | |
| 604 | assumes fin:"finite rhs" | |
| 605 |   shows "finite {r. Lam r \<in> rhs}"
 | |
| 606 | proof - | |
| 607 |   have "finite {Lam r | r. Lam r \<in> rhs}"
 | |
| 608 | by (rule rev_finite_subset[OF fin]) (auto) | |
| 609 |   then show "finite {r. Lam r \<in> rhs}"
 | |
| 81 | 610 | apply(simp add: image_Collect[symmetric]) | 
| 611 | apply(erule finite_imageD) | |
| 612 | apply(auto simp add: inj_on_def) | |
| 79 | 613 | done | 
| 42 | 614 | qed | 
| 615 | ||
| 616 | lemma rexp_of_empty: | |
| 617 | assumes finite:"finite rhs" | |
| 618 | and nonempty:"rhs_nonempty rhs" | |
| 79 | 619 |   shows "[] \<notin> L (\<Uplus> {r. Trn X r \<in> rhs})"
 | 
| 42 | 620 | using finite nonempty rhs_nonempty_def | 
| 79 | 621 | using finite_Trn[OF finite] | 
| 622 | by (auto) | |
| 42 | 623 | |
| 624 | lemma [intro!]: | |
| 625 | "P (Trn X r) \<Longrightarrow> (\<exists>a. (\<exists>r. a = Trn X r \<and> P a))" by auto | |
| 626 | ||
| 627 | lemma lang_of_rexp_of: | |
| 628 | assumes finite:"finite rhs" | |
| 79 | 629 |   shows "L ({Trn X r| r. Trn X r \<in> rhs}) = X ;; (L (\<Uplus>{r. Trn X r \<in> rhs}))"
 | 
| 42 | 630 | proof - | 
| 79 | 631 |   have "finite {r. Trn X r \<in> rhs}" 
 | 
| 632 | by (rule finite_Trn[OF finite]) | |
| 633 | then show ?thesis | |
| 634 | apply(auto simp add: Seq_def) | |
| 635 | apply(rule_tac x = "s\<^isub>1" in exI, rule_tac x = "s\<^isub>2" in exI, auto) | |
| 636 | apply(rule_tac x= "Trn X xa" in exI) | |
| 637 | apply(auto simp: Seq_def) | |
| 638 | done | |
| 42 | 639 | qed | 
| 640 | ||
| 641 | lemma rexp_of_lam_eq_lam_set: | |
| 79 | 642 | assumes fin: "finite rhs" | 
| 643 |   shows "L (\<Uplus>{r. Lam r \<in> rhs}) = L ({Lam r | r. Lam r \<in> rhs})"
 | |
| 42 | 644 | proof - | 
| 79 | 645 |   have "finite ({r. Lam r \<in> rhs})" using fin by (rule finite_Lam)
 | 
| 646 | then show ?thesis by auto | |
| 42 | 647 | qed | 
| 648 | ||
| 649 | lemma [simp]: | |
| 92 | 650 | "L (append_rexp r xb) = L xb ;; L r" | 
| 79 | 651 | apply (cases xb, auto simp: Seq_def) | 
| 54 | 652 | apply(rule_tac x = "s\<^isub>1 @ s\<^isub>1'" in exI, rule_tac x = "s\<^isub>2'" in exI) | 
| 653 | apply(auto simp: Seq_def) | |
| 654 | done | |
| 42 | 655 | |
| 656 | lemma lang_of_append_rhs: | |
| 657 | "L (append_rhs_rexp rhs r) = L rhs ;; L r" | |
| 658 | apply (auto simp:append_rhs_rexp_def image_def) | |
| 659 | apply (auto simp:Seq_def) | |
| 660 | apply (rule_tac x = "L xb ;; L r" in exI, auto simp add:Seq_def) | |
| 92 | 661 | by (rule_tac x = "append_rexp r xb" in exI, auto simp:Seq_def) | 
| 42 | 662 | |
| 663 | lemma classes_of_union_distrib: | |
| 664 | "classes_of A \<union> classes_of B = classes_of (A \<union> B)" | |
| 665 | by (auto simp add:classes_of_def) | |
| 666 | ||
| 667 | lemma lefts_of_union_distrib: | |
| 668 | "lefts_of A \<union> lefts_of B = lefts_of (A \<union> B)" | |
| 669 | by (auto simp:lefts_of_def) | |
| 670 | ||
| 671 | ||
| 672 | subsubsection {* Intialization *}
 | |
| 673 | ||
| 674 | text {*
 | |
| 86 | 675 |   The following several lemmas until @{text "init_ES_satisfy_invariant"} shows that
 | 
| 676 |   the initial equational system satisfies invariant @{text "invariant"}.
 | |
| 71 | 677 | *} | 
| 42 | 678 | |
| 679 | lemma defined_by_str: | |
| 680 |   "\<lbrakk>s \<in> X; X \<in> UNIV // (\<approx>Lang)\<rbrakk> \<Longrightarrow> X = (\<approx>Lang) `` {s}"
 | |
| 681 | by (auto simp:quotient_def Image_def str_eq_rel_def) | |
| 682 | ||
| 683 | lemma every_eqclass_has_transition: | |
| 684 | assumes has_str: "s @ [c] \<in> X" | |
| 685 | and in_CS: "X \<in> UNIV // (\<approx>Lang)" | |
| 686 |   obtains Y where "Y \<in> UNIV // (\<approx>Lang)" and "Y ;; {[c]} \<subseteq> X" and "s \<in> Y"
 | |
| 687 | proof - | |
| 688 |   def Y \<equiv> "(\<approx>Lang) `` {s}"
 | |
| 689 | have "Y \<in> UNIV // (\<approx>Lang)" | |
| 690 | unfolding Y_def quotient_def by auto | |
| 691 | moreover | |
| 692 |   have "X = (\<approx>Lang) `` {s @ [c]}" 
 | |
| 693 | using has_str in_CS defined_by_str by blast | |
| 694 |   then have "Y ;; {[c]} \<subseteq> X" 
 | |
| 695 | unfolding Y_def Image_def Seq_def | |
| 696 | unfolding str_eq_rel_def | |
| 697 | by clarsimp | |
| 698 | moreover | |
| 699 | have "s \<in> Y" unfolding Y_def | |
| 700 | unfolding Image_def str_eq_rel_def by simp | |
| 701 | ultimately show thesis by (blast intro: that) | |
| 702 | qed | |
| 703 | ||
| 704 | lemma l_eq_r_in_eqs: | |
| 705 | assumes X_in_eqs: "(X, xrhs) \<in> (eqs (UNIV // (\<approx>Lang)))" | |
| 706 | shows "X = L xrhs" | |
| 707 | proof | |
| 708 | show "X \<subseteq> L xrhs" | |
| 709 | proof | |
| 710 | fix x | |
| 711 | assume "(1)": "x \<in> X" | |
| 712 | show "x \<in> L xrhs" | |
| 713 | proof (cases "x = []") | |
| 714 | assume empty: "x = []" | |
| 715 | thus ?thesis using X_in_eqs "(1)" | |
| 716 | by (auto simp:eqs_def init_rhs_def) | |
| 717 | next | |
| 718 | assume not_empty: "x \<noteq> []" | |
| 719 | then obtain clist c where decom: "x = clist @ [c]" | |
| 720 | by (case_tac x rule:rev_cases, auto) | |
| 721 | have "X \<in> UNIV // (\<approx>Lang)" using X_in_eqs by (auto simp:eqs_def) | |
| 722 | then obtain Y | |
| 723 | where "Y \<in> UNIV // (\<approx>Lang)" | |
| 724 |         and "Y ;; {[c]} \<subseteq> X"
 | |
| 725 | and "clist \<in> Y" | |
| 726 | using decom "(1)" every_eqclass_has_transition by blast | |
| 727 | hence | |
| 92 | 728 |         "x \<in> L {Trn Y (CHAR c)| Y c. Y \<in> UNIV // (\<approx>Lang) \<and> Y \<Turnstile>c\<Rightarrow> X}"
 | 
| 71 | 729 | unfolding transition_def | 
| 730 | using "(1)" decom | |
| 42 | 731 | by (simp, rule_tac x = "Trn Y (CHAR c)" in exI, simp add:Seq_def) | 
| 71 | 732 | thus ?thesis using X_in_eqs "(1)" | 
| 733 | by (simp add: eqs_def init_rhs_def) | |
| 42 | 734 | qed | 
| 735 | qed | |
| 736 | next | |
| 737 | show "L xrhs \<subseteq> X" using X_in_eqs | |
| 71 | 738 | by (auto simp:eqs_def init_rhs_def transition_def) | 
| 42 | 739 | qed | 
| 740 | ||
| 741 | lemma finite_init_rhs: | |
| 742 | assumes finite: "finite CS" | |
| 743 | shows "finite (init_rhs CS X)" | |
| 744 | proof- | |
| 745 |   have "finite {Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" (is "finite ?A")
 | |
| 746 | proof - | |
| 747 |     def S \<equiv> "{(Y, c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" 
 | |
| 748 | def h \<equiv> "\<lambda> (Y, c). Trn Y (CHAR c)" | |
| 749 | have "finite (CS \<times> (UNIV::char set))" using finite by auto | |
| 750 | hence "finite S" using S_def | |
| 751 | by (rule_tac B = "CS \<times> UNIV" in finite_subset, auto) | |
| 752 | moreover have "?A = h ` S" by (auto simp: S_def h_def image_def) | |
| 753 | ultimately show ?thesis | |
| 754 | by auto | |
| 755 | qed | |
| 71 | 756 | thus ?thesis by (simp add:init_rhs_def transition_def) | 
| 42 | 757 | qed | 
| 758 | ||
| 86 | 759 | lemma init_ES_satisfy_invariant: | 
| 42 | 760 | assumes finite_CS: "finite (UNIV // (\<approx>Lang))" | 
| 86 | 761 | shows "invariant (eqs (UNIV // (\<approx>Lang)))" | 
| 42 | 762 | proof - | 
| 763 | have "finite (eqs (UNIV // (\<approx>Lang)))" using finite_CS | |
| 764 | by (simp add:eqs_def) | |
| 765 | moreover have "distinct_equas (eqs (UNIV // (\<approx>Lang)))" | |
| 766 | by (simp add:distinct_equas_def eqs_def) | |
| 767 | moreover have "ardenable (eqs (UNIV // (\<approx>Lang)))" | |
| 768 | by (auto simp add:ardenable_def eqs_def init_rhs_def rhs_nonempty_def del:L_rhs.simps) | |
| 769 | moreover have "valid_eqns (eqs (UNIV // (\<approx>Lang)))" | |
| 770 | using l_eq_r_in_eqs by (simp add:valid_eqns_def) | |
| 771 | moreover have "finite_rhs (eqs (UNIV // (\<approx>Lang)))" | |
| 772 | using finite_init_rhs[OF finite_CS] | |
| 773 | by (auto simp:finite_rhs_def eqs_def) | |
| 774 | moreover have "self_contained (eqs (UNIV // (\<approx>Lang)))" | |
| 775 | by (auto simp:self_contained_def eqs_def init_rhs_def classes_of_def lefts_of_def) | |
| 86 | 776 | ultimately show ?thesis by (simp add:invariant_def) | 
| 42 | 777 | qed | 
| 778 | ||
| 91 | 779 | subsubsection {* Interation step *}
 | 
| 42 | 780 | |
| 781 | text {*
 | |
| 91 | 782 |   From this point until @{text "iteration_step"}, 
 | 
| 783 |   the correctness of the iteration step @{text "iter X ES"} is proved.
 | |
| 71 | 784 | *} | 
| 785 | ||
| 94 | 786 | lemma Arden_keeps_eq: | 
| 42 | 787 | assumes l_eq_r: "X = L rhs" | 
| 79 | 788 |   and not_empty: "[] \<notin> L (\<Uplus>{r. Trn X r \<in> rhs})"
 | 
| 42 | 789 | and finite: "finite rhs" | 
| 94 | 790 | shows "X = L (Arden X rhs)" | 
| 42 | 791 | proof - | 
| 79 | 792 |   def A \<equiv> "L (\<Uplus>{r. Trn X r \<in> rhs})"
 | 
| 94 | 793 |   def b \<equiv> "rhs - {Trn X r | r. Trn X r \<in> rhs}"
 | 
| 42 | 794 | def B \<equiv> "L b" | 
| 795 | have "X = B ;; A\<star>" | |
| 796 | proof- | |
| 94 | 797 |     have "L rhs = L({Trn X r | r. Trn X r \<in> rhs} \<union> b)" by (auto simp: b_def)
 | 
| 79 | 798 | also have "\<dots> = X ;; A \<union> B" | 
| 799 | unfolding L_rhs_union_distrib[symmetric] | |
| 800 | by (simp only: lang_of_rexp_of finite B_def A_def) | |
| 801 | finally show ?thesis | |
| 42 | 802 | using l_eq_r not_empty | 
| 86 | 803 | apply(rule_tac arden[THEN iffD1]) | 
| 79 | 804 | apply(simp add: A_def) | 
| 805 | apply(simp) | |
| 806 | done | |
| 42 | 807 | qed | 
| 94 | 808 | moreover have "L (Arden X rhs) = B ;; A\<star>" | 
| 809 | by (simp only:Arden_def L_rhs_union_distrib lang_of_append_rhs | |
| 50 | 810 | B_def A_def b_def L_rexp.simps seq_union_distrib_left) | 
| 42 | 811 | ultimately show ?thesis by simp | 
| 812 | qed | |
| 813 | ||
| 814 | lemma append_keeps_finite: | |
| 815 | "finite rhs \<Longrightarrow> finite (append_rhs_rexp rhs r)" | |
| 816 | by (auto simp:append_rhs_rexp_def) | |
| 817 | ||
| 94 | 818 | lemma Arden_keeps_finite: | 
| 819 | "finite rhs \<Longrightarrow> finite (Arden X rhs)" | |
| 820 | by (auto simp:Arden_def append_keeps_finite) | |
| 42 | 821 | |
| 822 | lemma append_keeps_nonempty: | |
| 823 | "rhs_nonempty rhs \<Longrightarrow> rhs_nonempty (append_rhs_rexp rhs r)" | |
| 824 | apply (auto simp:rhs_nonempty_def append_rhs_rexp_def) | |
| 825 | by (case_tac x, auto simp:Seq_def) | |
| 826 | ||
| 827 | lemma nonempty_set_sub: | |
| 828 | "rhs_nonempty rhs \<Longrightarrow> rhs_nonempty (rhs - A)" | |
| 829 | by (auto simp:rhs_nonempty_def) | |
| 830 | ||
| 831 | lemma nonempty_set_union: | |
| 832 | "\<lbrakk>rhs_nonempty rhs; rhs_nonempty rhs'\<rbrakk> \<Longrightarrow> rhs_nonempty (rhs \<union> rhs')" | |
| 833 | by (auto simp:rhs_nonempty_def) | |
| 834 | ||
| 94 | 835 | lemma Arden_keeps_nonempty: | 
| 836 | "rhs_nonempty rhs \<Longrightarrow> rhs_nonempty (Arden X rhs)" | |
| 837 | by (simp only:Arden_def append_keeps_nonempty nonempty_set_sub) | |
| 42 | 838 | |
| 839 | ||
| 94 | 840 | lemma Subst_keeps_nonempty: | 
| 841 | "\<lbrakk>rhs_nonempty rhs; rhs_nonempty xrhs\<rbrakk> \<Longrightarrow> rhs_nonempty (Subst rhs X xrhs)" | |
| 842 | by (simp only:Subst_def append_keeps_nonempty nonempty_set_union nonempty_set_sub) | |
| 42 | 843 | |
| 94 | 844 | lemma Subst_keeps_eq: | 
| 42 | 845 | assumes substor: "X = L xrhs" | 
| 846 | and finite: "finite rhs" | |
| 94 | 847 | shows "L (Subst rhs X xrhs) = L rhs" (is "?Left = ?Right") | 
| 42 | 848 | proof- | 
| 94 | 849 |   def A \<equiv> "L (rhs - {Trn X r | r. Trn X r \<in> rhs})"
 | 
| 79 | 850 |   have "?Left = A \<union> L (append_rhs_rexp xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
 | 
| 94 | 851 | unfolding Subst_def | 
| 79 | 852 | unfolding L_rhs_union_distrib[symmetric] | 
| 853 | by (simp add: A_def) | |
| 854 |   moreover have "?Right = A \<union> L ({Trn X r | r. Trn X r \<in> rhs})"
 | |
| 42 | 855 | proof- | 
| 94 | 856 |     have "rhs = (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> ({Trn X r | r. Trn X r \<in> rhs})" by auto
 | 
| 79 | 857 | thus ?thesis | 
| 858 | unfolding A_def | |
| 859 | unfolding L_rhs_union_distrib | |
| 860 | by simp | |
| 42 | 861 | qed | 
| 79 | 862 |   moreover have "L (append_rhs_rexp xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L ({Trn X r | r. Trn X r \<in> rhs})" 
 | 
| 42 | 863 | using finite substor by (simp only:lang_of_append_rhs lang_of_rexp_of) | 
| 864 | ultimately show ?thesis by simp | |
| 865 | qed | |
| 866 | ||
| 94 | 867 | lemma Subst_keeps_finite_rhs: | 
| 868 | "\<lbrakk>finite rhs; finite yrhs\<rbrakk> \<Longrightarrow> finite (Subst rhs Y yrhs)" | |
| 869 | by (auto simp:Subst_def append_keeps_finite) | |
| 42 | 870 | |
| 94 | 871 | lemma Subst_all_keeps_finite: | 
| 42 | 872 | assumes finite:"finite (ES:: (string set \<times> rhs_item set) set)" | 
| 94 | 873 | shows "finite (Subst_all ES Y yrhs)" | 
| 42 | 874 | proof - | 
| 94 | 875 |   have "finite {(Ya, Subst yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \<in> ES}" 
 | 
| 42 | 876 | (is "finite ?A") | 
| 877 | proof- | |
| 878 |     def eqns' \<equiv> "{((Ya::string set), yrhsa)| Ya yrhsa. (Ya, yrhsa) \<in> ES}"
 | |
| 94 | 879 | def h \<equiv> "\<lambda> ((Ya::string set), yrhsa). (Ya, Subst yrhsa Y yrhs)" | 
| 42 | 880 | have "finite (h ` eqns')" using finite h_def eqns'_def by auto | 
| 881 | moreover have "?A = h ` eqns'" by (auto simp:h_def eqns'_def) | |
| 882 | ultimately show ?thesis by auto | |
| 883 | qed | |
| 94 | 884 | thus ?thesis by (simp add:Subst_all_def) | 
| 42 | 885 | qed | 
| 886 | ||
| 94 | 887 | lemma Subst_all_keeps_finite_rhs: | 
| 888 | "\<lbrakk>finite_rhs ES; finite yrhs\<rbrakk> \<Longrightarrow> finite_rhs (Subst_all ES Y yrhs)" | |
| 889 | by (auto intro:Subst_keeps_finite_rhs simp add:Subst_all_def finite_rhs_def) | |
| 42 | 890 | |
| 891 | lemma append_rhs_keeps_cls: | |
| 892 | "classes_of (append_rhs_rexp rhs r) = classes_of rhs" | |
| 893 | apply (auto simp:classes_of_def append_rhs_rexp_def) | |
| 894 | apply (case_tac xa, auto simp:image_def) | |
| 895 | by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "Trn x ra" in bexI, simp+) | |
| 896 | ||
| 94 | 897 | lemma Arden_removes_cl: | 
| 898 |   "classes_of (Arden Y yrhs) = classes_of yrhs - {Y}"
 | |
| 899 | apply (simp add:Arden_def append_rhs_keeps_cls) | |
| 42 | 900 | by (auto simp:classes_of_def) | 
| 901 | ||
| 902 | lemma lefts_of_keeps_cls: | |
| 94 | 903 | "lefts_of (Subst_all ES Y yrhs) = lefts_of ES" | 
| 904 | by (auto simp:lefts_of_def Subst_all_def) | |
| 42 | 905 | |
| 94 | 906 | lemma Subst_updates_cls: | 
| 42 | 907 | "X \<notin> classes_of xrhs \<Longrightarrow> | 
| 94 | 908 |       classes_of (Subst rhs X xrhs) = classes_of rhs \<union> classes_of xrhs - {X}"
 | 
| 909 | apply (simp only:Subst_def append_rhs_keeps_cls | |
| 42 | 910 | classes_of_union_distrib[THEN sym]) | 
| 94 | 911 | by (auto simp:classes_of_def) | 
| 42 | 912 | |
| 94 | 913 | lemma Subst_all_keeps_self_contained: | 
| 42 | 914 |   assumes sc: "self_contained (ES \<union> {(Y, yrhs)})" (is "self_contained ?A")
 | 
| 94 | 915 | shows "self_contained (Subst_all ES Y (Arden Y yrhs))" | 
| 42 | 916 | (is "self_contained ?B") | 
| 917 | proof- | |
| 918 |   { fix X xrhs'
 | |
| 919 | assume "(X, xrhs') \<in> ?B" | |
| 920 | then obtain xrhs | |
| 94 | 921 | where xrhs_xrhs': "xrhs' = Subst xrhs Y (Arden Y yrhs)" | 
| 922 | and X_in: "(X, xrhs) \<in> ES" by (simp add:Subst_all_def, blast) | |
| 42 | 923 | have "classes_of xrhs' \<subseteq> lefts_of ?B" | 
| 924 | proof- | |
| 94 | 925 | have "lefts_of ?B = lefts_of ES" by (auto simp add:lefts_of_def Subst_all_def) | 
| 42 | 926 | moreover have "classes_of xrhs' \<subseteq> lefts_of ES" | 
| 927 | proof- | |
| 928 | have "classes_of xrhs' \<subseteq> | |
| 94 | 929 |                         classes_of xrhs \<union> classes_of (Arden Y yrhs) - {Y}"
 | 
| 42 | 930 | proof- | 
| 94 | 931 | have "Y \<notin> classes_of (Arden Y yrhs)" | 
| 932 | using Arden_removes_cl by simp | |
| 933 | thus ?thesis using xrhs_xrhs' by (auto simp:Subst_updates_cls) | |
| 42 | 934 | qed | 
| 935 |         moreover have "classes_of xrhs \<subseteq> lefts_of ES \<union> {Y}" using X_in sc
 | |
| 936 | apply (simp only:self_contained_def lefts_of_union_distrib[THEN sym]) | |
| 937 | by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lefts_of_def) | |
| 94 | 938 |         moreover have "classes_of (Arden Y yrhs) \<subseteq> lefts_of ES \<union> {Y}" 
 | 
| 42 | 939 | using sc | 
| 94 | 940 | by (auto simp add:Arden_removes_cl self_contained_def lefts_of_def) | 
| 42 | 941 | ultimately show ?thesis by auto | 
| 942 | qed | |
| 943 | ultimately show ?thesis by simp | |
| 944 | qed | |
| 94 | 945 | } thus ?thesis by (auto simp only:Subst_all_def self_contained_def) | 
| 42 | 946 | qed | 
| 947 | ||
| 94 | 948 | lemma Subst_all_satisfy_invariant: | 
| 86 | 949 |   assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})"
 | 
| 94 | 950 | shows "invariant (Subst_all ES Y (Arden Y yrhs))" | 
| 42 | 951 | proof - | 
| 952 | have finite_yrhs: "finite yrhs" | |
| 86 | 953 | using invariant_ES by (auto simp:invariant_def finite_rhs_def) | 
| 42 | 954 | have nonempty_yrhs: "rhs_nonempty yrhs" | 
| 86 | 955 | using invariant_ES by (auto simp:invariant_def ardenable_def) | 
| 42 | 956 | have Y_eq_yrhs: "Y = L yrhs" | 
| 86 | 957 | using invariant_ES by (simp only:invariant_def valid_eqns_def, blast) | 
| 94 | 958 | have "distinct_equas (Subst_all ES Y (Arden Y yrhs))" | 
| 86 | 959 | using invariant_ES | 
| 94 | 960 | by (auto simp:distinct_equas_def Subst_all_def invariant_def) | 
| 961 | moreover have "finite (Subst_all ES Y (Arden Y yrhs))" | |
| 962 | using invariant_ES by (simp add:invariant_def Subst_all_keeps_finite) | |
| 963 | moreover have "finite_rhs (Subst_all ES Y (Arden Y yrhs))" | |
| 42 | 964 | proof- | 
| 86 | 965 | have "finite_rhs ES" using invariant_ES | 
| 966 | by (simp add:invariant_def finite_rhs_def) | |
| 94 | 967 | moreover have "finite (Arden Y yrhs)" | 
| 42 | 968 | proof - | 
| 86 | 969 | have "finite yrhs" using invariant_ES | 
| 970 | by (auto simp:invariant_def finite_rhs_def) | |
| 94 | 971 | thus ?thesis using Arden_keeps_finite by simp | 
| 42 | 972 | qed | 
| 973 | ultimately show ?thesis | |
| 94 | 974 | by (simp add:Subst_all_keeps_finite_rhs) | 
| 42 | 975 | qed | 
| 94 | 976 | moreover have "ardenable (Subst_all ES Y (Arden Y yrhs))" | 
| 42 | 977 | proof - | 
| 978 |     { fix X rhs
 | |
| 979 | assume "(X, rhs) \<in> ES" | |
| 86 | 980 | hence "rhs_nonempty rhs" using prems invariant_ES | 
| 981 | by (simp add:invariant_def ardenable_def) | |
| 42 | 982 | with nonempty_yrhs | 
| 94 | 983 | have "rhs_nonempty (Subst rhs Y (Arden Y yrhs))" | 
| 42 | 984 | by (simp add:nonempty_yrhs | 
| 94 | 985 | Subst_keeps_nonempty Arden_keeps_nonempty) | 
| 986 | } thus ?thesis by (auto simp add:ardenable_def Subst_all_def) | |
| 42 | 987 | qed | 
| 94 | 988 | moreover have "valid_eqns (Subst_all ES Y (Arden Y yrhs))" | 
| 42 | 989 | proof- | 
| 94 | 990 | have "Y = L (Arden Y yrhs)" | 
| 86 | 991 | using Y_eq_yrhs invariant_ES finite_yrhs nonempty_yrhs | 
| 94 | 992 | by (rule_tac Arden_keeps_eq, (simp add:rexp_of_empty)+) | 
| 86 | 993 | thus ?thesis using invariant_ES | 
| 42 | 994 | by (clarsimp simp add:valid_eqns_def | 
| 94 | 995 | Subst_all_def Subst_keeps_eq invariant_def finite_rhs_def | 
| 42 | 996 | simp del:L_rhs.simps) | 
| 997 | qed | |
| 998 | moreover | |
| 94 | 999 | have self_subst: "self_contained (Subst_all ES Y (Arden Y yrhs))" | 
| 1000 | using invariant_ES Subst_all_keeps_self_contained by (simp add:invariant_def) | |
| 86 | 1001 | ultimately show ?thesis using invariant_ES by (simp add:invariant_def) | 
| 42 | 1002 | qed | 
| 1003 | ||
| 94 | 1004 | lemma Subst_all_card_le: | 
| 42 | 1005 | assumes finite: "finite (ES::(string set \<times> rhs_item set) set)" | 
| 94 | 1006 | shows "card (Subst_all ES Y yrhs) <= card ES" | 
| 42 | 1007 | proof- | 
| 94 | 1008 | def f \<equiv> "\<lambda> x. ((fst x)::string set, Subst (snd x) Y yrhs)" | 
| 1009 | have "Subst_all ES Y yrhs = f ` ES" | |
| 1010 | apply (auto simp:Subst_all_def f_def image_def) | |
| 42 | 1011 | by (rule_tac x = "(Ya, yrhsa)" in bexI, simp+) | 
| 1012 | thus ?thesis using finite by (auto intro:card_image_le) | |
| 1013 | qed | |
| 1014 | ||
| 94 | 1015 | lemma Subst_all_cls_remains: | 
| 1016 | "(X, xrhs) \<in> ES \<Longrightarrow> \<exists> xrhs'. (X, xrhs') \<in> (Subst_all ES Y yrhs)" | |
| 1017 | by (auto simp:Subst_all_def) | |
| 42 | 1018 | |
| 1019 | lemma card_noteq_1_has_more: | |
| 1020 | assumes card:"card S \<noteq> 1" | |
| 1021 | and e_in: "e \<in> S" | |
| 1022 | and finite: "finite S" | |
| 1023 | obtains e' where "e' \<in> S \<and> e \<noteq> e'" | |
| 1024 | proof- | |
| 1025 |   have "card (S - {e}) > 0"
 | |
| 1026 | proof - | |
| 1027 | have "card S > 1" using card e_in finite | |
| 1028 | by (case_tac "card S", auto) | |
| 1029 | thus ?thesis using finite e_in by auto | |
| 1030 | qed | |
| 1031 |   hence "S - {e} \<noteq> {}" using finite by (rule_tac notI, simp)
 | |
| 1032 | thus "(\<And>e'. e' \<in> S \<and> e \<noteq> e' \<Longrightarrow> thesis) \<Longrightarrow> thesis" by auto | |
| 1033 | qed | |
| 1034 | ||
| 91 | 1035 | lemma iteration_step: | 
| 1036 | assumes Inv_ES: "invariant ES" | |
| 42 | 1037 | and X_in_ES: "(X, xrhs) \<in> ES" | 
| 1038 | and not_T: "card ES \<noteq> 1" | |
| 91 | 1039 | shows "(invariant (iter X ES) \<and> (\<exists> xrhs'.(X, xrhs') \<in> (iter X ES)) \<and> | 
| 1040 | (iter X ES, ES) \<in> measure card)" | |
| 42 | 1041 | proof - | 
| 91 | 1042 | have finite_ES: "finite ES" using Inv_ES by (simp add: invariant_def) | 
| 42 | 1043 | then obtain Y yrhs | 
| 1044 | where Y_in_ES: "(Y, yrhs) \<in> ES" and not_eq: "(X, xrhs) \<noteq> (Y, yrhs)" | |
| 1045 | using not_T X_in_ES by (drule_tac card_noteq_1_has_more, auto) | |
| 91 | 1046 | let ?ES' = "iter X ES" | 
| 1047 | show ?thesis | |
| 95 | 1048 | proof(unfold iter_def Remove_def, rule someI2 [where a = "(Y, yrhs)"], clarsimp) | 
| 91 | 1049 | from X_in_ES Y_in_ES and not_eq and Inv_ES | 
| 1050 | show "(Y, yrhs) \<in> ES \<and> X \<noteq> Y" | |
| 1051 | by (auto simp: invariant_def distinct_equas_def) | |
| 1052 | next | |
| 1053 | fix x | |
| 94 | 1054 |     let ?ES' = "let (Y, yrhs) = x in Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)"
 | 
| 91 | 1055 | assume prem: "case x of (Y, yrhs) \<Rightarrow> (Y, yrhs) \<in> ES \<and> (X \<noteq> Y)" | 
| 1056 | thus "invariant (?ES') \<and> (\<exists>xrhs'. (X, xrhs') \<in> ?ES') \<and> (?ES', ES) \<in> measure card" | |
| 1057 | proof(cases "x", simp) | |
| 1058 | fix Y yrhs | |
| 1059 | assume h: "(Y, yrhs) \<in> ES \<and> X \<noteq> Y" | |
| 94 | 1060 |       show "invariant (Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)) \<and>
 | 
| 1061 |              (\<exists>xrhs'. (X, xrhs') \<in> Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)) \<and>
 | |
| 1062 |              card (Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)) < card ES"
 | |
| 91 | 1063 | proof - | 
| 94 | 1064 |         have "invariant (Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs))" 
 | 
| 1065 | proof(rule Subst_all_satisfy_invariant) | |
| 91 | 1066 | from h have "(Y, yrhs) \<in> ES" by simp | 
| 1067 |           hence "ES - {(Y, yrhs)} \<union> {(Y, yrhs)} = ES" by auto
 | |
| 1068 |           with Inv_ES show "invariant (ES - {(Y, yrhs)} \<union> {(Y, yrhs)})" by auto
 | |
| 1069 | qed | |
| 1070 | moreover have | |
| 94 | 1071 |           "(\<exists>xrhs'. (X, xrhs') \<in> Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs))"
 | 
| 1072 | proof(rule Subst_all_cls_remains) | |
| 91 | 1073 | from X_in_ES and h | 
| 1074 |           show "(X, xrhs) \<in> ES - {(Y, yrhs)}" by auto
 | |
| 1075 | qed | |
| 1076 | moreover have | |
| 94 | 1077 |           "card (Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)) < card ES"
 | 
| 91 | 1078 | proof(rule le_less_trans) | 
| 1079 | show | |
| 94 | 1080 |             "card (Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)) \<le> 
 | 
| 91 | 1081 |                                                                 card (ES - {(Y, yrhs)})"
 | 
| 94 | 1082 | proof(rule Subst_all_card_le) | 
| 91 | 1083 |             show "finite (ES - {(Y, yrhs)})" using finite_ES by auto
 | 
| 1084 | qed | |
| 1085 | next | |
| 1086 |           show "card (ES - {(Y, yrhs)}) < card ES" using finite_ES h
 | |
| 1087 | by (auto simp:card_gt_0_iff intro:diff_Suc_less) | |
| 1088 | qed | |
| 1089 | ultimately show ?thesis | |
| 94 | 1090 | by (auto dest: Subst_all_card_le elim:le_less_trans) | 
| 91 | 1091 | qed | 
| 42 | 1092 | qed | 
| 1093 | qed | |
| 1094 | qed | |
| 1095 | ||
| 91 | 1096 | |
| 1097 | subsubsection {* Conclusion of the proof *}
 | |
| 42 | 1098 | |
| 1099 | text {*
 | |
| 1100 |   From this point until @{text "hard_direction"}, the hard direction is proved
 | |
| 1101 | through a simple application of the iteration principle. | |
| 1102 | *} | |
| 1103 | ||
| 91 | 1104 | lemma reduce_x: | 
| 1105 | assumes inv: "invariant ES" | |
| 1106 | and contain_x: "(X, xrhs) \<in> ES" | |
| 1107 |   shows "\<exists> xrhs'. reduce X ES = {(X, xrhs')} \<and> invariant(reduce X ES)"
 | |
| 1108 | proof - | |
| 1109 | let ?Inv = "\<lambda> ES. (invariant ES \<and> (\<exists> xrhs. (X, xrhs) \<in> ES))" | |
| 1110 | show ?thesis | |
| 1111 | proof (unfold reduce_def, | |
| 1112 | rule while_rule [where P = ?Inv and r = "measure card"]) | |
| 1113 | from inv and contain_x show "?Inv ES" by auto | |
| 1114 | next | |
| 1115 | show "wf (measure card)" by simp | |
| 1116 | next | |
| 1117 | fix ES | |
| 1118 | assume inv: "?Inv ES" and crd: "card ES \<noteq> 1" | |
| 1119 | show "(iter X ES, ES) \<in> measure card" | |
| 1120 | proof - | |
| 1121 | from inv obtain xrhs where x_in: "(X, xrhs) \<in> ES" by auto | |
| 1122 | from inv have "invariant ES" by simp | |
| 1123 | from iteration_step [OF this x_in crd] | |
| 1124 | show ?thesis by auto | |
| 1125 | qed | |
| 1126 | next | |
| 1127 | fix ES | |
| 1128 | assume inv: "?Inv ES" and crd: "card ES \<noteq> 1" | |
| 1129 | thus "?Inv (iter X ES)" | |
| 1130 | proof - | |
| 1131 | from inv obtain xrhs where x_in: "(X, xrhs) \<in> ES" by auto | |
| 1132 | from inv have "invariant ES" by simp | |
| 1133 | from iteration_step [OF this x_in crd] | |
| 1134 | show ?thesis by auto | |
| 1135 | qed | |
| 1136 | next | |
| 1137 | fix ES | |
| 1138 | assume "?Inv ES" and "\<not> card ES \<noteq> 1" | |
| 1139 |     thus "\<exists>xrhs'. ES = {(X, xrhs')} \<and> invariant ES"
 | |
| 1140 | apply (auto, rule_tac x = xrhs in exI) | |
| 1141 | by (auto simp: invariant_def dest!:card_Suc_Diff1 simp:card_eq_0_iff) | |
| 1142 | qed | |
| 42 | 1143 | qed | 
| 91 | 1144 | |
| 42 | 1145 | lemma last_cl_exists_rexp: | 
| 91 | 1146 |   assumes Inv_ES: "invariant {(X, xrhs)}"
 | 
| 42 | 1147 | shows "\<exists> (r::rexp). L r = X" (is "\<exists> r. ?P r") | 
| 1148 | proof- | |
| 94 | 1149 | def A \<equiv> "Arden X xrhs" | 
| 81 | 1150 |   have "?P (\<Uplus>{r. Lam r \<in> A})"
 | 
| 42 | 1151 | proof - | 
| 79 | 1152 |     have "L (\<Uplus>{r. Lam r \<in> A}) = L ({Lam r | r. Lam r \<in>  A})"
 | 
| 42 | 1153 | proof(rule rexp_of_lam_eq_lam_set) | 
| 79 | 1154 | show "finite A" | 
| 1155 | unfolding A_def | |
| 91 | 1156 | using Inv_ES | 
| 94 | 1157 | by (rule_tac Arden_keeps_finite) | 
| 86 | 1158 | (auto simp add: invariant_def finite_rhs_def) | 
| 42 | 1159 | qed | 
| 79 | 1160 | also have "\<dots> = L A" | 
| 42 | 1161 | proof- | 
| 80 | 1162 |       have "{Lam r | r. Lam r \<in> A} = A"
 | 
| 42 | 1163 | proof- | 
| 91 | 1164 |         have "classes_of A = {}" using Inv_ES 
 | 
| 79 | 1165 | unfolding A_def | 
| 94 | 1166 | by (simp add:Arden_removes_cl | 
| 86 | 1167 | self_contained_def invariant_def lefts_of_def) | 
| 79 | 1168 | thus ?thesis | 
| 1169 | unfolding A_def | |
| 80 | 1170 | by (auto simp only: classes_of_def, case_tac x, auto) | 
| 42 | 1171 | qed | 
| 80 | 1172 | thus ?thesis by simp | 
| 42 | 1173 | qed | 
| 1174 | also have "\<dots> = X" | |
| 79 | 1175 | unfolding A_def | 
| 94 | 1176 | proof(rule Arden_keeps_eq [THEN sym]) | 
| 91 | 1177 | show "X = L xrhs" using Inv_ES | 
| 1178 | by (auto simp only: invariant_def valid_eqns_def) | |
| 42 | 1179 | next | 
| 91 | 1180 |       from Inv_ES show "[] \<notin> L (\<Uplus>{r. Trn X r \<in> xrhs})"
 | 
| 1181 | by(simp add: invariant_def ardenable_def rexp_of_empty finite_rhs_def) | |
| 42 | 1182 | next | 
| 91 | 1183 | from Inv_ES show "finite xrhs" | 
| 1184 | by (simp add: invariant_def finite_rhs_def) | |
| 42 | 1185 | qed | 
| 81 | 1186 | finally show ?thesis by simp | 
| 42 | 1187 | qed | 
| 1188 | thus ?thesis by auto | |
| 1189 | qed | |
| 91 | 1190 | |
| 42 | 1191 | lemma every_eqcl_has_reg: | 
| 1192 | assumes finite_CS: "finite (UNIV // (\<approx>Lang))" | |
| 1193 | and X_in_CS: "X \<in> (UNIV // (\<approx>Lang))" | |
| 1194 | shows "\<exists> (reg::rexp). L reg = X" (is "\<exists> r. ?E r") | |
| 1195 | proof - | |
| 91 | 1196 | let ?ES = " eqs (UNIV // \<approx>Lang)" | 
| 1197 | from X_in_CS | |
| 1198 | obtain xrhs where "(X, xrhs) \<in> ?ES" | |
| 42 | 1199 | by (auto simp:eqs_def init_rhs_def) | 
| 91 | 1200 | from reduce_x [OF init_ES_satisfy_invariant [OF finite_CS] this] | 
| 1201 |   have "\<exists>xrhs'. reduce X ?ES = {(X, xrhs')} \<and> invariant (reduce X ?ES)" .
 | |
| 1202 |   then obtain xrhs' where "invariant {(X, xrhs')}" by auto
 | |
| 1203 | from last_cl_exists_rexp [OF this] | |
| 1204 | show ?thesis . | |
| 42 | 1205 | qed | 
| 1206 | ||
| 91 | 1207 | |
| 42 | 1208 | theorem hard_direction: | 
| 70 | 1209 | assumes finite_CS: "finite (UNIV // \<approx>A)" | 
| 1210 | shows "\<exists>r::rexp. A = L r" | |
| 42 | 1211 | proof - | 
| 70 | 1212 | have "\<forall> X \<in> (UNIV // \<approx>A). \<exists>reg::rexp. X = L reg" | 
| 42 | 1213 | using finite_CS every_eqcl_has_reg by blast | 
| 1214 | then obtain f | |
| 70 | 1215 | where f_prop: "\<forall> X \<in> (UNIV // \<approx>A). X = L ((f X)::rexp)" | 
| 1216 | by (auto dest: bchoice) | |
| 1217 | def rs \<equiv> "f ` (finals A)" | |
| 1218 | have "A = \<Union> (finals A)" using lang_is_union_of_finals by auto | |
| 76 | 1219 | also have "\<dots> = L (\<Uplus>rs)" | 
| 42 | 1220 | proof - | 
| 1221 | have "finite rs" | |
| 1222 | proof - | |
| 70 | 1223 | have "finite (finals A)" | 
| 1224 | using finite_CS finals_in_partitions[of "A"] | |
| 42 | 1225 | by (erule_tac finite_subset, simp) | 
| 1226 | thus ?thesis using rs_def by auto | |
| 1227 | qed | |
| 1228 | thus ?thesis | |
| 70 | 1229 | using f_prop rs_def finals_in_partitions[of "A"] by auto | 
| 42 | 1230 | qed | 
| 1231 | finally show ?thesis by blast | |
| 1232 | qed | |
| 1233 | ||
| 1234 | end |