author | Christian Urban <christian dot urban at kcl dot ac dot uk> |
Fri, 05 Jul 2013 17:19:17 +0100 | |
changeset 379 | 8c4b6fb43ebe |
parent 253 | bcef7669f55a |
permissions | -rw-r--r-- |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
1 |
(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
2 |
theory Closures |
203 | 3 |
imports Myhill "~~/src/HOL/Library/Infinite_Set" |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
4 |
begin |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
5 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
6 |
section {* Closure properties of regular languages *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
7 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
8 |
abbreviation |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
9 |
regular :: "'a lang \<Rightarrow> bool" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
10 |
where |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
11 |
"regular A \<equiv> \<exists>r. A = lang r" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
12 |
|
203 | 13 |
subsection {* Closure under @{text "\<union>"}, @{text "\<cdot>"} and @{text "\<star>"} *} |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
14 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
15 |
lemma closure_union [intro]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
16 |
assumes "regular A" "regular B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
17 |
shows "regular (A \<union> B)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
18 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
19 |
from assms obtain r1 r2::"'a rexp" where "lang r1 = A" "lang r2 = B" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
20 |
then have "A \<union> B = lang (Plus r1 r2)" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
21 |
then show "regular (A \<union> B)" by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
22 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
23 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
24 |
lemma closure_seq [intro]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
25 |
assumes "regular A" "regular B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
26 |
shows "regular (A \<cdot> B)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
27 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
28 |
from assms obtain r1 r2::"'a rexp" where "lang r1 = A" "lang r2 = B" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
29 |
then have "A \<cdot> B = lang (Times r1 r2)" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
30 |
then show "regular (A \<cdot> B)" by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
31 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
32 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
33 |
lemma closure_star [intro]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
34 |
assumes "regular A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
35 |
shows "regular (A\<star>)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
36 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
37 |
from assms obtain r::"'a rexp" where "lang r = A" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
38 |
then have "A\<star> = lang (Star r)" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
39 |
then show "regular (A\<star>)" by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
40 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
41 |
|
203 | 42 |
subsection {* Closure under complementation *} |
43 |
||
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
44 |
text {* Closure under complementation is proved via the |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
45 |
Myhill-Nerode theorem *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
46 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
47 |
lemma closure_complement [intro]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
48 |
fixes A::"('a::finite) lang" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
49 |
assumes "regular A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
50 |
shows "regular (- A)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
51 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
52 |
from assms have "finite (UNIV // \<approx>A)" by (simp add: Myhill_Nerode) |
181 | 53 |
then have "finite (UNIV // \<approx>(-A))" by (simp add: str_eq_def) |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
54 |
then show "regular (- A)" by (simp add: Myhill_Nerode) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
55 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
56 |
|
203 | 57 |
subsection {* Closure under @{text "-"} and @{text "\<inter>"} *} |
58 |
||
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
59 |
lemma closure_difference [intro]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
60 |
fixes A::"('a::finite) lang" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
61 |
assumes "regular A" "regular B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
62 |
shows "regular (A - B)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
63 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
64 |
have "A - B = - (- A \<union> B)" by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
65 |
moreover |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
66 |
have "regular (- (- A \<union> B))" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
67 |
using assms by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
68 |
ultimately show "regular (A - B)" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
69 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
70 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
71 |
lemma closure_intersection [intro]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
72 |
fixes A::"('a::finite) lang" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
73 |
assumes "regular A" "regular B" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
74 |
shows "regular (A \<inter> B)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
75 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
76 |
have "A \<inter> B = - (- A \<union> - B)" by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
77 |
moreover |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
78 |
have "regular (- (- A \<union> - B))" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
79 |
using assms by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
80 |
ultimately show "regular (A \<inter> B)" by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
81 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
82 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
83 |
subsection {* Closure under string reversal *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
84 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
85 |
fun |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
86 |
Rev :: "'a rexp \<Rightarrow> 'a rexp" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
87 |
where |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
88 |
"Rev Zero = Zero" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
89 |
| "Rev One = One" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
90 |
| "Rev (Atom c) = Atom c" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
91 |
| "Rev (Plus r1 r2) = Plus (Rev r1) (Rev r2)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
92 |
| "Rev (Times r1 r2) = Times (Rev r2) (Rev r1)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
93 |
| "Rev (Star r) = Star (Rev r)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
94 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
95 |
lemma rev_seq[simp]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
96 |
shows "rev ` (B \<cdot> A) = (rev ` A) \<cdot> (rev ` B)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
97 |
unfolding conc_def image_def |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
98 |
by (auto) (metis rev_append)+ |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
99 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
100 |
lemma rev_star1: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
101 |
assumes a: "s \<in> (rev ` A)\<star>" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
102 |
shows "s \<in> rev ` (A\<star>)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
103 |
using a |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
104 |
proof(induct rule: star_induct) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
105 |
case (append s1 s2) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
106 |
have inj: "inj (rev::'a list \<Rightarrow> 'a list)" unfolding inj_on_def by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
107 |
have "s1 \<in> rev ` A" "s2 \<in> rev ` (A\<star>)" by fact+ |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
108 |
then obtain x1 x2 where "x1 \<in> A" "x2 \<in> A\<star>" and eqs: "s1 = rev x1" "s2 = rev x2" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
109 |
then have "x1 \<in> A\<star>" "x2 \<in> A\<star>" by (auto) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
110 |
then have "x2 @ x1 \<in> A\<star>" by (auto) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
111 |
then have "rev (x2 @ x1) \<in> rev ` A\<star>" using inj by (simp only: inj_image_mem_iff) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
112 |
then show "s1 @ s2 \<in> rev ` A\<star>" using eqs by simp |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
113 |
qed (auto) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
114 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
115 |
lemma rev_star2: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
116 |
assumes a: "s \<in> A\<star>" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
117 |
shows "rev s \<in> (rev ` A)\<star>" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
118 |
using a |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
119 |
proof(induct rule: star_induct) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
120 |
case (append s1 s2) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
121 |
have inj: "inj (rev::'a list \<Rightarrow> 'a list)" unfolding inj_on_def by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
122 |
have "s1 \<in> A"by fact |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
123 |
then have "rev s1 \<in> rev ` A" using inj by (simp only: inj_image_mem_iff) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
124 |
then have "rev s1 \<in> (rev ` A)\<star>" by (auto) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
125 |
moreover |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
126 |
have "rev s2 \<in> (rev ` A)\<star>" by fact |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
127 |
ultimately show "rev (s1 @ s2) \<in> (rev ` A)\<star>" by (auto) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
128 |
qed (auto) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
129 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
130 |
lemma rev_star [simp]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
131 |
shows " rev ` (A\<star>) = (rev ` A)\<star>" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
132 |
using rev_star1 rev_star2 by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
133 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
134 |
lemma rev_lang: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
135 |
shows "rev ` (lang r) = lang (Rev r)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
136 |
by (induct r) (simp_all add: image_Un) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
137 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
138 |
lemma closure_reversal [intro]: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
139 |
assumes "regular A" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
140 |
shows "regular (rev ` A)" |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
141 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
142 |
from assms obtain r::"'a rexp" where "A = lang r" by auto |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
143 |
then have "lang (Rev r) = rev ` A" by (simp add: rev_lang) |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
144 |
then show "regular (rev` A)" by blast |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
145 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
146 |
|
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
147 |
subsection {* Closure under left-quotients *} |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
148 |
|
187 | 149 |
abbreviation |
203 | 150 |
"Deriv_lang A B \<equiv> \<Union>x \<in> A. Derivs x B" |
187 | 151 |
|
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
152 |
lemma closure_left_quotient: |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
153 |
assumes "regular A" |
203 | 154 |
shows "regular (Deriv_lang B A)" |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
155 |
proof - |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
156 |
from assms obtain r::"'a rexp" where eq: "lang r = A" by auto |
203 | 157 |
have fin: "finite (pderivs_lang B r)" by (rule finite_pderivs_lang) |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
158 |
|
379
8c4b6fb43ebe
polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
253
diff
changeset
|
159 |
have "Deriv_lang B (lang r) = (\<Union> (lang ` (pderivs_lang B r)))" |
203 | 160 |
by (simp add: Derivs_pderivs pderivs_lang_def) |
161 |
also have "\<dots> = lang (\<Uplus>(pderivs_lang B r))" using fin by simp |
|
162 |
finally have "Deriv_lang B A = lang (\<Uplus>(pderivs_lang B r))" using eq |
|
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
163 |
by simp |
203 | 164 |
then show "regular (Deriv_lang B A)" by auto |
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
165 |
qed |
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
166 |
|
203 | 167 |
subsection {* Finite and co-finite sets are regular *} |
200
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
168 |
|
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
169 |
lemma singleton_regular: |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
170 |
shows "regular {s}" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
171 |
proof (induct s) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
172 |
case Nil |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
173 |
have "{[]} = lang (One)" by simp |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
174 |
then show "regular {[]}" by blast |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
175 |
next |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
176 |
case (Cons c s) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
177 |
have "regular {s}" by fact |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
178 |
then obtain r where "{s} = lang r" by blast |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
179 |
then have "{c # s} = lang (Times (Atom c) r)" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
180 |
by (auto simp add: conc_def) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
181 |
then show "regular {c # s}" by blast |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
182 |
qed |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
183 |
|
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
184 |
lemma finite_regular: |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
185 |
assumes "finite A" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
186 |
shows "regular A" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
187 |
using assms |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
188 |
proof (induct) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
189 |
case empty |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
190 |
have "{} = lang (Zero)" by simp |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
191 |
then show "regular {}" by blast |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
192 |
next |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
193 |
case (insert s A) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
194 |
have "regular {s}" by (simp add: singleton_regular) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
195 |
moreover |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
196 |
have "regular A" by fact |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
197 |
ultimately have "regular ({s} \<union> A)" by (rule closure_union) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
198 |
then show "regular (insert s A)" by simp |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
199 |
qed |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
200 |
|
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
201 |
lemma cofinite_regular: |
226 | 202 |
fixes A::"'a::finite lang" |
200
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
203 |
assumes "finite (- A)" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
204 |
shows "regular A" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
205 |
proof - |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
206 |
from assms have "regular (- A)" by (simp add: finite_regular) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
207 |
then have "regular (-(- A))" by (rule closure_complement) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
208 |
then show "regular A" by simp |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
209 |
qed |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
210 |
|
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
211 |
|
241 | 212 |
subsection {* Continuation lemma for showing non-regularity of languages *} |
200
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
213 |
|
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
214 |
lemma continuation_lemma: |
224 | 215 |
fixes A B::"'a::finite lang" |
200
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
216 |
assumes reg: "regular A" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
217 |
and inf: "infinite B" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
218 |
shows "\<exists>x \<in> B. \<exists>y \<in> B. x \<noteq> y \<and> x \<approx>A y" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
219 |
proof - |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
220 |
def eqfun \<equiv> "\<lambda>A x::('a::finite list). (\<approx>A) `` {x}" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
221 |
have "finite (UNIV // \<approx>A)" using reg by (simp add: Myhill_Nerode) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
222 |
moreover |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
223 |
have "(eqfun A) ` B \<subseteq> UNIV // (\<approx>A)" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
224 |
unfolding eqfun_def quotient_def by auto |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
225 |
ultimately have "finite ((eqfun A) ` B)" by (rule rev_finite_subset) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
226 |
with inf have "\<exists>a \<in> B. infinite {b \<in> B. eqfun A b = eqfun A a}" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
227 |
by (rule pigeonhole_infinite) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
228 |
then obtain a where in_a: "a \<in> B" and "infinite {b \<in> B. eqfun A b = eqfun A a}" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
229 |
by blast |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
230 |
moreover |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
231 |
have "{b \<in> B. eqfun A b = eqfun A a} = {b \<in> B. b \<approx>A a}" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
232 |
unfolding eqfun_def Image_def str_eq_def by auto |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
233 |
ultimately have "infinite {b \<in> B. b \<approx>A a}" by simp |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
234 |
then have "infinite ({b \<in> B. b \<approx>A a} - {a})" by simp |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
235 |
moreover |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
236 |
have "{b \<in> B. b \<approx>A a} - {a} = {b \<in> B. b \<approx>A a \<and> b \<noteq> a}" by auto |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
237 |
ultimately have "infinite {b \<in> B. b \<approx>A a \<and> b \<noteq> a}" by simp |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
238 |
then have "{b \<in> B. b \<approx>A a \<and> b \<noteq> a} \<noteq> {}" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
239 |
by (metis finite.emptyI) |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
240 |
then obtain b where "b \<in> B" "b \<noteq> a" "b \<approx>A a" by blast |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
241 |
with in_a show "\<exists>x \<in> B. \<exists>y \<in> B. x \<noteq> y \<and> x \<approx>A y" |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
242 |
by blast |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
243 |
qed |
204856ef5573
added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents:
193
diff
changeset
|
244 |
|
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
245 |
|
241 | 246 |
subsection {* The language @{text "a\<^sup>n b\<^sup>n"} is not regular *} |
204 | 247 |
|
224 | 248 |
abbreviation |
249 |
replicate_rev ("_ ^^^ _" [100, 100] 100) |
|
250 |
where |
|
251 |
"a ^^^ n \<equiv> replicate n a" |
|
204 | 252 |
|
224 | 253 |
lemma an_bn_not_regular: |
254 |
shows "\<not> regular (\<Union>n. {CHR ''a'' ^^^ n @ CHR ''b'' ^^^ n})" |
|
255 |
proof |
|
256 |
def A\<equiv>"\<Union>n. {CHR ''a'' ^^^ n @ CHR ''b'' ^^^ n}" |
|
257 |
def B\<equiv>"\<Union>n. {CHR ''a'' ^^^ n}" |
|
258 |
assume as: "regular A" |
|
259 |
def B\<equiv>"\<Union>n. {CHR ''a'' ^^^ n}" |
|
252 | 260 |
|
253 | 261 |
have sameness: "\<And>i j. CHR ''a'' ^^^ i @ CHR ''b'' ^^^ j \<in> A \<longleftrightarrow> i = j" |
262 |
unfolding A_def |
|
263 |
apply auto |
|
264 |
apply(drule_tac f="\<lambda>s. length (filter (op= (CHR ''a'')) s) = length (filter (op= (CHR ''b'')) s)" |
|
265 |
in arg_cong) |
|
266 |
apply(simp) |
|
267 |
done |
|
252 | 268 |
|
224 | 269 |
have b: "infinite B" |
270 |
unfolding infinite_iff_countable_subset |
|
226 | 271 |
unfolding inj_on_def B_def |
272 |
by (rule_tac x="\<lambda>n. CHR ''a'' ^^^ n" in exI) (auto) |
|
224 | 273 |
moreover |
252 | 274 |
have "\<forall>x \<in> B. \<forall>y \<in> B. x \<noteq> y \<longrightarrow> \<not> (x \<approx>A y)" |
275 |
apply(auto) |
|
276 |
unfolding B_def |
|
224 | 277 |
apply(auto) |
252 | 278 |
apply(simp add: str_eq_def) |
279 |
apply(drule_tac x="CHR ''b'' ^^^ n" in spec) |
|
253 | 280 |
apply(simp add: sameness) |
224 | 281 |
done |
282 |
ultimately |
|
283 |
show "False" using continuation_lemma[OF as] by blast |
|
284 |
qed |
|
204 | 285 |
|
286 |
||
170
b1258b7d2789
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff
changeset
|
287 |
end |