thys/#Re1.thy#
changeset 64 4c6af8ca13ea
parent 63 498171d2379a
child 65 b31b224fa0e6
equal deleted inserted replaced
63:498171d2379a 64:4c6af8ca13ea
     1 
       
     2 theory Re1
       
     3   imports "Main" 
       
     4 begin
       
     5 
       
     6 section {* Sequential Composition of Sets *}
       
     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 seq_empty [simp]:
       
    16   shows "A ;; {[]} = A"
       
    17   and   "{[]} ;; A = A"
       
    18 by (simp_all add: Sequ_def)
       
    19 
       
    20 lemma seq_null [simp]:
       
    21   shows "A ;; {} = {}"
       
    22   and   "{} ;; A = {}"
       
    23 by (simp_all add: Sequ_def)
       
    24 
       
    25 section {* Regular Expressions *}
       
    26 
       
    27 datatype rexp =
       
    28   NULL
       
    29 | EMPTY
       
    30 | CHAR char
       
    31 | SEQ rexp rexp
       
    32 | ALT rexp rexp
       
    33 
       
    34 section {* Semantics of Regular Expressions *}
       
    35  
       
    36 fun
       
    37   L :: "rexp \<Rightarrow> string set"
       
    38 where
       
    39   "L (NULL) = {}"
       
    40 | "L (EMPTY) = {[]}"
       
    41 | "L (CHAR c) = {[c]}"
       
    42 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
       
    43 | "L (ALT r1 r2) = (L r1) \<union> (L r2)"
       
    44 
       
    45 value "L(CHAR c)"
       
    46 value "L(SEQ(CHAR c)(CHAR b))"
       
    47 
       
    48 
       
    49 section {* Values *}
       
    50 
       
    51 datatype val = 
       
    52   Void
       
    53 | Char char
       
    54 | Seq val val
       
    55 | Right val
       
    56 | Left val
       
    57 
       
    58 section {* Relation between values and regular expressions *}
       
    59 
       
    60 inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
       
    61 where
       
    62  "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
       
    63 | "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
       
    64 | "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
       
    65 | "\<turnstile> Void : EMPTY"
       
    66 | "\<turnstile> Char c : CHAR c"
       
    67 
       
    68 section {* The string behind a value *}
       
    69 
       
    70 fun flat :: "val \<Rightarrow> string"
       
    71 where
       
    72   "flat(Void) = []"
       
    73 | "flat(Char c) = [c]"
       
    74 | "flat(Left v) = flat(v)"
       
    75 | "flat(Right v) = flat(v)"
       
    76 | "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
       
    77 
       
    78 value "flat(Seq(Char c)(Char b))"
       
    79 value "flat(Right(Void))"
       
    80 
       
    81 fun flats :: "val \<Rightarrow> string list"
       
    82 where
       
    83   "flats(Void) = [[]]"
       
    84 | "flats(Char c) = [[c]]"
       
    85 | "flats(Left v) = flats(v)"
       
    86 | "flats(Right v) = flats(v)"
       
    87 | "flats(Seq v1 v2) = (flats v1) @ (flats v2)"
       
    88 
       
    89 value "flats(Seq(Char c)(Char b))"
       
    90 
       
    91 lemma Prf_flat_L:
       
    92   assumes "\<turnstile> v : r" shows "flat v \<in> L r"
       
    93 using assms
       
    94 apply(induct)
       
    95 apply(auto simp add: Sequ_def)
       
    96 done
       
    97 
       
    98 lemma L_flat_Prf:
       
    99   "L(r) = {flat v | v. \<turnstile> v : r}"
       
   100 apply(induct r)
       
   101 apply(auto dest: Prf_flat_L simp add: Sequ_def)
       
   102 apply (metis Prf.intros(4) flat.simps(1))
       
   103 apply (metis Prf.intros(5) flat.simps(2))
       
   104 apply (metis Prf.intros(1) flat.simps(5))
       
   105 apply (metis Prf.intros(2) flat.simps(3))
       
   106 apply (metis Prf.intros(3) flat.simps(4))
       
   107 apply(erule Prf.cases)
       
   108 apply(auto)
       
   109 done
       
   110 
       
   111 definition definition prefix :: :: "string \<Rightarrow> string \<Rightarrow> bool" ("_ \<sqsubset> _" [100, 100] 100)
       
   112 where
       
   113   "s1 \<sqsubset> s2 \<equiv> \<exists>s3. s1 @ s3 = s2"
       
   114 
       
   115 section {* Ordering of values *}
       
   116 
       
   117 inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
       
   118 where
       
   119   "\<lbrakk>v1 = v1'; v2 \<succ>r2 v2'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
       
   120 | "v1  \<succ>r1 v1' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
       
   121 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
       
   122 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
       
   123 | "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
       
   124 | "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
       
   125 | "Void \<succ>EMPTY Void"
       
   126 | "(Char c) \<succ>(CHAR c) (Char c)"
       
   127 
       
   128 section {* The ordering is reflexive *}
       
   129 
       
   130 lemma ValOrd_refl:
       
   131   assumes "\<turnstile> v : r"
       
   132   shows "v \<succ>r v"
       
   133 using assms
       
   134 apply(induct)
       
   135 apply(auto intro: ValOrd.intros)
       
   136 done
       
   137 
       
   138 lemma ValOrd_flats:
       
   139   assumes "v1 \<succ>r v2"
       
   140   shows "hd (flats v2) = hd (flats v1)"
       
   141 using assms
       
   142 apply(induct)
       
   143 apply(auto)
       
   144 oops
       
   145 
       
   146 
       
   147 section {* Posix definition *}
       
   148 
       
   149 definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
       
   150 where
       
   151   "POSIX v r \<equiv> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v')"
       
   152 
       
   153 (*
       
   154 an alternative definition: might cause problems
       
   155 with theorem mkeps_POSIX
       
   156 *)
       
   157 
       
   158 definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
       
   159 where
       
   160   "POSIX2 v r \<equiv> \<turnstile> v : r \<and> (\<forall>v'. \<turnstile> v' : r \<longrightarrow> v \<succ>r v')"
       
   161 
       
   162 definition POSIX3 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
       
   163 where
       
   164   "POSIX3 v r \<equiv> \<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> length (flat v') \<le> length(flat v)) \<longrightarrow> v \<succ>r v')"
       
   165 
       
   166 
       
   167 lemma POSIX_SEQ:
       
   168   assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
       
   169   shows "POSIX v1 r1 \<and> POSIX v2 r2"
       
   170 using assms
       
   171 unfolding POSIX_def
       
   172 apply(auto)
       
   173 apply(drule_tac x="Seq v' v2" in spec)
       
   174 apply(simp)
       
   175 apply(erule impE)
       
   176 apply(rule Prf.intros)
       
   177 apply(simp)
       
   178 apply(simp)
       
   179 apply(erule ValOrd.cases)
       
   180 apply(simp_all)
       
   181 apply(clarify)
       
   182 defer
       
   183 apply(drule_tac x="Seq v1 v'" in spec)
       
   184 apply(simp)
       
   185 apply(erule impE)
       
   186 apply(rule Prf.intros)
       
   187 apply(simp)
       
   188 apply(simp)
       
   189 apply(erule ValOrd.cases)
       
   190 apply(simp_all)
       
   191 apply(clarify)
       
   192 oops (*not true*)
       
   193 
       
   194 lemma POSIX_SEQ_I:
       
   195   assumes "POSIX v1 r1" "POSIX v2 r2" 
       
   196   shows "POSIX (Seq v1 v2) (SEQ r1 r2)" 
       
   197 using assms
       
   198 unfolding POSIX_def
       
   199 apply(auto)
       
   200 apply(rotate_tac 2)
       
   201 apply(erule Prf.cases)
       
   202 apply(simp_all)[5]
       
   203 apply(auto)[1]
       
   204 apply(rule ValOrd.intros)
       
   205 oops (* maybe also not true *)
       
   206 
       
   207 lemma POSIX3_SEQ_I:
       
   208   assumes "POSIX3 v1 r1" "POSIX3 v2 r2" 
       
   209   shows "POSIX3 (Seq v1 v2) (SEQ r1 r2)" 
       
   210 using assms
       
   211 unfolding POSIX3_def
       
   212 apply(auto)
       
   213 apply (metis Prf.intros(1))
       
   214 apply(rotate_tac 4)
       
   215 apply(erule Prf.cases)
       
   216 apply(simp_all)[5]
       
   217 apply(auto)[1]
       
   218 apply(case_tac "v1 = v1a")
       
   219 apply(auto)
       
   220 apply (metis ValOrd.intros(1))
       
   221 apply (rule ValOrd.intros(2))
       
   222 oops
       
   223 
       
   224 lemma POSIX_ALT2:
       
   225   assumes "POSIX (Left v1) (ALT r1 r2)"
       
   226   shows "POSIX v1 r1"
       
   227 using assms
       
   228 unfolding POSIX_def
       
   229 apply(auto)
       
   230 apply(drule_tac x="Left v'" in spec)
       
   231 apply(simp)
       
   232 apply(drule mp)
       
   233 apply(rule Prf.intros)
       
   234 apply(auto)
       
   235 apply(erule ValOrd.cases)
       
   236 apply(simp_all)
       
   237 done
       
   238 
       
   239 lemma POSIX2_ALT:
       
   240   assumes "POSIX2 (Left v1) (ALT r1 r2)"
       
   241   shows "POSIX2 v1 r1"
       
   242 using assms
       
   243 unfolding POSIX2_def
       
   244 apply(auto)
       
   245 oops
       
   246 
       
   247 lemma POSIX_ALT:
       
   248   assumes "POSIX (Left v1) (ALT r1 r2)"
       
   249   shows "POSIX v1 r1"
       
   250 using assms
       
   251 unfolding POSIX_def
       
   252 apply(auto)
       
   253 apply(drule_tac x="Left v'" in spec)
       
   254 apply(simp)
       
   255 apply(drule mp)
       
   256 apply(rule Prf.intros)
       
   257 apply(auto)
       
   258 apply(erule ValOrd.cases)
       
   259 apply(simp_all)
       
   260 done
       
   261 
       
   262 lemma POSIX2_ALT:
       
   263   assumes "POSIX2 (Left v1) (ALT r1 r2)"
       
   264   shows "POSIX2 v1 r1"
       
   265 using assms
       
   266 apply(simp add: POSIX2_def)
       
   267 apply(auto)
       
   268 apply(erule Prf.cases)
       
   269 apply(simp_all)[5]
       
   270 apply(drule_tac x="Left v'" in spec)
       
   271 apply(drule mp)
       
   272 apply(rule Prf.intros)
       
   273 apply(auto)
       
   274 apply(erule ValOrd.cases)
       
   275 apply(simp_all)
       
   276 done
       
   277 
       
   278 
       
   279 lemma POSIX_ALT1a:
       
   280   assumes "POSIX (Right v2) (ALT r1 r2)"
       
   281   shows "POSIX v2 r2"
       
   282 using assms
       
   283 unfolding POSIX_def
       
   284 apply(auto)
       
   285 apply(drule_tac x="Right v'" in spec)
       
   286 apply(simp)
       
   287 apply(drule mp)
       
   288 apply(rule Prf.intros)
       
   289 apply(auto)
       
   290 apply(erule ValOrd.cases)
       
   291 apply(simp_all)
       
   292 done
       
   293 
       
   294 lemma POSIX2_ALT1a:
       
   295   assumes "POSIX2 (Right v2) (ALT r1 r2)"
       
   296   shows "POSIX2 v2 r2"
       
   297 using assms
       
   298 unfolding POSIX2_def
       
   299 apply(auto)
       
   300 apply(erule Prf.cases)
       
   301 apply(simp_all)[5]
       
   302 apply(drule_tac x="Right v'" in spec)
       
   303 apply(drule mp)
       
   304 apply(rule Prf.intros)
       
   305 apply(auto)
       
   306 apply(erule ValOrd.cases)
       
   307 apply(simp_all)
       
   308 done
       
   309 
       
   310 
       
   311 lemma POSIX_ALT1b:
       
   312   assumes "POSIX (Right v2) (ALT r1 r2)"
       
   313   shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
       
   314 using assms
       
   315 apply(drule_tac POSIX_ALT1a)
       
   316 unfolding POSIX_def
       
   317 apply(auto)
       
   318 done
       
   319 
       
   320 lemma POSIX_ALT_I1:
       
   321   assumes "POSIX v1 r1" 
       
   322   shows "POSIX (Left v1) (ALT r1 r2)"
       
   323 using assms
       
   324 unfolding POSIX_def
       
   325 apply(auto)
       
   326 apply(rotate_tac 3)
       
   327 apply(erule Prf.cases)
       
   328 apply(simp_all)[5]
       
   329 apply(auto)
       
   330 apply(rule ValOrd.intros)
       
   331 apply(auto)
       
   332 apply(rule ValOrd.intros)
       
   333 by simp
       
   334 
       
   335 lemma POSIX2_ALT_I1:
       
   336   assumes "POSIX2 v1 r1" 
       
   337   shows "POSIX2 (Left v1) (ALT r1 r2)"
       
   338 using assms
       
   339 unfolding POSIX2_def
       
   340 apply(auto)
       
   341 apply(rule Prf.intros)
       
   342 apply(simp)
       
   343 apply(rotate_tac 2)
       
   344 apply(erule Prf.cases)
       
   345 apply(simp_all)[5]
       
   346 apply(auto)
       
   347 apply(rule ValOrd.intros)
       
   348 apply(auto)
       
   349 apply(rule ValOrd.intros)
       
   350 oops
       
   351 
       
   352 lemma POSIX_ALT_I2:
       
   353   assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
       
   354   shows "POSIX (Right v2) (ALT r1 r2)"
       
   355 using assms
       
   356 unfolding POSIX_def
       
   357 apply(auto)
       
   358 apply(rotate_tac 3)
       
   359 apply(erule Prf.cases)
       
   360 apply(simp_all)[5]
       
   361 apply(auto)
       
   362 apply(rule ValOrd.intros)
       
   363 apply metis
       
   364 done
       
   365 
       
   366 
       
   367 
       
   368 
       
   369 
       
   370 section {* The Matcher *}
       
   371 
       
   372 fun
       
   373  nullable :: "rexp \<Rightarrow> bool"
       
   374 where
       
   375   "nullable (NULL) = False"
       
   376 | "nullable (EMPTY) = True"
       
   377 | "nullable (CHAR c) = False"
       
   378 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
       
   379 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
       
   380 
       
   381 lemma nullable_correctness:
       
   382   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
       
   383 apply (induct r) 
       
   384 apply(auto simp add: Sequ_def) 
       
   385 done
       
   386 
       
   387 fun mkeps :: "rexp \<Rightarrow> val"
       
   388 where
       
   389   "mkeps(EMPTY) = Void"
       
   390 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
       
   391 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
       
   392 
       
   393 lemma mkeps_nullable:
       
   394   assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
       
   395 using assms
       
   396 apply(induct rule: nullable.induct)
       
   397 apply(auto intro: Prf.intros)
       
   398 done
       
   399 
       
   400 lemma mkeps_flat:
       
   401   assumes "nullable(r)" shows "flat (mkeps r) = []"
       
   402 using assms
       
   403 apply(induct rule: nullable.induct)
       
   404 apply(auto)
       
   405 done
       
   406 
       
   407 text {*
       
   408   The value mkeps returns is always the correct POSIX
       
   409   value.
       
   410 *}
       
   411 
       
   412 lemma mkeps_POSIX2:
       
   413   assumes "nullable r"
       
   414   shows "POSIX2 (mkeps r) r"
       
   415 using assms
       
   416 apply(induct r)
       
   417 apply(auto)[1]
       
   418 apply(simp add: POSIX2_def)
       
   419 oops
       
   420 
       
   421 lemma mkeps_POSIX3:
       
   422   assumes "nullable r"
       
   423   shows "POSIX3 (mkeps r) r"
       
   424 using assms
       
   425 apply(induct r)
       
   426 apply(auto)[1]
       
   427 apply(simp add: POSIX3_def)
       
   428 apply(auto)[1]
       
   429 apply (metis Prf.intros(4))
       
   430 apply(erule Prf.cases)
       
   431 apply(simp_all)[5]
       
   432 apply (metis ValOrd.intros)
       
   433 apply(simp add: POSIX3_def)
       
   434 apply(auto)[1]
       
   435 apply(simp add: POSIX3_def)
       
   436 apply(auto)[1]
       
   437 apply (metis mkeps.simps(2) mkeps_nullable nullable.simps(5))
       
   438 apply(rotate_tac 6)
       
   439 apply(erule Prf.cases)
       
   440 apply(simp_all)[5]
       
   441 apply (metis ValOrd.intros(2) add_leE gen_length_code(1) gen_length_def mkeps_flat)
       
   442 apply(auto)
       
   443 apply(simp add: POSIX3_def)
       
   444 apply(auto)
       
   445 apply (metis Prf.intros(2))
       
   446 apply(rotate_tac 4)
       
   447 apply(erule Prf.cases)
       
   448 apply(simp_all)[5]
       
   449 apply (metis ValOrd.intros(6))
       
   450 apply(auto)[1]
       
   451 apply (metis ValOrd.intros(3))
       
   452 apply(simp add: POSIX3_def)
       
   453 apply(auto)
       
   454 apply (metis Prf.intros(2))
       
   455 apply(rotate_tac 6)
       
   456 apply(erule Prf.cases)
       
   457 apply(simp_all)[5]
       
   458 apply (metis ValOrd.intros(6))
       
   459 apply (metis ValOrd.intros(3))
       
   460 apply(simp add: POSIX3_def)
       
   461 apply(auto)
       
   462 apply (metis Prf.intros(3))
       
   463 apply(rotate_tac 5)
       
   464 apply(erule Prf.cases)
       
   465 apply(simp_all)[5]
       
   466 apply (metis Prf_flat_L drop_0 drop_all list.size(3) mkeps_flat nullable_correctness)
       
   467 by (metis ValOrd.intros(5))
       
   468 
       
   469 
       
   470 lemma mkeps_POSIX:
       
   471   assumes "nullable r"
       
   472   shows "POSIX (mkeps r) r"
       
   473 using assms
       
   474 apply(induct r)
       
   475 apply(auto)[1]
       
   476 apply(simp add: POSIX_def)
       
   477 apply(auto)[1]
       
   478 apply(erule Prf.cases)
       
   479 apply(simp_all)[5]
       
   480 apply (metis ValOrd.intros)
       
   481 apply(simp add: POSIX_def)
       
   482 apply(auto)[1]
       
   483 apply(simp add: POSIX_def)
       
   484 apply(auto)[1]
       
   485 apply(erule Prf.cases)
       
   486 apply(simp_all)[5]
       
   487 apply(auto)
       
   488 apply (simp add: ValOrd.intros(2) mkeps_flat)
       
   489 apply(simp add: POSIX_def)
       
   490 apply(auto)[1]
       
   491 apply(erule Prf.cases)
       
   492 apply(simp_all)[5]
       
   493 apply(auto)
       
   494 apply (simp add: ValOrd.intros(6))
       
   495 apply (simp add: ValOrd.intros(3))
       
   496 apply(simp add: POSIX_def)
       
   497 apply(auto)[1]
       
   498 apply(erule Prf.cases)
       
   499 apply(simp_all)[5]
       
   500 apply(auto)
       
   501 apply (simp add: ValOrd.intros(6))
       
   502 apply (simp add: ValOrd.intros(3))
       
   503 apply(simp add: POSIX_def)
       
   504 apply(auto)[1]
       
   505 apply(erule Prf.cases)
       
   506 apply(simp_all)[5]
       
   507 apply(auto)
       
   508 apply (metis Prf_flat_L mkeps_flat nullable_correctness)
       
   509 by (simp add: ValOrd.intros(5))
       
   510 
       
   511 
       
   512 lemma mkeps_POSIX2:
       
   513   assumes "nullable r"
       
   514   shows "POSIX2 (mkeps r) r"
       
   515 using assms
       
   516 apply(induct r)
       
   517 apply(simp)
       
   518 apply(simp)
       
   519 apply(simp add: POSIX2_def)
       
   520 apply(rule conjI)
       
   521 apply(rule Prf.intros)
       
   522 apply(auto)[1]
       
   523 apply(erule Prf.cases)
       
   524 apply(simp_all)[5]
       
   525 apply(rule ValOrd.intros)
       
   526 apply(simp)
       
   527 apply(simp)
       
   528 apply(simp add: POSIX2_def)
       
   529 apply(rule conjI)
       
   530 apply(rule Prf.intros)
       
   531 apply(simp add: mkeps_nullable)
       
   532 apply(simp add: mkeps_nullable)
       
   533 apply(auto)[1]
       
   534 apply(rotate_tac 6)
       
   535 apply(erule Prf.cases)
       
   536 apply(simp_all)[5]
       
   537 apply(rule ValOrd.intros(2))
       
   538 apply(simp)
       
   539 apply(simp only: nullable.simps)
       
   540 apply(erule disjE)
       
   541 apply(simp)
       
   542 thm POSIX2_ALT1a
       
   543 apply(rule POSIX2_ALT)
       
   544 apply(simp add: POSIX2_def)
       
   545 apply(rule conjI)
       
   546 apply(rule Prf.intros)
       
   547 apply(simp add: mkeps_nullable)
       
   548 oops
       
   549 
       
   550 
       
   551 section {* Derivatives *}
       
   552 
       
   553 fun
       
   554  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
       
   555 where
       
   556   "der c (NULL) = NULL"
       
   557 | "der c (EMPTY) = NULL"
       
   558 | "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
       
   559 | "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
       
   560 | "der c (SEQ r1 r2) = 
       
   561      (if nullable r1
       
   562       then ALT (SEQ (der c r1) r2) (der c r2)
       
   563       else SEQ (der c r1) r2)"
       
   564 
       
   565 fun 
       
   566  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
       
   567 where
       
   568   "ders [] r = r"
       
   569 | "ders (c # s) r = ders s (der c r)"
       
   570 
       
   571 section {* Injection function *}
       
   572 
       
   573 fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
       
   574 where
       
   575   "injval (CHAR d) c Void = Char d"
       
   576 | "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
       
   577 | "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
       
   578 | "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
       
   579 | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
       
   580 | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
       
   581 
       
   582 section {* Projection function *}
       
   583 
       
   584 fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
       
   585 where
       
   586   "projval (CHAR d) c _ = Void"
       
   587 | "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
       
   588 | "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
       
   589 | "projval (SEQ r1 r2) c (Seq v1 v2) = 
       
   590      (if flat v1 = [] then Right(projval r2 c v2) 
       
   591       else if nullable r1 then Left (Seq (projval r1 c v1) v2)
       
   592                           else Seq (projval r1 c v1) v2)"
       
   593 
       
   594 text {*
       
   595   Injection value is related to r
       
   596 *}
       
   597 
       
   598 lemma v3:
       
   599   assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
       
   600 using assms
       
   601 apply(induct arbitrary: v rule: der.induct)
       
   602 apply(simp)
       
   603 apply(erule Prf.cases)
       
   604 apply(simp_all)[5]
       
   605 apply(simp)
       
   606 apply(erule Prf.cases)
       
   607 apply(simp_all)[5]
       
   608 apply(case_tac "c = c'")
       
   609 apply(simp)
       
   610 apply(erule Prf.cases)
       
   611 apply(simp_all)[5]
       
   612 apply (metis Prf.intros(5))
       
   613 apply(simp)
       
   614 apply(erule Prf.cases)
       
   615 apply(simp_all)[5]
       
   616 apply(simp)
       
   617 apply(erule Prf.cases)
       
   618 apply(simp_all)[5]
       
   619 apply (metis Prf.intros(2))
       
   620 apply (metis Prf.intros(3))
       
   621 apply(simp)
       
   622 apply(case_tac "nullable r1")
       
   623 apply(simp)
       
   624 apply(erule Prf.cases)
       
   625 apply(simp_all)[5]
       
   626 apply(auto)[1]
       
   627 apply(erule Prf.cases)
       
   628 apply(simp_all)[5]
       
   629 apply(auto)[1]
       
   630 apply (metis Prf.intros(1))
       
   631 apply(auto)[1]
       
   632 apply (metis Prf.intros(1) mkeps_nullable)
       
   633 apply(simp)
       
   634 apply(erule Prf.cases)
       
   635 apply(simp_all)[5]
       
   636 apply(auto)[1]
       
   637 apply(rule Prf.intros)
       
   638 apply(auto)[2]
       
   639 done
       
   640 
       
   641 text {*
       
   642   The string behin the injection value is an added c
       
   643 *}
       
   644 
       
   645 lemma v4:
       
   646   assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
       
   647 using assms
       
   648 apply(induct arbitrary: v rule: der.induct)
       
   649 apply(simp)
       
   650 apply(erule Prf.cases)
       
   651 apply(simp_all)[5]
       
   652 apply(simp)
       
   653 apply(erule Prf.cases)
       
   654 apply(simp_all)[5]
       
   655 apply(simp)
       
   656 apply(case_tac "c = c'")
       
   657 apply(simp)
       
   658 apply(auto)[1]
       
   659 apply(erule Prf.cases)
       
   660 apply(simp_all)[5]
       
   661 apply(simp)
       
   662 apply(erule Prf.cases)
       
   663 apply(simp_all)[5]
       
   664 apply(simp)
       
   665 apply(erule Prf.cases)
       
   666 apply(simp_all)[5]
       
   667 apply(simp)
       
   668 apply(case_tac "nullable r1")
       
   669 apply(simp)
       
   670 apply(erule Prf.cases)
       
   671 apply(simp_all)[5]
       
   672 apply(auto)[1]
       
   673 apply(erule Prf.cases)
       
   674 apply(simp_all)[5]
       
   675 apply(auto)[1]
       
   676 apply (metis mkeps_flat)
       
   677 apply(simp)
       
   678 apply(erule Prf.cases)
       
   679 apply(simp_all)[5]
       
   680 done
       
   681 
       
   682 text {*
       
   683   Injection followed by projection is the identity.
       
   684 *}
       
   685 
       
   686 lemma proj_inj_id:
       
   687   assumes "\<turnstile> v : der c r" 
       
   688   shows "projval r c (injval r c v) = v"
       
   689 using assms
       
   690 apply(induct r arbitrary: c v rule: rexp.induct)
       
   691 apply(simp)
       
   692 apply(erule Prf.cases)
       
   693 apply(simp_all)[5]
       
   694 apply(simp)
       
   695 apply(erule Prf.cases)
       
   696 apply(simp_all)[5]
       
   697 apply(simp)
       
   698 apply(case_tac "c = char")
       
   699 apply(simp)
       
   700 apply(erule Prf.cases)
       
   701 apply(simp_all)[5]
       
   702 apply(simp)
       
   703 apply(erule Prf.cases)
       
   704 apply(simp_all)[5]
       
   705 defer
       
   706 apply(simp)
       
   707 apply(erule Prf.cases)
       
   708 apply(simp_all)[5]
       
   709 apply(simp)
       
   710 apply(case_tac "nullable rexp1")
       
   711 apply(simp)
       
   712 apply(erule Prf.cases)
       
   713 apply(simp_all)[5]
       
   714 apply(auto)[1]
       
   715 apply(erule Prf.cases)
       
   716 apply(simp_all)[5]
       
   717 apply(auto)[1]
       
   718 apply (metis list.distinct(1) v4)
       
   719 apply(auto)[1]
       
   720 apply (metis mkeps_flat)
       
   721 apply(auto)
       
   722 apply(erule Prf.cases)
       
   723 apply(simp_all)[5]
       
   724 apply(auto)[1]
       
   725 apply(simp add: v4)
       
   726 done
       
   727 
       
   728 lemma "L r \<noteq> {} \<Longrightarrow> \<exists>v. POSIX3 v r"
       
   729 apply(induct r)
       
   730 apply(simp)
       
   731 apply(simp add: POSIX3_def)
       
   732 apply(rule_tac x="Void" in exI)
       
   733 apply(auto)[1]
       
   734 apply (metis Prf.intros(4))
       
   735 apply (metis POSIX3_def flat.simps(1) mkeps.simps(1) mkeps_POSIX3 nullable.simps(2) order_refl)
       
   736 apply(simp add: POSIX3_def)
       
   737 apply(rule_tac x="Char char" in exI)
       
   738 apply(auto)[1]
       
   739 apply (metis Prf.intros(5))
       
   740 apply(erule Prf.cases)
       
   741 apply(simp_all)[5]
       
   742 apply (metis ValOrd.intros(8))
       
   743 apply(simp add: Sequ_def)
       
   744 apply(auto)[1]
       
   745 apply(drule meta_mp)
       
   746 apply(auto)[2]
       
   747 apply(drule meta_mp)
       
   748 apply(auto)[2]
       
   749 apply(rule_tac x="Seq v va" in exI)
       
   750 apply(simp (no_asm) add: POSIX3_def)
       
   751 apply(auto)[1]
       
   752 apply (metis POSIX3_def Prf.intros(1))
       
   753 apply(erule Prf.cases)
       
   754 apply(simp_all)[5]
       
   755 apply(clarify)
       
   756 apply(case_tac "v  \<succ>r1a v1")
       
   757 apply(rule ValOrd.intros(2))
       
   758 apply(simp)
       
   759 apply(case_tac "v = v1")
       
   760 apply(rule ValOrd.intros(1))
       
   761 apply(simp)
       
   762 apply(simp)
       
   763 apply (metis ValOrd_refl)
       
   764 apply(simp add: POSIX3_def)
       
   765 
       
   766 
       
   767 lemma "\<exists>v. POSIX v r"
       
   768 apply(induct r)
       
   769 apply(rule exI)
       
   770 apply(simp add: POSIX_def)
       
   771 apply (metis (full_types) Prf_flat_L der.simps(1) der.simps(2) der.simps(3) flat.simps(1) nullable.simps(1) nullable_correctness proj_inj_id projval.simps(1) v3 v4)
       
   772 apply(rule_tac x = "Void" in exI)
       
   773 apply(simp add: POSIX_def)
       
   774 apply (metis POSIX_def flat.simps(1) mkeps.simps(1) mkeps_POSIX nullable.simps(2))
       
   775 apply(rule_tac x = "Char char" in exI)
       
   776 apply(simp add: POSIX_def)
       
   777 apply(auto) [1]
       
   778 apply(erule Prf.cases)
       
   779 apply(simp_all) [5]
       
   780 apply (metis ValOrd.intros(8))
       
   781 defer
       
   782 apply(auto)
       
   783 apply (metis POSIX_ALT_I1)
       
   784 (* maybe it is too early to instantiate this existential quantifier *)
       
   785 (* potentially this is the wrong POSIX value *)
       
   786 apply(case_tac "r1 = NULL")
       
   787 apply(simp add: POSIX_def)
       
   788 apply(auto)[1]
       
   789 apply (metis L.simps(1) L.simps(4) Prf_flat_L mkeps_flat nullable.simps(1) nullable.simps(2) nullable_correctness seq_null(2))
       
   790 apply(case_tac "r1 = EMPTY")
       
   791 apply(rule_tac x = "Seq Void va" in exI )
       
   792 apply(simp (no_asm) add: POSIX_def)
       
   793 apply(auto)
       
   794 apply(erule Prf.cases)
       
   795 apply(simp_all)
       
   796 apply(auto)[1]
       
   797 apply(erule Prf.cases)
       
   798 apply(simp_all)
       
   799 apply(rule ValOrd.intros(2))
       
   800 apply(rule ValOrd.intros)
       
   801 apply(case_tac "\<exists>c. r1 = CHAR c")
       
   802 apply(auto)
       
   803 apply(rule_tac x = "Seq (Char c) va" in exI )
       
   804 apply(simp (no_asm) add: POSIX_def)
       
   805 apply(auto)
       
   806 apply(erule Prf.cases)
       
   807 apply(simp_all)
       
   808 apply(auto)[1]
       
   809 apply(erule Prf.cases)
       
   810 apply(simp_all)
       
   811 apply(auto)[1]
       
   812 apply(rule ValOrd.intros(2))
       
   813 apply(rule ValOrd.intros)
       
   814 apply(case_tac "\<exists>r1a r1b. r1 = ALT r1a r1b")
       
   815 apply(auto)
       
   816 oops (* not sure if this can be proved by induction *)
       
   817 
       
   818 text {* 
       
   819 
       
   820   HERE: Crucial lemma that does not go through in the sequence case. 
       
   821 
       
   822 *}
       
   823 lemma v5:
       
   824   assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
       
   825   shows "POSIX (injval r c v) r"
       
   826 using assms
       
   827 apply(induct arbitrary: v rule: der.induct)
       
   828 apply(simp)
       
   829 apply(erule Prf.cases)
       
   830 apply(simp_all)[5]
       
   831 apply(simp)
       
   832 apply(erule Prf.cases)
       
   833 apply(simp_all)[5]
       
   834 apply(simp)
       
   835 apply(case_tac "c = c'")
       
   836 apply(auto simp add: POSIX_def)[1]
       
   837 apply(erule Prf.cases)
       
   838 apply(simp_all)[5]
       
   839 apply(erule Prf.cases)
       
   840 apply(simp_all)[5]
       
   841 using ValOrd.simps apply blast
       
   842 apply(auto)
       
   843 apply(erule Prf.cases)
       
   844 apply(simp_all)[5]
       
   845 (* base cases done *)
       
   846 (* ALT case *)
       
   847 apply(erule Prf.cases)
       
   848 apply(simp_all)[5]
       
   849 using POSIX_ALT POSIX_ALT_I1 apply blast
       
   850 apply(clarify)
       
   851 apply(subgoal_tac "POSIX v2 (der c r2)")
       
   852 prefer 2
       
   853 apply(auto simp add: POSIX_def)[1]
       
   854 apply (metis POSIX_ALT1a POSIX_def flat.simps(4))
       
   855 apply(frule POSIX_ALT1a)
       
   856 apply(drule POSIX_ALT1b)
       
   857 apply(rule POSIX_ALT_I2)
       
   858 apply(rotate_tac 1)
       
   859 apply(drule_tac x="v2" in meta_spec)
       
   860 apply(simp)
       
   861 apply(subgoal_tac "\<turnstile> Right (injval r2 c v2) : (ALT r1 r2)")
       
   862 prefer 2
       
   863 apply (metis Prf.intros(3) v3)
       
   864 
       
   865 apply auto[1]
       
   866 apply(subst v4)
       
   867 apply(auto)[2]
       
   868 apply(subst (asm) (4) POSIX_def)
       
   869 apply(subst (asm) v4)
       
   870 apply(drule_tac x="v2" in meta_spec)
       
   871 apply(simp)
       
   872 
       
   873 apply(auto)[2]
       
   874 
       
   875 thm POSIX_ALT_I2
       
   876 apply(rule POSIX_ALT_I2)
       
   877 
       
   878 apply(rule ccontr)
       
   879 apply(auto simp add: POSIX_def)[1]
       
   880 
       
   881 apply(rule allI)
       
   882 apply(rule impI)
       
   883 apply(erule conjE)
       
   884 thm POSIX_ALT_I2
       
   885 apply(frule POSIX_ALT1a)
       
   886 apply(drule POSIX_ALT1b)
       
   887 apply(rule POSIX_ALT_I2)
       
   888 apply auto[1]
       
   889 apply(subst v4)
       
   890 apply(auto)[2]
       
   891 apply(rotate_tac 1)
       
   892 apply(drule_tac x="v2" in meta_spec)
       
   893 apply(simp)
       
   894 apply(subst (asm) (4) POSIX_def)
       
   895 apply(subst (asm) v4)
       
   896 apply(auto)[2]
       
   897 (* stuck in the ALT case *)