theory Matcher
imports "Main"
begin
section {* Regular Expressions *}
datatype rexp =
ZERO
| ONE
| CHAR char
| SEQ rexp rexp
| ALT rexp rexp
| STAR rexp
section {* Sequential Composition of Sets *}
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 {* Two Simple Properties about Sequential Composition *}
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 {* Kleene Star for Sets *}
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 {* A Standard Property of Star *}
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 {* Semantics of Regular Expressions *}
fun
L :: "rexp \<Rightarrow> string set"
where
"L (ZERO) = {}"
| "L (ONE) = {[]}"
| "L (CHAR 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 {* The Matcher *}
fun
nullable :: "rexp \<Rightarrow> bool"
where
"nullable (ZERO) = False"
| "nullable (ONE) = True"
| "nullable (CHAR 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 {* Correctness Proof for Nullable *}
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
end