diff -r 0b63c0edfb09 -r ff6665581ced progs/Matcher2.thy --- /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 \ 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 \ string set \ string set" ("_ ;; _" [100,100] 100) +where + "A ;; B = {s1 @ s2 | s1 s2. s1 \ A \ s2 \ 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 \ C) = A ;; B \ A ;; C" +by (auto simp add: Seq_def) + +lemma seq_Union: + shows "A ;; (\x\B. C x) = (\x\B. A ;; C x)" +by (auto simp add: Seq_def) + +lemma seq_empty_in [simp]: + "[] \ A ;; B \ ([] \ A \ [] \ B)" +by (simp add: Seq_def) + +section {* Kleene Star for Sets *} + +inductive_set + Star :: "string set \ string set" ("_\" [101] 102) + for A :: "string set" +where + start[intro]: "[] \ A\" +| step[intro]: "\s1 \ A; s2 \ A\\ \ s1 @ s2 \ A\" + + +text {* A Standard Property of Star *} + +lemma star_cases: + shows "A\ = {[]} \ A ;; A\" +unfolding Seq_def +by (auto) (metis Star.simps) + +lemma star_decomp: + assumes a: "c # x \ A\" + shows "\a b. x = a @ b \ c # a \ A \ b \ A\" +using a +by (induct x\"c # x" rule: Star.induct) + (auto simp add: append_eq_Cons_conv) + +section {* Power for Sets *} + +fun + pow :: "string set \ nat \ string set" ("_ \ _" [101, 102] 101) +where + "A \ 0 = {[]}" +| "A \ (Suc n) = A ;; (A \ n)" + + +lemma pow_empty [simp]: + shows "[] \ A \ n \ (n = 0 \ [] \ A)" +by (induct n) (auto) + + +section {* Semantics of Regular Expressions *} + +fun + L :: "rexp \ 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) \ (L r2)" +| "L (STAR r) = (L r)\" +| "L (NOT r) = UNIV - (L r)" +| "L (PLUS r) = (L r) ;; ((L r)\)" +| "L (OPT r) = (L r) \ {[]}" +| "L (NTIMES r n) = (L r) \ n" +| "L (NMTIMES r n m) = (\i\ {n..n+m} . ((L r) \ i))" + + +section {* The Matcher *} + +fun + nullable :: "rexp \ bool" +where + "nullable (NULL) = False" +| "nullable (EMPTY) = True" +| "nullable (CHAR c) = False" +| "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (STAR r) = True" +| "nullable (NOT r) = (\(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 \ rexp \ 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 (\(c, r). M r)") +apply(simp_all) +done + +fun + ders :: "string \ rexp \ rexp" +where + "ders [] r = r" +| "ders (c # s) r = ders s (der c r)" + +fun + matcher :: "rexp \ string \ bool" +where + "matcher r s = nullable (ders s r)" + + +section {* Correctness Proof of the Matcher *} + +lemma nullable_correctness: + shows "nullable r \ [] \ (L r)" +by(induct r) (auto simp add: Seq_def) + + +section {* Left-Quotient of a Set *} + +definition + Der :: "char \ string set \ string set" +where + "Der c A \ {s. [c] @ s \ 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 \ B) = Der c A \ 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 \ (if [] \ 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\) = (Der c A) ;; A\" +proof - + have "Der c (A\) = Der c ({[]} \ A ;; A\)" + by (simp only: star_cases[symmetric]) + also have "... = Der c (A ;; A\)" + by (simp only: Der_union Der_empty) (simp) + also have "... = (Der c A) ;; A\ \ (if [] \ A then Der c (A\) else {})" + by simp + also have "... = (Der c A) ;; A\" + unfolding Seq_def Der_def + by (auto dest: star_decomp) + finally show "Der c (A\) = (Der c A) ;; A\" . +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 \ (Suc n)) = (Der c A) ;; (A \ n) \ (if [] \ A then Der c (A \ n) else {})" +unfolding Der_def +by(auto simp add: Cons_eq_append_conv Seq_def) + + +lemma Der_UNION [simp]: + shows "Der c (\x\A. B x) = (\x\A. Der c (B x))" +by (auto simp add: Der_def) + +lemma test: + "(\ x\Suc m. B x) = (B (Suc m) \ (\ x\m. B x))" +by (metis UN_insert atMost_Suc) + +lemma yy: + "(\x\{Suc n..Suc m}. B x) = (\x\{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 \ s \ L r" +by (induct s arbitrary: r) + (simp_all add: nullable_correctness der_correctness Der_def) + + +end \ No newline at end of file