progs/Matcher2.thy
changeset 971 51e00f223792
parent 456 2fddf8ab744f
child 972 ebb4a40d9bae
--- 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 \<open>Regular Expressions\<close>
 
 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 \<open>Sequential Composition of Sets\<close>
 
 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 *}
+text \<open>Two Simple Properties about Sequential Composition\<close>
 
 lemma seq_empty [simp]:
   shows "A ;; {[]} = A"
@@ -68,7 +68,7 @@
 done
 
 
-section {* Power for Sets *}
+section \<open>Power for Sets\<close>
 
 fun 
   pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _" [101, 102] 101)
@@ -84,7 +84,7 @@
   "A \<up> (n + m) = A \<up> n ;; A \<up> m"
 by (induct n) (simp_all add: seq_assoc)
 
-section {* Kleene Star for Sets *}
+section \<open>Kleene Star for Sets\<close>
 
 inductive_set
   Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
@@ -93,7 +93,7 @@
   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 *}
+text \<open>A Standard Property of Star\<close>
 
 lemma star_decomp: 
   assumes a: "c # x \<in> A\<star>" 
@@ -132,14 +132,14 @@
 
 
 
-section {* Semantics of Regular Expressions *}
+section \<open>Semantics of Regular Expressions\<close>
  
 fun
   L :: "rexp \<Rightarrow> 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) \<union> (L r2)"
 | "L (STAR r) = (L r)\<star>"
@@ -154,14 +154,14 @@
 apply(simp)
 done
 
-section {* The Matcher *}
+section \<open>The Matcher\<close>
 
 fun
  nullable :: "rexp \<Rightarrow> bool"
 where
   "nullable (NULL) = False"
 | "nullable (EMPTY) = True"
-| "nullable (CHAR c) = False"
+| "nullable (CH c) = False"
 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
 | "nullable (SEQ r1 r2) = (nullable r1 \<and> 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 \<open>Correctness Proof of the Matcher\<close>
 
 lemma nullable_correctness:
   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
@@ -282,7 +283,7 @@
 apply(auto simp add: Seq_def) 
 done
 
-section {* Left-Quotient of a Set *}
+section \<open>Left-Quotient of a Set\<close>
 
 definition
   Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
@@ -334,6 +335,31 @@
   finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
 qed
 
+lemma test:
+  assumes "[] \<in> A"
+  shows "Der c (A \<up> n) \<subseteq> (Der c A) ;; (A \<up> 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  \<up> (Suc n)) = (Der c A) ;; (A \<up> n)"
+proof -    
+  have "Der c (A  \<up> (Suc n)) = Der c (A ;; A \<up> n)"
+    by(simp)
+  also have "... = (Der c A) ;; (A \<up> n) \<union> (if [] \<in> A then Der c (A \<up> n) else {})"
+    by simp
+  also have "... =  (Der c A) ;; (A \<up> n)"
+    using test by force
+  finally show "Der c (A  \<up> (Suc n)) = (Der c A) ;; (A \<up> 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 (\<Union>x\<in>A. B x) = (\<Union>x\<in>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