updated
authorChristian Urban <urbanc@in.tum.de>
Sun, 17 Feb 2019 22:15:06 +0000
changeset 311 8b8db9558ecf
parent 310 c090baa7059d
child 312 8b0b414e71b0
updated
Literature/star-der-paper.pdf
exps/both.scala
thys/BitCoded.thy
thys/Lexer.thy
thys/PDerivs.thy
thys/Positions.thy
thys/Re.thy
thys/RegLangs.thy
thys/Spec.thy
Binary file Literature/star-der-paper.pdf has changed
--- a/exps/both.scala	Mon Feb 11 23:18:05 2019 +0000
+++ b/exps/both.scala	Sun Feb 17 22:15:06 2019 +0000
@@ -444,9 +444,6 @@
   case r => r
 }
 
-
-
-
 def bders_simp (s: List[Char], r: ARexp) : ARexp = s match {
   case Nil => r
   case c::s => bders_simp(s, bsimp(bder(c, r)))
@@ -463,7 +460,8 @@
  decode(r, blex_simp(internalise(r), s.toList))
 
 
-def btokenise_simp(r: Rexp, s: String) = env(blexing_simp(r, s)).map(esc2)
+def btokenise_simp(r: Rexp, s: String) = 
+  env(blexing_simp(r, s)).map(esc2)
 
 
 
@@ -503,6 +501,33 @@
 
 def btokenise_simp_full(r: Rexp, s: String) = env(blexing_simp_full(r, s)).map(esc2)
 
+// bders2 for strings in the ALTS case
+
+def bders2_simp(s: List[Char], a: ARexp) : ARexp = {
+  //println(s"s = ${s.length}   a = ${asize(a)}")
+  //Console.readLine
+  (s, a) match {
+  case (Nil, r) => r
+  case (s, AZERO) => AZERO
+  case (s, AONE(_)) => AZERO
+  case (s, APRED(bs, f, _)) => 
+    if (f(s.head) && s.tail == Nil) AONE(bs:::List(C(s.head))) else AZERO
+  case (s, AALTS(bs, rs)) => bsimp(AALTS(bs, rs.map(bders2_simp(s, _)))) 
+  case (c::s, r) => bders2_simp(s, bsimp(bder(c, r)))
+}}
+
+
+
+def blexing2_simp(r: Rexp, s: String) : Val = {
+  val bder = bders2_simp(s.toList, internalise(r))
+  if (bnullable(bder)) decode(r, bmkeps(bder)) else 
+  throw new Exception("Not matched")
+
+}
+
+def btokenise2_simp(r: Rexp, s: String) = 
+  env(blexing2_simp(r, s)).map(esc2)
+
 
 
 //   Testing
@@ -594,14 +619,22 @@
 println(astring(bders_simp("aaaaaabaaaabbbbbaaaaaaaaaaaaaaa".toList, internalise(re1))))
 
 
+for (i <- 0 to 100 by 5) {
+  //print("Old: " + time(tokenise_simp(re1, "a" * i)))
+  print(" Bit: " + time(btokenise_simp(re1, "a" * i)))
+  print(" Bit full simp: " + time(btokenise_simp_full(re1, "a" * i)))
+  println(" Bit2: " + time(btokenise2_simp(re1, "a" * i)))
+}
 
-
+Console.readLine
 
 
 // Bigger Tests
 //==============
 
 
+println("Big tests")
+
 val fib_prog = """
 write "Fib";
 read n;
@@ -627,12 +660,14 @@
   print("Old: " + time(tokenise_simp(WHILE_REGS, fib_prog * i)))
   print(" Bit: " + time(btokenise_simp(WHILE_REGS, fib_prog * i)))
   println(" Bit full simp: " + time(btokenise_simp_full(WHILE_REGS, fib_prog * i)))
+  //println(" Bit2: " + time(btokenise2_simp(WHILE_REGS, fib_prog * i)))
 }
 
 
 println("Original " + size(WHILE_REGS))
 println("Size Bit  " + asize(bders_simp((fib_prog * 1).toList, internalise(WHILE_REGS))))
 println("Size Bitf " + asize(bders_simp_full((fib_prog * 1).toList, internalise(WHILE_REGS))))
+println("Size Bit2 " + asize(bders2_simp((fib_prog * 1).toList, internalise(WHILE_REGS))))
 println("Size Old  " + size(ders_simp((fib_prog * 1).toList, WHILE_REGS)))
 println("Size Pder " + psize(pders_simp((fib_prog * 1).toList, WHILE_REGS)))
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/BitCoded.thy	Sun Feb 17 22:15:06 2019 +0000
@@ -0,0 +1,326 @@
+   
+theory BitCoded
+  imports "Lexer" 
+begin
+
+section {* Bit-Encodings *}
+
+datatype bit = Z | S
+
+fun 
+  code :: "val \<Rightarrow> bit list"
+where
+  "code Void = []"
+| "code (Char c) = []"
+| "code (Left v) = Z # (code v)"
+| "code (Right v) = S # (code v)"
+| "code (Seq v1 v2) = (code v1) @ (code v2)"
+| "code (Stars []) = [S]"
+| "code (Stars (v # vs)) =  (Z # code v) @ code (Stars vs)"
+
+
+fun 
+  Stars_add :: "val \<Rightarrow> val \<Rightarrow> val"
+where
+  "Stars_add v (Stars vs) = Stars (v # vs)"
+
+function
+  decode' :: "bit list \<Rightarrow> rexp \<Rightarrow> (val * bit list)"
+where
+  "decode' ds ZERO = (Void, [])"
+| "decode' ds ONE = (Void, ds)"
+| "decode' ds (CHAR d) = (Char d, ds)"
+| "decode' [] (ALT r1 r2) = (Void, [])"
+| "decode' (Z # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r1 in (Left v, ds'))"
+| "decode' (S # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r2 in (Right v, ds'))"
+| "decode' ds (SEQ r1 r2) = (let (v1, ds') = decode' ds r1 in
+                             let (v2, ds'') = decode' ds' r2 in (Seq v1 v2, ds''))"
+| "decode' [] (STAR r) = (Void, [])"
+| "decode' (S # ds) (STAR r) = (Stars [], ds)"
+| "decode' (Z # ds) (STAR r) = (let (v, ds') = decode' ds r in
+                                    let (vs, ds'') = decode' ds' (STAR r) 
+                                    in (Stars_add v vs, ds''))"
+by pat_completeness auto
+
+lemma decode'_smaller:
+  assumes "decode'_dom (ds, r)"
+  shows "length (snd (decode' ds r)) \<le> length ds"
+using assms
+apply(induct ds r)
+apply(auto simp add: decode'.psimps split: prod.split)
+using dual_order.trans apply blast
+by (meson dual_order.trans le_SucI)
+
+termination "decode'"  
+apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") 
+apply(auto dest!: decode'_smaller)
+by (metis less_Suc_eq_le snd_conv)
+
+definition
+  decode :: "bit list \<Rightarrow> rexp \<Rightarrow> val option"
+where
+  "decode ds r \<equiv> (let (v, ds') = decode' ds r 
+                  in (if ds' = [] then Some v else None))"
+
+lemma decode'_code_Stars:
+  assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> (\<forall>x. decode' (code v @ x) r = (v, x)) \<and> flat v \<noteq> []" 
+  shows "decode' (code (Stars vs) @ ds) (STAR r) = (Stars vs, ds)"
+  using assms
+  apply(induct vs)
+  apply(auto)
+  done
+
+lemma decode'_code:
+  assumes "\<Turnstile> v : r"
+  shows "decode' ((code v) @ ds) r = (v, ds)"
+using assms
+  apply(induct v r arbitrary: ds) 
+  apply(auto)
+  using decode'_code_Stars by blast
+
+lemma decode_code:
+  assumes "\<Turnstile> v : r"
+  shows "decode (code v) r = Some v"
+  using assms unfolding decode_def
+  by (smt append_Nil2 decode'_code old.prod.case)
+
+
+section {* Annotated Regular Expressions *}
+
+datatype arexp =
+  AZERO
+| AONE "bit list"
+| ACHAR "bit list" char
+| ASEQ "bit list" arexp arexp
+| AALTs "bit list" "arexp list"
+| ASTAR "bit list" arexp
+
+abbreviation
+  "AALT bs r1 r2 \<equiv> AALTs bs [r1, r2]"
+
+
+fun fuse :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp" where
+  "fuse bs AZERO = AZERO"
+| "fuse bs (AONE cs) = AONE (bs @ cs)" 
+| "fuse bs (ACHAR cs c) = ACHAR (bs @ cs) c"
+| "fuse bs (AALTs cs rs) = AALTs (bs @ cs) rs"
+| "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2"
+| "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r"
+
+fun intern :: "rexp \<Rightarrow> arexp" where
+  "intern ZERO = AZERO"
+| "intern ONE = AONE []"
+| "intern (CHAR c) = ACHAR [] c"
+| "intern (ALT r1 r2) = AALT [] (fuse [Z] (intern r1)) 
+                                (fuse [S]  (intern r2))"
+| "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)"
+| "intern (STAR r) = ASTAR [] (intern r)"
+
+
+fun retrieve :: "arexp \<Rightarrow> val \<Rightarrow> bit list" where
+  "retrieve (AONE bs) Void = bs"
+| "retrieve (ACHAR bs c) (Char d) = bs"
+| "retrieve (AALTs bs (r#rs)) (Left v) = bs @ retrieve r v"
+| "retrieve (AALTs bs (r#rs)) (Right v) = bs @ retrieve (AALTs [] rs) v"
+| "retrieve (ASEQ bs r1 r2) (Seq v1 v2) = bs @ retrieve r1 v1 @ retrieve r2 v2"
+| "retrieve (ASTAR bs r) (Stars []) = bs @ [S]"
+| "retrieve (ASTAR bs r) (Stars (v#vs)) = 
+     bs @ [Z] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)"
+
+fun 
+  erase :: "arexp \<Rightarrow> rexp"
+where
+  "erase AZERO = ZERO"
+| "erase (AONE _) = ONE"
+| "erase (ACHAR _ c) = CHAR c"
+| "erase (AALTs _ []) = ZERO"
+| "erase (AALTs _ [r]) = (erase r)"
+| "erase (AALTs _ (r#rs)) = ALT (erase r) (erase (AALTs [] rs))"
+| "erase (ASEQ _ r1 r2) = SEQ (erase r1) (erase r2)"
+| "erase (ASTAR _ r) = STAR (erase r)"
+
+fun
+ bnullable :: "arexp \<Rightarrow> bool"
+where
+  "bnullable (AZERO) = False"
+| "bnullable (AONE bs) = True"
+| "bnullable (ACHAR bs c) = False"
+| "bnullable (AALTs bs rs) = (\<exists>r \<in> set rs. bnullable r)"
+| "bnullable (ASEQ bs r1 r2) = (bnullable r1 \<and> bnullable r2)"
+| "bnullable (ASTAR bs r) = True"
+
+fun 
+  bmkeps :: "arexp \<Rightarrow> bit list"
+where
+  "bmkeps(AONE bs) = bs"
+| "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)"
+| "bmkeps(AALTs bs (r#rs)) = (if bnullable(r) then bs @ (bmkeps r) else (bmkeps (AALTs bs rs)))"
+| "bmkeps(ASTAR bs r) = bs @ [S]"
+
+
+fun
+ bder :: "char \<Rightarrow> arexp \<Rightarrow> arexp"
+where
+  "bder c (AZERO) = AZERO"
+| "bder c (AONE bs) = AZERO"
+| "bder c (ACHAR bs d) = (if c = d then AONE bs else AZERO)"
+| "bder c (AALTs bs rs) = AALTs bs (map (bder c) rs)"
+| "bder c (ASEQ bs r1 r2) = 
+     (if bnullable r1
+      then AALT bs (ASEQ [] (bder c r1) r2) (fuse (bmkeps r1) (bder c r2))
+      else ASEQ bs (bder c r1) r2)"
+| "bder c (ASTAR bs r) = ASEQ bs (fuse [Z] (bder c r)) (ASTAR [] r)"
+
+
+fun 
+  bders :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
+where
+  "bders r [] = r"
+| "bders r (c#s) = bders (bder c r) s"
+
+lemma bders_append:
+  "bders r (s1 @ s2) = bders (bders r s1) s2"
+  apply(induct s1 arbitrary: r s2)
+  apply(simp_all)
+  done
+
+lemma bnullable_correctness:
+  shows "nullable (erase r) = bnullable r"
+  apply(induct r rule: erase.induct)
+  apply(simp_all)
+  done
+
+lemma erase_fuse:
+  shows "erase (fuse bs r) = erase r"
+  apply(induct r rule: erase.induct)
+  apply(simp_all)
+  done
+
+lemma erase_intern [simp]:
+  shows "erase (intern r) = r"
+  apply(induct r)
+  apply(simp_all add: erase_fuse)
+  done
+
+lemma erase_bder [simp]:
+  shows "erase (bder a r) = der a (erase r)"
+  apply(induct r rule: erase.induct)
+  apply(simp_all add: erase_fuse bnullable_correctness)
+  done
+
+lemma erase_bders [simp]:
+  shows "erase (bders r s) = ders s (erase r)"
+  apply(induct s arbitrary: r )
+  apply(simp_all)
+  done
+
+lemma retrieve_encode_STARS:
+  assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> code v = retrieve (intern r) v"
+  shows "code (Stars vs) = retrieve (ASTAR [] (intern r)) (Stars vs)"
+  using assms
+  apply(induct vs)
+  apply(simp_all)
+  done
+
+lemma retrieve_fuse2:
+  assumes "\<Turnstile> v : (erase r)"
+  shows "retrieve (fuse bs r) v = bs @ retrieve r v"
+  using assms
+  apply(induct r arbitrary: v bs rule: erase.induct)
+         apply(auto elim: Prf_elims)[1]
+  sorry
+
+lemma retrieve_fuse:
+  assumes "\<Turnstile> v : r"
+  shows "retrieve (fuse bs (intern r)) v = bs @ retrieve (intern r) v"
+  using assms 
+  by (simp_all add: retrieve_fuse2)
+
+
+lemma retrieve_code:
+  assumes "\<Turnstile> v : r"
+  shows "code v = retrieve (intern r) v"
+  using assms
+  apply(induct v r )
+  apply(simp_all add: retrieve_fuse retrieve_encode_STARS)
+  sorry
+
+
+lemma bmkeps_retrieve:
+  assumes "nullable (erase r)"
+  shows "bmkeps r = retrieve r (mkeps (erase r))"
+  using assms
+  apply(induct r)
+       apply(auto simp add: bnullable_correctness)
+  sorry
+  
+lemma bder_retrieve:
+  assumes "\<Turnstile> v : der c (erase r)"
+  shows "retrieve (bder c r) v = retrieve r (injval (erase r) c v)"
+  using assms
+  apply(induct r arbitrary: v)
+  apply(auto elim!: Prf_elims simp add: retrieve_fuse2 bnullable_correctness bmkeps_retrieve)
+  sorry
+
+lemma MAIN_decode:
+  assumes "\<Turnstile> v : ders s r"
+  shows "Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r"
+  using assms
+proof (induct s arbitrary: v rule: rev_induct)
+  case Nil
+  have "\<Turnstile> v : ders [] r" by fact
+  then have "\<Turnstile> v : r" by simp
+  then have "Some v = decode (retrieve (intern r) v) r"
+    using decode_code retrieve_code by auto
+  then show "Some (flex r id [] v) = decode (retrieve (bders (intern r) []) v) r"
+    by simp
+next
+  case (snoc c s v)
+  have IH: "\<And>v. \<Turnstile> v : ders s r \<Longrightarrow> 
+     Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" by fact
+  have asm: "\<Turnstile> v : ders (s @ [c]) r" by fact
+  then have asm2: "\<Turnstile> injval (ders s r) c v : ders s r" 
+    by(simp add: Prf_injval ders_append)
+  have "Some (flex r id (s @ [c]) v) = Some (flex r id s (injval (ders s r) c v))"
+    by (simp add: flex_append)
+  also have "... = decode (retrieve (bders (intern r) s) (injval (ders s r) c v)) r"
+    using asm2 IH by simp
+  also have "... = decode (retrieve (bder c (bders (intern r) s)) v) r"
+    using asm by(simp_all add: bder_retrieve ders_append)
+  finally show "Some (flex r id (s @ [c]) v) = 
+                 decode (retrieve (bders (intern r) (s @ [c])) v) r" by (simp add: bders_append)
+qed
+
+
+definition blexer where
+ "blexer r s \<equiv> if bnullable (bders (intern r) s) then 
+                decode (bmkeps (bders (intern r) s)) r else None"
+
+lemma blexer_correctness:
+  shows "blexer r s = lexer r s"
+proof -
+  { define bds where "bds \<equiv> bders (intern r) s"
+    define ds  where "ds \<equiv> ders s r"
+    assume asm: "nullable ds"
+    have era: "erase bds = ds" 
+      unfolding ds_def bds_def by simp
+    have mke: "\<Turnstile> mkeps ds : ds"
+      using asm by (simp add: mkeps_nullable)
+    have "decode (bmkeps bds) r = decode (retrieve bds (mkeps ds)) r"
+      using bmkeps_retrieve
+      using asm era by (simp add: bmkeps_retrieve)
+    also have "... =  Some (flex r id s (mkeps ds))"
+      using mke by (simp_all add: MAIN_decode ds_def bds_def)
+    finally have "decode (bmkeps bds) r = Some (flex r id s (mkeps ds))" 
+      unfolding bds_def ds_def .
+  }
+  then show "blexer r s = lexer r s"
+    unfolding blexer_def lexer_flex
+    apply(subst bnullable_correctness[symmetric])
+    apply(simp)
+    done
+qed
+
+
+
+end
\ No newline at end of file
--- a/thys/Lexer.thy	Mon Feb 11 23:18:05 2019 +0000
+++ b/thys/Lexer.thy	Sun Feb 17 22:15:06 2019 +0000
@@ -3,8 +3,7 @@
   imports Spec 
 begin
 
-
-section {* The Lexer Functions by Sulzmann and Lu  *}
+section {* The Lexer Functions by Sulzmann and Lu  (without simplification) *}
 
 fun 
   mkeps :: "rexp \<Rightarrow> val"
@@ -274,7 +273,7 @@
 using Posix1(1) lexer_correct_None lexer_correct_Some by blast
 
 
-
+subsection {* A slight reformulation of the lexer algorithm using stacked functions*}
 
 fun flex :: "rexp \<Rightarrow> (val \<Rightarrow> val) => string \<Rightarrow> (val \<Rightarrow> val)"
   where
--- a/thys/PDerivs.thy	Mon Feb 11 23:18:05 2019 +0000
+++ b/thys/PDerivs.thy	Sun Feb 17 22:15:06 2019 +0000
@@ -340,5 +340,64 @@
   finally show "card (pders_Set A r) \<le> awidth r + 1" by simp
 qed
 
+(* tests *)
+
+lemma b: 
+  assumes "rd \<in> pder c r"
+  shows "size rd \<le> (Suc (size r)) * (Suc (size r))"
+  using assms
+  apply(induct r arbitrary: rd)
+  apply(auto)[3]
+  apply(case_tac "c = x")
+      apply(auto)[2]
+  prefer 2
+    apply(auto)[1]
+  oops
+  
+  
+
+lemma a: 
+  assumes "rd \<in> pders s (SEQ r1 r2)"
+  shows "size  rd \<le> Suc (size r1 + size r2)"
+  using assms
+  apply(induct s arbitrary: r1 r2 rd)
+   apply(simp)
+  apply(auto)
+  apply(case_tac "nullable r1")
+   apply(auto)
+  oops
+  
+
+lemma
+  shows "\<forall>rd \<in> (pders_Set UNIV r). size rd \<le> size r"
+  apply(induct r)
+       apply(auto)[1]
+       apply(simp add: pders_Set_def)
+      apply(simp add: pders_Set_def)
+     apply(simp add: pders_Set_def pders_CHAR)
+  using pders_CHAR apply fastforce
+    prefer 2
+    apply(simp add: pders_Set_def)
+    apply (meson Un_iff le_SucI trans_le_add1 trans_le_add2)
+  apply(simp add: pders_Set_def)
+   apply(auto)[1]
+  apply(case_tac y)
+  apply(simp)
+   apply(simp)
+   apply(auto)[1]
+   apply(case_tac "nullable r1")
+    apply(simp)
+    apply(auto)[1]
+     prefer 3
+     apply(simp)
+     apply(auto)[1]
+     apply(subgoal_tac "size xa \<le> size r1")
+      prefer 2
+      apply (metis UN_I pders.simps(1) pders_snoc singletonI)
+  oops
+  
+  
+  
+  
 
 end
\ No newline at end of file
--- a/thys/Positions.thy	Mon Feb 11 23:18:05 2019 +0000
+++ b/thys/Positions.thy	Sun Feb 17 22:15:06 2019 +0000
@@ -1,8 +1,10 @@
    
 theory Positions
-  imports "Spec" "Lexer" 
+  imports "Spec" "Lexer"
 begin
 
+chapter {* An alternative definition for POSIX values *}
+
 section {* Positions in Values *}
 
 fun 
--- a/thys/Re.thy	Mon Feb 11 23:18:05 2019 +0000
+++ b/thys/Re.thy	Sun Feb 17 22:15:06 2019 +0000
@@ -5,7 +5,7 @@
 
 
 section {* Sequential Composition of Sets *}
-m
+
 definition
   Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
 where 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/RegLangs.thy	Sun Feb 17 22:15:06 2019 +0000
@@ -0,0 +1,198 @@
+   
+theory RegLangs
+  imports Main "~~/src/HOL/Library/Sublist"
+begin
+
+section {* Sequential Composition of Languages *}
+
+definition
+  Sequ :: "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 *}
+
+lemma Sequ_empty_string [simp]:
+  shows "A ;; {[]} = A"
+  and   "{[]} ;; A = A"
+by (simp_all add: Sequ_def)
+
+lemma Sequ_empty [simp]:
+  shows "A ;; {} = {}"
+  and   "{} ;; A = {}"
+by (simp_all add: Sequ_def)
+
+
+section {* Semantic Derivative (Left Quotient) of Languages *}
+
+definition
+  Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
+where
+  "Der c A \<equiv> {s. c # s \<in> A}"
+
+definition
+  Ders :: "string \<Rightarrow> string set \<Rightarrow> string set"
+where
+  "Ders s A \<equiv> {s'. s @ 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_Sequ [simp]:
+  shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"
+unfolding Der_def Sequ_def
+by (auto simp add: Cons_eq_append_conv)
+
+
+section {* Kleene Star for Languages *}
+
+inductive_set
+  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
+  for A :: "string set"
+where
+  start[intro]: "[] \<in> A\<star>"
+| step[intro]:  "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"
+
+(* Arden's lemma *)
+
+lemma Star_cases:
+  shows "A\<star> = {[]} \<union> A ;; A\<star>"
+unfolding Sequ_def
+by (auto) (metis Star.simps)
+
+lemma Star_decomp: 
+  assumes "c # x \<in> A\<star>" 
+  shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A\<star>"
+using assms
+by (induct x\<equiv>"c # x" rule: Star.induct) 
+   (auto simp add: append_eq_Cons_conv)
+
+lemma Star_Der_Sequ: 
+  shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"
+unfolding Der_def Sequ_def
+by(auto simp add: Star_decomp)
+
+
+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>"
+    using Star_Der_Sequ by auto
+  finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
+qed
+
+lemma Star_concat:
+  assumes "\<forall>s \<in> set ss. s \<in> A"  
+  shows "concat ss \<in> A\<star>"
+using assms by (induct ss) (auto)
+
+lemma Star_split:
+  assumes "s \<in> A\<star>"
+  shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])"
+using assms
+  apply(induct rule: Star.induct)
+  using concat.simps(1) apply fastforce
+  apply(clarify)
+  by (metis append_Nil concat.simps(2) set_ConsD)
+
+
+
+section {* Regular Expressions *}
+
+datatype rexp =
+  ZERO
+| ONE
+| CHAR char
+| SEQ rexp rexp
+| ALT rexp rexp
+| STAR rexp
+
+section {* Semantics of Regular Expressions *}
+ 
+fun
+  L :: "rexp \<Rightarrow> string set"
+where
+  "L (ZERO) = {}"
+| "L (ONE) = {[]}"
+| "L (CHAR 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 {* Nullable, Derivatives *}
+
+fun
+ nullable :: "rexp \<Rightarrow> bool"
+where
+  "nullable (ZERO) = False"
+| "nullable (ONE) = True"
+| "nullable (CHAR c) = False"
+| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
+| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
+| "nullable (STAR r) = True"
+
+
+fun
+ der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
+where
+  "der c (ZERO) = ZERO"
+| "der c (ONE) = ZERO"
+| "der c (CHAR 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)"
+
+
+lemma nullable_correctness:
+  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
+by (induct r) (auto simp add: Sequ_def) 
+
+lemma der_correctness:
+  shows "L (der c r) = Der c (L r)"
+by (induct r) (simp_all add: nullable_correctness)
+
+lemma ders_correctness:
+  shows "L (ders s r) = Ders s (L r)"
+  by (induct s arbitrary: r)
+     (simp_all add: Ders_def der_correctness Der_def)
+
+lemma ders_append:
+  shows "ders (s1 @ s2) r = ders s2 (ders s1 r)"
+  by (induct s1 arbitrary: s2 r) (auto)
+
+
+
+end
\ No newline at end of file
--- a/thys/Spec.thy	Mon Feb 11 23:18:05 2019 +0000
+++ b/thys/Spec.thy	Sun Feb 17 22:15:06 2019 +0000
@@ -1,187 +1,10 @@
    
 theory Spec
-  imports Main "~~/src/HOL/Library/Sublist"
+  imports RegLangs
 begin
 
-section {* Sequential Composition of Languages *}
-
-definition
-  Sequ :: "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 *}
-
-lemma Sequ_empty_string [simp]:
-  shows "A ;; {[]} = A"
-  and   "{[]} ;; A = A"
-by (simp_all add: Sequ_def)
-
-lemma Sequ_empty [simp]:
-  shows "A ;; {} = {}"
-  and   "{} ;; A = {}"
-by (simp_all add: Sequ_def)
-
-
-section {* Semantic Derivative (Left Quotient) of Languages *}
-
-definition
-  Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
-where
-  "Der c A \<equiv> {s. c # s \<in> A}"
-
-definition
-  Ders :: "string \<Rightarrow> string set \<Rightarrow> string set"
-where
-  "Ders s A \<equiv> {s'. s @ 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_Sequ [simp]:
-  shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"
-unfolding Der_def Sequ_def
-by (auto simp add: Cons_eq_append_conv)
-
-
-section {* Kleene Star for Languages *}
-
-inductive_set
-  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
-  for A :: "string set"
-where
-  start[intro]: "[] \<in> A\<star>"
-| step[intro]:  "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"
-
-(* Arden's lemma *)
-
-lemma Star_cases:
-  shows "A\<star> = {[]} \<union> A ;; A\<star>"
-unfolding Sequ_def
-by (auto) (metis Star.simps)
-
-lemma Star_decomp: 
-  assumes "c # x \<in> A\<star>" 
-  shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A\<star>"
-using assms
-by (induct x\<equiv>"c # x" rule: Star.induct) 
-   (auto simp add: append_eq_Cons_conv)
-
-lemma Star_Der_Sequ: 
-  shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"
-unfolding Der_def Sequ_def
-by(auto simp add: Star_decomp)
-
 
-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>"
-    using Star_Der_Sequ by auto
-  finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
-qed
-
-
-section {* Regular Expressions *}
-
-datatype rexp =
-  ZERO
-| ONE
-| CHAR char
-| SEQ rexp rexp
-| ALT rexp rexp
-| STAR rexp
-
-section {* Semantics of Regular Expressions *}
- 
-fun
-  L :: "rexp \<Rightarrow> string set"
-where
-  "L (ZERO) = {}"
-| "L (ONE) = {[]}"
-| "L (CHAR 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 {* Nullable, Derivatives *}
-
-fun
- nullable :: "rexp \<Rightarrow> bool"
-where
-  "nullable (ZERO) = False"
-| "nullable (ONE) = True"
-| "nullable (CHAR c) = False"
-| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
-| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
-| "nullable (STAR r) = True"
-
-
-fun
- der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
-where
-  "der c (ZERO) = ZERO"
-| "der c (ONE) = ZERO"
-| "der c (CHAR 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)"
-
-
-lemma nullable_correctness:
-  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
-by (induct r) (auto simp add: Sequ_def) 
-
-lemma der_correctness:
-  shows "L (der c r) = Der c (L r)"
-by (induct r) (simp_all add: nullable_correctness)
-
-lemma ders_correctness:
-  shows "L (ders s r) = Ders s (L r)"
-by (induct s arbitrary: r)
-   (simp_all add: Ders_def der_correctness Der_def)
-
-lemma ders_append:
-  shows "ders (s1 @ s2) r = ders s2 (ders s1 r)"
-  apply(induct s1 arbitrary: s2 r)
-  apply(auto)
-  done
-
-
-section {* Values *}
+section {* "Plain" Values *}
 
 datatype val = 
   Void
@@ -212,28 +35,6 @@
  "flat (Stars vs) = flats vs"
 by (induct vs) (auto)
 
-lemma Star_concat:
-  assumes "\<forall>s \<in> set ss. s \<in> A"  
-  shows "concat ss \<in> A\<star>"
-using assms by (induct ss) (auto)
-
-lemma Star_cstring:
-  assumes "s \<in> A\<star>"
-  shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])"
-using assms
-apply(induct rule: Star.induct)
-apply(auto)[1]
-apply(rule_tac x="[]" in exI)
-apply(simp)
-apply(erule exE)
-apply(clarify)
-apply(case_tac "s1 = []")
-apply(rule_tac x="ss" in exI)
-apply(simp)
-apply(rule_tac x="s1#ss" in exI)
-apply(simp)
-done
-
 
 section {* Lexical Values *}
 
@@ -262,7 +63,7 @@
 by (auto intro: Prf.intros elim!: Prf_elims)
 
 
-lemma Star_cval:
+lemma flats_Prf_value:
   assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r"
   shows "\<exists>vs. flats vs = concat ss \<and> (\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])"
 using assms
@@ -293,9 +94,9 @@
   have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
   have "s \<in> L (STAR r)" by fact
   then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []"
-  using Star_cstring by auto  
+  using Star_split by auto  
   then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []"
-  using IH Star_cval by metis 
+  using IH flats_Prf_value by metis 
   then show "\<exists>v. \<Turnstile> v : STAR r \<and> flat v = s"
   using Prf.intros(6) flat_Stars by blast
 next 
@@ -455,7 +256,7 @@
 
 
 
-section {* Our POSIX Definition *}
+section {* Our inductive POSIX Definition *}
 
 inductive 
   Posix :: "string \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<in> _ \<rightarrow> _" [100, 100, 100] 100)
@@ -488,7 +289,8 @@
    (auto simp add: Sequ_def)
 
 text {*
-  Our Posix definition determines a unique value.
+  For a give value and string, our Posix definition 
+  determines a unique value.
 *}
 
 lemma Posix_determ: