3
|
1 |
theory Matcher
|
|
2 |
imports "Main"
|
|
3 |
begin
|
|
4 |
|
|
5 |
section {* Sequential Composition of Sets *}
|
|
6 |
|
154
|
7 |
definition
|
|
8 |
Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
|
3
|
9 |
where
|
154
|
10 |
"A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
|
|
11 |
|
|
12 |
|
|
13 |
text {* Two Simple Properties about Sequential Composition *}
|
|
14 |
|
|
15 |
lemma seq_empty [simp]:
|
|
16 |
shows "A ;; {[]} = A"
|
|
17 |
and "{[]} ;; A = A"
|
|
18 |
by (simp_all add: Seq_def)
|
|
19 |
|
|
20 |
lemma seq_null [simp]:
|
|
21 |
shows "A ;; {} = {}"
|
|
22 |
and "{} ;; A = {}"
|
|
23 |
by (simp_all add: Seq_def)
|
3
|
24 |
|
|
25 |
|
|
26 |
section {* Kleene Star for Sets *}
|
|
27 |
|
|
28 |
inductive_set
|
|
29 |
Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
|
154
|
30 |
for A :: "string set"
|
3
|
31 |
where
|
154
|
32 |
start[intro]: "[] \<in> A\<star>"
|
|
33 |
| step[intro]: "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"
|
3
|
34 |
|
|
35 |
|
154
|
36 |
text {* A Standard Property of Star *}
|
|
37 |
|
|
38 |
lemma star_cases:
|
|
39 |
shows "A\<star> = {[]} \<union> A ;; A\<star>"
|
|
40 |
unfolding Seq_def
|
|
41 |
by (auto) (metis Star.simps)
|
|
42 |
|
|
43 |
lemma star_decomp:
|
|
44 |
assumes a: "c # x \<in> A\<star>"
|
|
45 |
shows "\<exists>a b. x = a @ b \<and> c # a \<in> A \<and> b \<in> A\<star>"
|
|
46 |
using a
|
|
47 |
by (induct x\<equiv>"c # x" rule: Star.induct)
|
|
48 |
(auto simp add: append_eq_Cons_conv)
|
|
49 |
|
|
50 |
section {* Left-Quotient of a Set *}
|
|
51 |
|
|
52 |
definition
|
|
53 |
Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
|
|
54 |
where
|
|
55 |
"Der c A \<equiv> {s. [c] @ s \<in> A}"
|
|
56 |
|
|
57 |
lemma Der_null [simp]:
|
|
58 |
shows "Der c {} = {}"
|
|
59 |
unfolding Der_def
|
|
60 |
by auto
|
|
61 |
|
|
62 |
lemma Der_empty [simp]:
|
|
63 |
shows "Der c {[]} = {}"
|
|
64 |
unfolding Der_def
|
|
65 |
by auto
|
3
|
66 |
|
154
|
67 |
lemma Der_char [simp]:
|
|
68 |
shows "Der c {[d]} = (if c = d then {[]} else {})"
|
|
69 |
unfolding Der_def
|
|
70 |
by auto
|
|
71 |
|
|
72 |
lemma Der_union [simp]:
|
|
73 |
shows "Der c (A \<union> B) = Der c A \<union> Der c B"
|
|
74 |
unfolding Der_def
|
|
75 |
by auto
|
|
76 |
|
|
77 |
lemma Der_seq [simp]:
|
|
78 |
shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"
|
|
79 |
unfolding Der_def Seq_def
|
|
80 |
by (auto simp add: Cons_eq_append_conv)
|
|
81 |
|
|
82 |
lemma Der_star [simp]:
|
|
83 |
shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
|
|
84 |
proof -
|
|
85 |
have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"
|
|
86 |
by (simp only: star_cases[symmetric])
|
|
87 |
also have "... = Der c (A ;; A\<star>)"
|
|
88 |
by (simp only: Der_union Der_empty) (simp)
|
|
89 |
also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})"
|
|
90 |
by simp
|
|
91 |
also have "... = (Der c A) ;; A\<star>"
|
|
92 |
unfolding Seq_def Der_def
|
|
93 |
by (auto dest: star_decomp)
|
|
94 |
finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
|
|
95 |
qed
|
|
96 |
|
3
|
97 |
|
|
98 |
section {* Regular Expressions *}
|
|
99 |
|
|
100 |
datatype rexp =
|
|
101 |
NULL
|
|
102 |
| EMPTY
|
|
103 |
| CHAR char
|
|
104 |
| SEQ rexp rexp
|
|
105 |
| ALT rexp rexp
|
|
106 |
| STAR rexp
|
|
107 |
|
102
|
108 |
|
3
|
109 |
section {* Semantics of Regular Expressions *}
|
|
110 |
|
|
111 |
fun
|
|
112 |
L :: "rexp \<Rightarrow> string set"
|
|
113 |
where
|
|
114 |
"L (NULL) = {}"
|
|
115 |
| "L (EMPTY) = {[]}"
|
|
116 |
| "L (CHAR c) = {[c]}"
|
154
|
117 |
| "L (SEQ r1 r2) = (L r1) ;; (L r2)"
|
3
|
118 |
| "L (ALT r1 r2) = (L r1) \<union> (L r2)"
|
|
119 |
| "L (STAR r) = (L r)\<star>"
|
|
120 |
|
|
121 |
section {* The Matcher *}
|
|
122 |
|
|
123 |
fun
|
|
124 |
nullable :: "rexp \<Rightarrow> bool"
|
|
125 |
where
|
|
126 |
"nullable (NULL) = False"
|
|
127 |
| "nullable (EMPTY) = True"
|
|
128 |
| "nullable (CHAR c) = False"
|
|
129 |
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
|
|
130 |
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
|
|
131 |
| "nullable (STAR r) = True"
|
|
132 |
|
|
133 |
fun
|
|
134 |
der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
|
|
135 |
where
|
|
136 |
"der c (NULL) = NULL"
|
|
137 |
| "der c (EMPTY) = NULL"
|
5
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
138 |
| "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
|
3
|
139 |
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
|
|
140 |
| "der c (SEQ r1 r2) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)"
|
|
141 |
| "der c (STAR r) = SEQ (der c r) (STAR r)"
|
|
142 |
|
|
143 |
fun
|
|
144 |
derivative :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
|
|
145 |
where
|
|
146 |
"derivative [] r = r"
|
5
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
147 |
| "derivative (c # s) r = derivative s (der c r)"
|
3
|
148 |
|
|
149 |
fun
|
154
|
150 |
matcher :: "rexp \<Rightarrow> string \<Rightarrow> bool"
|
3
|
151 |
where
|
154
|
152 |
"matcher r s = nullable (derivative s r)"
|
3
|
153 |
|
|
154 |
|
|
155 |
section {* Correctness Proof of the Matcher *}
|
|
156 |
|
|
157 |
lemma nullable_correctness:
|
154
|
158 |
shows "nullable r \<longleftrightarrow> [] \<in> (L r)"
|
|
159 |
by (induct r) (auto simp add: Seq_def)
|
|
160 |
|
3
|
161 |
|
|
162 |
lemma der_correctness:
|
154
|
163 |
shows "L (der c r) = Der c (L r)"
|
|
164 |
by (induct r)
|
|
165 |
(simp_all add: nullable_correctness)
|
3
|
166 |
|
154
|
167 |
lemma matcher_correctness:
|
|
168 |
shows "matcher r s \<longleftrightarrow> s \<in> L r"
|
3
|
169 |
by (induct s arbitrary: r)
|
154
|
170 |
(simp_all add: nullable_correctness der_correctness Der_def)
|
3
|
171 |
|
|
172 |
section {* Examples *}
|
|
173 |
|
154
|
174 |
value "matcher NULL []"
|
|
175 |
value "matcher (CHAR (CHR ''a'')) [CHR ''a'']"
|
|
176 |
value "matcher (CHAR a) [a,a]"
|
|
177 |
value "matcher (STAR (CHAR a)) []"
|
|
178 |
value "matcher (STAR (CHAR a)) [a,a]"
|
|
179 |
value "matcher (SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ''abbbbc''"
|
|
180 |
value "matcher (SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ''abbcbbc''"
|
3
|
181 |
|
|
182 |
section {* Incorrect Matcher - fun-definition rejected *}
|
|
183 |
|
154
|
184 |
fun
|
3
|
185 |
match :: "rexp list \<Rightarrow> string \<Rightarrow> bool"
|
|
186 |
where
|
|
187 |
"match [] [] = True"
|
5
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
188 |
| "match [] (c # s) = False"
|
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
189 |
| "match (NULL # rs) s = False"
|
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
190 |
| "match (EMPTY # rs) s = match rs s"
|
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
191 |
| "match (CHAR c # rs) [] = False"
|
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
192 |
| "match (CHAR c # rs) (d # s) = (if c = d then match rs s else False)"
|
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
193 |
| "match (ALT r1 r2 # rs) s = (match (r1 # rs) s \<or> match (r2 # rs) s)"
|
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
194 |
| "match (SEQ r1 r2 # rs) s = match (r1 # r2 # rs) s"
|
074d9a4b2bc9
added a file about the easy closure properties of regular sets (the difficult parts, like complement, are missing)
urbanc
diff
changeset
|
195 |
| "match (STAR r # rs) s = (match rs s \<or> match (r # (STAR r) # rs) s)"
|
154
|
196 |
|
3
|
197 |
|
|
198 |
end |