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