diff -r 3b2f76950473 -r 5fcad75ade92 progs/Matcher.thy --- a/progs/Matcher.thy Sun Oct 02 08:42:01 2022 +0100 +++ b/progs/Matcher.thy Sun Oct 09 13:39:34 2022 +0100 @@ -3,25 +3,25 @@ begin -section {* Regular Expressions *} +section \Regular Expressions\ datatype rexp = ZERO | ONE -| CHAR char +| CH char | SEQ rexp rexp | ALT rexp rexp | STAR rexp -section {* Sequential Composition of Sets *} +section \Sequential Composition of Sets of Strings\ 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 *} +text \Two Simple Properties about Sequential Composition\ lemma seq_empty [simp]: shows "A ;; {[]} = A" @@ -33,7 +33,7 @@ and "{} ;; A = {}" by (simp_all add: Seq_def) -section {* Kleene Star for Sets *} +section \Kleene Star for Sets\ inductive_set Star :: "string set \ string set" ("_\" [101] 102) @@ -43,7 +43,7 @@ | step[intro]: "\s1 \ A; s2 \ A\\ \ s1 @ s2 \ A\" -text {* A Standard Property of Star *} +text \A Standard Property of Star\ lemma star_cases: shows "A\ = {[]} \ A ;; A\" @@ -58,32 +58,32 @@ (auto simp add: append_eq_Cons_conv) -section {* Semantics of Regular Expressions *} +section \Meaning of Regular Expressions\ fun L :: "rexp \ string set" where "L (ZERO) = {}" | "L (ONE) = {[]}" -| "L (CHAR c) = {[c]}" +| "L (CH c) = {[c]}" | "L (SEQ r1 r2) = (L r1) ;; (L r2)" | "L (ALT r1 r2) = (L r1) \ (L r2)" | "L (STAR r) = (L r)\" -section {* The Matcher *} +section \The Matcher\ fun nullable :: "rexp \ bool" where "nullable (ZERO) = False" | "nullable (ONE) = True" -| "nullable (CHAR c) = False" +| "nullable (CH c) = False" | "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" | "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" | "nullable (STAR r) = True" -section {* Correctness Proof for Nullable *} +section \Correctness Proof for Nullable\ lemma nullable_correctness: shows "nullable r \ [] \ (L r)" @@ -131,5 +131,91 @@ apply(simp_all add: Seq_def Star.start) done +section \Derivative Operation\ + +fun der :: "char \ rexp \ 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 \ 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)" + +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_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 \ s \ L r" +by (induct s arbitrary: r) + (simp_all add: nullable_correctness der_correctness Der_def) + + end \ No newline at end of file