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