thys/Re1.thy~
changeset 26 51444f205b5b
parent 25 a2a7f65f538a
child 27 378077bab5d2
equal deleted inserted replaced
25:a2a7f65f538a 26:51444f205b5b
     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 
       
    46 section {* Values *}
       
    47 
       
    48 datatype val = 
       
    49   Void
       
    50 | Char char
       
    51 | Seq val val
       
    52 | Right val
       
    53 | Left val
       
    54 
       
    55 section {* Relation between values and regular expressions *}
       
    56 
       
    57 inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
       
    58 where
       
    59  "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
       
    60 | "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
       
    61 | "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
       
    62 | "\<turnstile> Void : EMPTY"
       
    63 | "\<turnstile> Char c : CHAR c"
       
    64 
       
    65 section {* The string behind a value *}
       
    66 
       
    67 fun flat :: "val \<Rightarrow> string"
       
    68 where
       
    69   "flat(Void) = []"
       
    70 | "flat(Char c) = [c]"
       
    71 | "flat(Left v) = flat(v)"
       
    72 | "flat(Right v) = flat(v)"
       
    73 | "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
       
    74 
       
    75 
       
    76 lemma Prf_flat_L:
       
    77   assumes "\<turnstile> v : r" shows "flat v \<in> L r"
       
    78 using assms
       
    79 apply(induct)
       
    80 apply(auto simp add: Sequ_def)
       
    81 done
       
    82 
       
    83 lemma L_flat_Prf:
       
    84   "L(r) = {flat v | v. \<turnstile> v : r}"
       
    85 apply(induct r)
       
    86 apply(auto dest: Prf_flat_L simp add: Sequ_def)
       
    87 apply (metis Prf.intros(4) flat.simps(1))
       
    88 apply (metis Prf.intros(5) flat.simps(2))
       
    89 apply (metis Prf.intros(1) flat.simps(5))
       
    90 apply (metis Prf.intros(2) flat.simps(3))
       
    91 apply (metis Prf.intros(3) flat.simps(4))
       
    92 apply(erule Prf.cases)
       
    93 apply(auto)
       
    94 done
       
    95 
       
    96 section {* Ordering of values *}
       
    97 
       
    98 inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
       
    99 where
       
   100   "\<lbrakk>v1 = v1'; v2 \<succ>r2 v2'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
       
   101 | "v1  \<succ>r1 v1' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
       
   102 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
       
   103 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
       
   104 | "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
       
   105 | "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
       
   106 | "Void \<succ>EMPTY Void"
       
   107 | "(Char c) \<succ>(CHAR c) (Char c)"
       
   108 
       
   109 (*
       
   110 lemma
       
   111   assumes "r = SEQ (ALT EMPTY EMPTY) (ALT EMPTY (CHAR c))"
       
   112   shows "(Seq (Left Void) (Right (Char c))) \<succ>r (Seq (Left Void) (Left Void))"
       
   113 using assms
       
   114 apply(simp)
       
   115 apply(rule ValOrd.intros)
       
   116 apply(rule ValOrd.intros)
       
   117 apply(rule ValOrd.intros)
       
   118 apply(rule ValOrd.intros)
       
   119 apply(simp)
       
   120 done
       
   121 *)
       
   122 
       
   123 section {* Posix definition *}
       
   124 
       
   125 definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
       
   126 where
       
   127   "POSIX v r \<equiv> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v')"
       
   128 
       
   129 (*
       
   130 an alternative definition: might cause problems
       
   131 with theorem mkeps_POSIX
       
   132 *)
       
   133 
       
   134 definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
       
   135 where
       
   136   "POSIX2 v r \<equiv> \<turnstile> v : r \<and> (\<forall>v'. \<turnstile> v' : r \<longrightarrow> v \<succ>r v')"
       
   137 
       
   138 definition POSIX3 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
       
   139 where
       
   140   "POSIX3 v r \<equiv> \<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v \<ge> flat v') \<longrightarrow> v \<succ>r v')"
       
   141 
       
   142 
       
   143 (*
       
   144 lemma POSIX_SEQ:
       
   145   assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
       
   146   shows "POSIX v1 r1 \<and> POSIX v2 r2"
       
   147 using assms
       
   148 unfolding POSIX_def
       
   149 apply(auto)
       
   150 apply(drule_tac x="Seq v' v2" in spec)
       
   151 apply(simp)
       
   152 apply (smt Prf.intros(1) ValOrd.simps assms(3) rexp.inject(2) val.distinct(15) val.distinct(17) val.distinct(3) val.distinct(9) val.inject(2))
       
   153 apply(drule_tac x="Seq v1 v'" in spec)
       
   154 apply(simp)
       
   155 by (smt Prf.intros(1) ValOrd.simps rexp.inject(2) val.distinct(15) val.distinct(17) val.distinct(3) val.distinct(9) val.inject(2))
       
   156 *)
       
   157 
       
   158 (*
       
   159 lemma POSIX_SEQ_I:
       
   160   assumes "POSIX v1 r1" "POSIX v2 r2" 
       
   161   shows "POSIX (Seq v1 v2) (SEQ r1 r2)" 
       
   162 using assms
       
   163 unfolding POSIX_def
       
   164 apply(auto)
       
   165 apply(rotate_tac 2)
       
   166 apply(erule Prf.cases)
       
   167 apply(simp_all)[5]
       
   168 apply(auto)[1]
       
   169 apply(rule ValOrd.intros)
       
   170 
       
   171 apply(auto)
       
   172 done
       
   173 *)
       
   174 
       
   175 
       
   176 
       
   177 
       
   178 lemma POSIX_ALT2:
       
   179   assumes "POSIX (Left v1) (ALT r1 r2)"
       
   180   shows "POSIX v1 r1"
       
   181 using assms
       
   182 unfolding POSIX_def
       
   183 apply(auto)
       
   184 apply(drule_tac x="Left v'" in spec)
       
   185 apply(simp)
       
   186 apply(drule mp)
       
   187 apply(rule Prf.intros)
       
   188 apply(auto)
       
   189 apply(erule ValOrd.cases)
       
   190 apply(simp_all)
       
   191 done
       
   192 
       
   193 lemma POSIX2_ALT:
       
   194   assumes "POSIX2 (Left v1) (ALT r1 r2)"
       
   195   shows "POSIX2 v1 r1"
       
   196 using assms
       
   197 unfolding POSIX2_def
       
   198 apply(auto)
       
   199 
       
   200 done
       
   201 
       
   202 lemma POSIX_ALT:
       
   203   assumes "POSIX (Left v1) (ALT r1 r2)"
       
   204   shows "POSIX v1 r1"
       
   205 using assms
       
   206 unfolding POSIX_def
       
   207 apply(auto)
       
   208 apply(drule_tac x="Left v'" in spec)
       
   209 apply(simp)
       
   210 apply(drule mp)
       
   211 apply(rule Prf.intros)
       
   212 apply(auto)
       
   213 apply(erule ValOrd.cases)
       
   214 apply(simp_all)
       
   215 done
       
   216 
       
   217 lemma POSIX2_ALT:
       
   218   assumes "POSIX2 (Left v1) (ALT r1 r2)"
       
   219   shows "POSIX2 v1 r1"
       
   220 using assms
       
   221 apply(simp add: POSIX2_def)
       
   222 apply(auto)
       
   223 apply(erule Prf.cases)
       
   224 apply(simp_all)[5]
       
   225 apply(drule_tac x="Left v'" in spec)
       
   226 apply(drule mp)
       
   227 apply(rule Prf.intros)
       
   228 apply(auto)
       
   229 apply(erule ValOrd.cases)
       
   230 apply(simp_all)
       
   231 done
       
   232 
       
   233 
       
   234 lemma POSIX_ALT1a:
       
   235   assumes "POSIX (Right v2) (ALT r1 r2)"
       
   236   shows "POSIX v2 r2"
       
   237 using assms
       
   238 unfolding POSIX_def
       
   239 apply(auto)
       
   240 apply(drule_tac x="Right v'" in spec)
       
   241 apply(simp)
       
   242 apply(drule mp)
       
   243 apply(rule Prf.intros)
       
   244 apply(auto)
       
   245 apply(erule ValOrd.cases)
       
   246 apply(simp_all)
       
   247 done
       
   248 
       
   249 lemma POSIX2_ALT1a:
       
   250   assumes "POSIX2 (Right v2) (ALT r1 r2)"
       
   251   shows "POSIX2 v2 r2"
       
   252 using assms
       
   253 unfolding POSIX2_def
       
   254 apply(auto)
       
   255 apply(erule Prf.cases)
       
   256 apply(simp_all)[5]
       
   257 apply(drule_tac x="Right v'" in spec)
       
   258 apply(drule mp)
       
   259 apply(rule Prf.intros)
       
   260 apply(auto)
       
   261 apply(erule ValOrd.cases)
       
   262 apply(simp_all)
       
   263 done
       
   264 
       
   265 
       
   266 lemma POSIX_ALT1b:
       
   267   assumes "POSIX (Right v2) (ALT r1 r2)"
       
   268   shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
       
   269 using assms
       
   270 apply(drule_tac POSIX_ALT1a)
       
   271 unfolding POSIX_def
       
   272 apply(auto)
       
   273 done
       
   274 
       
   275 lemma POSIX_ALT_I1:
       
   276   assumes "POSIX v1 r1" 
       
   277   shows "POSIX (Left v1) (ALT r1 r2)"
       
   278 using assms
       
   279 unfolding POSIX_def
       
   280 apply(auto)
       
   281 apply(rotate_tac 3)
       
   282 apply(erule Prf.cases)
       
   283 apply(simp_all)[5]
       
   284 apply(auto)
       
   285 apply(rule ValOrd.intros)
       
   286 apply(auto)
       
   287 apply(rule ValOrd.intros)
       
   288 by simp
       
   289 
       
   290 lemma POSIX2_ALT_I1:
       
   291   assumes "POSIX2 v1 r1" 
       
   292   shows "POSIX2 (Left v1) (ALT r1 r2)"
       
   293 using assms
       
   294 unfolding POSIX2_def
       
   295 apply(auto)
       
   296 apply(rule Prf.intros)
       
   297 apply(simp)
       
   298 apply(rotate_tac 2)
       
   299 apply(erule Prf.cases)
       
   300 apply(simp_all)[5]
       
   301 apply(auto)
       
   302 apply(rule ValOrd.intros)
       
   303 apply(auto)
       
   304 apply(rule ValOrd.intros)
       
   305 
       
   306 
       
   307 by simp
       
   308 
       
   309 lemma POSIX_ALT_I2:
       
   310   assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
       
   311   shows "POSIX (Right v2) (ALT r1 r2)"
       
   312 using assms
       
   313 unfolding POSIX_def
       
   314 apply(auto)
       
   315 apply(rotate_tac 3)
       
   316 apply(erule Prf.cases)
       
   317 apply(simp_all)[5]
       
   318 apply(auto)
       
   319 apply(rule ValOrd.intros)
       
   320 apply metis
       
   321 done
       
   322 
       
   323 
       
   324 section {* The ordering is reflexive *}
       
   325 
       
   326 lemma ValOrd_refl:
       
   327   assumes "\<turnstile> v : r"
       
   328   shows "v \<succ>r v"
       
   329 using assms
       
   330 apply(induct)
       
   331 apply(auto intro: ValOrd.intros)
       
   332 done
       
   333 
       
   334 
       
   335 section {* The Matcher *}
       
   336 
       
   337 fun
       
   338  nullable :: "rexp \<Rightarrow> bool"
       
   339 where
       
   340   "nullable (NULL) = False"
       
   341 | "nullable (EMPTY) = True"
       
   342 | "nullable (CHAR c) = False"
       
   343 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
       
   344 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
       
   345 
       
   346 lemma nullable_correctness:
       
   347   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
       
   348 apply (induct r) 
       
   349 apply(auto simp add: Sequ_def) 
       
   350 done
       
   351 
       
   352 fun mkeps :: "rexp \<Rightarrow> val"
       
   353 where
       
   354   "mkeps(EMPTY) = Void"
       
   355 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
       
   356 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
       
   357 
       
   358 lemma mkeps_nullable:
       
   359   assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
       
   360 using assms
       
   361 apply(induct rule: nullable.induct)
       
   362 apply(auto intro: Prf.intros)
       
   363 done
       
   364 
       
   365 lemma mkeps_flat:
       
   366   assumes "nullable(r)" shows "flat (mkeps r) = []"
       
   367 using assms
       
   368 apply(induct rule: nullable.induct)
       
   369 apply(auto)
       
   370 done
       
   371 
       
   372 text {*
       
   373   The value mkeps returns is always the correct POSIX
       
   374   value.
       
   375 *}
       
   376 
       
   377 lemma mkeps_POSIX2:
       
   378   assumes "nullable r"
       
   379   shows "POSIX2 (mkeps r) r"
       
   380 using assms
       
   381 apply(induct r)
       
   382 apply(auto)[1]
       
   383 apply(simp add: POSIX2_def)
       
   384 
       
   385 lemma mkeps_POSIX:
       
   386   assumes "nullable r"
       
   387   shows "POSIX (mkeps r) r"
       
   388 using assms
       
   389 apply(induct r)
       
   390 apply(auto)[1]
       
   391 apply(simp add: POSIX_def)
       
   392 apply(auto)[1]
       
   393 apply(erule Prf.cases)
       
   394 apply(simp_all)[5]
       
   395 apply (metis ValOrd.intros)
       
   396 apply(simp add: POSIX_def)
       
   397 apply(auto)[1]
       
   398 apply(simp add: POSIX_def)
       
   399 apply(auto)[1]
       
   400 apply(erule Prf.cases)
       
   401 apply(simp_all)[5]
       
   402 apply(auto)
       
   403 apply (simp add: ValOrd.intros(2) mkeps_flat)
       
   404 apply(simp add: POSIX_def)
       
   405 apply(auto)[1]
       
   406 apply(erule Prf.cases)
       
   407 apply(simp_all)[5]
       
   408 apply(auto)
       
   409 apply (simp add: ValOrd.intros(6))
       
   410 apply (simp add: ValOrd.intros(3))
       
   411 apply(simp add: POSIX_def)
       
   412 apply(auto)[1]
       
   413 apply(erule Prf.cases)
       
   414 apply(simp_all)[5]
       
   415 apply(auto)
       
   416 apply (simp add: ValOrd.intros(6))
       
   417 apply (simp add: ValOrd.intros(3))
       
   418 apply(simp add: POSIX_def)
       
   419 apply(auto)[1]
       
   420 apply(erule Prf.cases)
       
   421 apply(simp_all)[5]
       
   422 apply(auto)
       
   423 apply (metis Prf_flat_L mkeps_flat nullable_correctness)
       
   424 by (simp add: ValOrd.intros(5))
       
   425 
       
   426 
       
   427 lemma mkeps_POSIX2:
       
   428   assumes "nullable r"
       
   429   shows "POSIX2 (mkeps r) r"
       
   430 using assms
       
   431 apply(induct r)
       
   432 apply(simp)
       
   433 apply(simp)
       
   434 apply(simp add: POSIX2_def)
       
   435 apply(rule conjI)
       
   436 apply(rule Prf.intros)
       
   437 apply(auto)[1]
       
   438 apply(erule Prf.cases)
       
   439 apply(simp_all)[5]
       
   440 apply(rule ValOrd.intros)
       
   441 apply(simp)
       
   442 apply(simp)
       
   443 apply(simp add: POSIX2_def)
       
   444 apply(rule conjI)
       
   445 apply(rule Prf.intros)
       
   446 apply(simp add: mkeps_nullable)
       
   447 apply(simp add: mkeps_nullable)
       
   448 apply(auto)[1]
       
   449 apply(rotate_tac 6)
       
   450 apply(erule Prf.cases)
       
   451 apply(simp_all)[5]
       
   452 apply(rule ValOrd.intros(2))
       
   453 apply(simp)
       
   454 apply(simp only: nullable.simps)
       
   455 apply(erule disjE)
       
   456 apply(simp)
       
   457 thm POSIX2_ALT1a
       
   458 apply(rule POSIX2_ALT)
       
   459 apply(simp add: POSIX2_def)
       
   460 apply(rule conjI)
       
   461 apply(rule Prf.intros)
       
   462 apply(simp add: mkeps_nullable)
       
   463 apply(auto)[1]
       
   464 apply(rotate_tac 4)
       
   465 apply(erule Prf.cases)
       
   466 apply(simp_all)[5]
       
   467 apply(rule ValOrd.intros)
       
   468 apply(simp)
       
   469 apply(rule ValOrd.intros)
       
   470 
       
   471 
       
   472 section {* Derivatives *}
       
   473 
       
   474 fun
       
   475  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
       
   476 where
       
   477   "der c (NULL) = NULL"
       
   478 | "der c (EMPTY) = NULL"
       
   479 | "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
       
   480 | "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
       
   481 | "der c (SEQ r1 r2) = 
       
   482      (if nullable r1
       
   483       then ALT (SEQ (der c r1) r2) (der c r2)
       
   484       else SEQ (der c r1) r2)"
       
   485 
       
   486 fun 
       
   487  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
       
   488 where
       
   489   "ders [] r = r"
       
   490 | "ders (c # s) r = ders s (der c r)"
       
   491 
       
   492 section {* Injection function *}
       
   493 
       
   494 fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
       
   495 where
       
   496   "injval (CHAR d) c Void = Char d"
       
   497 | "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
       
   498 | "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
       
   499 | "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
       
   500 | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
       
   501 | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
       
   502 
       
   503 section {* Projection function *}
       
   504 
       
   505 fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
       
   506 where
       
   507   "projval (CHAR d) c _ = Void"
       
   508 | "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
       
   509 | "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
       
   510 | "projval (SEQ r1 r2) c (Seq v1 v2) = 
       
   511      (if flat v1 = [] then Right(projval r2 c v2) 
       
   512       else if nullable r1 then Left (Seq (projval r1 c v1) v2)
       
   513                           else Seq (projval r1 c v1) v2)"
       
   514 
       
   515 text {*
       
   516   Injection value is related to r
       
   517 *}
       
   518 
       
   519 lemma v3:
       
   520   assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
       
   521 using assms
       
   522 apply(induct arbitrary: v rule: der.induct)
       
   523 apply(simp)
       
   524 apply(erule Prf.cases)
       
   525 apply(simp_all)[5]
       
   526 apply(erule Prf.cases)
       
   527 apply(simp_all)[5]
       
   528 apply(case_tac "c = c'")
       
   529 apply(simp)
       
   530 apply(erule Prf.cases)
       
   531 apply(simp_all)[5]
       
   532 apply (metis Prf.intros(5))
       
   533 apply(erule Prf.cases)
       
   534 apply(simp_all)[5]
       
   535 apply(erule Prf.cases)
       
   536 apply(simp_all)[5]
       
   537 apply (metis Prf.intros(2))
       
   538 apply (metis Prf.intros(3))
       
   539 apply(simp)
       
   540 apply(case_tac "nullable r1")
       
   541 apply(simp)
       
   542 apply(erule Prf.cases)
       
   543 apply(simp_all)[5]
       
   544 apply(auto)[1]
       
   545 apply(erule Prf.cases)
       
   546 apply(simp_all)[5]
       
   547 apply(auto)[1]
       
   548 apply (metis Prf.intros(1))
       
   549 apply(auto)[1]
       
   550 apply (metis Prf.intros(1) mkeps_nullable)
       
   551 apply(simp)
       
   552 apply(erule Prf.cases)
       
   553 apply(simp_all)[5]
       
   554 apply(auto)[1]
       
   555 apply(rule Prf.intros)
       
   556 apply(auto)[2]
       
   557 done
       
   558 
       
   559 text {*
       
   560   The string behin the injection value is an added c
       
   561 *}
       
   562 
       
   563 lemma v4:
       
   564   assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
       
   565 using assms
       
   566 apply(induct arbitrary: v rule: der.induct)
       
   567 apply(simp)
       
   568 apply(erule Prf.cases)
       
   569 apply(simp_all)[5]
       
   570 apply(simp)
       
   571 apply(erule Prf.cases)
       
   572 apply(simp_all)[5]
       
   573 apply(simp)
       
   574 apply(case_tac "c = c'")
       
   575 apply(simp)
       
   576 apply(auto)[1]
       
   577 apply(erule Prf.cases)
       
   578 apply(simp_all)[5]
       
   579 apply(simp)
       
   580 apply(erule Prf.cases)
       
   581 apply(simp_all)[5]
       
   582 apply(simp)
       
   583 apply(erule Prf.cases)
       
   584 apply(simp_all)[5]
       
   585 apply(simp)
       
   586 apply(case_tac "nullable r1")
       
   587 apply(simp)
       
   588 apply(erule Prf.cases)
       
   589 apply(simp_all)[5]
       
   590 apply(auto)[1]
       
   591 apply(erule Prf.cases)
       
   592 apply(simp_all)[5]
       
   593 apply(auto)[1]
       
   594 apply (metis mkeps_flat)
       
   595 apply(simp)
       
   596 apply(erule Prf.cases)
       
   597 apply(simp_all)[5]
       
   598 done
       
   599 
       
   600 text {*
       
   601   Injection followed by projection is the identity.
       
   602 *}
       
   603 
       
   604 lemma proj_inj_id:
       
   605   assumes "\<turnstile> v : der c r" 
       
   606   shows "projval r c (injval r c v) = v"
       
   607 using assms
       
   608 apply(induct r arbitrary: c v rule: rexp.induct)
       
   609 apply(simp)
       
   610 apply(erule Prf.cases)
       
   611 apply(simp_all)[5]
       
   612 apply(simp)
       
   613 apply(erule Prf.cases)
       
   614 apply(simp_all)[5]
       
   615 apply(simp)
       
   616 apply(case_tac "c = char")
       
   617 apply(simp)
       
   618 apply(erule Prf.cases)
       
   619 apply(simp_all)[5]
       
   620 apply(simp)
       
   621 apply(erule Prf.cases)
       
   622 apply(simp_all)[5]
       
   623 defer
       
   624 apply(simp)
       
   625 apply(erule Prf.cases)
       
   626 apply(simp_all)[5]
       
   627 apply(simp)
       
   628 apply(case_tac "nullable rexp1")
       
   629 apply(simp)
       
   630 apply(erule Prf.cases)
       
   631 apply(simp_all)[5]
       
   632 apply(auto)[1]
       
   633 apply(erule Prf.cases)
       
   634 apply(simp_all)[5]
       
   635 apply(auto)[1]
       
   636 apply (metis list.distinct(1) v4)
       
   637 apply(auto)[1]
       
   638 apply (metis mkeps_flat)
       
   639 apply(auto)
       
   640 apply(erule Prf.cases)
       
   641 apply(simp_all)[5]
       
   642 apply(auto)[1]
       
   643 apply(simp add: v4)
       
   644 done
       
   645 
       
   646 lemma "\<exists>v. POSIX v r"
       
   647 apply(induct r)
       
   648 apply(rule exI)
       
   649 apply(simp add: POSIX_def)
       
   650 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)
       
   651 apply(rule_tac x = "Void" in exI)
       
   652 apply(simp add: POSIX_def)
       
   653 apply (metis POSIX_def flat.simps(1) mkeps.simps(1) mkeps_POSIX nullable.simps(2))
       
   654 apply(rule_tac x = "Char char" in exI)
       
   655 apply(simp add: POSIX_def)
       
   656 apply(auto) [1]
       
   657 apply(erule Prf.cases)
       
   658 apply(simp_all) [5]
       
   659 apply (metis ValOrd.intros(8))
       
   660 defer
       
   661 apply(auto)
       
   662 apply (metis POSIX_ALT_I1)
       
   663 (* maybe it is too early to instantiate this existential quantifier *)
       
   664 (* potentially this is the wrong POSIX value *)
       
   665 apply(rule_tac x = "Seq v va" in exI )
       
   666 apply(simp (no_asm) add: POSIX_def)
       
   667 apply(auto)
       
   668 apply(erule Prf.cases)
       
   669 apply(simp_all)
       
   670 apply(case_tac "v \<succ>r1a v1")
       
   671 apply (metis ValOrd.intros(2))
       
   672 apply(simp add: POSIX_def)
       
   673 apply(case_tac "flat v = flat v1")
       
   674 apply(auto)[1]
       
   675 apply(simp only: append_eq_append_conv2)
       
   676 apply(auto)
       
   677 thm append_eq_append_conv2
       
   678 
       
   679 text {* 
       
   680 
       
   681   HERE: Crucial lemma that does not go through in the sequence case. 
       
   682 
       
   683 *}
       
   684 lemma v5:
       
   685   assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
       
   686   shows "POSIX (injval r c v) r"
       
   687 using assms
       
   688 apply(induct arbitrary: v rule: der.induct)
       
   689 apply(simp)
       
   690 apply(erule Prf.cases)
       
   691 apply(simp_all)[5]
       
   692 apply(simp)
       
   693 apply(erule Prf.cases)
       
   694 apply(simp_all)[5]
       
   695 apply(simp)
       
   696 apply(case_tac "c = c'")
       
   697 apply(auto simp add: POSIX_def)[1]
       
   698 apply(erule Prf.cases)
       
   699 apply(simp_all)[5]
       
   700 apply(erule Prf.cases)
       
   701 apply(simp_all)[5]
       
   702 using ValOrd.simps apply blast
       
   703 apply(auto)
       
   704 apply(erule Prf.cases)
       
   705 apply(simp_all)[5]
       
   706 (* base cases done *)
       
   707 (* ALT case *)
       
   708 apply(erule Prf.cases)
       
   709 apply(simp_all)[5]
       
   710 using POSIX_ALT POSIX_ALT_I1 apply blast
       
   711 apply(clarify)
       
   712 apply(subgoal_tac "POSIX v2 (der c r2)")
       
   713 prefer 2
       
   714 apply(auto simp add: POSIX_def)[1]
       
   715 apply (metis POSIX_ALT1a POSIX_def flat.simps(4))
       
   716 apply(rotate_tac 1)
       
   717 apply(drule_tac x="v2" in meta_spec)
       
   718 apply(simp)
       
   719 apply(subgoal_tac "\<turnstile> Right (injval r2 c v2) : (ALT r1 r2)")
       
   720 prefer 2
       
   721 apply (metis Prf.intros(3) v3)
       
   722 apply(rule ccontr)
       
   723 apply(auto simp add: POSIX_def)[1]
       
   724 
       
   725 apply(rule allI)
       
   726 apply(rule impI)
       
   727 apply(erule conjE)
       
   728 thm POSIX_ALT_I2
       
   729 apply(frule POSIX_ALT1a)
       
   730 apply(drule POSIX_ALT1b)
       
   731 apply(rule POSIX_ALT_I2)
       
   732 apply auto[1]
       
   733 apply(subst v4)
       
   734 apply(auto)[2]
       
   735 apply(rotate_tac 1)
       
   736 apply(drule_tac x="v2" in meta_spec)
       
   737 apply(simp)
       
   738 apply(subst (asm) (4) POSIX_def)
       
   739 apply(subst (asm) v4)
       
   740 apply(auto)[2]
       
   741 (* stuck in the ALT case *)