thys/Re.thy
changeset 82 26202889f829
parent 10 14d41b5b57b3
child 83 a8bcb5a0f9b9
equal deleted inserted replaced
81:7ac7782a7318 82:26202889f829
     1 (*test*)
     1    
     2 
       
     3 theory Re
     2 theory Re
     4   imports "Main" 
     3   imports "Main" 
     5 begin
     4 begin
       
     5 
     6 
     6 
     7 section {* Sequential Composition of Sets *}
     7 section {* Sequential Composition of Sets *}
     8 
     8 
     9 definition
     9 definition
    10   Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
    10   Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
    41 | "L (EMPTY) = {[]}"
    41 | "L (EMPTY) = {[]}"
    42 | "L (CHAR c) = {[c]}"
    42 | "L (CHAR c) = {[c]}"
    43 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
    43 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
    44 | "L (ALT r1 r2) = (L r1) \<union> (L r2)"
    44 | "L (ALT r1 r2) = (L r1) \<union> (L r2)"
    45 
    45 
       
    46 fun
       
    47  nullable :: "rexp \<Rightarrow> bool"
       
    48 where
       
    49   "nullable (NULL) = False"
       
    50 | "nullable (EMPTY) = True"
       
    51 | "nullable (CHAR c) = False"
       
    52 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
       
    53 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
       
    54 
       
    55 lemma nullable_correctness:
       
    56   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
       
    57 apply (induct r) 
       
    58 apply(auto simp add: Sequ_def) 
       
    59 done
    46 
    60 
    47 section {* Values *}
    61 section {* Values *}
    48 
    62 
    49 datatype val = 
    63 datatype val = 
    50   Void
    64   Void
    51 | Char char
    65 | Char char
    52 | Seq val val
    66 | Seq val val
    53 | Right val
    67 | Right val
    54 | Left val
    68 | Left val
    55 
    69 
       
    70 section {* The string behind a value *}
       
    71 
       
    72 fun 
       
    73   flat :: "val \<Rightarrow> string"
       
    74 where
       
    75   "flat(Void) = []"
       
    76 | "flat(Char c) = [c]"
       
    77 | "flat(Left v) = flat(v)"
       
    78 | "flat(Right v) = flat(v)"
       
    79 | "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
       
    80 
    56 section {* Relation between values and regular expressions *}
    81 section {* Relation between values and regular expressions *}
    57 
    82 
    58 inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
    83 inductive 
       
    84   Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
    59 where
    85 where
    60  "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
    86  "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
    61 | "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
    87 | "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
    62 | "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
    88 | "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
    63 | "\<turnstile> Void : EMPTY"
    89 | "\<turnstile> Void : EMPTY"
    64 | "\<turnstile> Char c : CHAR c"
    90 | "\<turnstile> Char c : CHAR c"
    65 
    91 
    66 thm Prf.intros(1)
    92 lemma not_nullable_flat:
    67 
    93   assumes "\<turnstile> v : r" "\<not>nullable r"
    68 section {* The string behind a value *}
    94   shows "flat v \<noteq> []"
    69 
    95 using assms
    70 fun flat :: "val \<Rightarrow> string"
    96 apply(induct)
    71 where
    97 apply(auto)
    72   "flat(Void) = []"
    98 done
    73 | "flat(Char c) = [c]"
       
    74 | "flat(Left v) = flat(v)"
       
    75 | "flat(Right v) = flat(v)"
       
    76 | "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
       
    77 
       
    78 
    99 
    79 lemma Prf_flat_L:
   100 lemma Prf_flat_L:
    80   assumes "\<turnstile> v : r" shows "flat v \<in> L r"
   101   assumes "\<turnstile> v : r" shows "flat v \<in> L r"
    81 using assms
   102 using assms
    82 apply(induct)
   103 apply(induct v r rule: Prf.induct)
    83 apply(auto simp add: Sequ_def)
   104 apply(auto simp add: Sequ_def)
    84 done
   105 done
    85 
   106 
    86 lemma L_flat_Prf:
   107 lemma L_flat_Prf:
    87   "L(r) = {flat v | v. \<turnstile> v : r}"
   108   "L(r) = {flat v | v. \<turnstile> v : r}"
    94 apply (metis Prf.intros(3) flat.simps(4))
   115 apply (metis Prf.intros(3) flat.simps(4))
    95 apply(erule Prf.cases)
   116 apply(erule Prf.cases)
    96 apply(auto)
   117 apply(auto)
    97 done
   118 done
    98 
   119 
    99 section {* Ordering of values *}
   120 section {* Greedy Ordering according to Frisch/Cardelli *}
       
   121 
       
   122 inductive 
       
   123   GrOrd :: "val \<Rightarrow> val \<Rightarrow> bool" ("_ gr\<succ> _")
       
   124 where 
       
   125   "v1 gr\<succ> v1' \<Longrightarrow> (Seq v1 v2) gr\<succ> (Seq v1' v2')"
       
   126 | "v2 gr\<succ> v2' \<Longrightarrow> (Seq v1 v2) gr\<succ> (Seq v1 v2')"
       
   127 | "v1 gr\<succ> v2 \<Longrightarrow> (Left v1) gr\<succ> (Left v2)"
       
   128 | "v1 gr\<succ> v2 \<Longrightarrow> (Right v1) gr\<succ> (Right v2)"
       
   129 | "(Left v2) gr\<succ>(Right v1)"
       
   130 | "(Char c) gr\<succ> (Char c)"
       
   131 | "(Void) gr\<succ> (Void)"
       
   132 
       
   133 lemma Gr_refl:
       
   134   assumes "\<turnstile> v : r"
       
   135   shows "v gr\<succ> v"
       
   136 using assms
       
   137 apply(induct)
       
   138 apply(auto intro: GrOrd.intros)
       
   139 done
       
   140 
       
   141 lemma Gr_total:
       
   142   assumes "\<turnstile> v1 : r" "\<turnstile> v2 : r"
       
   143   shows "v1 gr\<succ> v2 \<or> v2 gr\<succ> v1"
       
   144 using assms
       
   145 apply(induct v1 r arbitrary: v2 rule: Prf.induct)
       
   146 apply(rotate_tac 4)
       
   147 apply(erule Prf.cases)
       
   148 apply(simp_all)[5]
       
   149 apply(clarify)
       
   150 apply (metis GrOrd.intros(1) GrOrd.intros(2))
       
   151 apply(rotate_tac 2)
       
   152 apply(erule Prf.cases)
       
   153 apply(simp_all)
       
   154 apply(clarify)
       
   155 apply (metis GrOrd.intros(3))
       
   156 apply(clarify)
       
   157 apply (metis GrOrd.intros(5))
       
   158 apply(rotate_tac 2)
       
   159 apply(erule Prf.cases)
       
   160 apply(simp_all)
       
   161 apply(clarify)
       
   162 apply (metis GrOrd.intros(5))
       
   163 apply(clarify)
       
   164 apply (metis GrOrd.intros(4))
       
   165 apply(erule Prf.cases)
       
   166 apply(simp_all)
       
   167 apply (metis GrOrd.intros(7))
       
   168 apply(erule Prf.cases)
       
   169 apply(simp_all)
       
   170 apply (metis GrOrd.intros(6))
       
   171 done
       
   172 
       
   173 lemma Gr_trans: 
       
   174   assumes "v1 gr\<succ> v2" "v2 gr\<succ> v3" 
       
   175   and     "\<turnstile> v1 : r" "\<turnstile> v2 : r" "\<turnstile> v3 : r"
       
   176   shows "v1 gr\<succ> v3"
       
   177 using assms
       
   178 apply(induct r arbitrary: v1 v2 v3)
       
   179 apply(erule Prf.cases)
       
   180 apply(simp_all)[5]
       
   181 apply(erule Prf.cases)
       
   182 apply(simp_all)[5]
       
   183 apply(erule Prf.cases)
       
   184 apply(simp_all)[5]
       
   185 apply(erule Prf.cases)
       
   186 apply(simp_all)[5]
       
   187 apply(erule Prf.cases)
       
   188 apply(simp_all)[5]
       
   189 defer
       
   190 (* ALT case *)
       
   191 apply(erule Prf.cases)
       
   192 apply(simp_all (no_asm_use))[5]
       
   193 apply(erule Prf.cases)
       
   194 apply(simp_all (no_asm_use))[5]
       
   195 apply(erule Prf.cases)
       
   196 apply(simp_all (no_asm_use))[5]
       
   197 apply(clarify)
       
   198 apply(erule GrOrd.cases)
       
   199 apply(simp_all (no_asm_use))[7]
       
   200 apply(erule GrOrd.cases)
       
   201 apply(simp_all (no_asm_use))[7]
       
   202 apply (metis GrOrd.intros(3))
       
   203 apply(clarify)
       
   204 apply(erule GrOrd.cases)
       
   205 apply(simp_all (no_asm_use))[7]
       
   206 apply(erule GrOrd.cases)
       
   207 apply(simp_all (no_asm_use))[7]
       
   208 apply (metis GrOrd.intros(5))
       
   209 apply(erule Prf.cases)
       
   210 apply(simp_all (no_asm_use))[5]
       
   211 apply(clarify)
       
   212 apply(erule GrOrd.cases)
       
   213 apply(simp_all (no_asm_use))[7]
       
   214 apply(erule GrOrd.cases)
       
   215 apply(simp_all (no_asm_use))[7]
       
   216 apply (metis GrOrd.intros(5))
       
   217 apply(erule Prf.cases)
       
   218 apply(simp_all (no_asm_use))[5]
       
   219 apply(erule Prf.cases)
       
   220 apply(simp_all (no_asm_use))[5]
       
   221 apply(clarify)
       
   222 apply(erule GrOrd.cases)
       
   223 apply(simp_all (no_asm_use))[7]
       
   224 apply(clarify)
       
   225 apply(erule GrOrd.cases)
       
   226 apply(simp_all (no_asm_use))[7]
       
   227 apply(erule Prf.cases)
       
   228 apply(simp_all (no_asm_use))[5]
       
   229 apply(clarify)
       
   230 apply(erule GrOrd.cases)
       
   231 apply(simp_all (no_asm_use))[7]
       
   232 apply(erule GrOrd.cases)
       
   233 apply(simp_all (no_asm_use))[7]
       
   234 apply(clarify)
       
   235 apply(erule GrOrd.cases)
       
   236 apply(simp_all (no_asm_use))[7]
       
   237 apply(erule GrOrd.cases)
       
   238 apply(simp_all (no_asm_use))[7]
       
   239 apply (metis GrOrd.intros(4))
       
   240 (* SEQ case *)
       
   241 apply(erule Prf.cases)
       
   242 apply(simp_all (no_asm_use))[5]
       
   243 apply(erule Prf.cases)
       
   244 apply(simp_all (no_asm_use))[5]
       
   245 apply(erule Prf.cases)
       
   246 apply(simp_all (no_asm_use))[5]
       
   247 apply(clarify)
       
   248 apply(erule GrOrd.cases)
       
   249 apply(simp_all (no_asm_use))[7]
       
   250 apply(erule GrOrd.cases)
       
   251 apply(simp_all (no_asm_use))[7]
       
   252 apply(clarify)
       
   253 apply (metis GrOrd.intros(1))
       
   254 apply (metis GrOrd.intros(1))
       
   255 apply(erule GrOrd.cases)
       
   256 apply(simp_all (no_asm_use))[7]
       
   257 apply (metis GrOrd.intros(1))
       
   258 by (metis GrOrd.intros(1) Gr_refl)
       
   259 
       
   260 
       
   261 section {* Values Sets *}
       
   262 
       
   263 definition prefix :: "string \<Rightarrow> string \<Rightarrow> bool" ("_ \<sqsubseteq> _" [100, 100] 100)
       
   264 where
       
   265   "s1 \<sqsubseteq> s2 \<equiv> \<exists>s3. s1 @ s3 = s2"
       
   266 
       
   267 definition sprefix :: "string \<Rightarrow> string \<Rightarrow> bool" ("_ \<sqsubset> _" [100, 100] 100)
       
   268 where
       
   269   "s1 \<sqsubset> s2 \<equiv> (s1 \<sqsubseteq> s2 \<and> s1 \<noteq> s2)"
       
   270 
       
   271 lemma length_sprefix:
       
   272   "s1 \<sqsubset> s2 \<Longrightarrow> length s1 < length s2"
       
   273 unfolding sprefix_def prefix_def
       
   274 by (auto)
       
   275 
       
   276 definition Prefixes :: "string \<Rightarrow> string set" where
       
   277   "Prefixes s \<equiv> {sp. sp \<sqsubseteq> s}"
       
   278 
       
   279 definition Suffixes :: "string \<Rightarrow> string set" where
       
   280   "Suffixes s \<equiv> rev ` (Prefixes (rev s))"
       
   281 
       
   282 lemma Suffixes_in: 
       
   283   "\<exists>s1. s1 @ s2 = s3 \<Longrightarrow> s2 \<in> Suffixes s3"
       
   284 unfolding Suffixes_def Prefixes_def prefix_def image_def
       
   285 apply(auto)
       
   286 by (metis rev_rev_ident)
       
   287 
       
   288 lemma Prefixes_Cons:
       
   289   "Prefixes (c # s) = {[]} \<union> {c # sp | sp. sp \<in> Prefixes s}"
       
   290 unfolding Prefixes_def prefix_def
       
   291 apply(auto simp add: append_eq_Cons_conv) 
       
   292 done
       
   293 
       
   294 lemma finite_Prefixes:
       
   295   "finite (Prefixes s)"
       
   296 apply(induct s)
       
   297 apply(auto simp add: Prefixes_def prefix_def)[1]
       
   298 apply(simp add: Prefixes_Cons)
       
   299 done
       
   300 
       
   301 lemma finite_Suffixes:
       
   302   "finite (Suffixes s)"
       
   303 unfolding Suffixes_def
       
   304 apply(rule finite_imageI)
       
   305 apply(rule finite_Prefixes)
       
   306 done
       
   307 
       
   308 lemma prefix_Cons:
       
   309   "((c # s1) \<sqsubseteq> (c # s2)) = (s1 \<sqsubseteq> s2)"
       
   310 apply(auto simp add: prefix_def)
       
   311 done
       
   312 
       
   313 lemma prefix_append:
       
   314   "((s @ s1) \<sqsubseteq> (s @ s2)) = (s1 \<sqsubseteq> s2)"
       
   315 apply(induct s)
       
   316 apply(simp)
       
   317 apply(simp add: prefix_Cons)
       
   318 done
       
   319 
       
   320 
       
   321 definition Values :: "rexp \<Rightarrow> string \<Rightarrow> val set" where
       
   322   "Values r s \<equiv> {v. \<turnstile> v : r \<and> flat v \<sqsubseteq> s}"
       
   323 
       
   324 definition rest :: "val \<Rightarrow> string \<Rightarrow> string" where
       
   325   "rest v s \<equiv> drop (length (flat v)) s"
       
   326 
       
   327 lemma rest_Suffixes:
       
   328   "rest v s \<in> Suffixes s"
       
   329 unfolding rest_def
       
   330 by (metis Suffixes_in append_take_drop_id)
       
   331 
       
   332 
       
   333 lemma Values_recs:
       
   334   "Values (NULL) s = {}"
       
   335   "Values (EMPTY) s = {Void}"
       
   336   "Values (CHAR c) s = (if [c] \<sqsubseteq> s then {Char c} else {})" 
       
   337   "Values (ALT r1 r2) s = {Left v | v. v \<in> Values r1 s} \<union> {Right v | v. v \<in> Values r2 s}"
       
   338   "Values (SEQ r1 r2) s = {Seq v1 v2 | v1 v2. v1 \<in> Values r1 s \<and> v2 \<in> Values r2 (rest v1 s)}"
       
   339 unfolding Values_def
       
   340 apply(auto)
       
   341 (*NULL*)
       
   342 apply(erule Prf.cases)
       
   343 apply(simp_all)[5]
       
   344 (*EMPTY*)
       
   345 apply(erule Prf.cases)
       
   346 apply(simp_all)[5]
       
   347 apply(rule Prf.intros)
       
   348 apply (metis append_Nil prefix_def)
       
   349 (*CHAR*)
       
   350 apply(erule Prf.cases)
       
   351 apply(simp_all)[5]
       
   352 apply(rule Prf.intros)
       
   353 apply(erule Prf.cases)
       
   354 apply(simp_all)[5]
       
   355 (*ALT*)
       
   356 apply(erule Prf.cases)
       
   357 apply(simp_all)[5]
       
   358 apply (metis Prf.intros(2))
       
   359 apply (metis Prf.intros(3))
       
   360 (*SEQ*)
       
   361 apply(erule Prf.cases)
       
   362 apply(simp_all)[5]
       
   363 apply (simp add: append_eq_conv_conj prefix_def rest_def)
       
   364 apply (metis Prf.intros(1))
       
   365 apply (simp add: append_eq_conv_conj prefix_def rest_def)
       
   366 done
       
   367 
       
   368 lemma Values_finite:
       
   369   "finite (Values r s)"
       
   370 apply(induct r arbitrary: s)
       
   371 apply(simp_all add: Values_recs)
       
   372 thm finite_surj
       
   373 apply(rule_tac f="\<lambda>(x, y). Seq x y" and 
       
   374                A="{(v1, v2) | v1 v2. v1 \<in> Values r1 s \<and> v2 \<in> Values r2 (rest v1 s)}" in finite_surj)
       
   375 prefer 2
       
   376 apply(auto)[1]
       
   377 apply(rule_tac B="\<Union>sp \<in> Suffixes s. {(v1, v2). v1 \<in> Values r1 s \<and> v2 \<in> Values r2 sp}" in finite_subset)
       
   378 apply(auto)[1]
       
   379 apply (metis rest_Suffixes)
       
   380 apply(rule finite_UN_I)
       
   381 apply(rule finite_Suffixes)
       
   382 apply(simp)
       
   383 done
       
   384 
       
   385 section {* Sulzmann functions *}
       
   386 
       
   387 fun 
       
   388   mkeps :: "rexp \<Rightarrow> val"
       
   389 where
       
   390   "mkeps(EMPTY) = Void"
       
   391 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
       
   392 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
       
   393 
       
   394 lemma mkeps_nullable:
       
   395   assumes "nullable(r)" 
       
   396   shows "\<turnstile> mkeps r : r"
       
   397 using assms
       
   398 apply(induct rule: nullable.induct)
       
   399 apply(auto intro: Prf.intros)
       
   400 done
       
   401 
       
   402 lemma mkeps_flat:
       
   403   assumes "nullable(r)" 
       
   404   shows "flat (mkeps r) = []"
       
   405 using assms
       
   406 apply(induct rule: nullable.induct)
       
   407 apply(auto)
       
   408 done
       
   409 
       
   410 text {*
       
   411   The value mkeps returns is always the correct POSIX
       
   412   value.
       
   413 *}
       
   414 
       
   415 section {* Sulzmann's Ordering of values *}
   100 
   416 
   101 inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
   417 inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
   102 where
   418 where
   103   "\<lbrakk>v1 = v1'; v2 \<succ>r2 v2'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
   419   "v2 \<succ>r2 v2' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1 v2')" 
   104 | "v1  \<succ>r1 v1' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
   420 | "\<lbrakk>v1 \<succ>r1 v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
   105 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
   421 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
   106 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
   422 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
   107 | "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
   423 | "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
   108 | "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
   424 | "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
   109 | "Void \<succ>EMPTY Void"
   425 | "Void \<succ>EMPTY Void"
   110 | "(Char c) \<succ>(CHAR c) (Char c)"
   426 | "(Char c) \<succ>(CHAR c) (Char c)"
   111 
   427 
   112 (*
   428 inductive ValOrd2 :: "val \<Rightarrow> val \<Rightarrow> bool" ("_ 2\<succ> _" [100, 100] 100)
   113 lemma
   429 where
   114   assumes "r = SEQ (ALT EMPTY EMPTY) (ALT EMPTY (CHAR c))"
   430   "v2 2\<succ> v2' \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1 v2')" 
   115   shows "(Seq (Left Void) (Right (Char c))) \<succ>r (Seq (Left Void) (Left Void))"
   431 | "\<lbrakk>v1 2\<succ> v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1' v2')" 
   116 using assms
   432 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) 2\<succ> (Right v2)"
   117 apply(simp)
   433 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) 2\<succ> (Left v1)"
       
   434 | "v2 2\<succ> v2' \<Longrightarrow> (Right v2) 2\<succ> (Right v2')"
       
   435 | "v1 2\<succ> v1' \<Longrightarrow> (Left v1) 2\<succ> (Left v1')"
       
   436 | "Void 2\<succ> Void"
       
   437 | "(Char c) 2\<succ> (Char c)"
       
   438 
       
   439 lemma Ord1:
       
   440   "v1 \<succ>r v2 \<Longrightarrow> v1 2\<succ> v2"
       
   441 apply(induct rule: ValOrd.induct)
       
   442 apply(auto intro: ValOrd2.intros)
       
   443 done
       
   444 
       
   445 lemma Ord2:
       
   446   "v1 2\<succ> v2 \<Longrightarrow> \<exists>r. v1 \<succ>r v2"
       
   447 apply(induct v1 v2 rule: ValOrd2.induct)
       
   448 apply(auto intro: ValOrd.intros)
       
   449 done
       
   450 
       
   451 lemma Ord3:
       
   452   "\<lbrakk>v1 2\<succ> v2; \<turnstile> v1 : r\<rbrakk> \<Longrightarrow> v1 \<succ>r v2"
       
   453 apply(induct v1 v2 arbitrary: r rule: ValOrd2.induct)
       
   454 apply(auto intro: ValOrd.intros elim: Prf.cases)
       
   455 done
       
   456 
       
   457 lemma ValOrd_refl:
       
   458   assumes "\<turnstile> v : r"
       
   459   shows "v \<succ>r v"
       
   460 using assms
       
   461 apply(induct)
       
   462 apply(auto intro: ValOrd.intros)
       
   463 done
       
   464 
       
   465 lemma ValOrd_total:
       
   466   shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r\<rbrakk>  \<Longrightarrow> v1 \<succ>r v2 \<or> v2 \<succ>r v1"
       
   467 apply(induct r arbitrary: v1 v2)
       
   468 apply(auto)
       
   469 apply(erule Prf.cases)
       
   470 apply(simp_all)[5]
       
   471 apply(erule Prf.cases)
       
   472 apply(simp_all)[5]
       
   473 apply(erule Prf.cases)
       
   474 apply(simp_all)[5]
       
   475 apply (metis ValOrd.intros(7))
       
   476 apply(erule Prf.cases)
       
   477 apply(simp_all)[5]
       
   478 apply(erule Prf.cases)
       
   479 apply(simp_all)[5]
       
   480 apply (metis ValOrd.intros(8))
       
   481 apply(erule Prf.cases)
       
   482 apply(simp_all)[5]
       
   483 apply(erule Prf.cases)
       
   484 apply(simp_all)[5]
       
   485 apply(clarify)
       
   486 apply(case_tac "v1a = v1b")
       
   487 apply(simp)
       
   488 apply(rule ValOrd.intros(1))
       
   489 apply (metis ValOrd.intros(1))
       
   490 apply(rule ValOrd.intros(2))
       
   491 apply(auto)[2]
       
   492 apply(erule contrapos_np)
       
   493 apply(rule ValOrd.intros(2))
       
   494 apply(auto)
       
   495 apply(erule Prf.cases)
       
   496 apply(simp_all)[5]
       
   497 apply(erule Prf.cases)
       
   498 apply(simp_all)[5]
       
   499 apply (metis Ord1 Ord3 Prf.intros(2) ValOrd2.intros(6))
   118 apply(rule ValOrd.intros)
   500 apply(rule ValOrd.intros)
       
   501 apply(erule contrapos_np)
   119 apply(rule ValOrd.intros)
   502 apply(rule ValOrd.intros)
       
   503 apply (metis le_eq_less_or_eq neq_iff)
       
   504 apply(erule Prf.cases)
       
   505 apply(simp_all)[5]
   120 apply(rule ValOrd.intros)
   506 apply(rule ValOrd.intros)
       
   507 apply(erule contrapos_np)
   121 apply(rule ValOrd.intros)
   508 apply(rule ValOrd.intros)
   122 apply(simp)
   509 apply (metis le_eq_less_or_eq neq_iff)
   123 done
   510 apply(rule ValOrd.intros)
   124 *)
   511 apply(erule contrapos_np)
       
   512 apply(rule ValOrd.intros)
       
   513 by metis
       
   514 
       
   515 lemma ValOrd_anti:
       
   516   shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r; v1 \<succ>r v2; v2 \<succ>r v1\<rbrakk> \<Longrightarrow> v1 = v2"
       
   517 apply(induct r arbitrary: v1 v2)
       
   518 apply(erule Prf.cases)
       
   519 apply(simp_all)[5]
       
   520 apply(erule Prf.cases)
       
   521 apply(simp_all)[5]
       
   522 apply(erule Prf.cases)
       
   523 apply(simp_all)[5]
       
   524 apply(erule Prf.cases)
       
   525 apply(simp_all)[5]
       
   526 apply(erule Prf.cases)
       
   527 apply(simp_all)[5]
       
   528 apply(erule Prf.cases)
       
   529 apply(simp_all)[5]
       
   530 apply(erule Prf.cases)
       
   531 apply(simp_all)[5]
       
   532 apply(erule ValOrd.cases)
       
   533 apply(simp_all)[8]
       
   534 apply(erule ValOrd.cases)
       
   535 apply(simp_all)[8]
       
   536 apply(erule ValOrd.cases)
       
   537 apply(simp_all)[8]
       
   538 apply(erule Prf.cases)
       
   539 apply(simp_all)[5]
       
   540 apply(erule Prf.cases)
       
   541 apply(simp_all)[5]
       
   542 apply(erule ValOrd.cases)
       
   543 apply(simp_all)[8]
       
   544 apply(erule ValOrd.cases)
       
   545 apply(simp_all)[8]
       
   546 apply(erule ValOrd.cases)
       
   547 apply(simp_all)[8]
       
   548 apply(erule ValOrd.cases)
       
   549 apply(simp_all)[8]
       
   550 apply(erule Prf.cases)
       
   551 apply(simp_all)[5]
       
   552 apply(erule ValOrd.cases)
       
   553 apply(simp_all)[8]
       
   554 apply(erule ValOrd.cases)
       
   555 apply(simp_all)[8]
       
   556 apply(erule ValOrd.cases)
       
   557 apply(simp_all)[8]
       
   558 apply(erule ValOrd.cases)
       
   559 apply(simp_all)[8]
       
   560 done
       
   561 
       
   562 lemma refl_on_ValOrd:
       
   563   "refl_on (Values r s) {(v1, v2). v1 \<succ>r v2 \<and> v1 \<in> Values r s \<and> v2 \<in> Values r s}"
       
   564 unfolding refl_on_def
       
   565 apply(auto)
       
   566 apply(rule ValOrd_refl)
       
   567 apply(simp add: Values_def)
       
   568 done
       
   569 
   125 
   570 
   126 section {* Posix definition *}
   571 section {* Posix definition *}
   127 
   572 
   128 definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
   573 definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
   129 where
   574 where
   130   "POSIX v r \<equiv> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v')"
   575   "POSIX v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v'))"
   131 
   576 
   132 (*
   577 definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
   133 lemma POSIX_SEQ:
   578 where
       
   579   "POSIX2 v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v 2\<succ> v'))"
       
   580 
       
   581 lemma "POSIX v r = POSIX2 v r"
       
   582 unfolding POSIX_def POSIX2_def
       
   583 apply(auto)
       
   584 apply(rule Ord1)
       
   585 apply(auto)
       
   586 apply(rule Ord3)
       
   587 apply(auto)
       
   588 done
       
   589 
       
   590 section {* POSIX for some constructors *}
       
   591 
       
   592 lemma POSIX_SEQ1:
   134   assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
   593   assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
   135   shows "POSIX v1 r1 \<and> POSIX v2 r2"
   594   shows "POSIX v1 r1"
   136 using assms
   595 using assms
   137 unfolding POSIX_def
   596 unfolding POSIX_def
   138 apply(auto)
   597 apply(auto)
   139 apply(drule_tac x="Seq v' v2" in spec)
   598 apply(drule_tac x="Seq v' v2" in spec)
   140 apply(simp)
   599 apply(simp)
   141 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))
   600 apply(erule impE)
       
   601 apply(rule Prf.intros)
       
   602 apply(simp)
       
   603 apply(simp)
       
   604 apply(erule ValOrd.cases)
       
   605 apply(simp_all)
       
   606 apply(clarify)
       
   607 by (metis ValOrd_refl)
       
   608 
       
   609 lemma POSIX_SEQ2:
       
   610   assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2" 
       
   611   shows "POSIX v2 r2"
       
   612 using assms
       
   613 unfolding POSIX_def
       
   614 apply(auto)
   142 apply(drule_tac x="Seq v1 v'" in spec)
   615 apply(drule_tac x="Seq v1 v'" in spec)
   143 apply(simp)
   616 apply(simp)
   144 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))
   617 apply(erule impE)
   145 *)
   618 apply(rule Prf.intros)
   146 
   619 apply(simp)
   147 (*
   620 apply(simp)
   148 lemma POSIX_SEQ_I:
   621 apply(erule ValOrd.cases)
   149   assumes "POSIX v1 r1" "POSIX v2 r2" 
   622 apply(simp_all)
   150   shows "POSIX (Seq v1 v2) (SEQ r1 r2)" 
   623 done
   151 using assms
   624 
   152 unfolding POSIX_def
   625 lemma POSIX_ALT2:
   153 apply(auto)
       
   154 apply(rotate_tac 2)
       
   155 apply(erule Prf.cases)
       
   156 apply(simp_all)[5]
       
   157 apply(auto)[1]
       
   158 apply(rule ValOrd.intros)
       
   159 
       
   160 apply(auto)
       
   161 done
       
   162 *)
       
   163 
       
   164 lemma POSIX_ALT:
       
   165   assumes "POSIX (Left v1) (ALT r1 r2)"
   626   assumes "POSIX (Left v1) (ALT r1 r2)"
   166   shows "POSIX v1 r1"
   627   shows "POSIX v1 r1"
   167 using assms
   628 using assms
   168 unfolding POSIX_def
   629 unfolding POSIX_def
   169 apply(auto)
   630 apply(auto)
       
   631 apply(erule Prf.cases)
       
   632 apply(simp_all)[5]
   170 apply(drule_tac x="Left v'" in spec)
   633 apply(drule_tac x="Left v'" in spec)
   171 apply(simp)
   634 apply(simp)
   172 apply(drule mp)
   635 apply(drule mp)
   173 apply(rule Prf.intros)
   636 apply(rule Prf.intros)
   174 apply(auto)
   637 apply(auto)
   180   assumes "POSIX (Right v2) (ALT r1 r2)"
   643   assumes "POSIX (Right v2) (ALT r1 r2)"
   181   shows "POSIX v2 r2"
   644   shows "POSIX v2 r2"
   182 using assms
   645 using assms
   183 unfolding POSIX_def
   646 unfolding POSIX_def
   184 apply(auto)
   647 apply(auto)
       
   648 apply(erule Prf.cases)
       
   649 apply(simp_all)[5]
   185 apply(drule_tac x="Right v'" in spec)
   650 apply(drule_tac x="Right v'" in spec)
   186 apply(simp)
   651 apply(simp)
   187 apply(drule mp)
   652 apply(drule mp)
   188 apply(rule Prf.intros)
   653 apply(rule Prf.intros)
   189 apply(auto)
   654 apply(auto)
   190 apply(erule ValOrd.cases)
   655 apply(erule ValOrd.cases)
   191 apply(simp_all)
   656 apply(simp_all)
   192 done
   657 done
   193 
       
   194 
   658 
   195 lemma POSIX_ALT1b:
   659 lemma POSIX_ALT1b:
   196   assumes "POSIX (Right v2) (ALT r1 r2)"
   660   assumes "POSIX (Right v2) (ALT r1 r2)"
   197   shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
   661   shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
   198 using assms
   662 using assms
   205   assumes "POSIX v1 r1" 
   669   assumes "POSIX v1 r1" 
   206   shows "POSIX (Left v1) (ALT r1 r2)"
   670   shows "POSIX (Left v1) (ALT r1 r2)"
   207 using assms
   671 using assms
   208 unfolding POSIX_def
   672 unfolding POSIX_def
   209 apply(auto)
   673 apply(auto)
   210 apply(rotate_tac 3)
   674 apply (metis Prf.intros(2))
       
   675 apply(rotate_tac 2)
   211 apply(erule Prf.cases)
   676 apply(erule Prf.cases)
   212 apply(simp_all)[5]
   677 apply(simp_all)[5]
   213 apply(auto)
   678 apply(auto)
   214 apply(rule ValOrd.intros)
   679 apply(rule ValOrd.intros)
   215 apply(auto)
   680 apply(auto)
   220   assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
   685   assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
   221   shows "POSIX (Right v2) (ALT r1 r2)"
   686   shows "POSIX (Right v2) (ALT r1 r2)"
   222 using assms
   687 using assms
   223 unfolding POSIX_def
   688 unfolding POSIX_def
   224 apply(auto)
   689 apply(auto)
       
   690 apply (metis Prf.intros)
   225 apply(rotate_tac 3)
   691 apply(rotate_tac 3)
   226 apply(erule Prf.cases)
   692 apply(erule Prf.cases)
   227 apply(simp_all)[5]
   693 apply(simp_all)[5]
   228 apply(auto)
   694 apply(auto)
   229 apply(rule ValOrd.intros)
   695 apply(rule ValOrd.intros)
   230 apply metis
   696 apply metis
   231 done
   697 done
   232 
       
   233 
       
   234 section {* The ordering is reflexive *}
       
   235 
       
   236 lemma ValOrd_refl:
       
   237   assumes "\<turnstile> v : r"
       
   238   shows "v \<succ>r v"
       
   239 using assms
       
   240 apply(induct)
       
   241 apply(auto intro: ValOrd.intros)
       
   242 done
       
   243 
       
   244 
       
   245 section {* The Matcher *}
       
   246 
       
   247 fun
       
   248  nullable :: "rexp \<Rightarrow> bool"
       
   249 where
       
   250   "nullable (NULL) = False"
       
   251 | "nullable (EMPTY) = True"
       
   252 | "nullable (CHAR c) = False"
       
   253 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
       
   254 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
       
   255 
       
   256 lemma nullable_correctness:
       
   257   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
       
   258 apply (induct r) 
       
   259 apply(auto simp add: Sequ_def) 
       
   260 done
       
   261 
       
   262 fun mkeps :: "rexp \<Rightarrow> val"
       
   263 where
       
   264   "mkeps(EMPTY) = Void"
       
   265 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
       
   266 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
       
   267 
       
   268 lemma mkeps_nullable:
       
   269   assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
       
   270 using assms
       
   271 apply(induct rule: nullable.induct)
       
   272 apply(auto intro: Prf.intros)
       
   273 done
       
   274 
       
   275 lemma mkeps_flat:
       
   276   assumes "nullable(r)" shows "flat (mkeps r) = []"
       
   277 using assms
       
   278 apply(induct rule: nullable.induct)
       
   279 apply(auto)
       
   280 done
       
   281 
       
   282 text {*
       
   283   The value mkeps returns is always the correct POSIX
       
   284   value.
       
   285 *}
       
   286 
   698 
   287 lemma mkeps_POSIX:
   699 lemma mkeps_POSIX:
   288   assumes "nullable r"
   700   assumes "nullable r"
   289   shows "POSIX (mkeps r) r"
   701   shows "POSIX (mkeps r) r"
   290 using assms
   702 using assms
   291 apply(induct r)
   703 apply(induct r)
   292 apply(auto)[1]
   704 apply(auto)[1]
   293 apply(simp add: POSIX_def)
   705 apply(simp add: POSIX_def)
   294 apply(auto)[1]
   706 apply(auto)[1]
       
   707 apply (metis Prf.intros(4))
   295 apply(erule Prf.cases)
   708 apply(erule Prf.cases)
   296 apply(simp_all)[5]
   709 apply(simp_all)[5]
   297 apply (metis ValOrd.intros)
   710 apply (metis ValOrd.intros)
       
   711 apply(simp)
       
   712 apply(auto)[1]
   298 apply(simp add: POSIX_def)
   713 apply(simp add: POSIX_def)
   299 apply(auto)[1]
   714 apply(auto)[1]
   300 apply(simp add: POSIX_def)
   715 apply (metis mkeps.simps(2) mkeps_nullable nullable.simps(5))
   301 apply(auto)[1]
   716 apply(rotate_tac 6)
   302 apply(erule Prf.cases)
   717 apply(erule Prf.cases)
   303 apply(simp_all)[5]
   718 apply(simp_all)[5]
   304 apply(auto)
   719 apply (simp add: mkeps_flat)
   305 apply (simp add: ValOrd.intros(2) mkeps_flat)
   720 apply(case_tac "mkeps r1a = v1")
   306 apply(simp add: POSIX_def)
   721 apply(simp)
   307 apply(auto)[1]
   722 apply (metis ValOrd.intros(1))
   308 apply(erule Prf.cases)
   723 apply (rule ValOrd.intros(2))
   309 apply(simp_all)[5]
   724 apply metis
   310 apply(auto)
   725 apply(simp)
   311 apply (simp add: ValOrd.intros(6))
   726 (* ALT case *)
   312 apply (simp add: ValOrd.intros(3))
   727 thm mkeps.simps
   313 apply(simp add: POSIX_def)
   728 apply(simp)
   314 apply(auto)[1]
   729 apply(erule disjE)
   315 apply(erule Prf.cases)
   730 apply(simp)
   316 apply(simp_all)[5]
   731 apply (metis POSIX_ALT_I1)
   317 apply(auto)
   732 (* *)
   318 apply (simp add: ValOrd.intros(6))
   733 apply(auto)[1]
   319 apply (simp add: ValOrd.intros(3))
   734 thm  POSIX_ALT_I1
   320 apply(simp add: POSIX_def)
   735 apply (metis POSIX_ALT_I1)
   321 apply(auto)[1]
   736 apply(simp (no_asm) add: POSIX_def)
   322 apply(erule Prf.cases)
   737 apply(auto)[1]
   323 apply(simp_all)[5]
   738 apply(rule Prf.intros(3))
   324 apply(auto)
   739 apply(simp only: POSIX_def)
   325 apply (metis Prf_flat_L mkeps_flat nullable_correctness)
   740 apply(rotate_tac 4)
   326 by (simp add: ValOrd.intros(5))
   741 apply(erule Prf.cases)
       
   742 apply(simp_all)[5]
       
   743 thm mkeps_flat
       
   744 apply(simp add: mkeps_flat)
       
   745 apply(auto)[1]
       
   746 thm Prf_flat_L nullable_correctness
       
   747 apply (metis Prf_flat_L nullable_correctness)
       
   748 apply(rule ValOrd.intros)
       
   749 apply(subst (asm) POSIX_def)
       
   750 apply(clarify)
       
   751 apply(drule_tac x="v2" in spec)
       
   752 by simp
   327 
   753 
   328 
   754 
   329 section {* Derivatives *}
   755 section {* Derivatives *}
   330 
   756 
   331 fun
   757 fun
   344  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
   770  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
   345 where
   771 where
   346   "ders [] r = r"
   772   "ders [] r = r"
   347 | "ders (c # s) r = ders s (der c r)"
   773 | "ders (c # s) r = ders s (der c r)"
   348 
   774 
       
   775 
   349 section {* Injection function *}
   776 section {* Injection function *}
   350 
   777 
   351 fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
   778 fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
   352 where
   779 where
   353   "injval (CHAR d) c Void = Char d"
   780   "injval (EMPTY) c Void = Char c"
       
   781 | "injval (CHAR d) c Void = Char d"
       
   782 | "injval (CHAR d) c (Char c') = Seq (Char d) (Char c')"
   354 | "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
   783 | "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
   355 | "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
   784 | "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
       
   785 | "injval (SEQ r1 r2) c (Char c') = Seq (Char c) (Char c')"
   356 | "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
   786 | "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
   357 | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
   787 | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
   358 | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
   788 | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
   359 
   789 
       
   790 
   360 section {* Projection function *}
   791 section {* Projection function *}
   361 
   792 
   362 fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
   793 fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
   363 where
   794 where
   364   "projval (CHAR d) c _ = Void"
   795   "projval (CHAR d) c _ = Void"
   365 | "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
   796 | "projval (ALT r1 r2) c (Left v1) = Left (projval r1 c v1)"
   366 | "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
   797 | "projval (ALT r1 r2) c (Right v2) = Right (projval r2 c v2)"
   367 | "projval (SEQ r1 r2) c (Seq v1 v2) = 
   798 | "projval (SEQ r1 r2) c (Seq v1 v2) = 
   368      (if flat v1 = [] then Right(projval r2 c v2) 
   799      (if flat v1 = [] then Right(projval r2 c v2) 
   369       else if nullable r1 then Left (Seq (projval r1 c v1) v2)
   800       else if nullable r1 then Left (Seq (projval r1 c v1) v2)
   370                           else Seq (projval r1 c v1) v2)"
   801                           else Seq (projval r1 c v1) v2)"
   371 
   802 
   372 text {*
   803 text {*
   373   Injection value is related to r
   804   Injection value is related to r
   374 *}
   805 *}
   375 
   806 
   376 lemma v3:
   807 lemma v3:
   377   assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
   808   assumes "\<turnstile> v : der c r" 
       
   809   shows "\<turnstile> (injval r c v) : r"
   378 using assms
   810 using assms
   379 apply(induct arbitrary: v rule: der.induct)
   811 apply(induct arbitrary: v rule: der.induct)
   380 apply(simp)
   812 apply(simp)
   381 apply(erule Prf.cases)
   813 apply(erule Prf.cases)
   382 apply(simp_all)[5]
   814 apply(simp_all)[5]
       
   815 apply(simp)
   383 apply(erule Prf.cases)
   816 apply(erule Prf.cases)
   384 apply(simp_all)[5]
   817 apply(simp_all)[5]
   385 apply(case_tac "c = c'")
   818 apply(case_tac "c = c'")
   386 apply(simp)
   819 apply(simp)
   387 apply(erule Prf.cases)
   820 apply(erule Prf.cases)
   388 apply(simp_all)[5]
   821 apply(simp_all)[5]
   389 apply (metis Prf.intros(5))
   822 apply (metis Prf.intros(5))
   390 apply(erule Prf.cases)
   823 apply(simp)
   391 apply(simp_all)[5]
   824 apply(erule Prf.cases)
       
   825 apply(simp_all)[5]
       
   826 apply(simp)
   392 apply(erule Prf.cases)
   827 apply(erule Prf.cases)
   393 apply(simp_all)[5]
   828 apply(simp_all)[5]
   394 apply (metis Prf.intros(2))
   829 apply (metis Prf.intros(2))
   395 apply (metis Prf.intros(3))
   830 apply (metis Prf.intros(3))
   396 apply(simp)
   831 apply(simp)
   411 apply(auto)[1]
   846 apply(auto)[1]
   412 apply(rule Prf.intros)
   847 apply(rule Prf.intros)
   413 apply(auto)[2]
   848 apply(auto)[2]
   414 done
   849 done
   415 
   850 
       
   851 lemma v3_proj:
       
   852   assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s"
       
   853   shows "\<turnstile> (projval r c v) : der c r"
       
   854 using assms
       
   855 apply(induct rule: Prf.induct)
       
   856 prefer 4
       
   857 apply(simp)
       
   858 prefer 4
       
   859 apply(simp)
       
   860 apply (metis Prf.intros(4))
       
   861 prefer 2
       
   862 apply(simp)
       
   863 apply (metis Prf.intros(2))
       
   864 prefer 2
       
   865 apply(simp)
       
   866 apply (metis Prf.intros(3))
       
   867 apply(auto)
       
   868 apply(rule Prf.intros)
       
   869 apply(simp)
       
   870 apply (metis Prf_flat_L nullable_correctness)
       
   871 apply(rule Prf.intros)
       
   872 apply(rule Prf.intros)
       
   873 apply (metis Cons_eq_append_conv)
       
   874 apply(simp)
       
   875 apply(rule Prf.intros)
       
   876 apply (metis Cons_eq_append_conv)
       
   877 apply(simp)
       
   878 done
       
   879 
   416 text {*
   880 text {*
   417   The string behin the injection value is an added c
   881   The string behind the injection value is an added c
   418 *}
   882 *}
   419 
   883 
   420 lemma v4:
   884 lemma v4:
   421   assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
   885   assumes "\<turnstile> v : der c r" 
       
   886   shows "flat (injval r c v) = c # (flat v)"
   422 using assms
   887 using assms
   423 apply(induct arbitrary: v rule: der.induct)
   888 apply(induct arbitrary: v rule: der.induct)
   424 apply(simp)
   889 apply(simp)
   425 apply(erule Prf.cases)
   890 apply(erule Prf.cases)
   426 apply(simp_all)[5]
   891 apply(simp_all)[5]
   441 apply(simp_all)[5]
   906 apply(simp_all)[5]
   442 apply(simp)
   907 apply(simp)
   443 apply(case_tac "nullable r1")
   908 apply(case_tac "nullable r1")
   444 apply(simp)
   909 apply(simp)
   445 apply(erule Prf.cases)
   910 apply(erule Prf.cases)
   446 apply(simp_all)[5]
   911 apply(simp_all (no_asm_use))[5]
   447 apply(auto)[1]
   912 apply(auto)[1]
   448 apply(erule Prf.cases)
   913 apply(erule Prf.cases)
   449 apply(simp_all)[5]
   914 apply(simp_all)[5]
       
   915 apply(clarify)
       
   916 apply(simp only: injval.simps flat.simps)
   450 apply(auto)[1]
   917 apply(auto)[1]
   451 apply (metis mkeps_flat)
   918 apply (metis mkeps_flat)
   452 apply(simp)
   919 apply(simp)
   453 apply(erule Prf.cases)
   920 apply(erule Prf.cases)
   454 apply(simp_all)[5]
   921 apply(simp_all)[5]
   455 done
   922 done
       
   923 
       
   924 lemma v4_proj:
       
   925   assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s"
       
   926   shows "c # flat (projval r c v) = flat v"
       
   927 using assms
       
   928 apply(induct rule: Prf.induct)
       
   929 prefer 4
       
   930 apply(simp)
       
   931 prefer 4
       
   932 apply(simp)
       
   933 prefer 2
       
   934 apply(simp)
       
   935 prefer 2
       
   936 apply(simp)
       
   937 apply(auto)
       
   938 by (metis Cons_eq_append_conv)
       
   939 
       
   940 lemma v4_proj2:
       
   941   assumes "\<turnstile> v : r" and "(flat v) = c # s"
       
   942   shows "flat (projval r c v) = s"
       
   943 using assms
       
   944 by (metis list.inject v4_proj)
       
   945 
       
   946 lemma injval_inj: "inj_on (injval r c) {v. \<turnstile> v : der c r}"
       
   947 apply(induct c r rule: der.induct)
       
   948 unfolding inj_on_def
       
   949 apply(auto)[1]
       
   950 apply(erule Prf.cases)
       
   951 apply(simp_all)[5]
       
   952 apply(auto)[1]
       
   953 apply(erule Prf.cases)
       
   954 apply(simp_all)[5]
       
   955 apply(auto)[1]
       
   956 apply(erule Prf.cases)
       
   957 apply(simp_all)[5]
       
   958 apply(erule Prf.cases)
       
   959 apply(simp_all)[5]
       
   960 apply(erule Prf.cases)
       
   961 apply(simp_all)[5]
       
   962 apply(auto)[1]
       
   963 apply(erule Prf.cases)
       
   964 apply(simp_all)[5]
       
   965 apply(erule Prf.cases)
       
   966 apply(simp_all)[5]
       
   967 apply(erule Prf.cases)
       
   968 apply(simp_all)[5]
       
   969 apply(auto)[1]
       
   970 apply(erule Prf.cases)
       
   971 apply(simp_all)[5]
       
   972 apply(erule Prf.cases)
       
   973 apply(simp_all)[5]
       
   974 apply(clarify)
       
   975 apply(erule Prf.cases)
       
   976 apply(simp_all)[5]
       
   977 apply(erule Prf.cases)
       
   978 apply(simp_all)[5]
       
   979 apply(clarify)
       
   980 apply(erule Prf.cases)
       
   981 apply(simp_all)[5]
       
   982 apply(clarify)
       
   983 apply (metis list.distinct(1) mkeps_flat v4)
       
   984 apply(erule Prf.cases)
       
   985 apply(simp_all)[5]
       
   986 apply(clarify)
       
   987 apply(rotate_tac 6)
       
   988 apply(erule Prf.cases)
       
   989 apply(simp_all)[5]
       
   990 apply (metis list.distinct(1) mkeps_flat v4)
       
   991 apply(erule Prf.cases)
       
   992 apply(simp_all)[5]
       
   993 apply(erule Prf.cases)
       
   994 apply(simp_all)[5]
       
   995 done
       
   996 
       
   997 lemma Values_nullable:
       
   998   assumes "nullable r1"
       
   999   shows "mkeps r1 \<in> Values r1 s"
       
  1000 using assms
       
  1001 apply(induct r1 arbitrary: s)
       
  1002 apply(simp_all)
       
  1003 apply(simp add: Values_recs)
       
  1004 apply(simp add: Values_recs)
       
  1005 apply(simp add: Values_recs)
       
  1006 apply(auto)[1]
       
  1007 done
       
  1008 
       
  1009 lemma Values_injval:
       
  1010   assumes "v \<in> Values (der c r) s"
       
  1011   shows "injval r c v \<in> Values r (c#s)"
       
  1012 using assms
       
  1013 apply(induct c r arbitrary: v s rule: der.induct)
       
  1014 apply(simp add: Values_recs)
       
  1015 apply(simp add: Values_recs)
       
  1016 apply(case_tac "c = c'")
       
  1017 apply(simp)
       
  1018 apply(simp add: Values_recs)
       
  1019 apply(simp add: prefix_def)
       
  1020 apply(simp)
       
  1021 apply(simp add: Values_recs)
       
  1022 apply(simp)
       
  1023 apply(simp add: Values_recs)
       
  1024 apply(auto)[1]
       
  1025 apply(case_tac "nullable r1")
       
  1026 apply(simp)
       
  1027 apply(simp add: Values_recs)
       
  1028 apply(auto)[1]
       
  1029 apply(simp add: rest_def)
       
  1030 apply(subst v4)
       
  1031 apply(simp add: Values_def)
       
  1032 apply(simp add: Values_def)
       
  1033 apply(rule Values_nullable)
       
  1034 apply(assumption)
       
  1035 apply(simp add: rest_def)
       
  1036 apply(subst mkeps_flat)
       
  1037 apply(assumption)
       
  1038 apply(simp)
       
  1039 apply(simp)
       
  1040 apply(simp add: Values_recs)
       
  1041 apply(auto)[1]
       
  1042 apply(simp add: rest_def)
       
  1043 apply(subst v4)
       
  1044 apply(simp add: Values_def)
       
  1045 apply(simp add: Values_def)
       
  1046 done
       
  1047 
       
  1048 lemma Values_projval:
       
  1049   assumes "v \<in> Values r (c#s)" "\<exists>s. flat v = c # s"
       
  1050   shows "projval r c v \<in> Values (der c r) s"
       
  1051 using assms
       
  1052 apply(induct r arbitrary: v s c rule: rexp.induct)
       
  1053 apply(simp add: Values_recs)
       
  1054 apply(simp add: Values_recs)
       
  1055 apply(case_tac "c = char")
       
  1056 apply(simp)
       
  1057 apply(simp add: Values_recs)
       
  1058 apply(simp)
       
  1059 apply(simp add: Values_recs)
       
  1060 apply(simp add: prefix_def)
       
  1061 apply(case_tac "nullable rexp1")
       
  1062 apply(simp)
       
  1063 apply(simp add: Values_recs)
       
  1064 apply(auto)[1]
       
  1065 apply(simp add: rest_def)
       
  1066 apply (metis hd_Cons_tl hd_append2 list.sel(1))
       
  1067 apply(simp add: rest_def)
       
  1068 apply(simp add: append_eq_Cons_conv)
       
  1069 apply(auto)[1]
       
  1070 apply(subst v4_proj2)
       
  1071 apply(simp add: Values_def)
       
  1072 apply(assumption)
       
  1073 apply(simp)
       
  1074 apply(simp)
       
  1075 apply(simp add: Values_recs)
       
  1076 apply(auto)[1]
       
  1077 apply(auto simp add: Values_def not_nullable_flat)[1]
       
  1078 apply(simp add: append_eq_Cons_conv)
       
  1079 apply(auto)[1]
       
  1080 apply(simp add: append_eq_Cons_conv)
       
  1081 apply(auto)[1]
       
  1082 apply(simp add: rest_def)
       
  1083 apply(subst v4_proj2)
       
  1084 apply(simp add: Values_def)
       
  1085 apply(assumption)
       
  1086 apply(simp)
       
  1087 apply(simp add: Values_recs)
       
  1088 apply(auto)[1]
       
  1089 done
       
  1090 
       
  1091 
       
  1092 definition "MValue v r s \<equiv> (v \<in> Values r s \<and> (\<forall>v' \<in> Values r s. v 2\<succ> v'))"
       
  1093 
       
  1094 lemma MValue_ALTE:
       
  1095   assumes "MValue v (ALT r1 r2) s"
       
  1096   shows "(\<exists>vl. v = Left vl \<and> MValue vl r1 s \<and> (\<forall>vr \<in> Values r2 s. length (flat vr) \<le> length (flat vl))) \<or> 
       
  1097          (\<exists>vr. v = Right vr \<and> MValue vr r2 s \<and> (\<forall>vl \<in> Values r1 s. length (flat vl) < length (flat vr)))"
       
  1098 using assms
       
  1099 apply(simp add: MValue_def)
       
  1100 apply(simp add: Values_recs)
       
  1101 apply(auto)
       
  1102 apply(drule_tac x="Left x" in bspec)
       
  1103 apply(simp)
       
  1104 apply(erule ValOrd2.cases)
       
  1105 apply(simp_all)
       
  1106 apply(drule_tac x="Right vr" in bspec)
       
  1107 apply(simp)
       
  1108 apply(erule ValOrd2.cases)
       
  1109 apply(simp_all)
       
  1110 apply(drule_tac x="Right x" in bspec)
       
  1111 apply(simp)
       
  1112 apply(erule ValOrd2.cases)
       
  1113 apply(simp_all)
       
  1114 apply(drule_tac x="Left vl" in bspec)
       
  1115 apply(simp)
       
  1116 apply(erule ValOrd2.cases)
       
  1117 apply(simp_all)
       
  1118 done
       
  1119 
       
  1120 lemma MValue_ALTI1:
       
  1121   assumes "MValue vl r1 s"  "\<forall>vr \<in> Values r2 s. length (flat vr) \<le> length (flat vl)"
       
  1122   shows "MValue (Left vl) (ALT r1 r2) s"
       
  1123 using assms
       
  1124 apply(simp add: MValue_def)
       
  1125 apply(simp add: Values_recs)
       
  1126 apply(auto)
       
  1127 apply(rule ValOrd2.intros)
       
  1128 apply metis
       
  1129 apply(rule ValOrd2.intros)
       
  1130 apply metis
       
  1131 done
       
  1132 
       
  1133 lemma MValue_ALTI2:
       
  1134   assumes "MValue vr r2 s"  "\<forall>vl \<in> Values r1 s. length (flat vl) < length (flat vr)"
       
  1135   shows "MValue (Right vr) (ALT r1 r2) s"
       
  1136 using assms
       
  1137 apply(simp add: MValue_def)
       
  1138 apply(simp add: Values_recs)
       
  1139 apply(auto)
       
  1140 apply(rule ValOrd2.intros)
       
  1141 apply metis
       
  1142 apply(rule ValOrd2.intros)
       
  1143 apply metis
       
  1144 done
       
  1145 
       
  1146 lemma t: "(c#xs = c#ys) \<Longrightarrow> xs = ys"
       
  1147 by (metis list.sel(3))
       
  1148 
       
  1149 lemma t2: "(xs = ys) \<Longrightarrow> (c#xs) = (c#ys)"
       
  1150 by (metis)
       
  1151 
       
  1152 lemma "\<not>(nullable r) \<Longrightarrow> \<not>(\<exists>v. \<turnstile> v : r \<and> flat v = [])"
       
  1153 by (metis Prf_flat_L nullable_correctness)
       
  1154 
       
  1155 
       
  1156 lemma LeftRight:
       
  1157   assumes "(Left v1) \<succ>(der c (ALT r1 r2)) (Right v2)"
       
  1158   and "\<turnstile> v1 : der c r1" "\<turnstile> v2 : der c r2" 
       
  1159   shows "(injval (ALT r1 r2) c (Left v1)) \<succ>(ALT r1 r2) (injval (ALT r1 r2) c (Right v2))"
       
  1160 using assms
       
  1161 apply(simp)
       
  1162 apply(erule ValOrd.cases)
       
  1163 apply(simp_all)[8]
       
  1164 apply(rule ValOrd.intros)
       
  1165 apply(clarify)
       
  1166 apply(subst v4)
       
  1167 apply(simp)
       
  1168 apply(subst v4)
       
  1169 apply(simp)
       
  1170 apply(simp)
       
  1171 done
       
  1172 
       
  1173 lemma RightLeft:
       
  1174   assumes "(Right v1) \<succ>(der c (ALT r1 r2)) (Left v2)"
       
  1175   and "\<turnstile> v1 : der c r2" "\<turnstile> v2 : der c r1" 
       
  1176   shows "(injval (ALT r1 r2) c (Right v1)) \<succ>(ALT r1 r2) (injval (ALT r1 r2) c (Left v2))"
       
  1177 using assms
       
  1178 apply(simp)
       
  1179 apply(erule ValOrd.cases)
       
  1180 apply(simp_all)[8]
       
  1181 apply(rule ValOrd.intros)
       
  1182 apply(clarify)
       
  1183 apply(subst v4)
       
  1184 apply(simp)
       
  1185 apply(subst v4)
       
  1186 apply(simp)
       
  1187 apply(simp)
       
  1188 done
       
  1189 
       
  1190 lemma h: 
       
  1191   assumes "nullable r1" "\<turnstile> v1 : der c r1"
       
  1192   shows "injval r1 c v1 \<succ>r1 mkeps r1"
       
  1193 using assms
       
  1194 apply(induct r1 arbitrary: v1 rule: der.induct)
       
  1195 apply(simp)
       
  1196 apply(simp)
       
  1197 apply(erule Prf.cases)
       
  1198 apply(simp_all)[5]
       
  1199 apply(simp)
       
  1200 apply(simp)
       
  1201 apply(erule Prf.cases)
       
  1202 apply(simp_all)[5]
       
  1203 apply(clarify)
       
  1204 apply(auto)[1]
       
  1205 apply (metis ValOrd.intros(6))
       
  1206 apply (metis ValOrd.intros(6))
       
  1207 apply (metis ValOrd.intros(3) le_add2 list.size(3) mkeps_flat monoid_add_class.add.right_neutral)
       
  1208 apply(auto)[1]
       
  1209 apply (metis ValOrd.intros(4) length_greater_0_conv list.distinct(1) list.size(3) mkeps_flat v4)
       
  1210 apply (metis ValOrd.intros(4) length_greater_0_conv list.distinct(1) list.size(3) mkeps_flat v4)
       
  1211 apply (metis ValOrd.intros(5))
       
  1212 apply(simp)
       
  1213 apply(erule Prf.cases)
       
  1214 apply(simp_all)[5]
       
  1215 apply(clarify)
       
  1216 apply(erule Prf.cases)
       
  1217 apply(simp_all)[5]
       
  1218 apply(clarify)
       
  1219 apply (metis ValOrd.intros(2) list.distinct(1) mkeps_flat v4)
       
  1220 apply(clarify)
       
  1221 by (metis ValOrd.intros(1))
       
  1222 
       
  1223 lemma LeftRightSeq:
       
  1224   assumes "(Left (Seq v1 v2)) \<succ>(der c (SEQ r1 r2)) (Right v3)"
       
  1225   and "nullable r1" "\<turnstile> v1 : der c r1"
       
  1226   shows "(injval (SEQ r1 r2) c (Seq v1 v2)) \<succ>(SEQ r1 r2) (injval (SEQ r1 r2) c (Right v2))"
       
  1227 using assms
       
  1228 apply(simp)
       
  1229 apply(erule ValOrd.cases)
       
  1230 apply(simp_all)[8]
       
  1231 apply(clarify)
       
  1232 apply(simp)
       
  1233 apply(rule ValOrd.intros(2))
       
  1234 prefer 2
       
  1235 apply (metis list.distinct(1) mkeps_flat v4)
       
  1236 by (metis h)
       
  1237 
       
  1238 lemma rr1: 
       
  1239   assumes "\<turnstile> v : r" "\<not>nullable r" 
       
  1240   shows "flat v \<noteq> []"
       
  1241 using assms
       
  1242 by (metis Prf_flat_L nullable_correctness)
       
  1243 
       
  1244 (* HERE *)
       
  1245 
       
  1246 lemma Prf_inj_test:
       
  1247   assumes "v1 \<succ>(der c r) v2" 
       
  1248           "v1 \<in> Values (der c r) s"
       
  1249           "v2 \<in> Values (der c r) s"
       
  1250           "injval r c v1 \<in> Values r (c#s)"
       
  1251           "injval r c v2 \<in> Values r (c#s)"
       
  1252   shows "(injval r c v1) 2\<succ>  (injval r c v2)"
       
  1253 using assms
       
  1254 apply(induct c r arbitrary: v1 v2 s rule: der.induct)
       
  1255 (* NULL case *)
       
  1256 apply(simp add: Values_recs)
       
  1257 (* EMPTY case *)
       
  1258 apply(simp add: Values_recs)
       
  1259 (* CHAR case *)
       
  1260 apply(case_tac "c = c'")
       
  1261 apply(simp)
       
  1262 apply(simp add: Values_recs)
       
  1263 apply (metis ValOrd2.intros(8))
       
  1264 apply(simp add: Values_recs)
       
  1265 (* ALT case *)
       
  1266 apply(simp)
       
  1267 apply(simp add: Values_recs)
       
  1268 apply(auto)[1]
       
  1269 apply(erule ValOrd.cases)
       
  1270 apply(simp_all)[8]
       
  1271 apply (metis ValOrd2.intros(6))
       
  1272 apply(erule ValOrd.cases)
       
  1273 apply(simp_all)[8]
       
  1274 apply(rule ValOrd2.intros)
       
  1275 apply(subst v4)
       
  1276 apply(simp add: Values_def)
       
  1277 apply(subst v4)
       
  1278 apply(simp add: Values_def)
       
  1279 apply(simp)
       
  1280 apply(erule ValOrd.cases)
       
  1281 apply(simp_all)[8]
       
  1282 apply(rule ValOrd2.intros)
       
  1283 apply(subst v4)
       
  1284 apply(simp add: Values_def)
       
  1285 apply(subst v4)
       
  1286 apply(simp add: Values_def)
       
  1287 apply(simp)
       
  1288 apply(erule ValOrd.cases)
       
  1289 apply(simp_all)[8]
       
  1290 apply (metis ValOrd2.intros(5))
       
  1291 (* SEQ case*)
       
  1292 apply(simp)
       
  1293 apply(case_tac "nullable r1")
       
  1294 apply(simp)
       
  1295 defer
       
  1296 apply(simp)
       
  1297 apply(simp add: Values_recs)
       
  1298 apply(auto)[1]
       
  1299 apply(erule ValOrd.cases)
       
  1300 apply(simp_all)[8]
       
  1301 apply(clarify)
       
  1302 apply(rule ValOrd2.intros)
       
  1303 apply(simp)
       
  1304 apply (metis Ord1)
       
  1305 apply(clarify)
       
  1306 apply(rule ValOrd2.intros)
       
  1307 apply(subgoal_tac "rest v1 (flat v1 @ flat v2) = flat v2")
       
  1308 apply(simp)
       
  1309 apply(subgoal_tac "rest (injval r1 c v1) (c # flat v1 @ flat v2) = flat v2")
       
  1310 apply(simp)
       
  1311 oops
       
  1312 
       
  1313 lemma Prf_inj_test:
       
  1314   assumes "v1 \<succ>(der c r) v2" 
       
  1315           "v1 \<in> Values (der c r) s"
       
  1316           "v2 \<in> Values (der c r) s"
       
  1317           "injval r c v1 \<in> Values r (c#s)"
       
  1318           "injval r c v2 \<in> Values r (c#s)"
       
  1319   shows "(injval r c v1) 2\<succ>  (injval r c v2)"
       
  1320 using assms
       
  1321 apply(induct c r arbitrary: v1 v2 s rule: der.induct)
       
  1322 (* NULL case *)
       
  1323 apply(simp add: Values_recs)
       
  1324 (* EMPTY case *)
       
  1325 apply(simp add: Values_recs)
       
  1326 (* CHAR case *)
       
  1327 apply(case_tac "c = c'")
       
  1328 apply(simp)
       
  1329 apply(simp add: Values_recs)
       
  1330 apply (metis ValOrd2.intros(8))
       
  1331 apply(simp add: Values_recs)
       
  1332 (* ALT case *)
       
  1333 apply(simp)
       
  1334 apply(simp add: Values_recs)
       
  1335 apply(auto)[1]
       
  1336 apply(erule ValOrd.cases)
       
  1337 apply(simp_all)[8]
       
  1338 apply (metis ValOrd2.intros(6))
       
  1339 apply(erule ValOrd.cases)
       
  1340 apply(simp_all)[8]
       
  1341 apply(rule ValOrd2.intros)
       
  1342 apply(subst v4)
       
  1343 apply(simp add: Values_def)
       
  1344 apply(subst v4)
       
  1345 apply(simp add: Values_def)
       
  1346 apply(simp)
       
  1347 apply(erule ValOrd.cases)
       
  1348 apply(simp_all)[8]
       
  1349 apply(rule ValOrd2.intros)
       
  1350 apply(subst v4)
       
  1351 apply(simp add: Values_def)
       
  1352 apply(subst v4)
       
  1353 apply(simp add: Values_def)
       
  1354 apply(simp)
       
  1355 apply(erule ValOrd.cases)
       
  1356 apply(simp_all)[8]
       
  1357 apply (metis ValOrd2.intros(5))
       
  1358 (* SEQ case*)
       
  1359 apply(simp)
       
  1360 apply(case_tac "nullable r1")
       
  1361 apply(simp)
       
  1362 defer
       
  1363 apply(simp)
       
  1364 apply(simp add: Values_recs)
       
  1365 apply(auto)[1]
       
  1366 apply(erule ValOrd.cases)
       
  1367 apply(simp_all)[8]
       
  1368 apply(clarify)
       
  1369 apply(rule ValOrd2.intros)
       
  1370 apply(simp)
       
  1371 apply (metis Ord1)
       
  1372 apply(clarify)
       
  1373 apply(rule ValOrd2.intros)
       
  1374 apply metis
       
  1375 using injval_inj
       
  1376 apply(simp add: Values_def inj_on_def)
       
  1377 apply metis
       
  1378 apply(simp add: Values_recs)
       
  1379 apply(auto)[1]
       
  1380 apply(erule ValOrd.cases)
       
  1381 apply(simp_all)[8]
       
  1382 apply(clarify)
       
  1383 apply(erule ValOrd.cases)
       
  1384 apply(simp_all)[8]
       
  1385 apply(clarify)
       
  1386 apply (metis Ord1 ValOrd2.intros(1))
       
  1387 apply(clarify)
       
  1388 apply(rule ValOrd2.intros(2))
       
  1389 apply metis
       
  1390 using injval_inj
       
  1391 apply(simp add: Values_def inj_on_def)
       
  1392 apply metis
       
  1393 apply(erule ValOrd.cases)
       
  1394 apply(simp_all)[8]
       
  1395 apply(rule ValOrd2.intros(2))
       
  1396 thm h
       
  1397 apply(rule Ord1)
       
  1398 apply(rule h)
       
  1399 apply(simp)
       
  1400 apply(simp add: Values_def)
       
  1401 apply(simp add: Values_def)
       
  1402 apply (metis list.distinct(1) mkeps_flat v4)
       
  1403 apply(erule ValOrd.cases)
       
  1404 apply(simp_all)[8]
       
  1405 apply(clarify)
       
  1406 apply(simp add: Values_def)
       
  1407 defer
       
  1408 apply(erule ValOrd.cases)
       
  1409 apply(simp_all)[8]
       
  1410 apply(clarify)
       
  1411 apply(rule ValOrd2.intros(1))
       
  1412 apply(rotate_tac 1)
       
  1413 apply(drule_tac x="v2" in meta_spec)
       
  1414 apply(rotate_tac 8)
       
  1415 apply(drule_tac x="v2'" in meta_spec)
       
  1416 apply(rotate_tac 8)
       
  1417 oops
       
  1418 
       
  1419 lemma POSIX_der:
       
  1420   assumes "POSIX v (der c r)" "\<turnstile> v : der c r"
       
  1421   shows "POSIX (injval r c v) r"
       
  1422 using assms
       
  1423 unfolding POSIX_def
       
  1424 apply(auto)
       
  1425 thm v3
       
  1426 apply (erule v3)
       
  1427 thm v4
       
  1428 apply(subst (asm) v4)
       
  1429 apply(assumption)
       
  1430 apply(drule_tac x="projval r c v'" in spec)
       
  1431 apply(drule mp)
       
  1432 apply(rule conjI)
       
  1433 thm v3_proj
       
  1434 apply(rule v3_proj)
       
  1435 apply(simp)
       
  1436 apply(rule_tac x="flat v" in exI)
       
  1437 apply(simp)
       
  1438 thm t
       
  1439 apply(rule_tac c="c" in  t)
       
  1440 apply(simp)
       
  1441 thm v4_proj
       
  1442 apply(subst v4_proj)
       
  1443 apply(simp)
       
  1444 apply(rule_tac x="flat v" in exI)
       
  1445 apply(simp)
       
  1446 apply(simp)
       
  1447 oops
       
  1448 
       
  1449 lemma POSIX_der:
       
  1450   assumes "POSIX v (der c r)" "\<turnstile> v : der c r"
       
  1451   shows "POSIX (injval r c v) r"
       
  1452 using assms
       
  1453 apply(induct c r arbitrary: v rule: der.induct)
       
  1454 (* null case*)
       
  1455 apply(simp add: POSIX_def)
       
  1456 apply(auto)[1]
       
  1457 apply(erule Prf.cases)
       
  1458 apply(simp_all)[5]
       
  1459 apply(erule Prf.cases)
       
  1460 apply(simp_all)[5]
       
  1461 (* empty case *)
       
  1462 apply(simp add: POSIX_def)
       
  1463 apply(auto)[1]
       
  1464 apply(erule Prf.cases)
       
  1465 apply(simp_all)[5]
       
  1466 apply(erule Prf.cases)
       
  1467 apply(simp_all)[5]
       
  1468 (* char case *)
       
  1469 apply(simp add: POSIX_def)
       
  1470 apply(case_tac "c = c'")
       
  1471 apply(auto)[1]
       
  1472 apply(erule Prf.cases)
       
  1473 apply(simp_all)[5]
       
  1474 apply (metis Prf.intros(5))
       
  1475 apply(erule Prf.cases)
       
  1476 apply(simp_all)[5]
       
  1477 apply(erule Prf.cases)
       
  1478 apply(simp_all)[5]
       
  1479 apply (metis ValOrd.intros(8))
       
  1480 apply(auto)[1]
       
  1481 apply(erule Prf.cases)
       
  1482 apply(simp_all)[5]
       
  1483 apply(erule Prf.cases)
       
  1484 apply(simp_all)[5]
       
  1485 (* alt case *)
       
  1486 apply(erule Prf.cases)
       
  1487 apply(simp_all)[5]
       
  1488 apply(clarify)
       
  1489 apply(simp (no_asm) add: POSIX_def)
       
  1490 apply(auto)[1]
       
  1491 apply (metis Prf.intros(2) v3)
       
  1492 apply(rotate_tac 4)
       
  1493 apply(erule Prf.cases)
       
  1494 apply(simp_all)[5]
       
  1495 apply (metis POSIX_ALT2 POSIX_def ValOrd.intros(6))
       
  1496 apply (metis ValOrd.intros(3) order_refl)
       
  1497 apply(simp (no_asm) add: POSIX_def)
       
  1498 apply(auto)[1]
       
  1499 apply (metis Prf.intros(3) v3)
       
  1500 apply(rotate_tac 4)
       
  1501 apply(erule Prf.cases)
       
  1502 apply(simp_all)[5]
       
  1503 defer
       
  1504 apply (metis POSIX_ALT1a POSIX_def ValOrd.intros(5))
       
  1505 prefer 2
       
  1506 apply(subst (asm) (5) POSIX_def)
       
  1507 apply(auto)[1]
       
  1508 apply(rotate_tac 5)
       
  1509 apply(erule Prf.cases)
       
  1510 apply(simp_all)[5]
       
  1511 apply(rule ValOrd.intros)
       
  1512 apply(subst (asm) v4)
       
  1513 apply(simp)
       
  1514 apply(drule_tac x="Left (projval r1a c v1)" in spec)
       
  1515 apply(clarify)
       
  1516 apply(drule mp)
       
  1517 apply(rule conjI)
       
  1518 apply (metis Prf.intros(2) v3_proj)
       
  1519 apply(simp)
       
  1520 apply (metis v4_proj2)
       
  1521 apply(erule ValOrd.cases)
       
  1522 apply(simp_all)[8]
       
  1523 apply (metis less_not_refl v4_proj2)
       
  1524 (* seq case *)
       
  1525 apply(case_tac "nullable r1")
       
  1526 defer
       
  1527 apply(simp add: POSIX_def)
       
  1528 apply(auto)[1]
       
  1529 apply(erule Prf.cases)
       
  1530 apply(simp_all)[5]
       
  1531 apply (metis Prf.intros(1) v3)
       
  1532 apply(erule Prf.cases)
       
  1533 apply(simp_all)[5]
       
  1534 apply(erule Prf.cases)
       
  1535 apply(simp_all)[5]
       
  1536 apply(clarify)
       
  1537 apply(subst (asm) (3) v4)
       
  1538 apply(simp)
       
  1539 apply(simp)
       
  1540 apply(subgoal_tac "flat v1a \<noteq> []")
       
  1541 prefer 2
       
  1542 apply (metis Prf_flat_L nullable_correctness)
       
  1543 apply(subgoal_tac "\<exists>s. flat v1a = c # s")
       
  1544 prefer 2
       
  1545 apply (metis append_eq_Cons_conv)
       
  1546 apply(auto)[1]
       
  1547 oops
       
  1548 
       
  1549 
       
  1550 lemma POSIX_ex: "\<turnstile> v : r \<Longrightarrow> \<exists>v. POSIX v r"
       
  1551 apply(induct r arbitrary: v)
       
  1552 apply(erule Prf.cases)
       
  1553 apply(simp_all)[5]
       
  1554 apply(erule Prf.cases)
       
  1555 apply(simp_all)[5]
       
  1556 apply(rule_tac x="Void" in exI)
       
  1557 apply(simp add: POSIX_def)
       
  1558 apply(auto)[1]
       
  1559 apply (metis Prf.intros(4))
       
  1560 apply(erule Prf.cases)
       
  1561 apply(simp_all)[5]
       
  1562 apply (metis ValOrd.intros(7))
       
  1563 apply(erule Prf.cases)
       
  1564 apply(simp_all)[5]
       
  1565 apply(rule_tac x="Char c" in exI)
       
  1566 apply(simp add: POSIX_def)
       
  1567 apply(auto)[1]
       
  1568 apply (metis Prf.intros(5))
       
  1569 apply(erule Prf.cases)
       
  1570 apply(simp_all)[5]
       
  1571 apply (metis ValOrd.intros(8))
       
  1572 apply(erule Prf.cases)
       
  1573 apply(simp_all)[5]
       
  1574 apply(auto)[1]
       
  1575 apply(drule_tac x="v1" in meta_spec)
       
  1576 apply(drule_tac x="v2" in meta_spec)
       
  1577 apply(auto)[1]
       
  1578 defer
       
  1579 apply(erule Prf.cases)
       
  1580 apply(simp_all)[5]
       
  1581 apply(auto)[1]
       
  1582 apply (metis POSIX_ALT_I1)
       
  1583 apply (metis POSIX_ALT_I1 POSIX_ALT_I2)
       
  1584 apply(case_tac "nullable r1a")
       
  1585 apply(rule_tac x="Seq (mkeps r1a) va" in exI)
       
  1586 apply(auto simp add: POSIX_def)[1]
       
  1587 apply (metis Prf.intros(1) mkeps_nullable)
       
  1588 apply(simp add: mkeps_flat)
       
  1589 apply(rotate_tac 7)
       
  1590 apply(erule Prf.cases)
       
  1591 apply(simp_all)[5]
       
  1592 apply(case_tac "mkeps r1 = v1a")
       
  1593 apply(simp)
       
  1594 apply (rule ValOrd.intros(1))
       
  1595 apply (metis append_Nil mkeps_flat)
       
  1596 apply (rule ValOrd.intros(2))
       
  1597 apply(drule mkeps_POSIX)
       
  1598 apply(simp add: POSIX_def)
       
  1599 oops
       
  1600 
       
  1601 lemma POSIX_ex2: "\<turnstile> v : r \<Longrightarrow> \<exists>v. POSIX v r \<and> \<turnstile> v : r"
       
  1602 apply(induct r arbitrary: v)
       
  1603 apply(erule Prf.cases)
       
  1604 apply(simp_all)[5]
       
  1605 apply(erule Prf.cases)
       
  1606 apply(simp_all)[5]
       
  1607 apply(rule_tac x="Void" in exI)
       
  1608 apply(simp add: POSIX_def)
       
  1609 apply(auto)[1]
       
  1610 oops
       
  1611 
       
  1612 lemma POSIX_ALT_cases:
       
  1613   assumes "\<turnstile> v : (ALT r1 r2)" "POSIX v (ALT r1 r2)"
       
  1614   shows "(\<exists>v1. v = Left v1 \<and> POSIX v1 r1) \<or> (\<exists>v2. v = Right v2 \<and> POSIX v2 r2)"
       
  1615 using assms
       
  1616 apply(erule_tac Prf.cases)
       
  1617 apply(simp_all)
       
  1618 unfolding POSIX_def
       
  1619 apply(auto)
       
  1620 apply (metis POSIX_ALT2 POSIX_def assms(2))
       
  1621 by (metis POSIX_ALT1b assms(2))
       
  1622 
       
  1623 lemma POSIX_ALT_cases2:
       
  1624   assumes "POSIX v (ALT r1 r2)" "\<turnstile> v : (ALT r1 r2)" 
       
  1625   shows "(\<exists>v1. v = Left v1 \<and> POSIX v1 r1) \<or> (\<exists>v2. v = Right v2 \<and> POSIX v2 r2)"
       
  1626 using assms POSIX_ALT_cases by auto
       
  1627 
       
  1628 lemma Prf_flat_empty:
       
  1629   assumes "\<turnstile> v : r" "flat v = []"
       
  1630   shows "nullable r"
       
  1631 using assms
       
  1632 apply(induct)
       
  1633 apply(auto)
       
  1634 done
       
  1635 
       
  1636 lemma POSIX_proj:
       
  1637   assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
       
  1638   shows "POSIX (projval r c v) (der c r)"
       
  1639 using assms
       
  1640 apply(induct r c v arbitrary: rule: projval.induct)
       
  1641 defer
       
  1642 defer
       
  1643 defer
       
  1644 defer
       
  1645 apply(erule Prf.cases)
       
  1646 apply(simp_all)[5]
       
  1647 apply(erule Prf.cases)
       
  1648 apply(simp_all)[5]
       
  1649 apply(erule Prf.cases)
       
  1650 apply(simp_all)[5]
       
  1651 apply(erule Prf.cases)
       
  1652 apply(simp_all)[5]
       
  1653 apply(erule Prf.cases)
       
  1654 apply(simp_all)[5]
       
  1655 apply(erule Prf.cases)
       
  1656 apply(simp_all)[5]
       
  1657 apply(erule Prf.cases)
       
  1658 apply(simp_all)[5]
       
  1659 apply(erule Prf.cases)
       
  1660 apply(simp_all)[5]
       
  1661 apply(erule Prf.cases)
       
  1662 apply(simp_all)[5]
       
  1663 apply(erule Prf.cases)
       
  1664 apply(simp_all)[5]
       
  1665 apply(simp add: POSIX_def)
       
  1666 apply(auto)[1]
       
  1667 apply(erule Prf.cases)
       
  1668 apply(simp_all)[5]
       
  1669 oops
       
  1670 
       
  1671 lemma POSIX_proj:
       
  1672   assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
       
  1673   shows "POSIX (projval r c v) (der c r)"
       
  1674 using assms
       
  1675 apply(induct r arbitrary: c v rule: rexp.induct)
       
  1676 apply(erule Prf.cases)
       
  1677 apply(simp_all)[5]
       
  1678 apply(erule Prf.cases)
       
  1679 apply(simp_all)[5]
       
  1680 apply(erule Prf.cases)
       
  1681 apply(simp_all)[5]
       
  1682 apply(simp add: POSIX_def)
       
  1683 apply(auto)[1]
       
  1684 apply(erule Prf.cases)
       
  1685 apply(simp_all)[5]
       
  1686 oops
       
  1687 
       
  1688 lemma POSIX_proj:
       
  1689   assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
       
  1690   shows "POSIX (projval r c v) (der c r)"
       
  1691 using assms
       
  1692 apply(induct r c v arbitrary: rule: projval.induct)
       
  1693 defer
       
  1694 defer
       
  1695 defer
       
  1696 defer
       
  1697 apply(erule Prf.cases)
       
  1698 apply(simp_all)[5]
       
  1699 apply(erule Prf.cases)
       
  1700 apply(simp_all)[5]
       
  1701 apply(erule Prf.cases)
       
  1702 apply(simp_all)[5]
       
  1703 apply(erule Prf.cases)
       
  1704 apply(simp_all)[5]
       
  1705 apply(erule Prf.cases)
       
  1706 apply(simp_all)[5]
       
  1707 apply(erule Prf.cases)
       
  1708 apply(simp_all)[5]
       
  1709 apply(erule Prf.cases)
       
  1710 apply(simp_all)[5]
       
  1711 apply(erule Prf.cases)
       
  1712 apply(simp_all)[5]
       
  1713 apply(erule Prf.cases)
       
  1714 apply(simp_all)[5]
       
  1715 apply(erule Prf.cases)
       
  1716 apply(simp_all)[5]
       
  1717 apply(simp add: POSIX_def)
       
  1718 apply(auto)[1]
       
  1719 apply(erule Prf.cases)
       
  1720 apply(simp_all)[5]
       
  1721 oops
       
  1722 
       
  1723 lemma Prf_inj:
       
  1724   assumes "v1 \<succ>(der c r) v2" "\<turnstile> v1 : der c r" "\<turnstile> v2 : der c r" "flat v1 = flat v2"
       
  1725   shows "(injval r c v1) \<succ>r (injval r c v2)"
       
  1726 using assms
       
  1727 apply(induct arbitrary: v1 v2 rule: der.induct)
       
  1728 (* NULL case *)
       
  1729 apply(simp)
       
  1730 apply(erule ValOrd.cases)
       
  1731 apply(simp_all)[8]
       
  1732 (* EMPTY case *)
       
  1733 apply(erule ValOrd.cases)
       
  1734 apply(simp_all)[8]
       
  1735 (* CHAR case *)
       
  1736 apply(case_tac "c = c'")
       
  1737 apply(simp)
       
  1738 apply(erule ValOrd.cases)
       
  1739 apply(simp_all)[8]
       
  1740 apply(rule ValOrd.intros)
       
  1741 apply(simp)
       
  1742 apply(erule ValOrd.cases)
       
  1743 apply(simp_all)[8]
       
  1744 (* ALT case *)
       
  1745 apply(simp)
       
  1746 apply(erule ValOrd.cases)
       
  1747 apply(simp_all)[8]
       
  1748 apply(rule ValOrd.intros)
       
  1749 apply(subst v4)
       
  1750 apply(clarify)
       
  1751 apply(rotate_tac 3)
       
  1752 apply(erule Prf.cases)
       
  1753 apply(simp_all)[5]
       
  1754 apply(subst v4)
       
  1755 apply(clarify)
       
  1756 apply(rotate_tac 2)
       
  1757 apply(erule Prf.cases)
       
  1758 apply(simp_all)[5]
       
  1759 apply(simp)
       
  1760 apply(rule ValOrd.intros)
       
  1761 apply(clarify)
       
  1762 apply(rotate_tac 3)
       
  1763 apply(erule Prf.cases)
       
  1764 apply(simp_all)[5]
       
  1765 apply(clarify)
       
  1766 apply(erule Prf.cases)
       
  1767 apply(simp_all)[5]
       
  1768 apply(rule ValOrd.intros)
       
  1769 apply(clarify)
       
  1770 apply(erule Prf.cases)
       
  1771 apply(simp_all)[5]
       
  1772 apply(erule Prf.cases)
       
  1773 apply(simp_all)[5]
       
  1774 (* SEQ case*)
       
  1775 apply(simp)
       
  1776 apply(case_tac "nullable r1")
       
  1777 defer
       
  1778 apply(simp)
       
  1779 apply(erule ValOrd.cases)
       
  1780 apply(simp_all)[8]
       
  1781 apply(clarify)
       
  1782 apply(erule Prf.cases)
       
  1783 apply(simp_all)[5]
       
  1784 apply(erule Prf.cases)
       
  1785 apply(simp_all)[5]
       
  1786 apply(clarify)
       
  1787 apply(rule ValOrd.intros)
       
  1788 apply(simp)
       
  1789 oops
       
  1790 
   456 
  1791 
   457 text {*
  1792 text {*
   458   Injection followed by projection is the identity.
  1793   Injection followed by projection is the identity.
   459 *}
  1794 *}
   460 
  1795 
   508 lemma v5:
  1843 lemma v5:
   509   assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
  1844   assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
   510   shows "POSIX (injval r c v) r"
  1845   shows "POSIX (injval r c v) r"
   511 using assms
  1846 using assms
   512 apply(induct arbitrary: v rule: der.induct)
  1847 apply(induct arbitrary: v rule: der.induct)
   513 apply(simp)
  1848 (* NULL case *)
   514 apply(erule Prf.cases)
  1849 apply(simp)
   515 apply(simp_all)[5]
  1850 apply(erule Prf.cases)
   516 apply(simp)
  1851 apply(simp_all)[5]
   517 apply(erule Prf.cases)
  1852 (* EMPTY case *)
   518 apply(simp_all)[5]
  1853 apply(simp)
       
  1854 apply(erule Prf.cases)
       
  1855 apply(simp_all)[5]
       
  1856 (* CHAR case *)
   519 apply(simp)
  1857 apply(simp)
   520 apply(case_tac "c = c'")
  1858 apply(case_tac "c = c'")
   521 apply(auto simp add: POSIX_def)[1]
  1859 apply(auto simp add: POSIX_def)[1]
   522 apply(erule Prf.cases)
  1860 apply(erule Prf.cases)
   523 apply(simp_all)[5]
  1861 apply(simp_all)[5]
   524 apply(erule Prf.cases)
  1862 oops
   525 apply(simp_all)[5]
       
   526 using ValOrd.simps apply blast
       
   527 apply(auto)
       
   528 apply(erule Prf.cases)
       
   529 apply(simp_all)[5]
       
   530 (* base cases done *)
       
   531 (* ALT case *)
       
   532 apply(erule Prf.cases)
       
   533 apply(simp_all)[5]
       
   534 using POSIX_ALT POSIX_ALT_I1 apply blast
       
   535 apply(frule POSIX_ALT1a)
       
   536 apply(drule POSIX_ALT1b)
       
   537 apply(rule POSIX_ALT_I2)
       
   538 apply auto[1]
       
   539 apply(subst v4)
       
   540 apply(auto)[2]
       
   541 apply(rotate_tac 1)
       
   542 apply(drule_tac x="v2" in meta_spec)
       
   543 apply(simp)
       
   544 apply(subst (asm) (4) POSIX_def)
       
   545 apply(subst (asm) v4)
       
   546 apply(auto)[2]
       
   547 (* stuck in the ALT case *)