updated
authorChristian Urban <christian.urban@kcl.ac.uk>
Sun, 09 Oct 2022 13:39:34 +0100
changeset 882 5fcad75ade92
parent 881 3b2f76950473
child 883 740bb9557905
updated
progs/Matcher.thy
progs/matcher/re1.sc
progs/matcher/re3.sc
progs/pow.scala
slides/slides02.pdf
slides/slides02.tex
--- 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 \<open>Regular Expressions\<close>
 
 datatype rexp =
   ZERO
 | ONE
-| CHAR char
+| CH char
 | SEQ rexp rexp
 | ALT rexp rexp
 | STAR rexp
 
 
-section {* Sequential Composition of Sets *}
+section \<open>Sequential Composition of Sets of Strings\<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"
@@ -33,7 +33,7 @@
   and   "{} ;; A = {}"
 by (simp_all add: Seq_def)
 
-section {* Kleene Star for Sets *}
+section \<open>Kleene Star for Sets\<close>
 
 inductive_set
   Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
@@ -43,7 +43,7 @@
 | 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_cases:
   shows "A\<star> = {[]} \<union> A ;; A\<star>"
@@ -58,32 +58,32 @@
    (auto simp add: append_eq_Cons_conv)
 
 
-section {* Semantics of Regular Expressions *}
+section \<open>Meaning of Regular Expressions\<close>
  
 fun
   L :: "rexp \<Rightarrow> 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) \<union> (L r2)"
 | "L (STAR r) = (L r)\<star>"
 
-section {* The Matcher *}
+section \<open>The Matcher\<close>
 
 fun
  nullable :: "rexp \<Rightarrow> bool"
 where
   "nullable (ZERO) = False"
 | "nullable (ONE) = 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"
 
 
-section {* Correctness Proof for Nullable *}
+section \<open>Correctness Proof for Nullable\<close>
 
 lemma nullable_correctness:
   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
@@ -131,5 +131,91 @@
 apply(simp_all add: Seq_def Star.start)
 done
 
+section \<open>Derivative Operation\<close>
+
+fun der :: "char \<Rightarrow> rexp \<Rightarrow> 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 \<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)"
+
+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_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 \<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
--- a/progs/matcher/re1.sc	Sun Oct 02 08:42:01 2022 +0100
+++ b/progs/matcher/re1.sc	Sun Oct 09 13:39:34 2022 +0100
@@ -19,9 +19,9 @@
 case class SEQ(r1: Rexp, r2: Rexp) extends Rexp  // sequence
 case class STAR(r: Rexp) extends Rexp            // star
 
-
 // nullable function: tests whether a regular 
 // expression can recognise the empty string
+  
 def nullable(r: Rexp) : Boolean = r match {
   case ZERO => false
   case ONE => true
--- a/progs/matcher/re3.sc	Sun Oct 02 08:42:01 2022 +0100
+++ b/progs/matcher/re3.sc	Sun Oct 09 13:39:34 2022 +0100
@@ -65,6 +65,7 @@
 }
 
 
+
 // the derivative w.r.t. a string (iterates der)
 def ders(s: List[Char], r: Rexp) : Rexp = s match {
   case Nil => r
--- a/progs/pow.scala	Sun Oct 02 08:42:01 2022 +0100
+++ b/progs/pow.scala	Sun Oct 09 13:39:34 2022 +0100
@@ -3,16 +3,29 @@
 
 def pow(A: Set[String], n: Int) : Set[String] = n match {
   case 0 => Set("")
-  case n => concat(A, pow(A, n- 1))
+  case n => concat(A, pow(A, n - 1))
 }
 
+def powT(A: Set[String], n: Int, acc: Set[String] = Set("")) : Set[String] = 
+  n match {
+    case 0 => acc
+    case n => powT(A, n - 1, concat(acc, A))
+  }
+
+
 val A = Set("a", "b", "c", "d", "e")
 val B = Set("a", "b", "c", "d", "")
 pow(A, 4).size
 pow(B, 4).size
+powT(A, 4).size
+powT(B, 4).size
 
 
-val A = Set("aa", "a")
+val C = Set("a", "b")
+
+pow(C, 100).size
+powT(C, 100000)  
+
 val B = Set("aaa", "aaaa")
 concat(A, B).size                     // -> 3 
 
Binary file slides/slides02.pdf has changed
--- a/slides/slides02.tex	Sun Oct 02 08:42:01 2022 +0100
+++ b/slides/slides02.tex	Sun Oct 09 13:39:34 2022 +0100
@@ -36,16 +36,17 @@
   \begin{tabular}{@ {}c@ {}}
   \\[-3mm]
   \LARGE Compilers and \\[-1mm] 
-  \LARGE Formal Languages\\[3mm] 
+  \LARGE Formal Languages\\[-5mm] 
   \end{tabular}}
 
   \normalsize
   \begin{center}
   \begin{tabular}{ll}
-    Email:  & christian.urban at kcl.ac.uk\\
-    %Office Hours: & Thursdays 12 -- 14\\
-    %Location: & N7.07 (North Wing, Bush House)\\
-    Slides \& Progs: & KEATS (also homework is there)\\  
+  Email:  & christian.urban at kcl.ac.uk\\
+  Office Hour: & Fridays 11 -- 12\\
+  Location: & N7.07 (North Wing, Bush House)\\
+  Slides \& Progs: & KEATS\\
+  Pollev: & \texttt{\alert{https://pollev.com/cfltutoratki576}}\\  
   \end{tabular}
   \end{center}