thys/Re1.thy.orig
changeset 18 8c9349065477
equal deleted inserted replaced
17:a5427713eef4 18:8c9349065477
       
     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 
       
   139 (*
       
   140 lemma POSIX_SEQ:
       
   141   assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
       
   142   shows "POSIX v1 r1 \<and> POSIX v2 r2"
       
   143 using assms
       
   144 unfolding POSIX_def
       
   145 apply(auto)
       
   146 apply(drule_tac x="Seq v' v2" in spec)
       
   147 apply(simp)
       
   148 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))
       
   149 apply(drule_tac x="Seq v1 v'" in spec)
       
   150 apply(simp)
       
   151 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))
       
   152 *)
       
   153 
       
   154 (*
       
   155 lemma POSIX_SEQ_I:
       
   156   assumes "POSIX v1 r1" "POSIX v2 r2" 
       
   157   shows "POSIX (Seq v1 v2) (SEQ r1 r2)" 
       
   158 using assms
       
   159 unfolding POSIX_def
       
   160 apply(auto)
       
   161 apply(rotate_tac 2)
       
   162 apply(erule Prf.cases)
       
   163 apply(simp_all)[5]
       
   164 apply(auto)[1]
       
   165 apply(rule ValOrd.intros)
       
   166 
       
   167 apply(auto)
       
   168 done
       
   169 *)
       
   170 
       
   171 
       
   172 
       
   173 
       
   174 lemma POSIX_ALT2:
       
   175   assumes "POSIX (Left v1) (ALT r1 r2)"
       
   176   shows "POSIX v1 r1"
       
   177 using assms
       
   178 unfolding POSIX_def
       
   179 apply(auto)
       
   180 apply(drule_tac x="Left v'" in spec)
       
   181 apply(simp)
       
   182 apply(drule mp)
       
   183 apply(rule Prf.intros)
       
   184 apply(auto)
       
   185 apply(erule ValOrd.cases)
       
   186 apply(simp_all)
       
   187 done
       
   188 
       
   189 lemma POSIX2_ALT:
       
   190   assumes "POSIX2 (Left v1) (ALT r1 r2)"
       
   191   shows "POSIX2 v1 r1"
       
   192 using assms
       
   193 unfolding POSIX2_def
       
   194 apply(auto)
       
   195 
       
   196 done
       
   197 
       
   198 
       
   199 lemma POSIX_ALT2:
       
   200 lemma POSIX_ALT:
       
   201   assumes "POSIX (Left v1) (ALT r1 r2)"
       
   202   shows "POSIX v1 r1"
       
   203 using assms
       
   204 unfolding POSIX_def
       
   205 apply(auto)
       
   206 apply(drule_tac x="Left v'" in spec)
       
   207 apply(simp)
       
   208 apply(drule mp)
       
   209 apply(rule Prf.intros)
       
   210 apply(auto)
       
   211 apply(erule ValOrd.cases)
       
   212 apply(simp_all)
       
   213 done
       
   214 
       
   215 lemma POSIX_ALT1a:
       
   216   assumes "POSIX (Right v2) (ALT r1 r2)"
       
   217   shows "POSIX v2 r2"
       
   218 using assms
       
   219 unfolding POSIX_def
       
   220 apply(auto)
       
   221 apply(drule_tac x="Right v'" in spec)
       
   222 apply(simp)
       
   223 apply(drule mp)
       
   224 apply(rule Prf.intros)
       
   225 apply(auto)
       
   226 apply(erule ValOrd.cases)
       
   227 apply(simp_all)
       
   228 done
       
   229 
       
   230 
       
   231 lemma POSIX_ALT1b:
       
   232   assumes "POSIX (Right v2) (ALT r1 r2)"
       
   233   shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
       
   234 using assms
       
   235 apply(drule_tac POSIX_ALT1a)
       
   236 unfolding POSIX_def
       
   237 apply(auto)
       
   238 done
       
   239 
       
   240 lemma POSIX_ALT_I1:
       
   241   assumes "POSIX v1 r1" 
       
   242   shows "POSIX (Left v1) (ALT r1 r2)"
       
   243 using assms
       
   244 unfolding POSIX_def
       
   245 apply(auto)
       
   246 apply(rotate_tac 3)
       
   247 apply(erule Prf.cases)
       
   248 apply(simp_all)[5]
       
   249 apply(auto)
       
   250 apply(rule ValOrd.intros)
       
   251 apply(auto)
       
   252 apply(rule ValOrd.intros)
       
   253 by simp
       
   254 
       
   255 lemma POSIX_ALT_I2:
       
   256   assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
       
   257   shows "POSIX (Right v2) (ALT r1 r2)"
       
   258 using assms
       
   259 unfolding POSIX_def
       
   260 apply(auto)
       
   261 apply(rotate_tac 3)
       
   262 apply(erule Prf.cases)
       
   263 apply(simp_all)[5]
       
   264 apply(auto)
       
   265 apply(rule ValOrd.intros)
       
   266 apply metis
       
   267 done
       
   268 
       
   269 
       
   270 section {* The ordering is reflexive *}
       
   271 
       
   272 lemma ValOrd_refl:
       
   273   assumes "\<turnstile> v : r"
       
   274   shows "v \<succ>r v"
       
   275 using assms
       
   276 apply(induct)
       
   277 apply(auto intro: ValOrd.intros)
       
   278 done
       
   279 
       
   280 
       
   281 section {* The Matcher *}
       
   282 
       
   283 fun
       
   284  nullable :: "rexp \<Rightarrow> bool"
       
   285 where
       
   286   "nullable (NULL) = False"
       
   287 | "nullable (EMPTY) = True"
       
   288 | "nullable (CHAR c) = False"
       
   289 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
       
   290 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
       
   291 
       
   292 lemma nullable_correctness:
       
   293   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
       
   294 apply (induct r) 
       
   295 apply(auto simp add: Sequ_def) 
       
   296 done
       
   297 
       
   298 fun mkeps :: "rexp \<Rightarrow> val"
       
   299 where
       
   300   "mkeps(EMPTY) = Void"
       
   301 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
       
   302 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
       
   303 
       
   304 lemma mkeps_nullable:
       
   305   assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
       
   306 using assms
       
   307 apply(induct rule: nullable.induct)
       
   308 apply(auto intro: Prf.intros)
       
   309 done
       
   310 
       
   311 lemma mkeps_flat:
       
   312   assumes "nullable(r)" shows "flat (mkeps r) = []"
       
   313 using assms
       
   314 apply(induct rule: nullable.induct)
       
   315 apply(auto)
       
   316 done
       
   317 
       
   318 text {*
       
   319   The value mkeps returns is always the correct POSIX
       
   320   value.
       
   321 *}
       
   322 
       
   323 lemma mkeps_POSIX2:
       
   324   assumes "nullable r"
       
   325   shows "POSIX2 (mkeps r) r"
       
   326 using assms
       
   327 apply(induct r)
       
   328 apply(auto)[1]
       
   329 apply(simp add: POSIX2_def)
       
   330 
       
   331 lemma mkeps_POSIX:
       
   332   assumes "nullable r"
       
   333   shows "POSIX (mkeps r) r"
       
   334 using assms
       
   335 apply(induct r)
       
   336 apply(auto)[1]
       
   337 apply(simp add: POSIX_def)
       
   338 apply(auto)[1]
       
   339 apply(erule Prf.cases)
       
   340 apply(simp_all)[5]
       
   341 apply (metis ValOrd.intros)
       
   342 apply(simp add: POSIX_def)
       
   343 apply(auto)[1]
       
   344 apply(simp add: POSIX_def)
       
   345 apply(auto)[1]
       
   346 apply(erule Prf.cases)
       
   347 apply(simp_all)[5]
       
   348 apply(auto)
       
   349 apply (simp add: ValOrd.intros(2) mkeps_flat)
       
   350 apply(simp add: POSIX_def)
       
   351 apply(auto)[1]
       
   352 apply(erule Prf.cases)
       
   353 apply(simp_all)[5]
       
   354 apply(auto)
       
   355 apply (simp add: ValOrd.intros(6))
       
   356 apply (simp add: ValOrd.intros(3))
       
   357 apply(simp add: POSIX_def)
       
   358 apply(auto)[1]
       
   359 apply(erule Prf.cases)
       
   360 apply(simp_all)[5]
       
   361 apply(auto)
       
   362 apply (simp add: ValOrd.intros(6))
       
   363 apply (simp add: ValOrd.intros(3))
       
   364 apply(simp add: POSIX_def)
       
   365 apply(auto)[1]
       
   366 apply(erule Prf.cases)
       
   367 apply(simp_all)[5]
       
   368 apply(auto)
       
   369 apply (metis Prf_flat_L mkeps_flat nullable_correctness)
       
   370 by (simp add: ValOrd.intros(5))
       
   371 
       
   372 
       
   373 section {* Derivatives *}
       
   374 
       
   375 fun
       
   376  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
       
   377 where
       
   378   "der c (NULL) = NULL"
       
   379 | "der c (EMPTY) = NULL"
       
   380 | "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
       
   381 | "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
       
   382 | "der c (SEQ r1 r2) = 
       
   383      (if nullable r1
       
   384       then ALT (SEQ (der c r1) r2) (der c r2)
       
   385       else SEQ (der c r1) r2)"
       
   386 
       
   387 fun 
       
   388  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
       
   389 where
       
   390   "ders [] r = r"
       
   391 | "ders (c # s) r = ders s (der c r)"
       
   392 
       
   393 section {* Injection function *}
       
   394 
       
   395 fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
       
   396 where
       
   397   "injval (CHAR d) c Void = Char d"
       
   398 | "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
       
   399 | "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
       
   400 | "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
       
   401 | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
       
   402 | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
       
   403 
       
   404 section {* Projection function *}
       
   405 
       
   406 fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
       
   407 where
       
   408   "projval (CHAR d) c _ = Void"
       
   409 | "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
       
   410 | "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
       
   411 | "projval (SEQ r1 r2) c (Seq v1 v2) = 
       
   412      (if flat v1 = [] then Right(projval r2 c v2) 
       
   413       else if nullable r1 then Left (Seq (projval r1 c v1) v2)
       
   414                           else Seq (projval r1 c v1) v2)"
       
   415 
       
   416 text {*
       
   417   Injection value is related to r
       
   418 *}
       
   419 
       
   420 lemma v3:
       
   421   assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
       
   422 using assms
       
   423 apply(induct arbitrary: v rule: der.induct)
       
   424 apply(simp)
       
   425 apply(erule Prf.cases)
       
   426 apply(simp_all)[5]
       
   427 apply(erule Prf.cases)
       
   428 apply(simp_all)[5]
       
   429 apply(case_tac "c = c'")
       
   430 apply(simp)
       
   431 apply(erule Prf.cases)
       
   432 apply(simp_all)[5]
       
   433 apply (metis Prf.intros(5))
       
   434 apply(erule Prf.cases)
       
   435 apply(simp_all)[5]
       
   436 apply(erule Prf.cases)
       
   437 apply(simp_all)[5]
       
   438 apply (metis Prf.intros(2))
       
   439 apply (metis Prf.intros(3))
       
   440 apply(simp)
       
   441 apply(case_tac "nullable r1")
       
   442 apply(simp)
       
   443 apply(erule Prf.cases)
       
   444 apply(simp_all)[5]
       
   445 apply(auto)[1]
       
   446 apply(erule Prf.cases)
       
   447 apply(simp_all)[5]
       
   448 apply(auto)[1]
       
   449 apply (metis Prf.intros(1))
       
   450 apply(auto)[1]
       
   451 apply (metis Prf.intros(1) mkeps_nullable)
       
   452 apply(simp)
       
   453 apply(erule Prf.cases)
       
   454 apply(simp_all)[5]
       
   455 apply(auto)[1]
       
   456 apply(rule Prf.intros)
       
   457 apply(auto)[2]
       
   458 done
       
   459 
       
   460 text {*
       
   461   The string behin the injection value is an added c
       
   462 *}
       
   463 
       
   464 lemma v4:
       
   465   assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
       
   466 using assms
       
   467 apply(induct arbitrary: v rule: der.induct)
       
   468 apply(simp)
       
   469 apply(erule Prf.cases)
       
   470 apply(simp_all)[5]
       
   471 apply(simp)
       
   472 apply(erule Prf.cases)
       
   473 apply(simp_all)[5]
       
   474 apply(simp)
       
   475 apply(case_tac "c = c'")
       
   476 apply(simp)
       
   477 apply(auto)[1]
       
   478 apply(erule Prf.cases)
       
   479 apply(simp_all)[5]
       
   480 apply(simp)
       
   481 apply(erule Prf.cases)
       
   482 apply(simp_all)[5]
       
   483 apply(simp)
       
   484 apply(erule Prf.cases)
       
   485 apply(simp_all)[5]
       
   486 apply(simp)
       
   487 apply(case_tac "nullable r1")
       
   488 apply(simp)
       
   489 apply(erule Prf.cases)
       
   490 apply(simp_all)[5]
       
   491 apply(auto)[1]
       
   492 apply(erule Prf.cases)
       
   493 apply(simp_all)[5]
       
   494 apply(auto)[1]
       
   495 apply (metis mkeps_flat)
       
   496 apply(simp)
       
   497 apply(erule Prf.cases)
       
   498 apply(simp_all)[5]
       
   499 done
       
   500 
       
   501 text {*
       
   502   Injection followed by projection is the identity.
       
   503 *}
       
   504 
       
   505 lemma proj_inj_id:
       
   506   assumes "\<turnstile> v : der c r" 
       
   507   shows "projval r c (injval r c v) = v"
       
   508 using assms
       
   509 apply(induct r arbitrary: c v rule: rexp.induct)
       
   510 apply(simp)
       
   511 apply(erule Prf.cases)
       
   512 apply(simp_all)[5]
       
   513 apply(simp)
       
   514 apply(erule Prf.cases)
       
   515 apply(simp_all)[5]
       
   516 apply(simp)
       
   517 apply(case_tac "c = char")
       
   518 apply(simp)
       
   519 apply(erule Prf.cases)
       
   520 apply(simp_all)[5]
       
   521 apply(simp)
       
   522 apply(erule Prf.cases)
       
   523 apply(simp_all)[5]
       
   524 defer
       
   525 apply(simp)
       
   526 apply(erule Prf.cases)
       
   527 apply(simp_all)[5]
       
   528 apply(simp)
       
   529 apply(case_tac "nullable rexp1")
       
   530 apply(simp)
       
   531 apply(erule Prf.cases)
       
   532 apply(simp_all)[5]
       
   533 apply(auto)[1]
       
   534 apply(erule Prf.cases)
       
   535 apply(simp_all)[5]
       
   536 apply(auto)[1]
       
   537 apply (metis list.distinct(1) v4)
       
   538 apply(auto)[1]
       
   539 apply (metis mkeps_flat)
       
   540 apply(auto)
       
   541 apply(erule Prf.cases)
       
   542 apply(simp_all)[5]
       
   543 apply(auto)[1]
       
   544 apply(simp add: v4)
       
   545 done
       
   546 
       
   547 lemma "\<exists>v. POSIX v r"
       
   548 apply(induct r)
       
   549 apply(rule exI)
       
   550 apply(simp add: POSIX_def)
       
   551 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)
       
   552 apply(rule_tac x = "Void" in exI)
       
   553 apply(simp add: POSIX_def)
       
   554 apply (metis POSIX_def flat.simps(1) mkeps.simps(1) mkeps_POSIX nullable.simps(2))
       
   555 apply(rule_tac x = "Char char" in exI)
       
   556 apply(simp add: POSIX_def)
       
   557 apply(auto) [1]
       
   558 apply(erule Prf.cases)
       
   559 apply(simp_all) [5]
       
   560 apply (metis ValOrd.intros(8))
       
   561 defer
       
   562 apply(auto)
       
   563 apply (metis POSIX_ALT_I1)
       
   564 (* maybe it is too early to instantiate this existential quantifier *)
       
   565 (* potentially this is the wrong POSIX value *)
       
   566 apply(rule_tac x = "Seq v va" in exI )
       
   567 apply(simp (no_asm) add: POSIX_def)
       
   568 apply(auto)
       
   569 apply(erule Prf.cases)
       
   570 apply(simp_all)
       
   571 apply(case_tac "v \<succ>r1a v1")
       
   572 apply (metis ValOrd.intros(2))
       
   573 apply(simp add: POSIX_def)
       
   574 apply(case_tac "flat v = flat v1")
       
   575 apply(auto)[1]
       
   576 apply(simp only: append_eq_append_conv2)
       
   577 apply(auto)
       
   578 thm append_eq_append_conv2
       
   579 
       
   580 text {* 
       
   581 
       
   582   HERE: Crucial lemma that does not go through in the sequence case. 
       
   583 
       
   584 *}
       
   585 lemma v5:
       
   586   assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
       
   587   shows "POSIX (injval r c v) r"
       
   588 using assms
       
   589 apply(induct arbitrary: v rule: der.induct)
       
   590 apply(simp)
       
   591 apply(erule Prf.cases)
       
   592 apply(simp_all)[5]
       
   593 apply(simp)
       
   594 apply(erule Prf.cases)
       
   595 apply(simp_all)[5]
       
   596 apply(simp)
       
   597 apply(case_tac "c = c'")
       
   598 apply(auto simp add: POSIX_def)[1]
       
   599 apply(erule Prf.cases)
       
   600 apply(simp_all)[5]
       
   601 apply(erule Prf.cases)
       
   602 apply(simp_all)[5]
       
   603 using ValOrd.simps apply blast
       
   604 apply(auto)
       
   605 apply(erule Prf.cases)
       
   606 apply(simp_all)[5]
       
   607 (* base cases done *)
       
   608 (* ALT case *)
       
   609 apply(erule Prf.cases)
       
   610 apply(simp_all)[5]
       
   611 using POSIX_ALT POSIX_ALT_I1 apply blast
       
   612 apply(clarify)
       
   613 apply(subgoal_tac "POSIX v2 (der c r2)")
       
   614 prefer 2
       
   615 apply(auto simp add: POSIX_def)[1]
       
   616 apply (metis POSIX_ALT1a POSIX_def flat.simps(4))
       
   617 apply(rotate_tac 1)
       
   618 apply(drule_tac x="v2" in meta_spec)
       
   619 apply(simp)
       
   620 apply(subgoal_tac "\<turnstile> Right (injval r2 c v2) : (ALT r1 r2)")
       
   621 prefer 2
       
   622 apply (metis Prf.intros(3) v3)
       
   623 apply(rule ccontr)
       
   624 apply(auto simp add: POSIX_def)[1]
       
   625 
       
   626 apply(rule allI)
       
   627 apply(rule impI)
       
   628 apply(erule conjE)
       
   629 thm POSIX_ALT_I2
       
   630 apply(frule POSIX_ALT1a)
       
   631 apply(drule POSIX_ALT1b)
       
   632 apply(rule POSIX_ALT_I2)
       
   633 apply auto[1]
       
   634 apply(subst v4)
       
   635 apply(auto)[2]
       
   636 apply(rotate_tac 1)
       
   637 apply(drule_tac x="v2" in meta_spec)
       
   638 apply(simp)
       
   639 apply(subst (asm) (4) POSIX_def)
       
   640 apply(subst (asm) v4)
       
   641 apply(auto)[2]
       
   642 (* stuck in the ALT case *)