theory Matcher+ −
imports "Main" + −
begin+ −
+ −
+ −
section \<open>Regular Expressions\<close>+ −
+ −
datatype rexp =+ −
ZERO+ −
| ONE+ −
| CH char+ −
| SEQ rexp rexp+ −
| ALT rexp rexp+ −
| STAR rexp+ −
+ −
+ −
section \<open>Sequential Composition of Sets of Strings\<close>+ −
+ −
definition+ −
Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)+ −
where + −
"A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"+ −
+ −
text \<open>Two Simple Properties about Sequential Composition\<close>+ −
+ −
lemma seq_empty [simp]:+ −
shows "A ;; {[]} = A"+ −
and "{[]} ;; A = A"+ −
by (simp_all add: Seq_def)+ −
+ −
lemma seq_null [simp]:+ −
shows "A ;; {} = {}"+ −
and "{} ;; A = {}"+ −
by (simp_all add: Seq_def)+ −
+ −
section \<open>Kleene Star for Sets\<close>+ −
+ −
inductive_set+ −
Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)+ −
for A :: "string set"+ −
where+ −
start[intro]: "[] \<in> A\<star>"+ −
| step[intro]: "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"+ −
+ −
+ −
text \<open>A Standard Property of Star\<close>+ −
+ −
lemma star_cases:+ −
shows "A\<star> = {[]} \<union> A ;; A\<star>"+ −
unfolding Seq_def+ −
by (auto) (metis Star.simps)+ −
+ −
lemma star_decomp: + −
assumes a: "c # x \<in> A\<star>" + −
shows "\<exists>a b. x = a @ b \<and> c # a \<in> A \<and> b \<in> A\<star>"+ −
using a+ −
by (induct x\<equiv>"c # x" rule: Star.induct) + −
(auto simp add: append_eq_Cons_conv)+ −
+ −
+ −
section \<open>Meaning of Regular Expressions\<close>+ −
+ −
fun+ −
L :: "rexp \<Rightarrow> string set"+ −
where+ −
"L (ZERO) = {}"+ −
| "L (ONE) = {[]}"+ −
| "L (CH c) = {[c]}"+ −
| "L (SEQ r1 r2) = (L r1) ;; (L r2)"+ −
| "L (ALT r1 r2) = (L r1) \<union> (L r2)"+ −
| "L (STAR r) = (L r)\<star>"+ −
+ −
section \<open>The Matcher\<close>+ −
+ −
fun+ −
nullable :: "rexp \<Rightarrow> bool"+ −
where+ −
"nullable (ZERO) = False"+ −
| "nullable (ONE) = True"+ −
| "nullable (CH c) = False"+ −
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"+ −
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"+ −
| "nullable (STAR r) = True"+ −
+ −
+ −
section \<open>Correctness Proof for Nullable\<close>+ −
+ −
lemma nullable_correctness:+ −
shows "nullable r \<longleftrightarrow> [] \<in> (L r)"+ −
apply(induct r)+ −
(* ZERO case *)+ −
apply(simp only: nullable.simps)+ −
apply(simp only: L.simps)+ −
apply(simp)+ −
(* ONE case *)+ −
apply(simp only: nullable.simps)+ −
apply(simp only: L.simps)+ −
apply(simp)+ −
(* CHAR case *)+ −
apply(simp only: nullable.simps)+ −
apply(simp only: L.simps)+ −
apply(simp)+ −
prefer 2+ −
(* ALT case *)+ −
apply(simp (no_asm) only: nullable.simps)+ −
apply(simp only:)+ −
apply(simp only: L.simps)+ −
apply(simp)+ −
(* SEQ case *)+ −
oops+ −
+ −
lemma nullable_correctness:+ −
shows "nullable r \<longleftrightarrow> [] \<in> (L r)"+ −
apply(induct r)+ −
apply(simp_all)+ −
(* all easy subgoals are proved except the last 2 *)+ −
(* where the definition of Seq needs to be unfolded. *)+ −
oops+ −
+ −
lemma nullable_correctness:+ −
shows "nullable r \<longleftrightarrow> [] \<in> (L r)"+ −
apply(induct r)+ −
apply(simp_all add: Seq_def)+ −
(* except the star case every thing is proved *)+ −
(* we need to use the rule for Star.start *)+ −
oops+ −
+ −
lemma nullable_correctness:+ −
shows "nullable r \<longleftrightarrow> [] \<in> (L r)"+ −
apply(induct r)+ −
apply(simp_all add: Seq_def Star.start)+ −
done+ −
+ −
section \<open>Derivative Operation\<close>+ −
+ −
fun der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"+ −
where+ −
"der c (ZERO) = ZERO"+ −
| "der c (ONE) = ZERO"+ −
| "der c (CH d) = (if c = d then ONE else ZERO)"+ −
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"+ −
| "der c (SEQ r1 r2) = (if nullable r1 then ALT (SEQ (der c r1) r2) (der c r2)+ −
else SEQ (der c r1) r2)"+ −
| "der c (STAR r) = SEQ (der c r) (STAR r)"+ −
+ −
fun + −
ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"+ −
where+ −
"ders [] r = r"+ −
| "ders (c # s) r = ders s (der c r)"+ −
+ −
fun+ −
matcher :: "rexp \<Rightarrow> string \<Rightarrow> bool"+ −
where+ −
"matcher r s = nullable (ders s r)"+ −
+ −
definition+ −
Der :: "char \<Rightarrow> string set \<Rightarrow> string set"+ −
where+ −
"Der c A \<equiv> {s. [c] @ s \<in> A}"+ −
+ −
lemma Der_null [simp]:+ −
shows "Der c {} = {}"+ −
unfolding Der_def+ −
by auto+ −
+ −
lemma Der_empty [simp]:+ −
shows "Der c {[]} = {}"+ −
unfolding Der_def+ −
by auto+ −
+ −
lemma Der_char [simp]:+ −
shows "Der c {[d]} = (if c = d then {[]} else {})"+ −
unfolding Der_def+ −
by auto+ −
+ −
lemma Der_union [simp]:+ −
shows "Der c (A \<union> B) = Der c A \<union> Der c B"+ −
unfolding Der_def+ −
by auto+ −
+ −
lemma Der_insert_nil [simp]:+ −
shows "Der c (insert [] A) = Der c A"+ −
unfolding Der_def + −
by auto + −
+ −
lemma Der_seq [simp]:+ −
shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"+ −
unfolding Der_def Seq_def+ −
by (auto simp add: Cons_eq_append_conv)+ −
+ −
lemma Der_star [simp]:+ −
shows "Der c (A\<star>) = (Der c A) ;; A\<star>"+ −
proof - + −
have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"+ −
by (simp only: star_cases[symmetric])+ −
also have "... = Der c (A ;; A\<star>)"+ −
by (simp only: Der_union Der_empty) (simp)+ −
also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})"+ −
by simp+ −
also have "... = (Der c A) ;; A\<star>"+ −
unfolding Seq_def Der_def+ −
by (auto dest: star_decomp)+ −
finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .+ −
qed+ −
+ −
lemma der_correctness:+ −
shows "L (der c r) = Der c (L r)"+ −
apply(induct rule: der.induct) + −
apply(auto simp add: nullable_correctness)+ −
done+ −
+ −
+ −
lemma matcher_correctness:+ −
shows "matcher r s \<longleftrightarrow> s \<in> L r"+ −
by (induct s arbitrary: r)+ −
(simp_all add: nullable_correctness der_correctness Der_def)+ −
+ −
+ −
+ −
end + −