thys2/SpecAlts.thy
changeset 365 ec5e4fe4cc70
equal deleted inserted replaced
364:232aa2f19a75 365:ec5e4fe4cc70
       
     1    
       
     2 theory SpecAlts
       
     3   imports Main "~~/src/HOL/Library/Sublist"
       
     4 begin
       
     5 
       
     6 section {* Sequential Composition of Languages *}
       
     7 
       
     8 definition
       
     9   Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
       
    10 where 
       
    11   "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
       
    12 
       
    13 text {* Two Simple Properties about Sequential Composition *}
       
    14 
       
    15 lemma Sequ_empty_string [simp]:
       
    16   shows "A ;; {[]} = A"
       
    17   and   "{[]} ;; A = A"
       
    18 by (simp_all add: Sequ_def)
       
    19 
       
    20 lemma Sequ_empty [simp]:
       
    21   shows "A ;; {} = {}"
       
    22   and   "{} ;; A = {}"
       
    23 by (simp_all add: Sequ_def)
       
    24 
       
    25 
       
    26 section {* Semantic Derivative (Left Quotient) of Languages *}
       
    27 
       
    28 definition
       
    29   Der :: "char \<Rightarrow> string set \<Rightarrow> string set"
       
    30 where
       
    31   "Der c A \<equiv> {s. c # s \<in> A}"
       
    32 
       
    33 definition
       
    34   Ders :: "string \<Rightarrow> string set \<Rightarrow> string set"
       
    35 where
       
    36   "Ders s A \<equiv> {s'. s @ s' \<in> A}"
       
    37 
       
    38 lemma Der_null [simp]:
       
    39   shows "Der c {} = {}"
       
    40 unfolding Der_def
       
    41 by auto
       
    42 
       
    43 lemma Der_empty [simp]:
       
    44   shows "Der c {[]} = {}"
       
    45 unfolding Der_def
       
    46 by auto
       
    47 
       
    48 lemma Der_char [simp]:
       
    49   shows "Der c {[d]} = (if c = d then {[]} else {})"
       
    50 unfolding Der_def
       
    51 by auto
       
    52 
       
    53 lemma Der_union [simp]:
       
    54   shows "Der c (A \<union> B) = Der c A \<union> Der c B"
       
    55 unfolding Der_def
       
    56   by auto
       
    57 
       
    58 lemma Der_Union [simp]:
       
    59   shows "Der c (\<Union>B. A) = (\<Union>B. Der c A)"
       
    60 unfolding Der_def
       
    61 by auto
       
    62 
       
    63 lemma Der_Sequ [simp]:
       
    64   shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"
       
    65 unfolding Der_def Sequ_def
       
    66 by (auto simp add: Cons_eq_append_conv)
       
    67 
       
    68 
       
    69 section {* Kleene Star for Languages *}
       
    70 
       
    71 inductive_set
       
    72   Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
       
    73   for A :: "string set"
       
    74 where
       
    75   start[intro]: "[] \<in> A\<star>"
       
    76 | step[intro]:  "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"
       
    77 
       
    78 (* Arden's lemma *)
       
    79 
       
    80 lemma Star_cases:
       
    81   shows "A\<star> = {[]} \<union> A ;; A\<star>"
       
    82 unfolding Sequ_def
       
    83 by (auto) (metis Star.simps)
       
    84 
       
    85 lemma Star_decomp: 
       
    86   assumes "c # x \<in> A\<star>" 
       
    87   shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A\<star>"
       
    88 using assms
       
    89 by (induct x\<equiv>"c # x" rule: Star.induct) 
       
    90    (auto simp add: append_eq_Cons_conv)
       
    91 
       
    92 lemma Star_Der_Sequ: 
       
    93   shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"
       
    94 unfolding Der_def Sequ_def
       
    95 by(auto simp add: Star_decomp)
       
    96 
       
    97 
       
    98 lemma Der_star [simp]:
       
    99   shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
       
   100 proof -    
       
   101   have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"  
       
   102     by (simp only: Star_cases[symmetric])
       
   103   also have "... = Der c (A ;; A\<star>)"
       
   104     by (simp only: Der_union Der_empty) (simp)
       
   105   also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})"
       
   106     by simp
       
   107   also have "... =  (Der c A) ;; A\<star>"
       
   108     using Star_Der_Sequ by auto
       
   109   finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
       
   110 qed
       
   111 
       
   112 
       
   113 section {* Regular Expressions *}
       
   114 
       
   115 datatype rexp =
       
   116   ZERO
       
   117 | ONE
       
   118 | CHAR char
       
   119 | SEQ rexp rexp
       
   120 | ALTS "rexp list"
       
   121 | STAR rexp
       
   122 
       
   123 section {* Semantics of Regular Expressions *}
       
   124  
       
   125 fun
       
   126   L :: "rexp \<Rightarrow> string set"
       
   127 where
       
   128   "L (ZERO) = {}"
       
   129 | "L (ONE) = {[]}"
       
   130 | "L (CHAR c) = {[c]}"
       
   131 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
       
   132 | "L (ALTS rs) = (\<Union>r \<in> set rs. L r)"
       
   133 | "L (STAR r) = (L r)\<star>"
       
   134 
       
   135 
       
   136 section {* Nullable, Derivatives *}
       
   137 
       
   138 fun
       
   139  nullable :: "rexp \<Rightarrow> bool"
       
   140 where
       
   141   "nullable (ZERO) = False"
       
   142 | "nullable (ONE) = True"
       
   143 | "nullable (CHAR c) = False"
       
   144 | "nullable (ALTS rs) = (\<exists>r \<in> set rs. nullable r)"
       
   145 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
       
   146 | "nullable (STAR r) = True"
       
   147 
       
   148 
       
   149 fun
       
   150  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
       
   151 where
       
   152   "der c (ZERO) = ZERO"
       
   153 | "der c (ONE) = ZERO"
       
   154 | "der c (CHAR d) = (if c = d then ONE else ZERO)"
       
   155 | "der c (ALTS rs) = ALTS (map (der c) rs)"
       
   156 | "der c (SEQ r1 r2) = 
       
   157      (if nullable r1
       
   158       then ALTS [SEQ (der c r1) r2, der c r2]
       
   159       else SEQ (der c r1) r2)"
       
   160 | "der c (STAR r) = SEQ (der c r) (STAR r)"
       
   161 
       
   162 fun 
       
   163  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
       
   164 where
       
   165   "ders [] r = r"
       
   166 | "ders (c # s) r = ders s (der c r)"
       
   167 
       
   168 
       
   169 lemma nullable_correctness:
       
   170   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
       
   171 by (induct r) (auto simp add: Sequ_def) 
       
   172 
       
   173 lemma der_correctness:
       
   174   shows "L (der c r) = Der c (L r)"
       
   175   apply(induct r) 
       
   176        apply(simp_all add: nullable_correctness)
       
   177   apply(auto simp add: Der_def)
       
   178   done
       
   179 
       
   180 lemma ders_correctness:
       
   181   shows "L (ders s r) = Ders s (L r)"
       
   182 by (induct s arbitrary: r)
       
   183    (simp_all add: Ders_def der_correctness Der_def)
       
   184 
       
   185 fun flats :: "rexp list \<Rightarrow> rexp list"
       
   186   where
       
   187   "flats [] = []"
       
   188 | "flats (ZERO # rs1) = flats(rs1)"
       
   189 | "flats ((ALTS rs1) #rs2) = rs1 @ (flats rs2)"
       
   190 | "flats (r1 # rs2) = r1 # flats rs2"
       
   191 
       
   192 fun simp_SEQ where
       
   193   "simp_SEQ ONE r\<^sub>2 = r\<^sub>2"
       
   194 | "simp_SEQ r\<^sub>1 ONE = r\<^sub>1"
       
   195 | "simp_SEQ r\<^sub>1 r\<^sub>2 = SEQ r\<^sub>1 r\<^sub>2"  
       
   196  
       
   197 fun 
       
   198   simp :: "rexp \<Rightarrow> rexp"
       
   199 where
       
   200   "simp (ALTS rs) = ALTS (remdups (flats (map simp rs)))" 
       
   201 | "simp (SEQ r1 r2) = simp_SEQ (simp r1) (simp r2)" 
       
   202 | "simp r = r"
       
   203 
       
   204 lemma simp_SEQ_correctness:
       
   205   shows "L (simp_SEQ r1 r2) = L (SEQ r1 r2)"
       
   206   apply(induct r1 r2 rule: simp_SEQ.induct)
       
   207   apply(simp_all)
       
   208   done
       
   209 
       
   210 lemma flats_correctness:
       
   211   shows "(\<Union>r \<in> set (flats rs). L r) = L (ALTS rs)"
       
   212   apply(induct rs rule: flats.induct)
       
   213   apply(simp_all)
       
   214   done
       
   215 
       
   216 
       
   217 lemma simp_correctness:
       
   218   shows "L (simp r) = L r"
       
   219   apply(induct r)
       
   220   apply(simp_all)
       
   221   apply(simp add: simp_SEQ_correctness)
       
   222   apply(simp add: flats_correctness)
       
   223   done
       
   224 
       
   225 fun 
       
   226  ders2 :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
       
   227 where
       
   228   "ders2 [] r = r"
       
   229 | "ders2 (c # s) r = ders2 s (simp (der c r))"
       
   230 
       
   231 lemma ders2_ZERO:
       
   232   shows "ders2 s ZERO = ZERO"
       
   233   apply(induct s)
       
   234   apply(simp_all)
       
   235   done
       
   236 
       
   237 lemma ders2_ONE:
       
   238   shows "ders2 s ONE \<in> {ZERO, ONE}"
       
   239   apply(induct s)
       
   240   apply(simp_all)
       
   241   apply(auto)
       
   242   apply(case_tac s)
       
   243   apply(auto)
       
   244   apply(case_tac s)
       
   245   apply(auto)
       
   246   done
       
   247 
       
   248 lemma ders2_CHAR:
       
   249   shows "ders2 s (CHAR c) \<in> {ZERO, ONE, CHAR c}"
       
   250   apply(induct s)
       
   251   apply(simp_all)
       
   252   apply(auto simp add: ders2_ZERO)
       
   253   apply(case_tac s)
       
   254   apply(auto simp add: ders2_ZERO)
       
   255   using ders2_ONE
       
   256   apply(auto)[1]
       
   257   using ders2_ONE
       
   258   apply(auto)[1]
       
   259   done
       
   260 
       
   261 lemma remdup_size:
       
   262   shows "size_list f (remdups rs) \<le> size_list f rs"
       
   263   apply(induct rs)
       
   264    apply(simp_all)
       
   265   done
       
   266 
       
   267 lemma flats_append:
       
   268   shows "flats (rs1 @ rs2) = (flats rs1) @ (flats rs2)"
       
   269   apply(induct rs1 arbitrary: rs2)
       
   270    apply(auto)
       
   271   apply(case_tac a)
       
   272        apply(auto)
       
   273   done
       
   274 
       
   275 lemma flats_Cons:
       
   276   shows "flats (r # rs) = (flats [r]) @ (flats rs)"
       
   277   apply(subst flats_append[symmetric])
       
   278   apply(simp)
       
   279   done
       
   280 
       
   281 lemma flats_size:
       
   282   shows "size_list (\<lambda>x. size (ders2 s x)) (flats rs) \<le> size_list (\<lambda>x. size (ders2 s x))  rs"
       
   283   apply(induct rs arbitrary: s rule: flats.induct)
       
   284    apply(simp_all)
       
   285    apply(simp add: ders2_ZERO)
       
   286    apply (simp add: le_SucI)
       
   287   
       
   288    apply(subst flats_Cons)
       
   289   apply(simp)
       
   290   apply(case_tac a)
       
   291        apply(auto)
       
   292    apply(simp add: ders2_ZERO)
       
   293    apply (simp add: le_SucI)
       
   294   sorry
       
   295 
       
   296 lemma ders2_ALTS:
       
   297   shows "size (ders2 s (ALTS rs)) \<le> size (ALTS (map (ders2 s) rs))"
       
   298   apply(induct s arbitrary: rs)
       
   299    apply(simp_all)
       
   300   thm size_list_pointwise
       
   301   apply (simp add: size_list_pointwise)
       
   302   apply(drule_tac x="remdups (flats (map (simp \<circ> der a) rs))" in meta_spec)
       
   303   apply(rule le_trans)
       
   304    apply(assumption)
       
   305   apply(simp)
       
   306   apply(rule le_trans)
       
   307    apply(rule remdup_size)
       
   308   apply(simp add: comp_def)
       
   309   apply(rule le_trans)
       
   310   apply(rule flats_size)
       
   311   by (simp add: size_list_pointwise)
       
   312 
       
   313 definition
       
   314  "derss2 A r = {ders2 s r | s. s \<in> A}"
       
   315 
       
   316 lemma
       
   317   "\<forall>rd \<in> derss2 (UNIV) r. size rd \<le> Suc (size r)"
       
   318   apply(induct r)
       
   319   apply(auto simp add: derss2_def ders2_ZERO)[1]
       
   320       apply(auto simp add: derss2_def ders2_ZERO)[1]
       
   321   using ders2_ONE
       
   322       apply(auto)[1]
       
   323   apply (metis rexp.size(7) rexp.size(8) zero_le)
       
   324   using ders2_CHAR
       
   325      apply(auto)[1]
       
   326   apply (smt derss2_def le_SucI le_zero_eq mem_Collect_eq rexp.size(7) rexp.size(8) rexp.size(9))
       
   327     defer  
       
   328     apply(auto simp add: derss2_def)
       
   329     apply(rule le_trans)
       
   330      apply(rule ders2_ALTS)
       
   331     apply(simp)
       
   332     apply(simp add: comp_def)
       
   333   
       
   334     apply(simp add: size_list_pointwise)
       
   335     apply(case_tac s)
       
   336      apply(simp)
       
   337   apply(simp only:)
       
   338   apply(auto)[1]
       
   339    
       
   340   apply(case_tac s)
       
   341         apply(simp)
       
   342   apply(simp)
       
   343 
       
   344 section {* Values *}
       
   345 
       
   346 datatype val = 
       
   347   Void
       
   348 | Char char
       
   349 | Seq val val
       
   350 | Nth nat val
       
   351 | Stars "val list"
       
   352 
       
   353 
       
   354 section {* The string behind a value *}
       
   355 
       
   356 fun 
       
   357   flat :: "val \<Rightarrow> string"
       
   358 where
       
   359   "flat (Void) = []"
       
   360 | "flat (Char c) = [c]"
       
   361 | "flat (Nth n v) = flat v"
       
   362 | "flat (Seq v1 v2) = (flat v1) @ (flat v2)"
       
   363 | "flat (Stars []) = []"
       
   364 | "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" 
       
   365 
       
   366 abbreviation
       
   367   "flats vs \<equiv> concat (map flat vs)"
       
   368 
       
   369 lemma flat_Stars [simp]:
       
   370  "flat (Stars vs) = flats vs"
       
   371 by (induct vs) (auto)
       
   372 
       
   373 lemma Star_concat:
       
   374   assumes "\<forall>s \<in> set ss. s \<in> A"  
       
   375   shows "concat ss \<in> A\<star>"
       
   376 using assms by (induct ss) (auto)
       
   377 
       
   378 lemma Star_cstring:
       
   379   assumes "s \<in> A\<star>"
       
   380   shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])"
       
   381 using assms
       
   382 apply(induct rule: Star.induct)
       
   383 apply(auto)[1]
       
   384 apply(rule_tac x="[]" in exI)
       
   385 apply(simp)
       
   386 apply(erule exE)
       
   387 apply(clarify)
       
   388 apply(case_tac "s1 = []")
       
   389 apply(rule_tac x="ss" in exI)
       
   390 apply(simp)
       
   391 apply(rule_tac x="s1#ss" in exI)
       
   392 apply(simp)
       
   393 done
       
   394 
       
   395 
       
   396 section {* Lexical Values *}
       
   397 
       
   398 inductive 
       
   399   Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile> _ : _" [100, 100] 100)
       
   400 where
       
   401   "\<lbrakk>\<Turnstile> v1 : r1; \<Turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<Turnstile>  Seq v1 v2 : SEQ r1 r2"
       
   402 | "\<lbrakk>\<Turnstile> v1 : (nth rs n); n < length rs\<rbrakk> \<Longrightarrow> \<Turnstile> (Nth n v1) : ALTS rs"
       
   403 | "\<Turnstile> Void : ONE"
       
   404 | "\<Turnstile> Char c : CHAR c"
       
   405 | "\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> [] \<Longrightarrow> \<Turnstile> Stars vs : STAR r"
       
   406 
       
   407 inductive_cases Prf_elims:
       
   408   "\<Turnstile> v : ZERO"
       
   409   "\<Turnstile> v : SEQ r1 r2"
       
   410   "\<Turnstile> v : ALTS rs"
       
   411   "\<Turnstile> v : ONE"
       
   412   "\<Turnstile> v : CHAR c"
       
   413   "\<Turnstile> vs : STAR r"
       
   414 
       
   415 lemma Prf_Stars_appendE:
       
   416   assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r"
       
   417   shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r" 
       
   418 using assms
       
   419 by (auto intro: Prf.intros elim!: Prf_elims)
       
   420 
       
   421 
       
   422 lemma Star_cval:
       
   423   assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r"
       
   424   shows "\<exists>vs. flats vs = concat ss \<and> (\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])"
       
   425 using assms
       
   426 apply(induct ss)
       
   427 apply(auto)
       
   428 apply(rule_tac x="[]" in exI)
       
   429 apply(simp)
       
   430 apply(case_tac "flat v = []")
       
   431 apply(rule_tac x="vs" in exI)
       
   432 apply(simp)
       
   433 apply(rule_tac x="v#vs" in exI)
       
   434 apply(simp)
       
   435 done
       
   436 
       
   437 
       
   438 lemma L_flat_Prf1:
       
   439   assumes "\<Turnstile> v : r" 
       
   440   shows "flat v \<in> L r"
       
   441 using assms
       
   442   apply(induct) 
       
   443   apply(auto simp add: Sequ_def Star_concat)
       
   444   done  
       
   445 
       
   446 lemma L_flat_Prf2:
       
   447   assumes "s \<in> L r" 
       
   448   shows "\<exists>v. \<Turnstile> v : r \<and> flat v = s"
       
   449 using assms
       
   450 proof(induct r arbitrary: s)
       
   451   case (STAR r s)
       
   452   have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
       
   453   have "s \<in> L (STAR r)" by fact
       
   454   then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []"
       
   455   using Star_cstring by auto  
       
   456   then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []"
       
   457   using IH Star_cval by metis 
       
   458   then show "\<exists>v. \<Turnstile> v : STAR r \<and> flat v = s"
       
   459   using Prf.intros(5) flat_Stars by blast
       
   460 next 
       
   461   case (SEQ r1 r2 s)
       
   462   then show "\<exists>v. \<Turnstile> v : SEQ r1 r2 \<and> flat v = s"
       
   463   unfolding Sequ_def L.simps by (fastforce intro: Prf.intros)
       
   464 next
       
   465   case (ALTS rs s)
       
   466   then show "\<exists>v. \<Turnstile> v : ALTS rs \<and> flat v = s"
       
   467     unfolding L.simps 
       
   468     apply(auto)
       
   469     apply(case_tac rs)
       
   470      apply(simp)
       
   471     apply(simp)
       
   472     apply(auto)
       
   473      apply(drule_tac x="a" in meta_spec)
       
   474      apply(simp)
       
   475      apply(drule_tac x="s" in meta_spec)
       
   476      apply(simp)
       
   477      apply(erule exE)
       
   478      apply(rule_tac x="Nth 0 v" in exI)
       
   479      apply(simp)
       
   480      apply(rule Prf.intros)
       
   481       apply(simp)
       
   482      apply(simp)
       
   483     apply(drule_tac x="x" in meta_spec)
       
   484     apply(simp)
       
   485     apply(drule_tac x="s" in meta_spec)
       
   486     apply(simp)
       
   487     apply(erule exE)
       
   488     apply(subgoal_tac "\<exists>n. nth list n = x \<and> n < length list")
       
   489     apply(erule exE)
       
   490      apply(rule_tac x="Nth (Suc n) v" in exI)
       
   491      apply(simp)
       
   492      apply(rule Prf.intros)
       
   493       apply(simp)
       
   494      apply(simp)
       
   495     by (meson in_set_conv_nth)
       
   496 qed (auto intro: Prf.intros)
       
   497 
       
   498 
       
   499 lemma L_flat_Prf:
       
   500   shows "L(r) = {flat v | v. \<Turnstile> v : r}"
       
   501 using L_flat_Prf1 L_flat_Prf2 by blast
       
   502 
       
   503 
       
   504 
       
   505 section {* Sets of Lexical Values *}
       
   506 
       
   507 text {*
       
   508   Shows that lexical values are finite for a given regex and string.
       
   509 *}
       
   510 
       
   511 definition
       
   512   LV :: "rexp \<Rightarrow> string \<Rightarrow> val set"
       
   513 where  "LV r s \<equiv> {v. \<Turnstile> v : r \<and> flat v = s}"
       
   514 
       
   515 lemma LV_simps:
       
   516   shows "LV ZERO s = {}"
       
   517   and   "LV ONE s = (if s = [] then {Void} else {})"
       
   518   and   "LV (CHAR c) s = (if s = [c] then {Char c} else {})"
       
   519 unfolding LV_def
       
   520 by (auto intro: Prf.intros elim: Prf.cases)
       
   521 
       
   522 
       
   523 abbreviation
       
   524   "Prefixes s \<equiv> {s'. prefix s' s}"
       
   525 
       
   526 abbreviation
       
   527   "Suffixes s \<equiv> {s'. suffix s' s}"
       
   528 
       
   529 abbreviation
       
   530   "SSuffixes s \<equiv> {s'. strict_suffix s' s}"
       
   531 
       
   532 lemma Suffixes_cons [simp]:
       
   533   shows "Suffixes (c # s) = Suffixes s \<union> {c # s}"
       
   534 by (auto simp add: suffix_def Cons_eq_append_conv)
       
   535 
       
   536 
       
   537 lemma finite_Suffixes: 
       
   538   shows "finite (Suffixes s)"
       
   539 by (induct s) (simp_all)
       
   540 
       
   541 lemma finite_SSuffixes: 
       
   542   shows "finite (SSuffixes s)"
       
   543 proof -
       
   544   have "SSuffixes s \<subseteq> Suffixes s"
       
   545    unfolding strict_suffix_def suffix_def by auto
       
   546   then show "finite (SSuffixes s)"
       
   547    using finite_Suffixes finite_subset by blast
       
   548 qed
       
   549 
       
   550 lemma finite_Prefixes: 
       
   551   shows "finite (Prefixes s)"
       
   552 proof -
       
   553   have "finite (Suffixes (rev s))" 
       
   554     by (rule finite_Suffixes)
       
   555   then have "finite (rev ` Suffixes (rev s))" by simp
       
   556   moreover
       
   557   have "rev ` (Suffixes (rev s)) = Prefixes s"
       
   558   unfolding suffix_def prefix_def image_def
       
   559    by (auto)(metis rev_append rev_rev_ident)+
       
   560   ultimately show "finite (Prefixes s)" by simp
       
   561 qed
       
   562 
       
   563 lemma LV_STAR_finite:
       
   564   assumes "\<forall>s. finite (LV r s)"
       
   565   shows "finite (LV (STAR r) s)"
       
   566 proof(induct s rule: length_induct)
       
   567   fix s::"char list"
       
   568   assume "\<forall>s'. length s' < length s \<longrightarrow> finite (LV (STAR r) s')"
       
   569   then have IH: "\<forall>s' \<in> SSuffixes s. finite (LV (STAR r) s')"
       
   570     by (force simp add: strict_suffix_def suffix_def) 
       
   571   define f where "f \<equiv> \<lambda>(v, vs). Stars (v # vs)"
       
   572   define S1 where "S1 \<equiv> \<Union>s' \<in> Prefixes s. LV r s'"
       
   573   define S2 where "S2 \<equiv> \<Union>s2 \<in> SSuffixes s. Stars -` (LV (STAR r) s2)"
       
   574   have "finite S1" using assms
       
   575     unfolding S1_def by (simp_all add: finite_Prefixes)
       
   576   moreover 
       
   577   with IH have "finite S2" unfolding S2_def
       
   578     by (auto simp add: finite_SSuffixes inj_on_def finite_vimageI)
       
   579   ultimately 
       
   580   have "finite ({Stars []} \<union> f ` (S1 \<times> S2))" by simp
       
   581   moreover 
       
   582   have "LV (STAR r) s \<subseteq> {Stars []} \<union> f ` (S1 \<times> S2)" 
       
   583   unfolding S1_def S2_def f_def
       
   584   unfolding LV_def image_def prefix_def strict_suffix_def
       
   585   apply(auto)
       
   586   apply(case_tac x)
       
   587   apply(auto elim: Prf_elims)
       
   588   apply(erule Prf_elims)
       
   589   apply(auto)
       
   590   apply(case_tac vs)
       
   591   apply(auto intro: Prf.intros)  
       
   592   apply(rule exI)
       
   593   apply(rule conjI)
       
   594   apply(rule_tac x="flat a" in exI)
       
   595   apply(rule conjI)
       
   596   apply(rule_tac x="flats list" in exI)
       
   597   apply(simp)
       
   598    apply(blast)
       
   599   apply(simp add: suffix_def)
       
   600   using Prf.intros(5) by blast  
       
   601   ultimately
       
   602   show "finite (LV (STAR r) s)" by (simp add: finite_subset)
       
   603 qed  
       
   604     
       
   605 
       
   606 lemma LV_finite:
       
   607   shows "finite (LV r s)"
       
   608 proof(induct r arbitrary: s)
       
   609   case (ZERO s) 
       
   610   show "finite (LV ZERO s)" by (simp add: LV_simps)
       
   611 next
       
   612   case (ONE s)
       
   613   show "finite (LV ONE s)" by (simp add: LV_simps)
       
   614 next
       
   615   case (CHAR c s)
       
   616   show "finite (LV (CHAR c) s)" by (simp add: LV_simps)
       
   617 next 
       
   618   case (ALTS rs s)
       
   619   then show "finite (LV (ALTS rs) s)" 
       
   620     sorry
       
   621 next 
       
   622   case (SEQ r1 r2 s)
       
   623   define f where "f \<equiv> \<lambda>(v1, v2). Seq v1 v2"
       
   624   define S1 where "S1 \<equiv> \<Union>s' \<in> Prefixes s. LV r1 s'"
       
   625   define S2 where "S2 \<equiv> \<Union>s' \<in> Suffixes s. LV r2 s'"
       
   626   have IHs: "\<And>s. finite (LV r1 s)" "\<And>s. finite (LV r2 s)" by fact+
       
   627   then have "finite S1" "finite S2" unfolding S1_def S2_def
       
   628     by (simp_all add: finite_Prefixes finite_Suffixes)
       
   629   moreover
       
   630   have "LV (SEQ r1 r2) s \<subseteq> f ` (S1 \<times> S2)"
       
   631     unfolding f_def S1_def S2_def 
       
   632     unfolding LV_def image_def prefix_def suffix_def
       
   633     apply (auto elim!: Prf_elims)
       
   634     by (metis (mono_tags, lifting) mem_Collect_eq)  
       
   635   ultimately 
       
   636   show "finite (LV (SEQ r1 r2) s)"
       
   637     by (simp add: finite_subset)
       
   638 next
       
   639   case (STAR r s)
       
   640   then show "finite (LV (STAR r) s)" by (simp add: LV_STAR_finite)
       
   641 qed
       
   642 
       
   643 
       
   644 (*
       
   645 section {* Our POSIX Definition *}
       
   646 
       
   647 inductive 
       
   648   Posix :: "string \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<in> _ \<rightarrow> _" [100, 100, 100] 100)
       
   649 where
       
   650   Posix_ONE: "[] \<in> ONE \<rightarrow> Void"
       
   651 | Posix_CHAR: "[c] \<in> (CHAR c) \<rightarrow> (Char c)"
       
   652 | Posix_ALT1: "s \<in> r1 \<rightarrow> v \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Left v)"
       
   653 | Posix_ALT2: "\<lbrakk>s \<in> r2 \<rightarrow> v; s \<notin> L(r1)\<rbrakk> \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Right v)"
       
   654 | Posix_SEQ: "\<lbrakk>s1 \<in> r1 \<rightarrow> v1; s2 \<in> r2 \<rightarrow> v2;
       
   655     \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r1 \<and> s\<^sub>4 \<in> L r2)\<rbrakk> \<Longrightarrow> 
       
   656     (s1 @ s2) \<in> (SEQ r1 r2) \<rightarrow> (Seq v1 v2)"
       
   657 | Posix_STAR1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v \<noteq> [];
       
   658     \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk>
       
   659     \<Longrightarrow> (s1 @ s2) \<in> STAR r \<rightarrow> Stars (v # vs)"
       
   660 | Posix_STAR2: "[] \<in> STAR r \<rightarrow> Stars []"
       
   661 
       
   662 inductive_cases Posix_elims:
       
   663   "s \<in> ZERO \<rightarrow> v"
       
   664   "s \<in> ONE \<rightarrow> v"
       
   665   "s \<in> CHAR c \<rightarrow> v"
       
   666   "s \<in> ALT r1 r2 \<rightarrow> v"
       
   667   "s \<in> SEQ r1 r2 \<rightarrow> v"
       
   668   "s \<in> STAR r \<rightarrow> v"
       
   669 
       
   670 lemma Posix1:
       
   671   assumes "s \<in> r \<rightarrow> v"
       
   672   shows "s \<in> L r" "flat v = s"
       
   673 using assms
       
   674 by (induct s r v rule: Posix.induct)
       
   675    (auto simp add: Sequ_def)
       
   676 
       
   677 text {*
       
   678   Our Posix definition determines a unique value.
       
   679 *}
       
   680 
       
   681 lemma Posix_determ:
       
   682   assumes "s \<in> r \<rightarrow> v1" "s \<in> r \<rightarrow> v2"
       
   683   shows "v1 = v2"
       
   684 using assms
       
   685 proof (induct s r v1 arbitrary: v2 rule: Posix.induct)
       
   686   case (Posix_ONE v2)
       
   687   have "[] \<in> ONE \<rightarrow> v2" by fact
       
   688   then show "Void = v2" by cases auto
       
   689 next 
       
   690   case (Posix_CHAR c v2)
       
   691   have "[c] \<in> CHAR c \<rightarrow> v2" by fact
       
   692   then show "Char c = v2" by cases auto
       
   693 next 
       
   694   case (Posix_ALT1 s r1 v r2 v2)
       
   695   have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact
       
   696   moreover
       
   697   have "s \<in> r1 \<rightarrow> v" by fact
       
   698   then have "s \<in> L r1" by (simp add: Posix1)
       
   699   ultimately obtain v' where eq: "v2 = Left v'" "s \<in> r1 \<rightarrow> v'" by cases auto 
       
   700   moreover
       
   701   have IH: "\<And>v2. s \<in> r1 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact
       
   702   ultimately have "v = v'" by simp
       
   703   then show "Left v = v2" using eq by simp
       
   704 next 
       
   705   case (Posix_ALT2 s r2 v r1 v2)
       
   706   have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact
       
   707   moreover
       
   708   have "s \<notin> L r1" by fact
       
   709   ultimately obtain v' where eq: "v2 = Right v'" "s \<in> r2 \<rightarrow> v'" 
       
   710     by cases (auto simp add: Posix1) 
       
   711   moreover
       
   712   have IH: "\<And>v2. s \<in> r2 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact
       
   713   ultimately have "v = v'" by simp
       
   714   then show "Right v = v2" using eq by simp
       
   715 next
       
   716   case (Posix_SEQ s1 r1 v1 s2 r2 v2 v')
       
   717   have "(s1 @ s2) \<in> SEQ r1 r2 \<rightarrow> v'" 
       
   718        "s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2"
       
   719        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by fact+
       
   720   then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \<in> r1 \<rightarrow> v1'" "s2 \<in> r2 \<rightarrow> v2'"
       
   721   apply(cases) apply (auto simp add: append_eq_append_conv2)
       
   722   using Posix1(1) by fastforce+
       
   723   moreover
       
   724   have IHs: "\<And>v1'. s1 \<in> r1 \<rightarrow> v1' \<Longrightarrow> v1 = v1'"
       
   725             "\<And>v2'. s2 \<in> r2 \<rightarrow> v2' \<Longrightarrow> v2 = v2'" by fact+
       
   726   ultimately show "Seq v1 v2 = v'" by simp
       
   727 next
       
   728   case (Posix_STAR1 s1 r v s2 vs v2)
       
   729   have "(s1 @ s2) \<in> STAR r \<rightarrow> v2" 
       
   730        "s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" "flat v \<noteq> []"
       
   731        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact+
       
   732   then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (Stars vs')"
       
   733   apply(cases) apply (auto simp add: append_eq_append_conv2)
       
   734   using Posix1(1) apply fastforce
       
   735   apply (metis Posix1(1) Posix_STAR1.hyps(6) append_Nil append_Nil2)
       
   736   using Posix1(2) by blast
       
   737   moreover
       
   738   have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
       
   739             "\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
       
   740   ultimately show "Stars (v # vs) = v2" by auto
       
   741 next
       
   742   case (Posix_STAR2 r v2)
       
   743   have "[] \<in> STAR r \<rightarrow> v2" by fact
       
   744   then show "Stars [] = v2" by cases (auto simp add: Posix1)
       
   745 qed
       
   746 
       
   747 
       
   748 text {*
       
   749   Our POSIX value is a lexical value.
       
   750 *}
       
   751 
       
   752 lemma Posix_LV:
       
   753   assumes "s \<in> r \<rightarrow> v"
       
   754   shows "v \<in> LV r s"
       
   755 using assms unfolding LV_def
       
   756 apply(induct rule: Posix.induct)
       
   757 apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)
       
   758 done
       
   759 *)
       
   760 
       
   761 
       
   762 end