diff -r 1d4659dd83fe -r 51e00f223792 progs/Matcher2.thy --- a/progs/Matcher2.thy Fri Oct 18 05:59:04 2024 +0100 +++ b/progs/Matcher2.thy Fri Oct 25 18:54:08 2024 +0100 @@ -11,12 +11,12 @@ by (metis UN_extend_simps(10) image_Suc_atLeastAtMost) -section {* Regular Expressions *} +section \Regular Expressions\ datatype rexp = NULL | EMPTY -| CHAR char +| CH char | SEQ rexp rexp | ALT rexp rexp | STAR rexp @@ -28,14 +28,14 @@ | UPNTIMES rexp nat -section {* Sequential Composition of Sets *} +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 *} +text \Two Simple Properties about Sequential Composition\ lemma seq_empty [simp]: shows "A ;; {[]} = A" @@ -68,7 +68,7 @@ done -section {* Power for Sets *} +section \Power for Sets\ fun pow :: "string set \ nat \ string set" ("_ \ _" [101, 102] 101) @@ -84,7 +84,7 @@ "A \ (n + m) = A \ n ;; A \ m" by (induct n) (simp_all add: seq_assoc) -section {* Kleene Star for Sets *} +section \Kleene Star for Sets\ inductive_set Star :: "string set \ string set" ("_\" [101] 102) @@ -93,7 +93,7 @@ start[intro]: "[] \ A\" | step[intro]: "\s1 \ A; s2 \ A\\ \ s1 @ s2 \ A\" -text {* A Standard Property of Star *} +text \A Standard Property of Star\ lemma star_decomp: assumes a: "c # x \ A\" @@ -132,14 +132,14 @@ -section {* Semantics of Regular Expressions *} +section \Semantics of Regular Expressions\ fun L :: "rexp \ string set" where "L (NULL) = {}" | "L (EMPTY) = {[]}" -| "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)\" @@ -154,14 +154,14 @@ apply(simp) done -section {* The Matcher *} +section \The Matcher\ fun nullable :: "rexp \ bool" where "nullable (NULL) = False" | "nullable (EMPTY) = 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" @@ -176,7 +176,7 @@ where "M (NULL) = 0" | "M (EMPTY) = 0" -| "M (CHAR char) = 0" +| "M (CH 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)" @@ -191,7 +191,7 @@ where "der c (NULL) = NULL" | "der c (EMPTY) = NULL" -| "der c (CHAR d) = (if c = d then EMPTY else NULL)" +| "der c (CH 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)" @@ -232,7 +232,8 @@ defer apply(subgoal_tac "Suc (Suc m) - Suc (Suc n) < Suc (Suc m) - Suc n") prefer 2 -apply(auto)[1] + apply(auto)[1] + (* apply (metis Suc_mult_less_cancel1 mult.assoc numeral_eq_Suc) apply(subgoal_tac "0 < (Suc (Suc m) - Suc n)") @@ -274,7 +275,7 @@ "matcher r s = nullable (ders s r)" -section {* Correctness Proof of the Matcher *} +section \Correctness Proof of the Matcher\ lemma nullable_correctness: shows "nullable r \ [] \ (L r)" @@ -282,7 +283,7 @@ apply(auto simp add: Seq_def) done -section {* Left-Quotient of a Set *} +section \Left-Quotient of a Set\ definition Der :: "char \ string set \ string set" @@ -334,6 +335,31 @@ finally show "Der c (A\) = (Der c A) ;; A\" . qed +lemma test: + assumes "[] \ A" + shows "Der c (A \ n) \ (Der c A) ;; (A \ n)" + using assms + apply(induct n) + apply(simp) + apply(simp) + apply(auto simp add: Der_def Seq_def) + apply blast + by force + +lemma Der_ntimes [simp]: + shows "Der c (A \ (Suc n)) = (Der c A) ;; (A \ n)" +proof - + have "Der c (A \ (Suc n)) = Der c (A ;; A \ n)" + by(simp) + also have "... = (Der c A) ;; (A \ n) \ (if [] \ A then Der c (A \ n) else {})" + by simp + also have "... = (Der c A) ;; (A \ n)" + using test by force + finally show "Der c (A \ (Suc n)) = (Der c A) ;; (A \ n)" . +qed + + + lemma Der_UNIV [simp]: "Der c (UNIV - A) = UNIV - Der c A" unfolding Der_def @@ -346,7 +372,67 @@ lemma Der_UNION [simp]: shows "Der c (\x\A. B x) = (\x\A. Der c (B x))" -by (auto simp add: Der_def) + by (auto simp add: Der_def) + +lemma if_f: + shows "f(if B then C else D) = (if B then f(C) else f(D))" + by simp + + +lemma der_correctness: + shows "L (der c r) = Der c (L r)" +proof(induct r) + case NULL + then show ?case by simp +next + case EMPTY + then show ?case by simp +next + case (CH x) + then show ?case by simp +next + case (SEQ r1 r2) + then show ?case + by (simp add: nullable_correctness) +next + case (ALT r1 r2) + then show ?case by simp +next + case (STAR r) + then show ?case + by simp +next + case (NOT r) + then show ?case by simp +next + case (PLUS r) + then show ?case by simp +next + case (OPT r) + then show ?case by simp +next + case (NTIMES r n) + then show ?case + apply(induct n) + apply(simp) + apply(simp only: L.simps) + apply(simp only: Der_pow) + apply(simp only: der.simps L.simps) + apply(simp only: nullable_correctness) + apply(simp only: if_f) + by simp +next + case (NMTIMES r n m) + then show ?case + apply(case_tac n) + sorry +next + case (UPNTIMES r x2) + then show ?case sorry +qed + + + lemma der_correctness: shows "L (der c r) = Der c (L r)" @@ -379,7 +465,6 @@ by simp lemma "L(der c (UPNTIMES r (Suc n))) = L(SEQ (der c r) (UPNTIMES r n))" -using assms proof(induct n) case 0 show ?case by simp next