2
+ − 1
theory MyhilNerode
0
+ − 2
imports "Main"
+ − 3
begin
+ − 4
+ − 5
text {* sequential composition of languages *}
+ − 6
+ − 7
definition
+ − 8
lang_seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ; _" [100,100] 100)
+ − 9
where
+ − 10
"L1 ; L2 = {s1@s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
+ − 11
+ − 12
lemma lang_seq_empty:
+ − 13
shows "{[]} ; L = L"
+ − 14
and "L ; {[]} = L"
+ − 15
unfolding lang_seq_def by auto
+ − 16
+ − 17
lemma lang_seq_null:
+ − 18
shows "{} ; L = {}"
+ − 19
and "L ; {} = {}"
+ − 20
unfolding lang_seq_def by auto
+ − 21
+ − 22
lemma lang_seq_append:
+ − 23
assumes a: "s1 \<in> L1"
+ − 24
and b: "s2 \<in> L2"
+ − 25
shows "s1@s2 \<in> L1 ; L2"
+ − 26
unfolding lang_seq_def
+ − 27
using a b by auto
+ − 28
+ − 29
lemma lang_seq_union:
+ − 30
shows "(L1 \<union> L2); L3 = (L1; L3) \<union> (L2; L3)"
+ − 31
and "L1; (L2 \<union> L3) = (L1; L2) \<union> (L1; L3)"
+ − 32
unfolding lang_seq_def by auto
+ − 33
+ − 34
lemma lang_seq_assoc:
+ − 35
shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)"
+ − 36
by (simp add: lang_seq_def Collect_def mem_def expand_fun_eq)
+ − 37
(metis append_assoc)
+ − 38
+ − 39
lemma lang_seq_minus:
+ − 40
shows "(L1; L2) - {[]} =
+ − 41
(if [] \<in> L1 then L2 - {[]} else {}) \<union>
+ − 42
(if [] \<in> L2 then L1 - {[]} else {}) \<union> ((L1 - {[]}); (L2 - {[]}))"
+ − 43
apply(auto simp add: lang_seq_def)
+ − 44
apply(metis mem_def self_append_conv)
+ − 45
apply(metis mem_def self_append_conv2)
+ − 46
apply(metis mem_def self_append_conv2)
+ − 47
apply(metis mem_def self_append_conv)
+ − 48
done
+ − 49
+ − 50
section {* Kleene star for languages defined as least fixed point *}
+ − 51
+ − 52
inductive_set
+ − 53
Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
+ − 54
for L :: "string set"
+ − 55
where
+ − 56
start[intro]: "[] \<in> L\<star>"
+ − 57
| step[intro]: "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>"
+ − 58
+ − 59
lemma lang_star_empty:
+ − 60
shows "{}\<star> = {[]}"
+ − 61
by (auto elim: Star.cases)
+ − 62
+ − 63
lemma lang_star_cases:
+ − 64
shows "L\<star> = {[]} \<union> L ; L\<star>"
+ − 65
proof
+ − 66
{ fix x
+ − 67
have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L ; L\<star>"
+ − 68
unfolding lang_seq_def
+ − 69
by (induct rule: Star.induct) (auto)
+ − 70
}
+ − 71
then show "L\<star> \<subseteq> {[]} \<union> L ; L\<star>" by auto
+ − 72
next
+ − 73
show "{[]} \<union> L ; L\<star> \<subseteq> L\<star>"
+ − 74
unfolding lang_seq_def by auto
+ − 75
qed
+ − 76
+ − 77
lemma lang_star_cases':
+ − 78
shows "L\<star> = {[]} \<union> L\<star> ; L"
+ − 79
proof
+ − 80
{ fix x
+ − 81
have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L\<star> ; L"
+ − 82
unfolding lang_seq_def
+ − 83
apply (induct rule: Star.induct)
+ − 84
apply simp
+ − 85
apply simp
+ − 86
apply (erule disjE)
+ − 87
apply (auto)[1]
+ − 88
apply (erule exE | erule conjE)+
+ − 89
apply (rule disjI2)
+ − 90
apply (rule_tac x = "s1 @ s1a" in exI)
+ − 91
by auto
+ − 92
}
+ − 93
then show "L\<star> \<subseteq> {[]} \<union> L\<star> ; L" by auto
+ − 94
next
+ − 95
show "{[]} \<union> L\<star> ; L \<subseteq> L\<star>"
+ − 96
unfolding lang_seq_def
+ − 97
apply auto
+ − 98
apply (erule Star.induct)
+ − 99
apply auto
+ − 100
apply (drule step[of _ _ "[]"])
+ − 101
by (auto intro:start)
+ − 102
qed
+ − 103
+ − 104
lemma lang_star_simple:
+ − 105
shows "L \<subseteq> L\<star>"
+ − 106
by (subst lang_star_cases)
+ − 107
(auto simp only: lang_seq_def)
+ − 108
+ − 109
lemma lang_star_prop0_aux:
+ − 110
"s2 \<in> L\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L \<longrightarrow> (\<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4)"
+ − 111
apply (erule Star.induct)
+ − 112
apply (clarify, rule_tac x = "[]" in exI, rule_tac x = s1 in exI, simp add:start)
+ − 113
apply (clarify, drule_tac x = s1 in spec)
+ − 114
apply (drule mp, simp, clarify)
+ − 115
apply (rule_tac x = "s1a @ s3" in exI, rule_tac x = s4 in exI)
+ − 116
by auto
+ − 117
+ − 118
lemma lang_star_prop0:
+ − 119
"\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> \<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4"
+ − 120
by (auto dest:lang_star_prop0_aux)
+ − 121
+ − 122
lemma lang_star_prop1:
+ − 123
assumes asm: "L1; L2 \<subseteq> L2"
+ − 124
shows "L1\<star>; L2 \<subseteq> L2"
+ − 125
proof -
+ − 126
{ fix s1 s2
+ − 127
assume minor: "s2 \<in> L2"
+ − 128
assume major: "s1 \<in> L1\<star>"
+ − 129
then have "s1@s2 \<in> L2"
+ − 130
proof(induct rule: Star.induct)
+ − 131
case start
+ − 132
show "[]@s2 \<in> L2" using minor by simp
+ − 133
next
+ − 134
case (step s1 s1')
+ − 135
have "s1 \<in> L1" by fact
+ − 136
moreover
+ − 137
have "s1'@s2 \<in> L2" by fact
+ − 138
ultimately have "s1@(s1'@s2) \<in> L1; L2" by (auto simp add: lang_seq_def)
+ − 139
with asm have "s1@(s1'@s2) \<in> L2" by auto
+ − 140
then show "(s1@s1')@s2 \<in> L2" by simp
+ − 141
qed
+ − 142
}
+ − 143
then show "L1\<star>; L2 \<subseteq> L2" by (auto simp add: lang_seq_def)
+ − 144
qed
+ − 145
+ − 146
lemma lang_star_prop2_aux:
+ − 147
"s2 \<in> L2\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L1 \<and> L1 ; L2 \<subseteq> L1 \<longrightarrow> s1 @ s2 \<in> L1"
+ − 148
apply (erule Star.induct, simp)
+ − 149
apply (clarify, drule_tac x = "s1a @ s1" in spec)
+ − 150
by (auto simp:lang_seq_def)
+ − 151
+ − 152
lemma lang_star_prop2:
+ − 153
"L1; L2 \<subseteq> L1 \<Longrightarrow> L1 ; L2\<star> \<subseteq> L1"
+ − 154
by (auto dest!:lang_star_prop2_aux simp:lang_seq_def)
+ − 155
+ − 156
lemma lang_star_seq_subseteq:
+ − 157
shows "L ; L\<star> \<subseteq> L\<star>"
+ − 158
using lang_star_cases by blast
+ − 159
+ − 160
lemma lang_star_double:
+ − 161
shows "L\<star>; L\<star> = L\<star>"
+ − 162
proof
+ − 163
show "L\<star> ; L\<star> \<subseteq> L\<star>"
+ − 164
using lang_star_prop1 lang_star_seq_subseteq by blast
+ − 165
next
+ − 166
have "L\<star> \<subseteq> L\<star> \<union> L\<star>; (L; L\<star>)" by auto
+ − 167
also have "\<dots> = L\<star>;{[]} \<union> L\<star>; (L; L\<star>)" by (simp add: lang_seq_empty)
+ − 168
also have "\<dots> = L\<star>; ({[]} \<union> L; L\<star>)" by (simp only: lang_seq_union)
+ − 169
also have "\<dots> = L\<star>; L\<star>" using lang_star_cases by simp
+ − 170
finally show "L\<star> \<subseteq> L\<star> ; L\<star>" by simp
+ − 171
qed
+ − 172
+ − 173
lemma lang_star_seq_subseteq':
+ − 174
shows "L\<star>; L \<subseteq> L\<star>"
+ − 175
proof -
+ − 176
have "L \<subseteq> L\<star>" by (rule lang_star_simple)
+ − 177
then have "L\<star>; L \<subseteq> L\<star>; L\<star>" by (auto simp add: lang_seq_def)
+ − 178
then show "L\<star>; L \<subseteq> L\<star>" using lang_star_double by blast
+ − 179
qed
+ − 180
+ − 181
lemma
+ − 182
shows "L\<star> \<subseteq> L\<star>\<star>"
+ − 183
by (rule lang_star_simple)
+ − 184
+ − 185
section {* tricky section *}
+ − 186
+ − 187
lemma k1:
+ − 188
assumes b: "s \<in> L\<star>"
+ − 189
and a: "s \<noteq> []"
+ − 190
shows "s \<in> (L - {[]}); L\<star>"
+ − 191
using b a
+ − 192
apply(induct rule: Star.induct)
+ − 193
apply(simp)
+ − 194
apply(case_tac "s1=[]")
+ − 195
apply(simp)
+ − 196
apply(simp add: lang_seq_def)
+ − 197
apply(blast)
+ − 198
done
+ − 199
+ − 200
section {* (relies on lemma k1) *}
+ − 201
+ − 202
lemma zzz:
+ − 203
shows "{s. c#s \<in> L1\<star>} = {s. c#s \<in> L1} ; (L1\<star>)"
+ − 204
apply(auto simp add: lang_seq_def Cons_eq_append_conv)
+ − 205
apply(drule k1)
+ − 206
apply(auto)[1]
+ − 207
apply(auto simp add: lang_seq_def)[1]
+ − 208
apply(rule_tac x="tl s1" in exI)
+ − 209
apply(rule_tac x="s2" in exI)
+ − 210
apply(auto)[1]
+ − 211
apply(auto simp add: Cons_eq_append_conv)[2]
+ − 212
apply(drule lang_seq_append)
+ − 213
apply(assumption)
+ − 214
apply(rotate_tac 1)
+ − 215
apply(drule rev_subsetD)
+ − 216
apply(rule lang_star_seq_subseteq)
+ − 217
apply(simp)
+ − 218
done
+ − 219
+ − 220
+ − 221
+ − 222
section {* regular expressions *}
+ − 223
+ − 224
datatype rexp =
+ − 225
NULL
+ − 226
| EMPTY
+ − 227
| CHAR char
+ − 228
| SEQ rexp rexp
+ − 229
| ALT rexp rexp
+ − 230
| STAR rexp
+ − 231
+ − 232
+ − 233
consts L:: "'a \<Rightarrow> string set"
+ − 234
+ − 235
fun
+ − 236
L_rexp :: "rexp \<Rightarrow> string set"
+ − 237
where
+ − 238
"L_rexp (NULL) = {}"
+ − 239
| "L_rexp (EMPTY) = {[]}"
+ − 240
| "L_rexp (CHAR c) = {[c]}"
+ − 241
| "L_rexp (SEQ r1 r2) = (L_rexp r1) ; (L_rexp r2)"
+ − 242
| "L_rexp (ALT r1 r2) = (L_rexp r1) \<union> (L_rexp r2)"
+ − 243
| "L_rexp (STAR r) = (L_rexp r)\<star>"
+ − 244
+ − 245
defs (overloaded)
+ − 246
l_rexp_abs: "L rexp \<equiv> L_rexp rexp"
+ − 247
+ − 248
declare L_rexp.simps [simp del] L_rexp.simps [folded l_rexp_abs, simp add]
+ − 249
+ − 250
definition
+ − 251
Ls :: "rexp set \<Rightarrow> string set"
+ − 252
where
+ − 253
"Ls R = (\<Union>r\<in>R. (L r))"
+ − 254
+ − 255
lemma Ls_union:
+ − 256
"Ls (R1 \<union> R2) = (Ls R1) \<union> (Ls R2)"
+ − 257
unfolding Ls_def by auto
+ − 258
+ − 259
text {* helper function for termination proofs *}
+ − 260
fun
+ − 261
Left :: "rexp \<Rightarrow> rexp"
+ − 262
where
+ − 263
"Left (SEQ r1 r2) = r1"
+ − 264
+ − 265
text {* dagger function *}
+ − 266
+ − 267
function
+ − 268
dagger :: "rexp \<Rightarrow> char \<Rightarrow> rexp list" ("_ \<dagger> _")
+ − 269
where
+ − 270
c1: "(NULL \<dagger> c) = []"
+ − 271
| c2: "(EMPTY) \<dagger> c = []"
+ − 272
| c3: "(CHAR c') \<dagger> c = (if c = c' then [EMPTY] else [])"
+ − 273
| c4: "(ALT r1 r2) \<dagger> c = r1 \<dagger> c @ r2 \<dagger> c"
+ − 274
| c5: "(SEQ NULL r2) \<dagger> c = []"
+ − 275
| c6: "(SEQ EMPTY r2) \<dagger> c = r2 \<dagger> c"
+ − 276
| c7: "(SEQ (CHAR c') r2) \<dagger> c = (if c = c' then [r2] else [])"
+ − 277
| c8: "(SEQ (SEQ r11 r12) r2) \<dagger> c = (SEQ r11 (SEQ r12 r2)) \<dagger> c"
+ − 278
| c9: "(SEQ (ALT r11 r12) r2) \<dagger> c = (SEQ r11 r2) \<dagger> c @ (SEQ r12 r2) \<dagger> c"
+ − 279
| c10: "(SEQ (STAR r1) r2) \<dagger> c = r2 \<dagger> c @ [SEQ r' (SEQ (STAR r1) r2). r' \<leftarrow> r1 \<dagger> c]"
+ − 280
| c11: "(STAR r) \<dagger> c = [SEQ r' (STAR r) . r' \<leftarrow> r \<dagger> c]"
+ − 281
by (pat_completeness) (auto)
+ − 282
+ − 283
termination dagger
+ − 284
by (relation "measures [\<lambda>(r, c). size r, \<lambda>(r, c). size (Left r)]") (simp_all)
+ − 285
+ − 286
lemma dagger_correctness:
+ − 287
"Ls (set r \<dagger> c) = {s. c#s \<in> L r}"
+ − 288
proof (induct rule: dagger.induct)
+ − 289
case (1 c)
+ − 290
show "Ls (set NULL \<dagger> c) = {s. c#s \<in> L NULL}" by (simp add: Ls_def)
+ − 291
next
+ − 292
case (2 c)
+ − 293
show "Ls (set EMPTY \<dagger> c) = {s. c#s \<in> L EMPTY}" by (simp add: Ls_def)
+ − 294
next
+ − 295
case (3 c' c)
+ − 296
show "Ls (set CHAR c' \<dagger> c) = {s. c#s \<in> L (CHAR c')}" by (simp add: Ls_def)
+ − 297
next
+ − 298
case (4 r1 r2 c)
+ − 299
have ih1: "Ls (set r1 \<dagger> c) = {s. c#s \<in> L r1}" by fact
+ − 300
have ih2: "Ls (set r2 \<dagger> c) = {s. c#s \<in> L r2}" by fact
+ − 301
show "Ls (set ALT r1 r2 \<dagger> c) = {s. c#s \<in> L (ALT r1 r2)}"
+ − 302
by (simp add: Ls_union ih1 ih2 Collect_disj_eq)
+ − 303
next
+ − 304
case (5 r2 c)
+ − 305
show "Ls (set SEQ NULL r2 \<dagger> c) = {s. c#s \<in> L (SEQ NULL r2)}" by (simp add: Ls_def lang_seq_null)
+ − 306
next
+ − 307
case (6 r2 c)
+ − 308
have ih: "Ls (set r2 \<dagger> c) = {s. c#s \<in> L r2}" by fact
+ − 309
show "Ls (set SEQ EMPTY r2 \<dagger> c) = {s. c#s \<in> L (SEQ EMPTY r2)}"
+ − 310
by (simp add: ih lang_seq_empty)
+ − 311
next
+ − 312
case (7 c' r2 c)
+ − 313
show "Ls (set SEQ (CHAR c') r2 \<dagger> c) = {s. c#s \<in> L (SEQ (CHAR c') r2)}"
+ − 314
by (simp add: Ls_def lang_seq_def)
+ − 315
next
+ − 316
case (8 r11 r12 r2 c)
+ − 317
have ih: "Ls (set SEQ r11 (SEQ r12 r2) \<dagger> c) = {s. c#s \<in> L (SEQ r11 (SEQ r12 r2))}" by fact
+ − 318
show "Ls (set SEQ (SEQ r11 r12) r2 \<dagger> c) = {s. c#s \<in> L (SEQ (SEQ r11 r12) r2)}"
+ − 319
by (simp add: ih lang_seq_assoc)
+ − 320
next
+ − 321
case (9 r11 r12 r2 c)
+ − 322
have ih1: "Ls (set SEQ r11 r2 \<dagger> c) = {s. c#s \<in> L (SEQ r11 r2)}" by fact
+ − 323
have ih2: "Ls (set SEQ r12 r2 \<dagger> c) = {s. c#s \<in> L (SEQ r12 r2)}" by fact
+ − 324
show "Ls (set SEQ (ALT r11 r12) r2 \<dagger> c) = {s. c#s \<in> L (SEQ (ALT r11 r12) r2)}"
+ − 325
by (simp add: Ls_union ih1 ih2 lang_seq_union Collect_disj_eq)
+ − 326
next
+ − 327
case (10 r1 r2 c)
+ − 328
have ih2: "Ls (set r2 \<dagger> c) = {s. c#s \<in> L r2}" by fact
+ − 329
have ih1: "Ls (set r1 \<dagger> c) = {s. c#s \<in> L r1}" by fact
+ − 330
have "Ls (set SEQ (STAR r1) r2 \<dagger> c) = Ls (set r2 \<dagger> c) \<union> (Ls (set r1 \<dagger> c); ((L r1)\<star> ; L r2))"
+ − 331
by (auto simp add: lang_seq_def Ls_def)
+ − 332
also have "\<dots> = {s. c#s \<in> L r2} \<union> ({s. c#s \<in> L r1} ; ((L r1)\<star> ; L r2))" using ih1 ih2 by simp
+ − 333
also have "\<dots> = {s. c#s \<in> L r2} \<union> ({s. c#s \<in> L r1} ; (L r1)\<star>) ; L r2" by (simp add: lang_seq_assoc)
+ − 334
also have "\<dots> = {s. c#s \<in> L r2} \<union> {s. c#s \<in> (L r1)\<star>} ; L r2" by (simp add: zzz)
+ − 335
also have "\<dots> = {s. c#s \<in> L r2} \<union> {s. c#s \<in> (L r1)\<star> ; L r2}"
+ − 336
by (auto simp add: lang_seq_def Cons_eq_append_conv)
+ − 337
also have "\<dots> = {s. c#s \<in> (L r1)\<star> ; L r2}"
+ − 338
by (force simp add: lang_seq_def)
+ − 339
finally show "Ls (set SEQ (STAR r1) r2 \<dagger> c) = {s. c#s \<in> L (SEQ (STAR r1) r2)}" by simp
+ − 340
next
+ − 341
case (11 r c)
+ − 342
have ih: "Ls (set r \<dagger> c) = {s. c#s \<in> L r}" by fact
+ − 343
have "Ls (set (STAR r) \<dagger> c) = Ls (set r \<dagger> c) ; (L r)\<star>"
+ − 344
by (auto simp add: lang_seq_def Ls_def)
+ − 345
also have "\<dots> = {s. c#s \<in> L r} ; (L r)\<star>" using ih by simp
+ − 346
also have "\<dots> = {s. c#s \<in> (L r)\<star>}" using zzz by simp
+ − 347
finally show "Ls (set (STAR r) \<dagger> c) = {s. c#s \<in> L (STAR r)}" by simp
+ − 348
qed
+ − 349
+ − 350
+ − 351
text {* matcher function (based on the "list"-dagger function) *}
+ − 352
fun
+ − 353
first_True :: "bool list \<Rightarrow> bool"
+ − 354
where
+ − 355
"first_True [] = False"
+ − 356
| "first_True (x#xs) = (if x then True else first_True xs)"
+ − 357
+ − 358
lemma not_first_True[simp]:
+ − 359
shows "(\<not>(first_True xs)) = (\<forall>x \<in> set xs. \<not>x)"
+ − 360
by (induct xs) (auto)
+ − 361
+ − 362
lemma first_True:
+ − 363
shows "(first_True xs) = (\<exists>x \<in> set xs. x)"
+ − 364
by (induct xs) (auto)
+ − 365
+ − 366
text {* matcher function *}
+ − 367
+ − 368
function
+ − 369
matcher :: "rexp \<Rightarrow> string \<Rightarrow> bool" ("_ ! _")
+ − 370
where
+ − 371
"NULL ! s = False"
+ − 372
| "EMPTY ! s = (s =[])"
+ − 373
| "CHAR c ! s = (s = [c])"
+ − 374
| "ALT r1 r2 ! s = (r1 ! s \<or> r2 ! s)"
+ − 375
| "STAR r ! [] = True"
+ − 376
| "STAR r ! c#s = first_True [SEQ (r') (STAR r) ! s. r' \<leftarrow> r \<dagger> c]"
+ − 377
| "SEQ r1 r2 ! [] = (r1 ! [] \<and> r2 ! [])"
+ − 378
| "SEQ NULL r2 ! (c#s) = False"
+ − 379
| "SEQ EMPTY r2 ! (c#s) = (r2 ! c#s)"
+ − 380
| "SEQ (CHAR c') r2 ! (c#s) = (if c'=c then r2 ! s else False)"
+ − 381
| "SEQ (SEQ r11 r12) r2 ! (c#s) = (SEQ r11 (SEQ r12 r2) ! c#s)"
+ − 382
| "SEQ (ALT r11 r12) r2 ! (c#s) = ((SEQ r11 r2) ! (c#s) \<or> (SEQ r12 r2) ! (c#s))"
+ − 383
| "SEQ (STAR r1) r2 ! (c#s) = (r2 ! (c#s) \<or> first_True [SEQ r' (SEQ (STAR r1) r2) ! s. r' \<leftarrow> r1 \<dagger> c])"
+ − 384
by (pat_completeness) (auto)
+ − 385
+ − 386
termination matcher
+ − 387
by(relation "measures [\<lambda>(r,s). length s, \<lambda>(r,s). size r, \<lambda>(r,s). size (Left r)]") (simp_all)
+ − 388
+ − 389
text {* positive correctness of the matcher *}
+ − 390
lemma matcher1:
+ − 391
shows "r ! s \<Longrightarrow> s \<in> L r"
+ − 392
proof (induct rule: matcher.induct)
+ − 393
case (1 s)
+ − 394
have "NULL ! s" by fact
+ − 395
then show "s \<in> L NULL" by simp
+ − 396
next
+ − 397
case (2 s)
+ − 398
have "EMPTY ! s" by fact
+ − 399
then show "s \<in> L EMPTY" by simp
+ − 400
next
+ − 401
case (3 c s)
+ − 402
have "CHAR c ! s" by fact
+ − 403
then show "s \<in> L (CHAR c)" by simp
+ − 404
next
+ − 405
case (4 r1 r2 s)
+ − 406
have ih1: "r1 ! s \<Longrightarrow> s \<in> L r1" by fact
+ − 407
have ih2: "r2 ! s \<Longrightarrow> s \<in> L r2" by fact
+ − 408
have "ALT r1 r2 ! s" by fact
+ − 409
with ih1 ih2 show "s \<in> L (ALT r1 r2)" by auto
+ − 410
next
+ − 411
case (5 r)
+ − 412
have "STAR r ! []" by fact
+ − 413
then show "[] \<in> L (STAR r)" by auto
+ − 414
next
+ − 415
case (6 r c s)
+ − 416
have ih1: "\<And>rx. \<lbrakk>rx \<in> set r \<dagger> c; SEQ rx (STAR r) ! s\<rbrakk> \<Longrightarrow> s \<in> L (SEQ rx (STAR r))" by fact
+ − 417
have as: "STAR r ! c#s" by fact
+ − 418
then obtain r' where imp1: "r' \<in> set r \<dagger> c" and imp2: "SEQ r' (STAR r) ! s"
+ − 419
by (auto simp add: first_True)
+ − 420
from imp2 imp1 have "s \<in> L (SEQ r' (STAR r))" using ih1 by simp
+ − 421
then have "s \<in> L r' ; (L r)\<star>" by simp
+ − 422
then have "s \<in> Ls (set r \<dagger> c) ; (L r)\<star>" using imp1 by (auto simp add: Ls_def lang_seq_def)
+ − 423
then have "s \<in> {s. c#s \<in> L r} ; (L r)\<star>" by (auto simp add: dagger_correctness)
+ − 424
then have "s \<in> {s. c#s \<in> (L r)\<star>}" by (simp add: zzz)
+ − 425
then have "c#s \<in> {[c]}; {s. c#s \<in> (L r)\<star>}" by (auto simp add: lang_seq_def)
+ − 426
then have "c#s \<in> (L r)\<star>" by (auto simp add: lang_seq_def)
+ − 427
then show "c#s \<in> L (STAR r)" by simp
+ − 428
next
+ − 429
case (7 r1 r2)
+ − 430
have ih1: "r1 ! [] \<Longrightarrow> [] \<in> L r1" by fact
+ − 431
have ih2: "r2 ! [] \<Longrightarrow> [] \<in> L r2" by fact
+ − 432
have as: "SEQ r1 r2 ! []" by fact
+ − 433
then have "r1 ! [] \<and> r2 ! []" by simp
+ − 434
then show "[] \<in> L (SEQ r1 r2)" using ih1 ih2 by (simp add: lang_seq_def)
+ − 435
next
+ − 436
case (8 r2 c s)
+ − 437
have "SEQ NULL r2 ! c#s" by fact
+ − 438
then show "c#s \<in> L (SEQ NULL r2)" by simp
+ − 439
next
+ − 440
case (9 r2 c s)
+ − 441
have ih1: "r2 ! c#s \<Longrightarrow> c#s \<in> L r2" by fact
+ − 442
have "SEQ EMPTY r2 ! c#s" by fact
+ − 443
then show "c#s \<in> L (SEQ EMPTY r2)" using ih1 by (simp add: lang_seq_def)
+ − 444
next
+ − 445
case (10 c' r2 c s)
+ − 446
have ih1: "\<lbrakk>c' = c; r2 ! s\<rbrakk> \<Longrightarrow> s \<in> L r2" by fact
+ − 447
have "SEQ (CHAR c') r2 ! c#s" by fact
+ − 448
then show "c#s \<in> L (SEQ (CHAR c') r2)"
+ − 449
using ih1 by (auto simp add: lang_seq_def split: if_splits)
+ − 450
next
+ − 451
case (11 r11 r12 r2 c s)
+ − 452
have ih1: "SEQ r11 (SEQ r12 r2) ! c#s \<Longrightarrow> c#s \<in> L (SEQ r11 (SEQ r12 r2))" by fact
+ − 453
have "SEQ (SEQ r11 r12) r2 ! c#s" by fact
+ − 454
then have "c#s \<in> L (SEQ r11 (SEQ r12 r2))" using ih1 by simp
+ − 455
then show "c#s \<in> L (SEQ (SEQ r11 r12) r2)" by (simp add: lang_seq_assoc)
+ − 456
next
+ − 457
case (12 r11 r12 r2 c s)
+ − 458
have ih1: "SEQ r11 r2 ! c#s \<Longrightarrow> c#s \<in> L (SEQ r11 r2)" by fact
+ − 459
have ih2: "SEQ r12 r2 ! c#s \<Longrightarrow> c#s \<in> L (SEQ r12 r2)" by fact
+ − 460
have "SEQ (ALT r11 r12) r2 ! c#s" by fact
+ − 461
then show "c#s \<in> L (SEQ (ALT r11 r12) r2)"
+ − 462
using ih1 ih2 by (auto simp add: lang_seq_union)
+ − 463
next
+ − 464
case (13 r1 r2 c s)
+ − 465
have ih1: "r2 ! c#s \<Longrightarrow> c#s \<in> L r2" by fact
+ − 466
have ih2: "\<And>r'. \<lbrakk>r' \<in> set r1 \<dagger> c; SEQ r' (SEQ (STAR r1) r2) ! s\<rbrakk> \<Longrightarrow>
+ − 467
s \<in> L (SEQ r' (SEQ (STAR r1) r2))" by fact
+ − 468
have "SEQ (STAR r1) r2 ! c#s" by fact
+ − 469
then have "(r2 ! c#s) \<or> (\<exists>r' \<in> set r1 \<dagger> c. SEQ r' (SEQ (STAR r1) r2) ! s)" by (auto simp add: first_True)
+ − 470
moreover
+ − 471
{ assume "r2 ! c#s"
+ − 472
with ih1 have "c#s \<in> L r2" by simp
+ − 473
then have "c # s \<in> L r1\<star> ; L r2"
+ − 474
by (auto simp add: lang_seq_def)
+ − 475
then have "c#s \<in> L (SEQ (STAR r1) r2)" by simp
+ − 476
}
+ − 477
moreover
+ − 478
{ assume "\<exists>r' \<in> set r1 \<dagger> c. SEQ r' (SEQ (STAR r1) r2) ! s"
+ − 479
then obtain r' where imp1: "r' \<in> set r1 \<dagger> c" and imp2: "SEQ r' (SEQ (STAR r1) r2) ! s" by blast
+ − 480
from imp2 imp1 have "s \<in> L (SEQ r' (SEQ (STAR r1) r2))" using ih2 by simp
+ − 481
then have "s \<in> L r' ; ((L r1)\<star> ; L r2)" by simp
+ − 482
then have "s \<in> Ls (set r1 \<dagger> c) ; ((L r1)\<star> ; L r2)" using imp1 by (auto simp add: Ls_def lang_seq_def)
+ − 483
then have "s \<in> {s. c#s \<in> L r1} ; ((L r1)\<star> ; L r2)" by (simp add: dagger_correctness)
+ − 484
then have "s \<in> ({s. c#s \<in> L r1} ; (L r1)\<star>) ; L r2" by (simp add: lang_seq_assoc)
+ − 485
then have "s \<in> {s. c#s \<in> (L r1)\<star>} ; L r2" by (simp add: zzz)
+ − 486
then have "c#s \<in> {[c]}; ({s. c#s \<in> (L r1)\<star>}; L r2)" by (auto simp add: lang_seq_def)
+ − 487
then have "c#s \<in> ({[c]}; {s. c#s \<in> (L r1)\<star>}) ; L r2" by (simp add: lang_seq_assoc)
+ − 488
then have "c#s \<in> (L r1)\<star>; L r2" by (auto simp add: lang_seq_def)
+ − 489
then have "c#s \<in> L (SEQ (STAR r1) r2)" by simp
+ − 490
}
+ − 491
ultimately show "c#s \<in> L (SEQ (STAR r1) r2)" by blast
+ − 492
qed
+ − 493
+ − 494
text {* negative correctness of the matcher *}
+ − 495
lemma matcher2:
+ − 496
shows "\<not> r ! s \<Longrightarrow> s \<notin> L r"
+ − 497
proof (induct rule: matcher.induct)
+ − 498
case (1 s)
+ − 499
have "\<not> NULL ! s" by fact
+ − 500
then show "s \<notin> L NULL" by simp
+ − 501
next
+ − 502
case (2 s)
+ − 503
have "\<not> EMPTY ! s" by fact
+ − 504
then show "s \<notin> L EMPTY" by simp
+ − 505
next
+ − 506
case (3 c s)
+ − 507
have "\<not> CHAR c ! s" by fact
+ − 508
then show "s \<notin> L (CHAR c)" by simp
+ − 509
next
+ − 510
case (4 r1 r2 s)
+ − 511
have ih2: "\<not> r1 ! s \<Longrightarrow> s \<notin> L r1" by fact
+ − 512
have ih4: "\<not> r2 ! s \<Longrightarrow> s \<notin> L r2" by fact
+ − 513
have "\<not> ALT r1 r2 ! s" by fact
+ − 514
then show "s \<notin> L (ALT r1 r2)" by (simp add: ih2 ih4)
+ − 515
next
+ − 516
case (5 r)
+ − 517
have "\<not> STAR r ! []" by fact
+ − 518
then show "[] \<notin> L (STAR r)" by simp
+ − 519
next
+ − 520
case (6 r c s)
+ − 521
have ih: "\<And>rx. \<lbrakk>rx \<in> set r \<dagger> c; \<not>SEQ rx (STAR r) ! s\<rbrakk> \<Longrightarrow> s \<notin> L (SEQ rx (STAR r))" by fact
+ − 522
have as: "\<not> STAR r ! c#s" by fact
+ − 523
then have "\<forall>r'\<in> set r \<dagger> c. \<not> (SEQ r' (STAR r) ! s)" by simp
+ − 524
then have "\<forall>r'\<in> set r \<dagger> c. s \<notin> L (SEQ r' (STAR r))" using ih by auto
+ − 525
then have "\<forall>r'\<in> set r \<dagger> c. s \<notin> L r' ; ((L r)\<star>)" by simp
+ − 526
then have "s \<notin> (Ls (set r \<dagger> c)) ; ((L r)\<star>)" by (auto simp add: Ls_def lang_seq_def)
+ − 527
then have "s \<notin> {s. c#s \<in> L r} ; ((L r)\<star>)" by (simp add: dagger_correctness)
+ − 528
then have "s \<notin> {s. c#s \<in> (L r)\<star>}" by (simp add: zzz)
+ − 529
then have "c#s \<notin> {[c]} ; {s. c#s \<in> (L r)\<star>}" by (auto simp add: lang_seq_assoc lang_seq_def)
+ − 530
then have "c#s \<notin> (L r)\<star>" by (simp add: lang_seq_def)
+ − 531
then show "c#s \<notin> L (STAR r)" by simp
+ − 532
next
+ − 533
case (7 r1 r2)
+ − 534
have ih2: "\<not> r1 ! [] \<Longrightarrow> [] \<notin> L r1" by fact
+ − 535
have ih4: "\<not> r2 ! [] \<Longrightarrow> [] \<notin> L r2" by fact
+ − 536
have "\<not> SEQ r1 r2 ! []" by fact
+ − 537
then have "\<not> r1 ! [] \<or> \<not> r2 ! []" by simp
+ − 538
then show "[] \<notin> L (SEQ r1 r2)" using ih2 ih4
+ − 539
by (auto simp add: lang_seq_def)
+ − 540
next
+ − 541
case (8 r2 c s)
+ − 542
have "\<not> SEQ NULL r2 ! c#s" by fact
+ − 543
then show "c#s \<notin> L (SEQ NULL r2)" by (simp add: lang_seq_null)
+ − 544
next
+ − 545
case (9 r2 c s)
+ − 546
have ih1: "\<not> r2 ! c#s \<Longrightarrow> c#s \<notin> L r2" by fact
+ − 547
have "\<not> SEQ EMPTY r2 ! c#s" by fact
+ − 548
then show "c#s \<notin> L (SEQ EMPTY r2)"
+ − 549
using ih1 by (simp add: lang_seq_def)
+ − 550
next
+ − 551
case (10 c' r2 c s)
+ − 552
have ih2: "\<lbrakk>c' = c; \<not>r2 ! s\<rbrakk> \<Longrightarrow> s \<notin> L r2" by fact
+ − 553
have "\<not> SEQ (CHAR c') r2 ! c#s" by fact
+ − 554
then show "c#s \<notin> L (SEQ (CHAR c') r2)"
+ − 555
using ih2 by (auto simp add: lang_seq_def)
+ − 556
next
+ − 557
case (11 r11 r12 r2 c s)
+ − 558
have ih2: "\<not> SEQ r11 (SEQ r12 r2) ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r11 (SEQ r12 r2))" by fact
+ − 559
have "\<not> SEQ (SEQ r11 r12) r2 ! c#s" by fact
+ − 560
then show "c#s \<notin> L (SEQ (SEQ r11 r12) r2)"
+ − 561
using ih2 by (auto simp add: lang_seq_def)
+ − 562
next
+ − 563
case (12 r11 r12 r2 c s)
+ − 564
have ih2: "\<not> SEQ r11 r2 ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r11 r2)" by fact
+ − 565
have ih4: "\<not> SEQ r12 r2 ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r12 r2)" by fact
+ − 566
have "\<not> SEQ (ALT r11 r12) r2 ! c#s" by fact
+ − 567
then show " c#s \<notin> L (SEQ (ALT r11 r12) r2)"
+ − 568
using ih2 ih4 by (simp add: lang_seq_union)
+ − 569
next
+ − 570
case (13 r1 r2 c s)
+ − 571
have ih1: "\<not>r2 ! c#s \<Longrightarrow> c#s \<notin> L r2" by fact
+ − 572
have ih2: "\<And>rx. \<lbrakk>rx \<in> set r1 \<dagger> c; \<not> SEQ rx (SEQ (STAR r1) r2) ! s\<rbrakk>
+ − 573
\<Longrightarrow> s \<notin> L (SEQ rx (SEQ (STAR r1) r2))" by fact
+ − 574
have as: "\<not> SEQ (STAR r1) r2 ! c#s" by fact
+ − 575
then have as1: "\<not>r2 ! c#s" and as2: "\<forall>r1'\<in>set r1 \<dagger> c. \<not> SEQ r1' (SEQ (STAR r1) r2) ! s" by simp_all
+ − 576
from as1 have bs1: "c#s \<notin> L r2" using ih1 by simp
+ − 577
from as2 have "\<forall>r1'\<in>set r1 \<dagger> c. \<not> SEQ r1' (SEQ (STAR r1) r2) ! s" by simp
+ − 578
then have "\<forall>r1'\<in>set r1 \<dagger> c. s \<notin> L (SEQ r1' (SEQ (STAR r1) r2))" using ih2 by simp
+ − 579
then have "\<forall>r1'\<in>set r1 \<dagger> c. s \<notin> L r1'; ((L r1)\<star>; L r2)" by simp
+ − 580
then have "s \<notin> (Ls (set r1 \<dagger> c)) ; ((L r1)\<star>; L r2)" by (auto simp add: Ls_def lang_seq_def)
+ − 581
then have "s \<notin> {s. c#s \<in> L r1} ; ((L r1)\<star>; L r2)" by (simp add: dagger_correctness)
+ − 582
then have "s \<notin> ({s. c#s \<in> L r1} ; (L r1)\<star>); L r2" by (simp add: lang_seq_assoc)
+ − 583
then have "s \<notin> {s. c#s \<in> (L r1)\<star>}; L r2" by (simp add: zzz)
+ − 584
then have "c#s \<notin> {[c]}; ({s. c#s \<in> (L r1)\<star>}; L r2)" by (auto simp add: lang_seq_def)
+ − 585
then have "c#s \<notin> (L r1)\<star>; L r2" using bs1 by (auto simp add: lang_seq_def Cons_eq_append_conv)
+ − 586
then show "c#s \<notin> L (SEQ (STAR r1) r2)" by simp
+ − 587
qed
+ − 588
+ − 589
section {* Questions *}
+ − 590
+ − 591
text {*
+ − 592
- Why was the key lemma k1 omitted; is there an easy, non-induction
+ − 593
way for obtaining this property?
+ − 594
- Why was False included in the definition of the STAR-clause in
+ − 595
the matcher? Has this something to do with executing the code?
+ − 596
+ − 597
*}
+ − 598
+ − 599
section {* Code *}
+ − 600
+ − 601
export_code dagger in SML module_name Dagger file -
+ − 602
export_code matcher in SML module_name Dagger file -
+ − 603
+ − 604
section {* Examples *}
+ − 605
+ − 606
text {* since now everything is based on lists, the evaluation is quite fast *}
+ − 607
+ − 608
value "NULL ! []"
+ − 609
value "(CHAR (CHR ''a'')) ! [CHR ''a'']"
+ − 610
value "((CHAR a) ! [a,a])"
+ − 611
value "(STAR (CHAR a)) ! []"
+ − 612
value "(STAR (CHAR a)) ! [a,a]"
+ − 613
value "(SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ! ''abbbbc''"
+ − 614
value "(SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ! ''abbcbbc''"
+ − 615
+ − 616
section {* Slind et al's matcher based on derivatives *}
+ − 617
+ − 618
fun
+ − 619
nullable :: "rexp \<Rightarrow> bool"
+ − 620
where
+ − 621
"nullable (NULL) = False"
+ − 622
| "nullable (EMPTY) = True"
+ − 623
| "nullable (CHAR c) = False"
+ − 624
| "nullable (ALT r1 r2) = ((nullable r1) \<or> (nullable r2))"
+ − 625
| "nullable (SEQ r1 r2) = ((nullable r1) \<and> (nullable r2))"
+ − 626
| "nullable (STAR r) = True"
+ − 627
+ − 628
lemma nullable:
+ − 629
shows "([] \<in> L r) = nullable r"
+ − 630
by (induct r)
+ − 631
(auto simp add: lang_seq_def)
+ − 632
+ − 633
fun
+ − 634
der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
+ − 635
where
+ − 636
"der c (NULL) = NULL"
+ − 637
| "der c (EMPTY) = NULL"
+ − 638
| "der c (CHAR c') = (if c=c' then EMPTY else NULL)"
+ − 639
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
+ − 640
| "der c (SEQ r1 r2) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)"
+ − 641
| "der c (STAR r) = SEQ (der c r) (STAR r)"
+ − 642
+ − 643
lemma k2:
+ − 644
assumes b: "s \<in> L1\<star>"
+ − 645
and a: "s \<noteq> []"
+ − 646
shows "s \<in> (L1; (L1\<star>))"
+ − 647
using b a
+ − 648
apply(induct rule: Star.induct)
+ − 649
apply(simp)
+ − 650
apply(case_tac "s1=[]")
+ − 651
apply(simp)
+ − 652
apply(simp add: lang_seq_def)
+ − 653
apply(blast)
+ − 654
done
+ − 655
+ − 656
+ − 657
lemma der_correctness:
+ − 658
shows "(s \<in> L (der c r)) = ((c#s) \<in> L r)"
+ − 659
proof (induct r arbitrary: s)
+ − 660
case (NULL s)
+ − 661
show "(s \<in> L (der c NULL)) = (c#s \<in> L NULL)" by simp
+ − 662
next
+ − 663
case (EMPTY s)
+ − 664
show "(s \<in> L (der c EMPTY)) = (c#s \<in> L EMPTY)" by simp
+ − 665
next
+ − 666
case (CHAR c' s)
+ − 667
show "(s \<in> L (der c (CHAR c'))) = (c#s \<in> L (CHAR c'))" by simp
+ − 668
next
+ − 669
case (SEQ r1 r2 s)
+ − 670
have ih1: "\<And>s. (s \<in> L (der c r1)) = (c#s \<in> L r1)" by fact
+ − 671
have ih2: "\<And>s. (s \<in> L (der c r2)) = (c#s \<in> L r2)" by fact
+ − 672
show "(s \<in> L (der c (SEQ r1 r2))) = (c#s \<in> L (SEQ r1 r2))"
+ − 673
using ih1 ih2
+ − 674
by (auto simp add: nullable[symmetric] lang_seq_def Cons_eq_append_conv)
+ − 675
next
+ − 676
case (ALT r1 r2 s)
+ − 677
have ih1: "\<And>s. (s \<in> L (der c r1)) = (c#s \<in> L r1)" by fact
+ − 678
have ih2: "\<And>s. (s \<in> L (der c r2)) = (c#s \<in> L r2)" by fact
+ − 679
show "(s \<in> L (der c (ALT r1 r2))) = (c#s \<in> L (ALT r1 r2))"
+ − 680
using ih1 ih2 by (auto simp add: lang_seq_def)
+ − 681
next
+ − 682
case (STAR r s)
+ − 683
have ih1: "\<And>s. (s \<in> L (der c r)) = (c#s \<in> L r)" by fact
+ − 684
show "(s \<in> L (der c (STAR r))) = (c#s \<in> L (STAR r))"
+ − 685
using ih1
+ − 686
apply(simp)
+ − 687
apply(auto simp add: lang_seq_def)
+ − 688
apply(drule lang_seq_append)
+ − 689
apply(assumption)
+ − 690
apply(simp)
+ − 691
apply(subst lang_star_cases)
+ − 692
apply(simp)
+ − 693
thm k1
+ − 694
apply(drule k2)
+ − 695
apply(simp)
+ − 696
apply(simp add: lang_seq_def)
+ − 697
apply(erule exE)+
+ − 698
apply(erule conjE)+
+ − 699
apply(auto simp add: lang_seq_def Cons_eq_append_conv)
+ − 700
apply(drule k1)
+ − 701
apply(simp)
+ − 702
apply(simp add: lang_seq_def)
+ − 703
apply(erule exE)+
+ − 704
apply(erule conjE)+
+ − 705
apply(auto simp add: lang_seq_def Cons_eq_append_conv)
+ − 706
done
+ − 707
qed
+ − 708
+ − 709
fun
+ − 710
derivative :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
+ − 711
where
+ − 712
"derivative [] r = r"
+ − 713
| "derivative (c#s) r = derivative s (der c r)"
+ − 714
+ − 715
fun
+ − 716
slind_matcher :: "rexp \<Rightarrow> string \<Rightarrow> bool"
+ − 717
where
+ − 718
"slind_matcher r s = nullable (derivative s r)"
+ − 719
+ − 720
lemma slind_matcher:
+ − 721
shows "slind_matcher r s = (s \<in> L r)"
+ − 722
by (induct s arbitrary: r)
+ − 723
(auto simp add: nullable der_correctness)
+ − 724
+ − 725
export_code slind_matcher in SML module_name Slind file -
+ − 726
+ − 727
+ − 728
(* ******************************************** now is the codes writen by chunhan ************************************* *)
+ − 729
+ − 730
section {* Arden's Lemma revised *}
+ − 731
+ − 732
lemma arden_aux1:
+ − 733
assumes a: "X \<subseteq> X ; A \<union> B"
+ − 734
and b: "[] \<notin> A"
+ − 735
shows "x \<in> X \<Longrightarrow> x \<in> B ; A\<star>"
+ − 736
apply (induct x taking:length rule:measure_induct)
+ − 737
apply (subgoal_tac "x \<in> X ; A \<union> B")
+ − 738
defer
+ − 739
using a
+ − 740
apply (auto)[1]
+ − 741
apply simp
+ − 742
apply (erule disjE)
+ − 743
defer
+ − 744
apply (auto simp add:lang_seq_def) [1]
+ − 745
apply (subgoal_tac "\<exists> x1 x2. x = x1 @ x2 \<and> x1 \<in> X \<and> x2 \<in> A")
+ − 746
defer
+ − 747
apply (auto simp add:lang_seq_def) [1]
+ − 748
apply (erule exE | erule conjE)+
+ − 749
apply simp
+ − 750
apply (drule_tac x = x1 in spec)
+ − 751
apply (simp)
+ − 752
using b
+ − 753
apply -
+ − 754
apply (auto)[1]
+ − 755
apply (subgoal_tac "x1 @ x2 \<in> (B ; A\<star>) ; A")
+ − 756
defer
+ − 757
apply (auto simp add:lang_seq_def)[1]
+ − 758
by (metis Un_absorb1 lang_seq_assoc lang_seq_union(2) lang_star_double lang_star_simple mem_def sup1CI)
+ − 759
+ − 760
theorem ardens_revised:
+ − 761
assumes nemp: "[] \<notin> A"
+ − 762
shows "(X = X ; A \<union> B) \<longleftrightarrow> (X = B ; A\<star>)"
+ − 763
apply(rule iffI)
+ − 764
defer
+ − 765
apply(simp)
+ − 766
apply(subst lang_star_cases')
+ − 767
apply(subst lang_seq_union)
+ − 768
apply(simp add: lang_seq_empty)
+ − 769
apply(simp add: lang_seq_assoc)
+ − 770
apply(auto)[1]
+ − 771
proof -
+ − 772
assume "X = X ; A \<union> B"
+ − 773
then have as1: "X ; A \<union> B \<subseteq> X" and as2: "X \<subseteq> X ; A \<union> B" by simp_all
+ − 774
from as1 have a: "X ; A \<subseteq> X" and b: "B \<subseteq> X" by simp_all
+ − 775
from b have "B; A\<star> \<subseteq> X ; A\<star>" by (auto simp add: lang_seq_def)
+ − 776
moreover
+ − 777
from a have "X ; A\<star> \<subseteq> X"
+ − 778
+ − 779
by (rule lang_star_prop2)
+ − 780
ultimately have f1: "B ; A\<star> \<subseteq> X" by simp
+ − 781
from as2 nemp
+ − 782
have f2: "X \<subseteq> B; A\<star>" using arden_aux1 by auto
+ − 783
from f1 f2 show "X = B; A\<star>" by auto
+ − 784
qed
+ − 785
+ − 786
section {* equiv class' definition *}
+ − 787
+ − 788
definition
+ − 789
equiv_str :: "string \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> bool" ("_ \<equiv>_\<equiv> _" [100, 100, 100] 100)
+ − 790
where
+ − 791
"x \<equiv>L'\<equiv> y \<longleftrightarrow> (\<forall>z. x@z \<in> L' \<longleftrightarrow> y@z \<in> L')"
+ − 792
+ − 793
definition
+ − 794
equiv_class :: "string \<Rightarrow> (string set) \<Rightarrow> string set" ("\<lbrakk>_\<rbrakk>_" [100, 100] 100)
+ − 795
where
+ − 796
"\<lbrakk>x\<rbrakk>L' \<equiv> {y. x \<equiv>L'\<equiv> y}"
+ − 797
+ − 798
text {* Chunhan modifies Q to Quo *}
+ − 799
definition
+ − 800
quot :: "string set \<Rightarrow> (string set) \<Rightarrow> (string set) set" ("_ Quo _" [100, 100] 100)
+ − 801
where
+ − 802
"L' Quo R \<equiv> { \<lbrakk>x\<rbrakk>R | x. x \<in> L'}"
+ − 803
+ − 804
lemma lang_eqs_union_of_eqcls:
+ − 805
"Lang = \<Union> {X. X \<in> (UNIV Quo Lang) \<and> (\<forall> x \<in> X. x \<in> Lang)}"
+ − 806
proof
+ − 807
show "Lang \<subseteq> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}"
+ − 808
proof
+ − 809
fix x
+ − 810
assume "x \<in> Lang"
+ − 811
thus "x \<in> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}"
+ − 812
proof (simp add:quot_def)
+ − 813
assume "(1)": "x \<in> Lang"
+ − 814
show "\<exists>xa. (\<exists>x. xa = \<lbrakk>x\<rbrakk>Lang) \<and> (\<forall>x\<in>xa. x \<in> Lang) \<and> x \<in> xa" (is "\<exists>xa.?P xa")
+ − 815
proof -
+ − 816
have "?P (\<lbrakk>x\<rbrakk>Lang)" using "(1)"
+ − 817
by (auto simp:equiv_class_def equiv_str_def dest: spec[where x = "[]"])
+ − 818
thus ?thesis by blast
+ − 819
qed
+ − 820
qed
+ − 821
qed
+ − 822
next
+ − 823
show "\<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<subseteq> Lang"
+ − 824
by auto
+ − 825
qed
+ − 826
+ − 827
lemma empty_notin_CS: "{} \<notin> UNIV Quo Lang"
+ − 828
apply (clarsimp simp:quot_def equiv_class_def)
+ − 829
by (rule_tac x = x in exI, auto simp:equiv_str_def)
+ − 830
+ − 831
lemma no_two_cls_inters:
+ − 832
"\<lbrakk>X \<in> UNIV Quo Lang; Y \<in> UNIV Quo Lang; X \<noteq> Y\<rbrakk> \<Longrightarrow> X \<inter> Y = {}"
+ − 833
by (auto simp:quot_def equiv_class_def equiv_str_def)
+ − 834
+ − 835
text {* equiv_class transition *}
+ − 836
definition
+ − 837
CT :: "string set \<Rightarrow> char \<Rightarrow> string set \<Rightarrow> bool" ("_-_\<rightarrow>_" [99,99]99)
+ − 838
where
+ − 839
"X-c\<rightarrow>Y \<equiv> ((X;{[c]}) \<subseteq> Y)"
+ − 840
+ − 841
types t_equa_rhs = "(string set \<times> rexp) set"
+ − 842
+ − 843
types t_equa = "(string set \<times> t_equa_rhs)"
+ − 844
+ − 845
types t_equas = "t_equa set"
+ − 846
+ − 847
text {* "empty_rhs" generates "\<lambda>" for init-state, just like "\<lambda>" for final states in Brzozowski method.
+ − 848
But if the init-state is "{[]}" ("\<lambda>" itself) then empty set is returned, see definition of "equation_rhs" *}
+ − 849
definition
+ − 850
empty_rhs :: "string set \<Rightarrow> t_equa_rhs"
+ − 851
where
+ − 852
"empty_rhs X \<equiv> if ([] \<in> X) then {({[]}, EMPTY)} else {}"
+ − 853
+ − 854
definition
+ − 855
folds :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"
+ − 856
where
+ − 857
"folds f z S \<equiv> SOME x. fold_graph f z S x"
+ − 858
+ − 859
definition
+ − 860
equation_rhs :: "(string set) set \<Rightarrow> string set \<Rightarrow> t_equa_rhs"
+ − 861
where
+ − 862
"equation_rhs CS X \<equiv> if (X = {[]}) then {({[]}, EMPTY)}
+ − 863
else {(S, folds ALT NULL {CHAR c| c. S-c\<rightarrow>X} )| S. S \<in> CS} \<union> empty_rhs X"
+ − 864
+ − 865
definition
+ − 866
equations :: "(string set) set \<Rightarrow> t_equas"
+ − 867
where
+ − 868
"equations CS \<equiv> {(X, equation_rhs CS X) | X. X \<in> CS}"
+ − 869
+ − 870
definition
+ − 871
L_rhs :: "t_equa_rhs \<Rightarrow> string set"
+ − 872
where
+ − 873
"L_rhs rhs \<equiv> {x. \<exists> X r. (X, r) \<in> rhs \<and> x \<in> X;(L r)}"
+ − 874
+ − 875
defs (overloaded)
+ − 876
l_rhs_abs: "L rhs \<equiv> L_rhs rhs"
+ − 877
+ − 878
lemmas L_def = L_rhs_def [folded l_rhs_abs] L_rexp.simps (* ??? is this OK ?? *)
+ − 879
+ − 880
definition
+ − 881
distinct_rhs :: "t_equa_rhs \<Rightarrow> bool"
+ − 882
where
+ − 883
"distinct_rhs rhs \<equiv> \<forall> X reg\<^isub>1 reg\<^isub>2. (X, reg\<^isub>1) \<in> rhs \<and> (X, reg\<^isub>2) \<in> rhs \<longrightarrow> reg\<^isub>1 = reg\<^isub>2"
+ − 884
+ − 885
definition
+ − 886
distinct_equas_rhs :: "t_equas \<Rightarrow> bool"
+ − 887
where
+ − 888
"distinct_equas_rhs equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> distinct_rhs rhs"
+ − 889
+ − 890
definition
+ − 891
distinct_equas :: "t_equas \<Rightarrow> bool"
+ − 892
where
+ − 893
"distinct_equas equas \<equiv> \<forall> X rhs rhs'. (X, rhs) \<in> equas \<and> (X, rhs') \<in> equas \<longrightarrow> rhs = rhs'"
+ − 894
+ − 895
definition
+ − 896
seq_rhs_r :: "t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"
+ − 897
where
+ − 898
"seq_rhs_r rhs r \<equiv> (\<lambda>(X, reg). (X, SEQ reg r)) ` rhs"
+ − 899
+ − 900
definition
+ − 901
del_x_paired :: "('a \<times> 'b) set \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'b) set"
+ − 902
where
+ − 903
"del_x_paired S x \<equiv> S - {X. X \<in> S \<and> fst X = x}"
+ − 904
+ − 905
definition
+ − 906
arden_variate :: "string set \<Rightarrow> rexp \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"
+ − 907
where
+ − 908
"arden_variate X r rhs \<equiv> seq_rhs_r (del_x_paired rhs X) (STAR r)"
+ − 909
+ − 910
definition
+ − 911
no_EMPTY_rhs :: "t_equa_rhs \<Rightarrow> bool"
+ − 912
where
+ − 913
"no_EMPTY_rhs rhs \<equiv> \<forall> X r. (X, r) \<in> rhs \<and> X \<noteq> {[]} \<longrightarrow> [] \<notin> L r"
+ − 914
+ − 915
definition
+ − 916
no_EMPTY_equas :: "t_equas \<Rightarrow> bool"
+ − 917
where
+ − 918
"no_EMPTY_equas equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> no_EMPTY_rhs rhs"
+ − 919
+ − 920
lemma fold_alt_null_eqs:
+ − 921
"finite rS \<Longrightarrow> x \<in> L (folds ALT NULL rS) = (\<exists> r \<in> rS. x \<in> L r)"
+ − 922
apply (simp add:folds_def)
+ − 923
apply (rule someI2_ex)
+ − 924
apply (erule finite_imp_fold_graph)
+ − 925
apply (erule fold_graph.induct)
+ − 926
by auto (*??? how do this be in Isar ?? *)
+ − 927
+ − 928
lemma seq_rhs_r_prop1:
+ − 929
"L (seq_rhs_r rhs r) = (L rhs);(L r)"
+ − 930
apply (rule set_ext, rule iffI)
+ − 931
apply (auto simp:L_def seq_rhs_r_def image_def lang_seq_def)
+ − 932
apply (rule_tac x = "s1 @ s1a" in exI, rule_tac x = "s2a" in exI, simp)
+ − 933
apply (rule_tac x = a in exI, rule_tac x = b in exI, simp)
+ − 934
apply (rule_tac x = s1 in exI, rule_tac x = s1a in exI, simp)
+ − 935
apply (rule_tac x = X in exI, rule_tac x = "SEQ ra r" in exI, simp)
+ − 936
apply (rule conjI)
+ − 937
apply (rule_tac x = "(X, ra)" in bexI, simp+)
+ − 938
apply (rule_tac x = s1a in exI, rule_tac x = "s2a @ s2" in exI, simp)
+ − 939
apply (simp add:lang_seq_def)
+ − 940
by (rule_tac x = s2a in exI, rule_tac x = s2 in exI, simp)
+ − 941
+ − 942
lemma del_x_paired_prop1:
+ − 943
"\<lbrakk>distinct_rhs rhs; (X, r) \<in> rhs\<rbrakk> \<Longrightarrow> X ; L r \<union> L (del_x_paired rhs X) = L rhs"
+ − 944
apply (simp add:L_def del_x_paired_def)
+ − 945
apply (rule set_ext, rule iffI, simp)
+ − 946
apply (erule disjE, rule_tac x = X in exI, rule_tac x = r in exI, simp)
+ − 947
apply (clarify, rule_tac x = Xa in exI, rule_tac x = ra in exI, simp)
+ − 948
apply (clarsimp, drule_tac x = Xa in spec, drule_tac x = ra in spec)
+ − 949
apply (auto simp:distinct_rhs_def)
+ − 950
done
+ − 951
+ − 952
lemma arden_variate_prop:
+ − 953
assumes "(X, rx) \<in> rhs"
+ − 954
shows "(\<forall> Y. Y \<noteq> X \<longrightarrow> (\<exists> r. (Y, r) \<in> rhs) = (\<exists> r. (Y, r) \<in> (arden_variate X rx rhs)))"
+ − 955
proof (rule allI, rule impI)
+ − 956
fix Y
+ − 957
assume "(1)": "Y \<noteq> X"
+ − 958
show "(\<exists>r. (Y, r) \<in> rhs) = (\<exists>r. (Y, r) \<in> arden_variate X rx rhs)"
+ − 959
proof
+ − 960
assume "(1_1)": "\<exists>r. (Y, r) \<in> rhs"
+ − 961
show "\<exists>r. (Y, r) \<in> arden_variate X rx rhs" (is "\<exists>r. ?P r")
+ − 962
proof -
+ − 963
from "(1_1)" obtain r where "(Y, r) \<in> rhs" ..
+ − 964
hence "?P (SEQ r (STAR rx))"
+ − 965
proof (simp add:arden_variate_def image_def)
+ − 966
have "(Y, r) \<in> rhs \<Longrightarrow> (Y, r) \<in> del_x_paired rhs X"
+ − 967
by (auto simp:del_x_paired_def "(1)")
+ − 968
thus "(Y, r) \<in> rhs \<Longrightarrow> (Y, SEQ r (STAR rx)) \<in> seq_rhs_r (del_x_paired rhs X) (STAR rx)"
+ − 969
by (auto simp:seq_rhs_r_def)
+ − 970
qed
+ − 971
thus ?thesis by blast
+ − 972
qed
+ − 973
next
+ − 974
assume "(2_1)": "\<exists>r. (Y, r) \<in> arden_variate X rx rhs"
+ − 975
thus "\<exists>r. (Y, r) \<in> rhs" (is "\<exists> r. ?P r")
+ − 976
by (auto simp:arden_variate_def del_x_paired_def seq_rhs_r_def image_def)
+ − 977
qed
+ − 978
qed
+ − 979
+ − 980
text {*
+ − 981
arden_variate_valid: proves variation from "X = X;r + Y;ry + \<dots>" to "X = Y;(SEQ ry (STAR r)) + \<dots>" holds the law of "language of left equiv language of right"
+ − 982
*}
+ − 983
lemma arden_variate_valid:
+ − 984
assumes X_not_empty: "X \<noteq> {[]}"
+ − 985
and l_eq_r: "X = L rhs"
+ − 986
and dist: "distinct_rhs rhs"
+ − 987
and no_empty: "no_EMPTY_rhs rhs"
+ − 988
and self_contained: "(X, r) \<in> rhs"
+ − 989
shows "X = L (arden_variate X r rhs)"
+ − 990
proof -
+ − 991
have "[] \<notin> L r" using no_empty X_not_empty self_contained
+ − 992
by (auto simp:no_EMPTY_rhs_def)
+ − 993
hence ardens: "X = X;(L r) \<union> (L (del_x_paired rhs X)) \<longleftrightarrow> X = (L (del_x_paired rhs X)) ; (L r)\<star>"
+ − 994
by (rule ardens_revised)
+ − 995
have del_x: "X = X ; L r \<union> L (del_x_paired rhs X) \<longleftrightarrow> X = L rhs" using dist l_eq_r self_contained
+ − 996
by (auto dest:del_x_paired_prop1)
+ − 997
show ?thesis
+ − 998
proof
+ − 999
show "X \<subseteq> L (arden_variate X r rhs)"
+ − 1000
proof
+ − 1001
fix x
+ − 1002
assume "(1_1)": "x \<in> X" with l_eq_r ardens del_x
+ − 1003
show "x \<in> L (arden_variate X r rhs)"
+ − 1004
by (simp add:arden_variate_def seq_rhs_r_prop1)
+ − 1005
qed
+ − 1006
next
+ − 1007
show "L (arden_variate X r rhs) \<subseteq> X"
+ − 1008
proof
+ − 1009
fix x
+ − 1010
assume "(2_1)": "x \<in> L (arden_variate X r rhs)" with ardens del_x l_eq_r
+ − 1011
show "x \<in> X"
+ − 1012
by (simp add:arden_variate_def seq_rhs_r_prop1)
+ − 1013
qed
+ − 1014
qed
+ − 1015
qed
+ − 1016
+ − 1017
text {* merge_rhs {(X1, r1), (x2, r2}, (x4, r4), \<dots>} {(x1, r1'), (x3, r3'), \<dots>} = {(x1, ALT r1 r1'}, (x2, r2), (x3, r3'), (x4, r4), \<dots>} *}
+ − 1018
definition
+ − 1019
merge_rhs :: "t_equa_rhs \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"
+ − 1020
where
+ − 1021
"merge_rhs rhs rhs' \<equiv> {(X, r). (\<exists> r1 r2. ((X,r1) \<in> rhs \<and> (X, r2) \<in> rhs') \<and> r = ALT r1 r2) \<or>
+ − 1022
(\<exists> r1. (X, r1) \<in> rhs \<and> (\<not> (\<exists> r2. (X, r2) \<in> rhs')) \<and> r = r1) \<or>
+ − 1023
(\<exists> r2. (X, r2) \<in> rhs' \<and> (\<not> (\<exists> r1. (X, r1) \<in> rhs)) \<and> r = r2) }"
+ − 1024
+ − 1025
+ − 1026
text {* rhs_subst rhs X=xrhs r: substitude all occurence of X in rhs((X,r) \<in> rhs) with xrhs *}
+ − 1027
definition
+ − 1028
rhs_subst :: "t_equa_rhs \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"
+ − 1029
where
+ − 1030
"rhs_subst rhs X xrhs r \<equiv> merge_rhs (del_x_paired rhs X) (seq_rhs_r xrhs r)"
+ − 1031
+ − 1032
definition
+ − 1033
equas_subst_f :: "string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa \<Rightarrow> t_equa"
+ − 1034
where
+ − 1035
"equas_subst_f X xrhs equa \<equiv> let (Y, rhs) = equa in
+ − 1036
if (\<exists> r. (X, r) \<in> rhs)
+ − 1037
then (Y, rhs_subst rhs X xrhs (SOME r. (X, r) \<in> rhs))
+ − 1038
else equa"
+ − 1039
+ − 1040
definition
+ − 1041
equas_subst :: "t_equas \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equas"
+ − 1042
where
+ − 1043
"equas_subst ES X xrhs \<equiv> del_x_paired (equas_subst_f X xrhs ` ES) X"
+ − 1044
+ − 1045
lemma lang_seq_prop1:
+ − 1046
"x \<in> X ; L r \<Longrightarrow> x \<in> X ; (L r \<union> L r')"
+ − 1047
by (auto simp:lang_seq_def)
+ − 1048
+ − 1049
lemma lang_seq_prop1':
+ − 1050
"x \<in> X; L r \<Longrightarrow> x \<in> X ; (L r' \<union> L r)"
+ − 1051
by (auto simp:lang_seq_def)
+ − 1052
+ − 1053
lemma lang_seq_prop2:
+ − 1054
"x \<in> X; (L r \<union> L r') \<Longrightarrow> x \<in> X;L r \<or> x \<in> X;L r'"
+ − 1055
by (auto simp:lang_seq_def)
+ − 1056
+ − 1057
lemma merge_rhs_prop1:
+ − 1058
shows "L (merge_rhs rhs rhs') = L rhs \<union> L rhs' "
+ − 1059
apply (auto simp add:merge_rhs_def L_def dest!:lang_seq_prop2 intro:lang_seq_prop1)
+ − 1060
apply (rule_tac x = X in exI, rule_tac x = r1 in exI, simp)
+ − 1061
apply (case_tac "\<exists> r2. (X, r2) \<in> rhs'")
+ − 1062
apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r r2" in exI, simp add:lang_seq_prop1)
+ − 1063
apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)
+ − 1064
apply (case_tac "\<exists> r1. (X, r1) \<in> rhs")
+ − 1065
apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r1 r" in exI, simp add:lang_seq_prop1')
+ − 1066
apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)
+ − 1067
done
+ − 1068
+ − 1069
lemma no_EMPTY_rhss_imp_merge_no_EMPTY:
+ − 1070
"\<lbrakk>no_EMPTY_rhs rhs; no_EMPTY_rhs rhs'\<rbrakk> \<Longrightarrow> no_EMPTY_rhs (merge_rhs rhs rhs')"
+ − 1071
apply (simp add:no_EMPTY_rhs_def merge_rhs_def)
+ − 1072
apply (clarify, (erule conjE | erule exE | erule disjE)+)
+ − 1073
by auto
+ − 1074
+ − 1075
lemma distinct_rhs_prop:
+ − 1076
"\<lbrakk>distinct_rhs rhs; (X, r1) \<in> rhs; (X, r2) \<in> rhs\<rbrakk> \<Longrightarrow> r1 = r2"
+ − 1077
by (auto simp:distinct_rhs_def)
+ − 1078
+ − 1079
lemma merge_rhs_prop2:
+ − 1080
assumes dist_rhs: "distinct_rhs rhs"
+ − 1081
and dist_rhs':"distinct_rhs rhs'"
+ − 1082
shows "distinct_rhs (merge_rhs rhs rhs')"
+ − 1083
apply (auto simp:merge_rhs_def distinct_rhs_def)
+ − 1084
using dist_rhs
+ − 1085
apply (drule distinct_rhs_prop, simp+)
+ − 1086
using dist_rhs'
+ − 1087
apply (drule distinct_rhs_prop, simp+)
+ − 1088
using dist_rhs
+ − 1089
apply (drule distinct_rhs_prop, simp+)
+ − 1090
using dist_rhs'
+ − 1091
apply (drule distinct_rhs_prop, simp+)
+ − 1092
done
+ − 1093
+ − 1094
lemma seq_rhs_r_holds_distinct:
+ − 1095
"distinct_rhs rhs \<Longrightarrow> distinct_rhs (seq_rhs_r rhs r)"
+ − 1096
by (auto simp:distinct_rhs_def seq_rhs_r_def)
+ − 1097
+ − 1098
lemma seq_rhs_r_prop0:
+ − 1099
assumes l_eq_r: "X = L xrhs"
+ − 1100
shows "L (seq_rhs_r xrhs r) = X ; L r "
+ − 1101
using l_eq_r
+ − 1102
by (auto simp:seq_rhs_r_prop1)
+ − 1103
+ − 1104
lemma rhs_subst_prop1:
+ − 1105
assumes l_eq_r: "X = L xrhs"
+ − 1106
and dist: "distinct_rhs rhs"
+ − 1107
shows "(X, r) \<in> rhs \<Longrightarrow> L rhs = L (rhs_subst rhs X xrhs r)"
+ − 1108
apply (simp add:rhs_subst_def merge_rhs_prop1)
+ − 1109
using l_eq_r
+ − 1110
apply (drule_tac r = r in seq_rhs_r_prop0, simp)
+ − 1111
using dist
+ − 1112
apply (auto dest:del_x_paired_prop1)
+ − 1113
done
+ − 1114
+ − 1115
lemma del_x_paired_holds_distinct_rhs:
+ − 1116
"distinct_rhs rhs \<Longrightarrow> distinct_rhs (del_x_paired rhs x)"
+ − 1117
by (auto simp:distinct_rhs_def del_x_paired_def)
+ − 1118
+ − 1119
lemma rhs_subst_holds_distinct_rhs:
+ − 1120
"\<lbrakk>distinct_rhs rhs; distinct_rhs xrhs\<rbrakk> \<Longrightarrow> distinct_rhs (rhs_subst rhs X xrhs r)"
+ − 1121
apply (drule_tac r = r and rhs = xrhs in seq_rhs_r_holds_distinct)
+ − 1122
apply (drule_tac x = X in del_x_paired_holds_distinct_rhs)
+ − 1123
by (auto dest:merge_rhs_prop2[where rhs = "del_x_paired rhs X"] simp:rhs_subst_def)
+ − 1124
+ − 1125
section {* myhill-nerode theorem *}
+ − 1126
+ − 1127
definition left_eq_cls :: "t_equas \<Rightarrow> (string set) set"
+ − 1128
where
+ − 1129
"left_eq_cls ES \<equiv> {X. \<exists> rhs. (X, rhs) \<in> ES} "
+ − 1130
+ − 1131
definition right_eq_cls :: "t_equas \<Rightarrow> (string set) set"
+ − 1132
where
+ − 1133
"right_eq_cls ES \<equiv> {Y. \<exists> X rhs r. (X, rhs) \<in> ES \<and> (Y, r) \<in> rhs }"
+ − 1134
+ − 1135
definition rhs_eq_cls :: "t_equa_rhs \<Rightarrow> (string set) set"
+ − 1136
where
+ − 1137
"rhs_eq_cls rhs \<equiv> {Y. \<exists> r. (Y, r) \<in> rhs}"
+ − 1138
+ − 1139
definition ardenable :: "t_equa \<Rightarrow> bool"
+ − 1140
where
+ − 1141
"ardenable equa \<equiv> let (X, rhs) = equa in
+ − 1142
distinct_rhs rhs \<and> no_EMPTY_rhs rhs \<and> X = L rhs"
+ − 1143
+ − 1144
text {*
+ − 1145
Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds.
+ − 1146
*}
+ − 1147
definition Inv :: "string set \<Rightarrow> t_equas \<Rightarrow> bool"
+ − 1148
where
+ − 1149
"Inv X ES \<equiv> finite ES \<and> (\<exists> rhs. (X, rhs) \<in> ES) \<and> distinct_equas ES \<and>
+ − 1150
(\<forall> X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs) \<and> X \<noteq> {} \<and> rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls ES))"
+ − 1151
+ − 1152
text {*
+ − 1153
TCon: Termination Condition of the equation-system decreasion.
+ − 1154
*}
+ − 1155
definition TCon:: "'a set \<Rightarrow> bool"
+ − 1156
where
+ − 1157
"TCon ES \<equiv> card ES = 1"
+ − 1158
+ − 1159
+ − 1160
text {*
+ − 1161
The following is a iteration principle, and is the main framework for the proof:
+ − 1162
1: We can form the initial equation-system using "equations" defined above, and prove it has invariance Inv by lemma "init_ES_satisfy_Inv";
+ − 1163
2: We can decrease the number of the equation-system using ardens_lemma_revised and substitution ("equas_subst", defined above),
+ − 1164
and prove it holds the property "step" in "wf_iter" by lemma "iteration_step"
+ − 1165
and finally using property P and Q to prove the myhill-nerode theorem
+ − 1166
+ − 1167
*}
+ − 1168
lemma wf_iter [rule_format]:
+ − 1169
fixes f
+ − 1170
assumes step: "\<And> e. \<lbrakk>P e; \<not> Q e\<rbrakk> \<Longrightarrow> (\<exists> e'. P e' \<and> (f(e'), f(e)) \<in> less_than)"
+ − 1171
shows pe: "P e \<longrightarrow> (\<exists> e'. P e' \<and> Q e')"
+ − 1172
proof(induct e rule: wf_induct
+ − 1173
[OF wf_inv_image[OF wf_less_than, where f = "f"]], clarify)
+ − 1174
fix x
+ − 1175
assume h [rule_format]:
+ − 1176
"\<forall>y. (y, x) \<in> inv_image less_than f \<longrightarrow> P y \<longrightarrow> (\<exists>e'. P e' \<and> Q e')"
+ − 1177
and px: "P x"
+ − 1178
show "\<exists>e'. P e' \<and> Q e'"
+ − 1179
proof(cases "Q x")
+ − 1180
assume "Q x" with px show ?thesis by blast
+ − 1181
next
+ − 1182
assume nq: "\<not> Q x"
+ − 1183
from step [OF px nq]
+ − 1184
obtain e' where pe': "P e'" and ltf: "(f e', f x) \<in> less_than" by auto
+ − 1185
show ?thesis
+ − 1186
proof(rule h)
+ − 1187
from ltf show "(e', x) \<in> inv_image less_than f"
+ − 1188
by (simp add:inv_image_def)
+ − 1189
next
+ − 1190
from pe' show "P e'" .
+ − 1191
qed
+ − 1192
qed
+ − 1193
qed
+ − 1194
+ − 1195
+ − 1196
(* ********************************* BEGIN: proving the initial equation-system satisfies Inv **************************************** *)
+ − 1197
+ − 1198
lemma distinct_rhs_equations:
+ − 1199
"(X, xrhs) \<in> equations (UNIV Quo Lang) \<Longrightarrow> distinct_rhs xrhs"
+ − 1200
by (auto simp: equations_def equation_rhs_def distinct_rhs_def empty_rhs_def dest:no_two_cls_inters)
+ − 1201
+ − 1202
lemma every_nonempty_eqclass_has_strings:
+ − 1203
"\<lbrakk>X \<in> (UNIV Quo Lang); X \<noteq> {[]}\<rbrakk> \<Longrightarrow> \<exists> clist. clist \<in> X \<and> clist \<noteq> []"
+ − 1204
by (auto simp:quot_def equiv_class_def equiv_str_def)
+ − 1205
+ − 1206
lemma every_eqclass_is_derived_from_empty:
+ − 1207
assumes not_empty: "X \<noteq> {[]}"
+ − 1208
shows "X \<in> (UNIV Quo Lang) \<Longrightarrow> \<exists> clist. {[]};{clist} \<subseteq> X \<and> clist \<noteq> []"
+ − 1209
using not_empty
+ − 1210
apply (drule_tac every_nonempty_eqclass_has_strings, simp)
+ − 1211
by (auto intro:exI[where x = clist] simp:lang_seq_def)
+ − 1212
+ − 1213
lemma equiv_str_in_CS:
+ − 1214
"\<lbrakk>clist\<rbrakk>Lang \<in> (UNIV Quo Lang)"
+ − 1215
by (auto simp:quot_def)
+ − 1216
+ − 1217
lemma has_str_imp_defined_by_str:
+ − 1218
"\<lbrakk>str \<in> X; X \<in> UNIV Quo Lang\<rbrakk> \<Longrightarrow> X = \<lbrakk>str\<rbrakk>Lang"
+ − 1219
by (auto simp:quot_def equiv_class_def equiv_str_def)
+ − 1220
+ − 1221
lemma every_eqclass_has_ascendent:
+ − 1222
assumes has_str: "clist @ [c] \<in> X"
+ − 1223
and in_CS: "X \<in> UNIV Quo Lang"
+ − 1224
shows "\<exists> Y. Y \<in> UNIV Quo Lang \<and> Y-c\<rightarrow>X \<and> clist \<in> Y" (is "\<exists> Y. ?P Y")
+ − 1225
proof -
+ − 1226
have "?P (\<lbrakk>clist\<rbrakk>Lang)"
+ − 1227
proof -
+ − 1228
have "\<lbrakk>clist\<rbrakk>Lang \<in> UNIV Quo Lang"
+ − 1229
by (simp add:quot_def, rule_tac x = clist in exI, simp)
+ − 1230
moreover have "\<lbrakk>clist\<rbrakk>Lang-c\<rightarrow>X"
+ − 1231
proof -
+ − 1232
have "X = \<lbrakk>(clist @ [c])\<rbrakk>Lang" using has_str in_CS
+ − 1233
by (auto intro!:has_str_imp_defined_by_str)
+ − 1234
moreover have "\<forall> sl. sl \<in> \<lbrakk>clist\<rbrakk>Lang \<longrightarrow> sl @ [c] \<in> \<lbrakk>(clist @ [c])\<rbrakk>Lang"
+ − 1235
by (auto simp:equiv_class_def equiv_str_def)
+ − 1236
ultimately show ?thesis unfolding CT_def lang_seq_def
+ − 1237
by auto
+ − 1238
qed
+ − 1239
moreover have "clist \<in> \<lbrakk>clist\<rbrakk>Lang"
+ − 1240
by (auto simp:equiv_str_def equiv_class_def)
+ − 1241
ultimately show "?P (\<lbrakk>clist\<rbrakk>Lang)" by simp
+ − 1242
qed
+ − 1243
thus ?thesis by blast
+ − 1244
qed
+ − 1245
+ − 1246
lemma finite_charset_rS:
+ − 1247
"finite {CHAR c |c. Y-c\<rightarrow>X}"
+ − 1248
by (rule_tac A = UNIV and f = CHAR in finite_surj, auto)
+ − 1249
+ − 1250
lemma l_eq_r_in_equations:
+ − 1251
assumes X_in_equas: "(X, xrhs) \<in> equations (UNIV Quo Lang)"
+ − 1252
shows "X = L xrhs"
+ − 1253
proof (cases "X = {[]}")
+ − 1254
case True
+ − 1255
thus ?thesis using X_in_equas
+ − 1256
by (simp add:equations_def equation_rhs_def L_def lang_seq_def)
+ − 1257
next
+ − 1258
case False
+ − 1259
show ?thesis
+ − 1260
proof
+ − 1261
show "X \<subseteq> L xrhs"
+ − 1262
proof
+ − 1263
fix x
+ − 1264
assume "(1)": "x \<in> X"
+ − 1265
show "x \<in> L xrhs"
+ − 1266
proof (cases "x = []")
+ − 1267
assume empty: "x = []"
+ − 1268
hence "x \<in> L (empty_rhs X)" using "(1)"
+ − 1269
by (auto simp:empty_rhs_def L_def lang_seq_def)
+ − 1270
thus ?thesis using X_in_equas False empty "(1)"
+ − 1271
unfolding equations_def equation_rhs_def by (auto simp:L_def)
+ − 1272
next
+ − 1273
assume not_empty: "x \<noteq> []"
+ − 1274
hence "\<exists> clist c. x = clist @ [c]" by (case_tac x rule:rev_cases, auto)
+ − 1275
then obtain clist c where decom: "x = clist @ [c]" by blast
+ − 1276
moreover have "\<And> Y. \<lbrakk>Y \<in> UNIV Quo Lang; Y-c\<rightarrow>X; clist \<in> Y\<rbrakk>\<Longrightarrow> [c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})"
+ − 1277
proof -
+ − 1278
fix Y
+ − 1279
assume Y_is_eq_cl: "Y \<in> UNIV Quo Lang"
+ − 1280
and Y_CT_X: "Y-c\<rightarrow>X"
+ − 1281
and clist_in_Y: "clist \<in> Y"
+ − 1282
with finite_charset_rS
+ − 1283
show "[c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})"
+ − 1284
by (auto simp :fold_alt_null_eqs)
+ − 1285
qed
+ − 1286
hence "\<exists>Xa. Xa \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})"
+ − 1287
using X_in_equas False not_empty "(1)" decom
+ − 1288
by (auto dest!:every_eqclass_has_ascendent simp:equations_def equation_rhs_def L_def lang_seq_def)
+ − 1289
then obtain Xa where "Xa \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})" by blast
+ − 1290
hence "x \<in> L {(S, folds ALT NULL {CHAR c |c. S-c\<rightarrow>X}) |S. S \<in> UNIV Quo Lang}" using X_in_equas "(1)" decom
+ − 1291
by (auto simp add:L_def equations_def equation_rhs_def intro!:exI[where x = Xa])
+ − 1292
thus "x \<in> L xrhs" using X_in_equas False not_empty unfolding equations_def equation_rhs_def
+ − 1293
by (auto simp:L_def)
+ − 1294
qed
+ − 1295
qed
+ − 1296
next
+ − 1297
show "L xrhs \<subseteq> X"
+ − 1298
proof
+ − 1299
fix x
+ − 1300
assume "(2)": "x \<in> L xrhs"
+ − 1301
have "(2_1)": "\<And> s1 s2 r Xa. \<lbrakk>s1 \<in> Xa; s2 \<in> L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X"
+ − 1302
using finite_charset_rS
+ − 1303
by (auto simp:CT_def lang_seq_def fold_alt_null_eqs)
+ − 1304
have "(2_2)": "\<And> s1 s2 Xa r.\<lbrakk>s1 \<in> Xa; s2 \<in> L r; (Xa, r) \<in> empty_rhs X\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X"
+ − 1305
by (simp add:empty_rhs_def split:if_splits)
+ − 1306
show "x \<in> X" using X_in_equas False "(2)"
+ − 1307
by (auto intro:"(2_1)" "(2_2)" simp:equations_def equation_rhs_def L_def lang_seq_def)
+ − 1308
qed
+ − 1309
qed
+ − 1310
qed
+ − 1311
+ − 1312
lemma finite_CT_chars:
+ − 1313
"finite {CHAR c |c. Xa-c\<rightarrow>X}"
+ − 1314
by (auto)
+ − 1315
+ − 1316
lemma no_EMPTY_equations:
+ − 1317
"(X, xrhs) \<in> equations CS \<Longrightarrow> no_EMPTY_rhs xrhs"
+ − 1318
apply (clarsimp simp add:equations_def equation_rhs_def)
+ − 1319
apply (simp add:no_EMPTY_rhs_def empty_rhs_def, auto)
+ − 1320
apply (subgoal_tac "finite {CHAR c |c. Xa-c\<rightarrow>X}", drule_tac x = "[]" in fold_alt_null_eqs, clarsimp, rule finite_CT_chars)+
+ − 1321
done
+ − 1322
+ − 1323
lemma init_ES_satisfy_ardenable:
+ − 1324
"(X, xrhs) \<in> equations (UNIV Quo Lang) \<Longrightarrow> ardenable (X, xrhs)"
+ − 1325
unfolding ardenable_def
+ − 1326
by (auto intro:distinct_rhs_equations no_EMPTY_equations simp:l_eq_r_in_equations)
+ − 1327
+ − 1328
lemma init_ES_satisfy_Inv:
+ − 1329
assumes finite_CS: "finite (UNIV Quo Lang)"
+ − 1330
and X_in_eq_cls: "X \<in> UNIV Quo Lang"
+ − 1331
shows "Inv X (equations (UNIV Quo Lang))"
+ − 1332
proof -
+ − 1333
have "finite (equations (UNIV Quo Lang))" using finite_CS
+ − 1334
by (auto simp:equations_def)
+ − 1335
moreover have "\<exists>rhs. (X, rhs) \<in> equations (UNIV Quo Lang)" using X_in_eq_cls
+ − 1336
by (simp add:equations_def)
+ − 1337
moreover have "distinct_equas (equations (UNIV Quo Lang))"
+ − 1338
by (auto simp:distinct_equas_def equations_def)
+ − 1339
moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow>
+ − 1340
rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equations (UNIV Quo Lang)))"
+ − 1341
apply (simp add:left_eq_cls_def equations_def rhs_eq_cls_def equation_rhs_def)
+ − 1342
by (auto simp:empty_rhs_def split:if_splits)
+ − 1343
moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> X \<noteq> {}"
+ − 1344
by (clarsimp simp:equations_def empty_notin_CS intro:classical)
+ − 1345
moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> ardenable (X, xrhs)"
+ − 1346
by (auto intro!:init_ES_satisfy_ardenable)
+ − 1347
ultimately show ?thesis by (simp add:Inv_def)
+ − 1348
qed
+ − 1349
+ − 1350
+ − 1351
(* ********************************* END: proving the initial equation-system satisfies Inv **************************************** *)
+ − 1352
+ − 1353
+ − 1354
(* ***************************** BEGIN: proving every equation-system's iteration step satisfies Inv ******************************* *)
+ − 1355
+ − 1356
lemma not_T_aux: "\<lbrakk>\<not> TCon (insert a A); x = a\<rbrakk>
+ − 1357
\<Longrightarrow> \<exists>y. x \<noteq> y \<and> y \<in> insert a A "
+ − 1358
apply (case_tac "insert a A = {a}")
+ − 1359
by (auto simp:TCon_def)
+ − 1360
+ − 1361
lemma not_T_atleast_2[rule_format]:
+ − 1362
"finite S \<Longrightarrow> \<forall> x. x \<in> S \<and> (\<not> TCon S)\<longrightarrow> (\<exists> y. x \<noteq> y \<and> y \<in> S)"
+ − 1363
apply (erule finite.induct, simp)
+ − 1364
apply (clarify, case_tac "x = a")
+ − 1365
by (erule not_T_aux, auto)
+ − 1366
+ − 1367
lemma exist_another_equa:
+ − 1368
"\<lbrakk>\<not> TCon ES; finite ES; distinct_equas ES; (X, rhl) \<in> ES\<rbrakk> \<Longrightarrow> \<exists> Y yrhl. (Y, yrhl) \<in> ES \<and> X \<noteq> Y"
+ − 1369
apply (drule not_T_atleast_2, simp)
+ − 1370
apply (clarsimp simp:distinct_equas_def)
+ − 1371
apply (drule_tac x= X in spec, drule_tac x = rhl in spec, drule_tac x = b in spec)
+ − 1372
by auto
+ − 1373
+ − 1374
lemma Inv_mono_with_lambda:
+ − 1375
assumes Inv_ES: "Inv X ES"
+ − 1376
and X_noteq_Y: "X \<noteq> {[]}"
+ − 1377
shows "Inv X (ES - {({[]}, yrhs)})"
+ − 1378
proof -
+ − 1379
have "finite (ES - {({[]}, yrhs)})" using Inv_ES
+ − 1380
by (simp add:Inv_def)
+ − 1381
moreover have "\<exists>rhs. (X, rhs) \<in> ES - {({[]}, yrhs)}" using Inv_ES X_noteq_Y
+ − 1382
by (simp add:Inv_def)
+ − 1383
moreover have "distinct_equas (ES - {({[]}, yrhs)})" using Inv_ES X_noteq_Y
+ − 1384
apply (clarsimp simp:Inv_def distinct_equas_def)
+ − 1385
by (drule_tac x = Xa in spec, simp)
+ − 1386
moreover have "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow>
+ − 1387
ardenable (X, xrhs) \<and> X \<noteq> {}" using Inv_ES
+ − 1388
by (clarify, simp add:Inv_def)
+ − 1389
moreover
+ − 1390
have "insert {[]} (left_eq_cls (ES - {({[]}, yrhs)})) = insert {[]} (left_eq_cls ES)"
+ − 1391
by (auto simp:left_eq_cls_def)
+ − 1392
hence "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow>
+ − 1393
rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (ES - {({[]}, yrhs)}))"
+ − 1394
using Inv_ES by (auto simp:Inv_def)
+ − 1395
ultimately show ?thesis by (simp add:Inv_def)
+ − 1396
qed
+ − 1397
+ − 1398
lemma non_empty_card_prop:
+ − 1399
"finite ES \<Longrightarrow> \<forall>e. e \<in> ES \<longrightarrow> card ES - Suc 0 < card ES"
+ − 1400
apply (erule finite.induct, simp)
+ − 1401
apply (case_tac[!] "a \<in> A")
+ − 1402
by (auto simp:insert_absorb)
+ − 1403
+ − 1404
lemma ardenable_prop:
+ − 1405
assumes not_lambda: "Y \<noteq> {[]}"
+ − 1406
and ardable: "ardenable (Y, yrhs)"
+ − 1407
shows "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" (is "\<exists> yrhs'. ?P yrhs'")
+ − 1408
proof (cases "(\<exists> reg. (Y, reg) \<in> yrhs)")
+ − 1409
case True
+ − 1410
thus ?thesis
+ − 1411
proof
+ − 1412
fix reg
+ − 1413
assume self_contained: "(Y, reg) \<in> yrhs"
+ − 1414
show ?thesis
+ − 1415
proof -
+ − 1416
have "?P (arden_variate Y reg yrhs)"
+ − 1417
proof -
+ − 1418
have "Y = L (arden_variate Y reg yrhs)"
+ − 1419
using self_contained not_lambda ardable
+ − 1420
by (rule_tac arden_variate_valid, simp_all add:ardenable_def)
+ − 1421
moreover have "distinct_rhs (arden_variate Y reg yrhs)"
+ − 1422
using ardable
+ − 1423
by (auto simp:distinct_rhs_def arden_variate_def seq_rhs_r_def del_x_paired_def ardenable_def)
+ − 1424
moreover have "rhs_eq_cls (arden_variate Y reg yrhs) = rhs_eq_cls yrhs - {Y}"
+ − 1425
proof -
+ − 1426
have "\<And> rhs r. rhs_eq_cls (seq_rhs_r rhs r) = rhs_eq_cls rhs"
+ − 1427
apply (auto simp:rhs_eq_cls_def seq_rhs_r_def image_def)
+ − 1428
by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "(x, ra)" in bexI, simp+)
+ − 1429
moreover have "\<And> rhs X. rhs_eq_cls (del_x_paired rhs X) = rhs_eq_cls rhs - {X}"
+ − 1430
by (auto simp:rhs_eq_cls_def del_x_paired_def)
+ − 1431
ultimately show ?thesis by (simp add:arden_variate_def)
+ − 1432
qed
+ − 1433
ultimately show ?thesis by simp
+ − 1434
qed
+ − 1435
thus ?thesis by (rule_tac x= "arden_variate Y reg yrhs" in exI, simp)
+ − 1436
qed
+ − 1437
qed
+ − 1438
next
+ − 1439
case False
+ − 1440
hence "(2)": "rhs_eq_cls yrhs - {Y} = rhs_eq_cls yrhs"
+ − 1441
by (auto simp:rhs_eq_cls_def)
+ − 1442
show ?thesis
+ − 1443
proof -
+ − 1444
have "?P yrhs" using False ardable "(2)"
+ − 1445
by (simp add:ardenable_def)
+ − 1446
thus ?thesis by blast
+ − 1447
qed
+ − 1448
qed
+ − 1449
+ − 1450
lemma equas_subst_f_del_no_other:
+ − 1451
assumes self_contained: "(Y, rhs) \<in> ES"
+ − 1452
shows "\<exists> rhs'. (Y, rhs') \<in> (equas_subst_f X xrhs ` ES)" (is "\<exists> rhs'. ?P rhs'")
+ − 1453
proof -
+ − 1454
have "\<exists> rhs'. equas_subst_f X xrhs (Y, rhs) = (Y, rhs')"
+ − 1455
by (auto simp:equas_subst_f_def)
+ − 1456
then obtain rhs' where "equas_subst_f X xrhs (Y, rhs) = (Y, rhs')" by blast
+ − 1457
hence "?P rhs'" unfolding image_def using self_contained
+ − 1458
by (auto intro:bexI[where x = "(Y, rhs)"])
+ − 1459
thus ?thesis by blast
+ − 1460
qed
+ − 1461
+ − 1462
lemma del_x_paired_del_only_x:
+ − 1463
"\<lbrakk>X \<noteq> Y; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> (X, rhs) \<in> del_x_paired ES Y"
+ − 1464
by (auto simp:del_x_paired_def)
+ − 1465
+ − 1466
lemma del_x_paired_del_only_x':
+ − 1467
"(X, rhs) \<in> del_x_paired ES Y \<Longrightarrow> X \<noteq> Y \<and> (X, rhs) \<in> ES"
+ − 1468
by (auto simp:del_x_paired_def)
+ − 1469
+ − 1470
lemma equas_subst_del_no_other:
+ − 1471
"\<lbrakk>(X, rhs) \<in> ES; X \<noteq> Y\<rbrakk> \<Longrightarrow> (\<exists>rhs. (X, rhs) \<in> equas_subst ES Y rhs')"
+ − 1472
unfolding equas_subst_def
+ − 1473
apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other)
+ − 1474
by (erule exE, drule del_x_paired_del_only_x, auto)
+ − 1475
+ − 1476
lemma equas_subst_holds_distinct:
+ − 1477
"distinct_equas ES \<Longrightarrow> distinct_equas (equas_subst ES Y rhs')"
+ − 1478
apply (clarsimp simp add:equas_subst_def distinct_equas_def del_x_paired_def equas_subst_f_def)
+ − 1479
by (auto split:if_splits)
+ − 1480
+ − 1481
lemma del_x_paired_dels:
+ − 1482
"(X, rhs) \<in> ES \<Longrightarrow> {Y. Y \<in> ES \<and> fst Y = X} \<inter> ES \<noteq> {}"
+ − 1483
by (auto)
+ − 1484
+ − 1485
lemma del_x_paired_subset:
+ − 1486
"(X, rhs) \<in> ES \<Longrightarrow> ES - {Y. Y \<in> ES \<and> fst Y = X} \<subset> ES"
+ − 1487
apply (drule del_x_paired_dels)
+ − 1488
by auto
+ − 1489
+ − 1490
lemma del_x_paired_card_less:
+ − 1491
"\<lbrakk>finite ES; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> card (del_x_paired ES X) < card ES"
+ − 1492
apply (simp add:del_x_paired_def)
+ − 1493
apply (drule del_x_paired_subset)
+ − 1494
by (auto intro:psubset_card_mono)
+ − 1495
+ − 1496
lemma equas_subst_card_less:
+ − 1497
"\<lbrakk>finite ES; (Y, yrhs) \<in> ES\<rbrakk> \<Longrightarrow> card (equas_subst ES Y rhs') < card ES"
+ − 1498
apply (simp add:equas_subst_def)
+ − 1499
apply (frule_tac h = "equas_subst_f Y rhs'" in finite_imageI)
+ − 1500
apply (drule_tac f = "equas_subst_f Y rhs'" in Finite_Set.card_image_le)
+ − 1501
apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other,erule exE)
+ − 1502
by (drule del_x_paired_card_less, auto)
+ − 1503
+ − 1504
lemma equas_subst_holds_distinct_rhs:
+ − 1505
assumes dist': "distinct_rhs yrhs'"
+ − 1506
and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+ − 1507
and X_in : "(X, xrhs) \<in> equas_subst ES Y yrhs'"
+ − 1508
shows "distinct_rhs xrhs"
+ − 1509
using X_in history
+ − 1510
apply (clarsimp simp:equas_subst_def del_x_paired_def)
+ − 1511
apply (drule_tac x = a in spec, drule_tac x = b in spec)
+ − 1512
apply (simp add:ardenable_def equas_subst_f_def)
+ − 1513
by (auto intro:rhs_subst_holds_distinct_rhs simp:dist' split:if_splits)
+ − 1514
+ − 1515
lemma r_no_EMPTY_imp_seq_rhs_r_no_EMPTY:
+ − 1516
"[] \<notin> L r \<Longrightarrow> no_EMPTY_rhs (seq_rhs_r rhs r)"
+ − 1517
by (auto simp:no_EMPTY_rhs_def seq_rhs_r_def lang_seq_def)
+ − 1518
+ − 1519
lemma del_x_paired_holds_no_EMPTY:
+ − 1520
"no_EMPTY_rhs yrhs \<Longrightarrow> no_EMPTY_rhs (del_x_paired yrhs Y)"
+ − 1521
by (auto simp:no_EMPTY_rhs_def del_x_paired_def)
+ − 1522
+ − 1523
lemma rhs_subst_holds_no_EMPTY:
+ − 1524
"\<lbrakk>no_EMPTY_rhs yrhs; (Y, r) \<in> yrhs; Y \<noteq> {[]}\<rbrakk> \<Longrightarrow> no_EMPTY_rhs (rhs_subst yrhs Y rhs' r)"
+ − 1525
apply (auto simp:rhs_subst_def intro!:no_EMPTY_rhss_imp_merge_no_EMPTY r_no_EMPTY_imp_seq_rhs_r_no_EMPTY del_x_paired_holds_no_EMPTY)
+ − 1526
by (auto simp:no_EMPTY_rhs_def)
+ − 1527
+ − 1528
lemma equas_subst_holds_no_EMPTY:
+ − 1529
assumes substor: "Y \<noteq> {[]}"
+ − 1530
and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+ − 1531
and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'"
+ − 1532
shows "no_EMPTY_rhs xrhs"
+ − 1533
proof-
+ − 1534
from X_in have "\<exists> (Z, zrhs) \<in> ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)"
+ − 1535
by (auto simp add:equas_subst_def del_x_paired_def)
+ − 1536
then obtain Z zrhs where Z_in: "(Z, zrhs) \<in> ES"
+ − 1537
and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast
+ − 1538
hence dist_zrhs: "distinct_rhs zrhs" using history
+ − 1539
by (auto simp:ardenable_def)
+ − 1540
show ?thesis
+ − 1541
proof (cases "\<exists> r. (Y, r) \<in> zrhs")
+ − 1542
case True
+ − 1543
then obtain r where Y_in_zrhs: "(Y, r) \<in> zrhs" ..
+ − 1544
hence some: "(SOME r. (Y, r) \<in> zrhs) = r" using Z_in dist_zrhs
+ − 1545
by (auto simp:distinct_rhs_def)
+ − 1546
hence "no_EMPTY_rhs (rhs_subst zrhs Y yrhs' r)"
+ − 1547
using substor Y_in_zrhs history Z_in
+ − 1548
by (rule_tac rhs_subst_holds_no_EMPTY, auto simp:ardenable_def)
+ − 1549
thus ?thesis using X_Z True some
+ − 1550
by (simp add:equas_subst_def equas_subst_f_def)
+ − 1551
next
+ − 1552
case False
+ − 1553
hence "(X, xrhs) = (Z, zrhs)" using Z_in X_Z
+ − 1554
by (simp add:equas_subst_f_def)
+ − 1555
thus ?thesis using history Z_in
+ − 1556
by (auto simp:ardenable_def)
+ − 1557
qed
+ − 1558
qed
+ − 1559
+ − 1560
lemma equas_subst_f_holds_left_eq_right:
+ − 1561
assumes substor: "Y = L rhs'"
+ − 1562
and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> distinct_rhs xrhs \<and> X = L xrhs"
+ − 1563
and subst: "(X, xrhs) = equas_subst_f Y rhs' (Z, zrhs)"
+ − 1564
and self_contained: "(Z, zrhs) \<in> ES"
+ − 1565
shows "X = L xrhs"
+ − 1566
proof (cases "\<exists> r. (Y, r) \<in> zrhs")
+ − 1567
case True
+ − 1568
from True obtain r where "(1)":"(Y, r) \<in> zrhs" ..
+ − 1569
show ?thesis
+ − 1570
proof -
+ − 1571
from history self_contained
+ − 1572
have dist: "distinct_rhs zrhs" by auto
+ − 1573
hence "(SOME r. (Y, r) \<in> zrhs) = r" using self_contained "(1)"
+ − 1574
using distinct_rhs_def by (auto intro:some_equality)
+ − 1575
moreover have "L zrhs = L (rhs_subst zrhs Y rhs' r)" using substor dist "(1)" self_contained
+ − 1576
by (rule_tac rhs_subst_prop1, auto simp:distinct_equas_rhs_def)
+ − 1577
ultimately show ?thesis using subst history self_contained
+ − 1578
by (auto simp:equas_subst_f_def split:if_splits)
+ − 1579
qed
+ − 1580
next
+ − 1581
case False
+ − 1582
thus ?thesis using history subst self_contained
+ − 1583
by (auto simp:equas_subst_f_def)
+ − 1584
qed
+ − 1585
+ − 1586
lemma equas_subst_holds_left_eq_right:
+ − 1587
assumes history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+ − 1588
and substor: "Y = L rhs'"
+ − 1589
and X_in : "(X, xrhs) \<in> equas_subst ES Y yrhs'"
+ − 1590
shows "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y rhs' \<longrightarrow> X = L xrhs"
+ − 1591
apply (clarsimp simp add:equas_subst_def del_x_paired_def)
+ − 1592
using substor
+ − 1593
apply (drule_tac equas_subst_f_holds_left_eq_right)
+ − 1594
using history
+ − 1595
by (auto simp:ardenable_def)
+ − 1596
+ − 1597
lemma equas_subst_holds_ardenable:
+ − 1598
assumes substor: "Y = L yrhs'"
+ − 1599
and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+ − 1600
and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'"
+ − 1601
and dist': "distinct_rhs yrhs'"
+ − 1602
and not_lambda: "Y \<noteq> {[]}"
+ − 1603
shows "ardenable (X, xrhs)"
+ − 1604
proof -
+ − 1605
have "distinct_rhs xrhs" using history X_in dist'
+ − 1606
by (auto dest:equas_subst_holds_distinct_rhs)
+ − 1607
moreover have "no_EMPTY_rhs xrhs" using history X_in not_lambda
+ − 1608
by (auto intro:equas_subst_holds_no_EMPTY)
+ − 1609
moreover have "X = L xrhs" using history substor X_in
+ − 1610
by (auto dest: equas_subst_holds_left_eq_right)
+ − 1611
ultimately show ?thesis using ardenable_def by simp
+ − 1612
qed
+ − 1613
+ − 1614
lemma equas_subst_holds_cls_defined:
+ − 1615
assumes X_in: "(X, xrhs) \<in> equas_subst ES Y yrhs'"
+ − 1616
and Inv_ES: "Inv X' ES"
+ − 1617
and subst: "(Y, yrhs) \<in> ES"
+ − 1618
and cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}"
+ − 1619
shows "rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))"
+ − 1620
proof-
+ − 1621
have tac: "\<lbrakk> A \<subseteq> B; C \<subseteq> D; E \<subseteq> A \<union> B\<rbrakk> \<Longrightarrow> E \<subseteq> B \<union> D" by auto
+ − 1622
from X_in have "\<exists> (Z, zrhs) \<in> ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)"
+ − 1623
by (auto simp add:equas_subst_def del_x_paired_def)
+ − 1624
then obtain Z zrhs where Z_in: "(Z, zrhs) \<in> ES"
+ − 1625
and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast
+ − 1626
hence "rhs_eq_cls zrhs \<subseteq> insert {[]} (left_eq_cls ES)" using Inv_ES
+ − 1627
by (auto simp:Inv_def)
+ − 1628
moreover have "rhs_eq_cls yrhs' \<subseteq> insert {[]} (left_eq_cls ES) - {Y}"
+ − 1629
using Inv_ES subst cls_holds_but_Y
+ − 1630
by (auto simp:Inv_def)
+ − 1631
moreover have "rhs_eq_cls xrhs \<subseteq> rhs_eq_cls zrhs \<union> rhs_eq_cls yrhs' - {Y}"
+ − 1632
using X_Z cls_holds_but_Y
+ − 1633
apply (clarsimp simp add:equas_subst_f_def rhs_subst_def split:if_splits)
+ − 1634
by (auto simp:rhs_eq_cls_def merge_rhs_def del_x_paired_def seq_rhs_r_def)
+ − 1635
moreover have "left_eq_cls (equas_subst ES Y yrhs') = left_eq_cls ES - {Y}" using subst
+ − 1636
by (auto simp: left_eq_cls_def equas_subst_def del_x_paired_def equas_subst_f_def
+ − 1637
dest: equas_subst_f_del_no_other
+ − 1638
split: if_splits)
+ − 1639
ultimately show ?thesis by blast
+ − 1640
qed
+ − 1641
+ − 1642
lemma iteration_step:
+ − 1643
assumes Inv_ES: "Inv X ES"
+ − 1644
and not_T: "\<not> TCon ES"
+ − 1645
shows "(\<exists> ES'. Inv X ES' \<and> (card ES', card ES) \<in> less_than)"
+ − 1646
proof -
+ − 1647
from Inv_ES not_T have another: "\<exists>Y yrhs. (Y, yrhs) \<in> ES \<and> X \<noteq> Y" unfolding Inv_def
+ − 1648
by (clarify, rule_tac exist_another_equa[where X = X], auto)
+ − 1649
then obtain Y yrhs where subst: "(Y, yrhs) \<in> ES" and not_X: " X \<noteq> Y" by blast
+ − 1650
show ?thesis (is "\<exists> ES'. ?P ES'")
+ − 1651
proof (cases "Y = {[]}")
+ − 1652
case True
+ − 1653
--"in this situation, we pick a \"\<lambda>\" equation, thus directly remove it from the equation-system"
+ − 1654
have "?P (ES - {(Y, yrhs)})"
+ − 1655
proof
+ − 1656
show "Inv X (ES - {(Y, yrhs)})" using True not_X
+ − 1657
by (simp add:Inv_ES Inv_mono_with_lambda)
+ − 1658
next
+ − 1659
show "(card (ES - {(Y, yrhs)}), card ES) \<in> less_than" using Inv_ES subst
+ − 1660
by (auto elim:non_empty_card_prop[rule_format] simp:Inv_def)
+ − 1661
qed
+ − 1662
thus ?thesis by blast
+ − 1663
next
+ − 1664
case False
+ − 1665
--"in this situation, we pick a equation and using ardenable to get a rhs without itself in it, then use equas_subst to form a new equation-system"
+ − 1666
hence "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" using subst Inv_ES
+ − 1667
by (auto intro:ardenable_prop simp:Inv_def)
+ − 1668
then obtain yrhs' where Y'_l_eq_r: "Y = L yrhs'"
+ − 1669
and dist_Y': "distinct_rhs yrhs'"
+ − 1670
and cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" by blast
+ − 1671
hence "?P (equas_subst ES Y yrhs')"
+ − 1672
proof -
+ − 1673
have finite_del: "\<And> S x. finite S \<Longrightarrow> finite (del_x_paired S x)"
+ − 1674
apply (rule_tac A = "del_x_paired S x" in finite_subset)
+ − 1675
by (auto simp:del_x_paired_def)
+ − 1676
have "finite (equas_subst ES Y yrhs')" using Inv_ES
+ − 1677
by (auto intro!:finite_del simp:equas_subst_def Inv_def)
+ − 1678
moreover have "\<exists>rhs. (X, rhs) \<in> equas_subst ES Y yrhs'" using Inv_ES not_X
+ − 1679
by (auto intro:equas_subst_del_no_other simp:Inv_def)
+ − 1680
moreover have "distinct_equas (equas_subst ES Y yrhs')" using Inv_ES
+ − 1681
by (auto intro:equas_subst_holds_distinct simp:Inv_def)
+ − 1682
moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> ardenable (X, xrhs)"
+ − 1683
using Inv_ES dist_Y' False Y'_l_eq_r
+ − 1684
apply (clarsimp simp:Inv_def)
+ − 1685
by (rule equas_subst_holds_ardenable, simp_all)
+ − 1686
moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> X \<noteq> {}" using Inv_ES
+ − 1687
by (clarsimp simp:equas_subst_def Inv_def del_x_paired_def equas_subst_f_def split:if_splits, auto)
+ − 1688
moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow>
+ − 1689
rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))"
+ − 1690
using Inv_ES subst cls_holds_but_Y
+ − 1691
apply (rule_tac impI | rule_tac allI)+
+ − 1692
by (erule equas_subst_holds_cls_defined, auto)
+ − 1693
moreover have "(card (equas_subst ES Y yrhs'), card ES) \<in> less_than"using Inv_ES subst
+ − 1694
by (simp add:equas_subst_card_less Inv_def)
+ − 1695
ultimately show "?P (equas_subst ES Y yrhs')" by (auto simp:Inv_def)
+ − 1696
qed
+ − 1697
thus ?thesis by blast
+ − 1698
qed
+ − 1699
qed
+ − 1700
+ − 1701
(* ****************************** END: proving every equation-system's iteration step satisfies Inv ******************************* *)
+ − 1702
+ − 1703
lemma iteration_conc:
+ − 1704
assumes history: "Inv X ES"
+ − 1705
shows "\<exists> ES'. Inv X ES' \<and> TCon ES'" (is "\<exists> ES'. ?P ES'")
+ − 1706
proof (cases "TCon ES")
+ − 1707
case True
+ − 1708
hence "?P ES" using history by simp
+ − 1709
thus ?thesis by blast
+ − 1710
next
+ − 1711
case False
+ − 1712
thus ?thesis using history iteration_step
+ − 1713
by (rule_tac f = card in wf_iter, simp_all)
+ − 1714
qed
+ − 1715
+ − 1716
lemma eqset_imp_iff': "A = B \<Longrightarrow> \<forall> x. x \<in> A \<longleftrightarrow> x \<in> B"
+ − 1717
apply (auto simp:mem_def)
+ − 1718
done
+ − 1719
+ − 1720
lemma set_cases2:
+ − 1721
"\<lbrakk>(A = {} \<Longrightarrow> R A); \<And> x. (A = {x}) \<Longrightarrow> R A; \<And> x y. \<lbrakk>x \<noteq> y; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> R A\<rbrakk> \<Longrightarrow> R A"
+ − 1722
apply (case_tac "A = {}", simp)
+ − 1723
by (case_tac "\<exists> x. A = {x}", clarsimp, blast)
+ − 1724
+ − 1725
lemma rhs_aux:"\<lbrakk>distinct_rhs rhs; {Y. \<exists>r. (Y, r) \<in> rhs} = {X}\<rbrakk> \<Longrightarrow> (\<exists>r. rhs = {(X, r)})"
+ − 1726
apply (rule_tac A = rhs in set_cases2, simp)
+ − 1727
apply (drule_tac x = X in eqset_imp_iff, clarsimp)
+ − 1728
apply (drule eqset_imp_iff',clarsimp)
+ − 1729
apply (frule_tac x = a in spec, drule_tac x = aa in spec)
+ − 1730
by (auto simp:distinct_rhs_def)
+ − 1731
+ − 1732
lemma every_eqcl_has_reg:
+ − 1733
assumes finite_CS: "finite (UNIV Quo Lang)"
+ − 1734
and X_in_CS: "X \<in> (UNIV Quo Lang)"
+ − 1735
shows "\<exists> (reg::rexp). L reg = X" (is "\<exists> r. ?E r")
+ − 1736
proof-
+ − 1737
have "\<exists>ES'. Inv X ES' \<and> TCon ES'" using finite_CS X_in_CS
+ − 1738
by (auto intro:init_ES_satisfy_Inv iteration_conc)
+ − 1739
then obtain ES' where Inv_ES': "Inv X ES'"
+ − 1740
and TCon_ES': "TCon ES'" by blast
+ − 1741
from Inv_ES' TCon_ES'
+ − 1742
have "\<exists> rhs. ES' = {(X, rhs)}"
+ − 1743
apply (clarsimp simp:Inv_def TCon_def)
+ − 1744
apply (rule_tac x = rhs in exI)
+ − 1745
by (auto dest!:card_Suc_Diff1 simp:card_eq_0_iff)
+ − 1746
then obtain rhs where ES'_single_equa: "ES' = {(X, rhs)}" ..
+ − 1747
hence X_ardenable: "ardenable (X, rhs)" using Inv_ES'
+ − 1748
by (simp add:Inv_def)
+ − 1749
+ − 1750
from X_ardenable have X_l_eq_r: "X = L rhs"
+ − 1751
by (simp add:ardenable_def)
+ − 1752
hence rhs_not_empty: "rhs \<noteq> {}" using Inv_ES' ES'_single_equa
+ − 1753
by (auto simp:Inv_def ardenable_def L_def)
+ − 1754
have rhs_eq_cls: "rhs_eq_cls rhs \<subseteq> {X, {[]}}"
+ − 1755
using Inv_ES' ES'_single_equa
+ − 1756
by (auto simp:Inv_def ardenable_def left_eq_cls_def)
+ − 1757
have X_not_empty: "X \<noteq> {}" using Inv_ES' ES'_single_equa
+ − 1758
by (auto simp:Inv_def)
+ − 1759
show ?thesis
+ − 1760
proof (cases "X = {[]}")
+ − 1761
case True
+ − 1762
hence "?E EMPTY" by (simp add:L_def)
+ − 1763
thus ?thesis by blast
+ − 1764
next
+ − 1765
case False with X_ardenable
+ − 1766
have "\<exists> rhs'. X = L rhs' \<and> rhs_eq_cls rhs' = rhs_eq_cls rhs - {X} \<and> distinct_rhs rhs'"
+ − 1767
by (drule_tac ardenable_prop, auto)
+ − 1768
then obtain rhs' where X_eq_rhs': "X = L rhs'"
+ − 1769
and rhs'_eq_cls: "rhs_eq_cls rhs' = rhs_eq_cls rhs - {X}"
+ − 1770
and rhs'_dist : "distinct_rhs rhs'" by blast
+ − 1771
have "rhs_eq_cls rhs' \<subseteq> {{[]}}" using rhs_eq_cls False rhs'_eq_cls rhs_not_empty
+ − 1772
by blast
+ − 1773
hence "rhs_eq_cls rhs' = {{[]}}" using X_not_empty X_eq_rhs'
+ − 1774
by (auto simp:L_def rhs_eq_cls_def)
+ − 1775
hence "\<exists> r. rhs' = {({[]}, r)}" using rhs'_dist
+ − 1776
by (auto intro:rhs_aux simp:rhs_eq_cls_def)
+ − 1777
then obtain r where "rhs' = {({[]}, r)}" ..
+ − 1778
hence "?E r" using X_eq_rhs' by (auto simp add:L_def lang_seq_def)
+ − 1779
thus ?thesis by blast
+ − 1780
qed
+ − 1781
qed
+ − 1782
+ − 1783
theorem myhill_nerode:
+ − 1784
assumes finite_CS: "finite (UNIV Quo Lang)"
+ − 1785
shows "\<exists> (reg::rexp). Lang = L reg" (is "\<exists> r. ?P r")
+ − 1786
proof -
+ − 1787
have has_r_each: "\<forall>C\<in>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists>(r::rexp). C = L r" using finite_CS
+ − 1788
by (auto dest:every_eqcl_has_reg)
+ − 1789
have "\<exists> (rS::rexp set). finite rS \<and>
+ − 1790
(\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> rS. C = L r) \<and>
+ − 1791
(\<forall> r \<in> rS. \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r)"
+ − 1792
(is "\<exists> rS. ?Q rS")
+ − 1793
proof-
+ − 1794
have "\<And> C. C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<Longrightarrow> C = L (SOME (ra::rexp). C = L ra)"
+ − 1795
using has_r_each
+ − 1796
apply (erule_tac x = C in ballE, erule_tac exE)
+ − 1797
by (rule_tac a = r in someI2, simp+)
+ − 1798
hence "?Q ((\<lambda> C. SOME r. C = L r) ` {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang})" using has_r_each
+ − 1799
using finite_CS by auto
+ − 1800
thus ?thesis by blast
+ − 1801
qed
+ − 1802
then obtain rS where finite_rS : "finite rS"
+ − 1803
and has_r_each': "\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> (rS::rexp set). C = L r"
+ − 1804
and has_cl_each: "\<forall> r \<in> (rS::rexp set). \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r" by blast
+ − 1805
have "?P (folds ALT NULL rS)"
+ − 1806
proof
+ − 1807
show "Lang \<subseteq> L (folds ALT NULL rS)" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_r_each'
+ − 1808
apply (clarsimp simp:fold_alt_null_eqs) by blast
+ − 1809
next
+ − 1810
show "L (folds ALT NULL rS) \<subseteq> Lang" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_cl_each
+ − 1811
by (clarsimp simp:fold_alt_null_eqs)
+ − 1812
qed
+ − 1813
thus ?thesis by blast
+ − 1814
qed
+ − 1815
+ − 1816
end
+ − 1817
+ − 1818
+ − 1819
+ − 1820
+ − 1821
+ − 1822
+ − 1823
+ − 1824
+ − 1825
+ − 1826