--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/progs/Matcher2.thy Thu Nov 14 20:10:06 2013 +0000
@@ -0,0 +1,283 @@
+theory Matcher2
+ imports "Main"
+begin
+
+section {* Regular Expressions *}
+
+datatype rexp =
+ NULL
+| EMPTY
+| CHAR char
+| SEQ rexp rexp
+| ALT rexp rexp
+| STAR rexp
+| NOT rexp
+| PLUS rexp
+| OPT rexp
+| NTIMES rexp nat
+| NMTIMES rexp nat nat
+
+fun M :: "rexp \<Rightarrow> nat"
+where
+ "M (NULL) = 0"
+| "M (EMPTY) = 0"
+| "M (CHAR char) = 0"
+| "M (SEQ r1 r2) = Suc ((M r1) + (M r2))"
+| "M (ALT r1 r2) = Suc ((M r1) + (M r2))"
+| "M (STAR r) = Suc (M r)"
+| "M (NOT r) = Suc (M r)"
+| "M (PLUS r) = Suc (M r)"
+| "M (OPT r) = Suc (M r)"
+| "M (NTIMES r n) = Suc (M r) * 2 * (Suc n)"
+| "M (NMTIMES r n m) = Suc (M r) * 2 * (Suc n + Suc m)"
+
+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)
+
+lemma seq_union:
+ shows "A ;; (B \<union> C) = A ;; B \<union> A ;; C"
+by (auto simp add: Seq_def)
+
+lemma seq_Union:
+ shows "A ;; (\<Union>x\<in>B. C x) = (\<Union>x\<in>B. A ;; C x)"
+by (auto simp add: Seq_def)
+
+lemma seq_empty_in [simp]:
+ "[] \<in> A ;; B \<longleftrightarrow> ([] \<in> A \<and> [] \<in> B)"
+by (simp 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 {* Power for Sets *}
+
+fun
+ pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _" [101, 102] 101)
+where
+ "A \<up> 0 = {[]}"
+| "A \<up> (Suc n) = A ;; (A \<up> n)"
+
+
+lemma pow_empty [simp]:
+ shows "[] \<in> A \<up> n \<longleftrightarrow> (n = 0 \<or> [] \<in> A)"
+by (induct n) (auto)
+
+
+section {* Semantics of Regular Expressions *}
+
+fun
+ L :: "rexp \<Rightarrow> string set"
+where
+ "L (NULL) = {}"
+| "L (EMPTY) = {[]}"
+| "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>"
+| "L (NOT r) = UNIV - (L r)"
+| "L (PLUS r) = (L r) ;; ((L r)\<star>)"
+| "L (OPT r) = (L r) \<union> {[]}"
+| "L (NTIMES r n) = (L r) \<up> n"
+| "L (NMTIMES r n m) = (\<Union>i\<in> {n..n+m} . ((L r) \<up> i))"
+
+
+section {* The Matcher *}
+
+fun
+ nullable :: "rexp \<Rightarrow> bool"
+where
+ "nullable (NULL) = False"
+| "nullable (EMPTY) = 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"
+| "nullable (NOT r) = (\<not>(nullable r))"
+| "nullable (PLUS r) = (nullable r)"
+| "nullable (OPT r) = True"
+| "nullable (NTIMES r n) = (if n = 0 then True else nullable r)"
+| "nullable (NMTIMES r n m) = (if n = 0 then True else nullable r)"
+
+function
+ der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
+where
+ "der c (NULL) = NULL"
+| "der c (EMPTY) = NULL"
+| "der c (CHAR d) = (if c = d then EMPTY else NULL)"
+| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
+| "der c (SEQ r1 r2) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)"
+| "der c (STAR r) = SEQ (der c r) (STAR r)"
+| "der c (NOT r) = NOT(der c r)"
+| "der c (PLUS r) = SEQ (der c r) (STAR r)"
+| "der c (OPT r) = der c r"
+| "der c (NTIMES r 0) = NULL"
+| "der c (NTIMES r (Suc n)) = der c (SEQ r (NTIMES r n))"
+| "der c (NMTIMES r 0 0) = NULL"
+| "der c (NMTIMES r 0 (Suc m)) = ALT (der c (NTIMES r (Suc m))) (der c (NMTIMES r 0 m))"
+| "der c (NMTIMES r (Suc n) m) = der c (SEQ r (NMTIMES r n m))"
+by pat_completeness auto
+
+termination der
+apply(relation "measure (\<lambda>(c, r). M r)")
+apply(simp_all)
+done
+
+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)"
+
+
+section {* Correctness Proof of the Matcher *}
+
+lemma nullable_correctness:
+ shows "nullable r \<longleftrightarrow> [] \<in> (L r)"
+by(induct r) (auto simp add: Seq_def)
+
+
+section {* Left-Quotient of a Set *}
+
+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_UNIV [simp]:
+ "Der c (UNIV - A) = UNIV - Der c A"
+unfolding Der_def
+by (auto)
+
+lemma Der_pow [simp]:
+ shows "Der c (A \<up> (Suc n)) = (Der c A) ;; (A \<up> n) \<union> (if [] \<in> A then Der c (A \<up> n) else {})"
+unfolding Der_def
+by(auto simp add: Cons_eq_append_conv Seq_def)
+
+
+lemma Der_UNION [simp]:
+ shows "Der c (\<Union>x\<in>A. B x) = (\<Union>x\<in>A. Der c (B x))"
+by (auto simp add: Der_def)
+
+lemma test:
+ "(\<Union> x\<le>Suc m. B x) = (B (Suc m) \<union> (\<Union> x\<le>m. B x))"
+by (metis UN_insert atMost_Suc)
+
+lemma yy:
+ "(\<Union>x\<in>{Suc n..Suc m}. B x) = (\<Union>x\<in>{n..m}. B (Suc x))"
+by (metis UN_extend_simps(10) image_Suc_atLeastAtMost)
+
+lemma uu:
+ "(Suc n) + m = Suc (n + m)"
+by simp
+
+lemma der_correctness:
+ shows "L (der c r) = Der c (L r)"
+apply(induct rule: der.induct)
+apply(simp_all add: nullable_correctness)[12]
+apply(simp only: L.simps der.simps)
+apply(simp only: Der_UNION)
+apply(simp del: pow.simps Der_pow)
+apply(simp only: atLeast0AtMost)
+apply(simp only: test)
+apply(simp only: L.simps der.simps)
+apply(simp only: Der_UNION)
+apply(simp only: yy add_Suc)
+apply(simp only: seq_Union)
+apply(simp only: Der_UNION)
+apply(simp only: pow.simps)
+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
\ No newline at end of file