cleaned up version of Re1
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Thu, 17 Dec 2015 14:16:24 +0000
changeset 82 26202889f829
parent 81 7ac7782a7318
child 83 a8bcb5a0f9b9
cleaned up version of Re1
thys/Re.thy
--- a/thys/Re.thy	Thu Dec 17 13:14:36 2015 +0000
+++ b/thys/Re.thy	Thu Dec 17 14:16:24 2015 +0000
@@ -1,9 +1,9 @@
-(*test*)
-
+   
 theory Re
   imports "Main" 
 begin
 
+
 section {* Sequential Composition of Sets *}
 
 definition
@@ -43,6 +43,20 @@
 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
 | "L (ALT r1 r2) = (L r1) \<union> (L r2)"
 
+fun
+ nullable :: "rexp \<Rightarrow> bool"
+where
+  "nullable (NULL) = False"
+| "nullable (EMPTY) = True"
+| "nullable (CHAR c) = False"
+| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
+| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
+
+lemma nullable_correctness:
+  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
+apply (induct r) 
+apply(auto simp add: Sequ_def) 
+done
 
 section {* Values *}
 
@@ -53,9 +67,21 @@
 | Right val
 | Left val
 
+section {* The string behind a value *}
+
+fun 
+  flat :: "val \<Rightarrow> string"
+where
+  "flat(Void) = []"
+| "flat(Char c) = [c]"
+| "flat(Left v) = flat(v)"
+| "flat(Right v) = flat(v)"
+| "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
+
 section {* Relation between values and regular expressions *}
 
-inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
+inductive 
+  Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
 where
  "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
 | "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
@@ -63,23 +89,18 @@
 | "\<turnstile> Void : EMPTY"
 | "\<turnstile> Char c : CHAR c"
 
-thm Prf.intros(1)
-
-section {* The string behind a value *}
-
-fun flat :: "val \<Rightarrow> string"
-where
-  "flat(Void) = []"
-| "flat(Char c) = [c]"
-| "flat(Left v) = flat(v)"
-| "flat(Right v) = flat(v)"
-| "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
-
+lemma not_nullable_flat:
+  assumes "\<turnstile> v : r" "\<not>nullable r"
+  shows "flat v \<noteq> []"
+using assms
+apply(induct)
+apply(auto)
+done
 
 lemma Prf_flat_L:
   assumes "\<turnstile> v : r" shows "flat v \<in> L r"
 using assms
-apply(induct)
+apply(induct v r rule: Prf.induct)
 apply(auto simp add: Sequ_def)
 done
 
@@ -96,12 +117,307 @@
 apply(auto)
 done
 
-section {* Ordering of values *}
+section {* Greedy Ordering according to Frisch/Cardelli *}
+
+inductive 
+  GrOrd :: "val \<Rightarrow> val \<Rightarrow> bool" ("_ gr\<succ> _")
+where 
+  "v1 gr\<succ> v1' \<Longrightarrow> (Seq v1 v2) gr\<succ> (Seq v1' v2')"
+| "v2 gr\<succ> v2' \<Longrightarrow> (Seq v1 v2) gr\<succ> (Seq v1 v2')"
+| "v1 gr\<succ> v2 \<Longrightarrow> (Left v1) gr\<succ> (Left v2)"
+| "v1 gr\<succ> v2 \<Longrightarrow> (Right v1) gr\<succ> (Right v2)"
+| "(Left v2) gr\<succ>(Right v1)"
+| "(Char c) gr\<succ> (Char c)"
+| "(Void) gr\<succ> (Void)"
+
+lemma Gr_refl:
+  assumes "\<turnstile> v : r"
+  shows "v gr\<succ> v"
+using assms
+apply(induct)
+apply(auto intro: GrOrd.intros)
+done
+
+lemma Gr_total:
+  assumes "\<turnstile> v1 : r" "\<turnstile> v2 : r"
+  shows "v1 gr\<succ> v2 \<or> v2 gr\<succ> v1"
+using assms
+apply(induct v1 r arbitrary: v2 rule: Prf.induct)
+apply(rotate_tac 4)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply (metis GrOrd.intros(1) GrOrd.intros(2))
+apply(rotate_tac 2)
+apply(erule Prf.cases)
+apply(simp_all)
+apply(clarify)
+apply (metis GrOrd.intros(3))
+apply(clarify)
+apply (metis GrOrd.intros(5))
+apply(rotate_tac 2)
+apply(erule Prf.cases)
+apply(simp_all)
+apply(clarify)
+apply (metis GrOrd.intros(5))
+apply(clarify)
+apply (metis GrOrd.intros(4))
+apply(erule Prf.cases)
+apply(simp_all)
+apply (metis GrOrd.intros(7))
+apply(erule Prf.cases)
+apply(simp_all)
+apply (metis GrOrd.intros(6))
+done
+
+lemma Gr_trans: 
+  assumes "v1 gr\<succ> v2" "v2 gr\<succ> v3" 
+  and     "\<turnstile> v1 : r" "\<turnstile> v2 : r" "\<turnstile> v3 : r"
+  shows "v1 gr\<succ> v3"
+using assms
+apply(induct r arbitrary: v1 v2 v3)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+defer
+(* ALT case *)
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply (metis GrOrd.intros(3))
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply (metis GrOrd.intros(5))
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply (metis GrOrd.intros(5))
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply (metis GrOrd.intros(4))
+(* SEQ case *)
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(erule Prf.cases)
+apply(simp_all (no_asm_use))[5]
+apply(clarify)
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply(clarify)
+apply (metis GrOrd.intros(1))
+apply (metis GrOrd.intros(1))
+apply(erule GrOrd.cases)
+apply(simp_all (no_asm_use))[7]
+apply (metis GrOrd.intros(1))
+by (metis GrOrd.intros(1) Gr_refl)
+
+
+section {* Values Sets *}
+
+definition prefix :: "string \<Rightarrow> string \<Rightarrow> bool" ("_ \<sqsubseteq> _" [100, 100] 100)
+where
+  "s1 \<sqsubseteq> s2 \<equiv> \<exists>s3. s1 @ s3 = s2"
+
+definition sprefix :: "string \<Rightarrow> string \<Rightarrow> bool" ("_ \<sqsubset> _" [100, 100] 100)
+where
+  "s1 \<sqsubset> s2 \<equiv> (s1 \<sqsubseteq> s2 \<and> s1 \<noteq> s2)"
+
+lemma length_sprefix:
+  "s1 \<sqsubset> s2 \<Longrightarrow> length s1 < length s2"
+unfolding sprefix_def prefix_def
+by (auto)
+
+definition Prefixes :: "string \<Rightarrow> string set" where
+  "Prefixes s \<equiv> {sp. sp \<sqsubseteq> s}"
+
+definition Suffixes :: "string \<Rightarrow> string set" where
+  "Suffixes s \<equiv> rev ` (Prefixes (rev s))"
+
+lemma Suffixes_in: 
+  "\<exists>s1. s1 @ s2 = s3 \<Longrightarrow> s2 \<in> Suffixes s3"
+unfolding Suffixes_def Prefixes_def prefix_def image_def
+apply(auto)
+by (metis rev_rev_ident)
+
+lemma Prefixes_Cons:
+  "Prefixes (c # s) = {[]} \<union> {c # sp | sp. sp \<in> Prefixes s}"
+unfolding Prefixes_def prefix_def
+apply(auto simp add: append_eq_Cons_conv) 
+done
+
+lemma finite_Prefixes:
+  "finite (Prefixes s)"
+apply(induct s)
+apply(auto simp add: Prefixes_def prefix_def)[1]
+apply(simp add: Prefixes_Cons)
+done
+
+lemma finite_Suffixes:
+  "finite (Suffixes s)"
+unfolding Suffixes_def
+apply(rule finite_imageI)
+apply(rule finite_Prefixes)
+done
+
+lemma prefix_Cons:
+  "((c # s1) \<sqsubseteq> (c # s2)) = (s1 \<sqsubseteq> s2)"
+apply(auto simp add: prefix_def)
+done
+
+lemma prefix_append:
+  "((s @ s1) \<sqsubseteq> (s @ s2)) = (s1 \<sqsubseteq> s2)"
+apply(induct s)
+apply(simp)
+apply(simp add: prefix_Cons)
+done
+
+
+definition Values :: "rexp \<Rightarrow> string \<Rightarrow> val set" where
+  "Values r s \<equiv> {v. \<turnstile> v : r \<and> flat v \<sqsubseteq> s}"
+
+definition rest :: "val \<Rightarrow> string \<Rightarrow> string" where
+  "rest v s \<equiv> drop (length (flat v)) s"
+
+lemma rest_Suffixes:
+  "rest v s \<in> Suffixes s"
+unfolding rest_def
+by (metis Suffixes_in append_take_drop_id)
+
+
+lemma Values_recs:
+  "Values (NULL) s = {}"
+  "Values (EMPTY) s = {Void}"
+  "Values (CHAR c) s = (if [c] \<sqsubseteq> s then {Char c} else {})" 
+  "Values (ALT r1 r2) s = {Left v | v. v \<in> Values r1 s} \<union> {Right v | v. v \<in> Values r2 s}"
+  "Values (SEQ r1 r2) s = {Seq v1 v2 | v1 v2. v1 \<in> Values r1 s \<and> v2 \<in> Values r2 (rest v1 s)}"
+unfolding Values_def
+apply(auto)
+(*NULL*)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+(*EMPTY*)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(rule Prf.intros)
+apply (metis append_Nil prefix_def)
+(*CHAR*)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(rule Prf.intros)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+(*ALT*)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis Prf.intros(2))
+apply (metis Prf.intros(3))
+(*SEQ*)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (simp add: append_eq_conv_conj prefix_def rest_def)
+apply (metis Prf.intros(1))
+apply (simp add: append_eq_conv_conj prefix_def rest_def)
+done
+
+lemma Values_finite:
+  "finite (Values r s)"
+apply(induct r arbitrary: s)
+apply(simp_all add: Values_recs)
+thm finite_surj
+apply(rule_tac f="\<lambda>(x, y). Seq x y" and 
+               A="{(v1, v2) | v1 v2. v1 \<in> Values r1 s \<and> v2 \<in> Values r2 (rest v1 s)}" in finite_surj)
+prefer 2
+apply(auto)[1]
+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)
+apply(auto)[1]
+apply (metis rest_Suffixes)
+apply(rule finite_UN_I)
+apply(rule finite_Suffixes)
+apply(simp)
+done
+
+section {* Sulzmann functions *}
+
+fun 
+  mkeps :: "rexp \<Rightarrow> val"
+where
+  "mkeps(EMPTY) = Void"
+| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
+| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
+
+lemma mkeps_nullable:
+  assumes "nullable(r)" 
+  shows "\<turnstile> mkeps r : r"
+using assms
+apply(induct rule: nullable.induct)
+apply(auto intro: Prf.intros)
+done
+
+lemma mkeps_flat:
+  assumes "nullable(r)" 
+  shows "flat (mkeps r) = []"
+using assms
+apply(induct rule: nullable.induct)
+apply(auto)
+done
+
+text {*
+  The value mkeps returns is always the correct POSIX
+  value.
+*}
+
+section {* Sulzmann's Ordering of values *}
 
 inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
 where
-  "\<lbrakk>v1 = v1'; v2 \<succ>r2 v2'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
-| "v1  \<succ>r1 v1' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
+  "v2 \<succ>r2 v2' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1 v2')" 
+| "\<lbrakk>v1 \<succ>r1 v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" 
 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
 | "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
@@ -109,64 +425,211 @@
 | "Void \<succ>EMPTY Void"
 | "(Char c) \<succ>(CHAR c) (Char c)"
 
-(*
-lemma
-  assumes "r = SEQ (ALT EMPTY EMPTY) (ALT EMPTY (CHAR c))"
-  shows "(Seq (Left Void) (Right (Char c))) \<succ>r (Seq (Left Void) (Left Void))"
+inductive ValOrd2 :: "val \<Rightarrow> val \<Rightarrow> bool" ("_ 2\<succ> _" [100, 100] 100)
+where
+  "v2 2\<succ> v2' \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1 v2')" 
+| "\<lbrakk>v1 2\<succ> v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1' v2')" 
+| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) 2\<succ> (Right v2)"
+| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) 2\<succ> (Left v1)"
+| "v2 2\<succ> v2' \<Longrightarrow> (Right v2) 2\<succ> (Right v2')"
+| "v1 2\<succ> v1' \<Longrightarrow> (Left v1) 2\<succ> (Left v1')"
+| "Void 2\<succ> Void"
+| "(Char c) 2\<succ> (Char c)"
+
+lemma Ord1:
+  "v1 \<succ>r v2 \<Longrightarrow> v1 2\<succ> v2"
+apply(induct rule: ValOrd.induct)
+apply(auto intro: ValOrd2.intros)
+done
+
+lemma Ord2:
+  "v1 2\<succ> v2 \<Longrightarrow> \<exists>r. v1 \<succ>r v2"
+apply(induct v1 v2 rule: ValOrd2.induct)
+apply(auto intro: ValOrd.intros)
+done
+
+lemma Ord3:
+  "\<lbrakk>v1 2\<succ> v2; \<turnstile> v1 : r\<rbrakk> \<Longrightarrow> v1 \<succ>r v2"
+apply(induct v1 v2 arbitrary: r rule: ValOrd2.induct)
+apply(auto intro: ValOrd.intros elim: Prf.cases)
+done
+
+lemma ValOrd_refl:
+  assumes "\<turnstile> v : r"
+  shows "v \<succ>r v"
 using assms
+apply(induct)
+apply(auto intro: ValOrd.intros)
+done
+
+lemma ValOrd_total:
+  shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r\<rbrakk>  \<Longrightarrow> v1 \<succ>r v2 \<or> v2 \<succ>r v1"
+apply(induct r arbitrary: v1 v2)
+apply(auto)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis ValOrd.intros(7))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis ValOrd.intros(8))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(case_tac "v1a = v1b")
 apply(simp)
+apply(rule ValOrd.intros(1))
+apply (metis ValOrd.intros(1))
+apply(rule ValOrd.intros(2))
+apply(auto)[2]
+apply(erule contrapos_np)
+apply(rule ValOrd.intros(2))
+apply(auto)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis Ord1 Ord3 Prf.intros(2) ValOrd2.intros(6))
 apply(rule ValOrd.intros)
+apply(erule contrapos_np)
 apply(rule ValOrd.intros)
+apply (metis le_eq_less_or_eq neq_iff)
+apply(erule Prf.cases)
+apply(simp_all)[5]
 apply(rule ValOrd.intros)
+apply(erule contrapos_np)
+apply(rule ValOrd.intros)
+apply (metis le_eq_less_or_eq neq_iff)
+apply(rule ValOrd.intros)
+apply(erule contrapos_np)
 apply(rule ValOrd.intros)
-apply(simp)
+by metis
+
+lemma ValOrd_anti:
+  shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r; v1 \<succ>r v2; v2 \<succ>r v1\<rbrakk> \<Longrightarrow> v1 = v2"
+apply(induct r arbitrary: v1 v2)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
 done
-*)
+
+lemma refl_on_ValOrd:
+  "refl_on (Values r s) {(v1, v2). v1 \<succ>r v2 \<and> v1 \<in> Values r s \<and> v2 \<in> Values r s}"
+unfolding refl_on_def
+apply(auto)
+apply(rule ValOrd_refl)
+apply(simp add: Values_def)
+done
+
 
 section {* Posix definition *}
 
 definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
 where
-  "POSIX v r \<equiv> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v')"
+  "POSIX v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v'))"
+
+definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" 
+where
+  "POSIX2 v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v 2\<succ> v'))"
 
-(*
-lemma POSIX_SEQ:
+lemma "POSIX v r = POSIX2 v r"
+unfolding POSIX_def POSIX2_def
+apply(auto)
+apply(rule Ord1)
+apply(auto)
+apply(rule Ord3)
+apply(auto)
+done
+
+section {* POSIX for some constructors *}
+
+lemma POSIX_SEQ1:
   assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
-  shows "POSIX v1 r1 \<and> POSIX v2 r2"
+  shows "POSIX v1 r1"
 using assms
 unfolding POSIX_def
 apply(auto)
 apply(drule_tac x="Seq v' v2" in spec)
 apply(simp)
-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))
-apply(drule_tac x="Seq v1 v'" in spec)
+apply(erule impE)
+apply(rule Prf.intros)
+apply(simp)
 apply(simp)
-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))
-*)
+apply(erule ValOrd.cases)
+apply(simp_all)
+apply(clarify)
+by (metis ValOrd_refl)
 
-(*
-lemma POSIX_SEQ_I:
-  assumes "POSIX v1 r1" "POSIX v2 r2" 
-  shows "POSIX (Seq v1 v2) (SEQ r1 r2)" 
+lemma POSIX_SEQ2:
+  assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2" 
+  shows "POSIX v2 r2"
 using assms
 unfolding POSIX_def
 apply(auto)
-apply(rotate_tac 2)
-apply(erule Prf.cases)
-apply(simp_all)[5]
-apply(auto)[1]
-apply(rule ValOrd.intros)
+apply(drule_tac x="Seq v1 v'" in spec)
+apply(simp)
+apply(erule impE)
+apply(rule Prf.intros)
+apply(simp)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)
+done
 
-apply(auto)
-done
-*)
-
-lemma POSIX_ALT:
+lemma POSIX_ALT2:
   assumes "POSIX (Left v1) (ALT r1 r2)"
   shows "POSIX v1 r1"
 using assms
 unfolding POSIX_def
 apply(auto)
+apply(erule Prf.cases)
+apply(simp_all)[5]
 apply(drule_tac x="Left v'" in spec)
 apply(simp)
 apply(drule mp)
@@ -182,6 +645,8 @@
 using assms
 unfolding POSIX_def
 apply(auto)
+apply(erule Prf.cases)
+apply(simp_all)[5]
 apply(drule_tac x="Right v'" in spec)
 apply(simp)
 apply(drule mp)
@@ -191,7 +656,6 @@
 apply(simp_all)
 done
 
-
 lemma POSIX_ALT1b:
   assumes "POSIX (Right v2) (ALT r1 r2)"
   shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
@@ -207,7 +671,8 @@
 using assms
 unfolding POSIX_def
 apply(auto)
-apply(rotate_tac 3)
+apply (metis Prf.intros(2))
+apply(rotate_tac 2)
 apply(erule Prf.cases)
 apply(simp_all)[5]
 apply(auto)
@@ -222,6 +687,7 @@
 using assms
 unfolding POSIX_def
 apply(auto)
+apply (metis Prf.intros)
 apply(rotate_tac 3)
 apply(erule Prf.cases)
 apply(simp_all)[5]
@@ -230,60 +696,6 @@
 apply metis
 done
 
-
-section {* The ordering is reflexive *}
-
-lemma ValOrd_refl:
-  assumes "\<turnstile> v : r"
-  shows "v \<succ>r v"
-using assms
-apply(induct)
-apply(auto intro: ValOrd.intros)
-done
-
-
-section {* The Matcher *}
-
-fun
- nullable :: "rexp \<Rightarrow> bool"
-where
-  "nullable (NULL) = False"
-| "nullable (EMPTY) = True"
-| "nullable (CHAR c) = False"
-| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
-| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
-
-lemma nullable_correctness:
-  shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
-apply (induct r) 
-apply(auto simp add: Sequ_def) 
-done
-
-fun mkeps :: "rexp \<Rightarrow> val"
-where
-  "mkeps(EMPTY) = Void"
-| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
-| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
-
-lemma mkeps_nullable:
-  assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
-using assms
-apply(induct rule: nullable.induct)
-apply(auto intro: Prf.intros)
-done
-
-lemma mkeps_flat:
-  assumes "nullable(r)" shows "flat (mkeps r) = []"
-using assms
-apply(induct rule: nullable.induct)
-apply(auto)
-done
-
-text {*
-  The value mkeps returns is always the correct POSIX
-  value.
-*}
-
 lemma mkeps_POSIX:
   assumes "nullable r"
   shows "POSIX (mkeps r) r"
@@ -292,38 +704,52 @@
 apply(auto)[1]
 apply(simp add: POSIX_def)
 apply(auto)[1]
+apply (metis Prf.intros(4))
 apply(erule Prf.cases)
 apply(simp_all)[5]
 apply (metis ValOrd.intros)
-apply(simp add: POSIX_def)
+apply(simp)
 apply(auto)[1]
 apply(simp add: POSIX_def)
 apply(auto)[1]
-apply(erule Prf.cases)
-apply(simp_all)[5]
-apply(auto)
-apply (simp add: ValOrd.intros(2) mkeps_flat)
-apply(simp add: POSIX_def)
-apply(auto)[1]
+apply (metis mkeps.simps(2) mkeps_nullable nullable.simps(5))
+apply(rotate_tac 6)
 apply(erule Prf.cases)
 apply(simp_all)[5]
-apply(auto)
-apply (simp add: ValOrd.intros(6))
-apply (simp add: ValOrd.intros(3))
-apply(simp add: POSIX_def)
+apply (simp add: mkeps_flat)
+apply(case_tac "mkeps r1a = v1")
+apply(simp)
+apply (metis ValOrd.intros(1))
+apply (rule ValOrd.intros(2))
+apply metis
+apply(simp)
+(* ALT case *)
+thm mkeps.simps
+apply(simp)
+apply(erule disjE)
+apply(simp)
+apply (metis POSIX_ALT_I1)
+(* *)
 apply(auto)[1]
+thm  POSIX_ALT_I1
+apply (metis POSIX_ALT_I1)
+apply(simp (no_asm) add: POSIX_def)
+apply(auto)[1]
+apply(rule Prf.intros(3))
+apply(simp only: POSIX_def)
+apply(rotate_tac 4)
 apply(erule Prf.cases)
 apply(simp_all)[5]
-apply(auto)
-apply (simp add: ValOrd.intros(6))
-apply (simp add: ValOrd.intros(3))
-apply(simp add: POSIX_def)
+thm mkeps_flat
+apply(simp add: mkeps_flat)
 apply(auto)[1]
-apply(erule Prf.cases)
-apply(simp_all)[5]
-apply(auto)
-apply (metis Prf_flat_L mkeps_flat nullable_correctness)
-by (simp add: ValOrd.intros(5))
+thm Prf_flat_L nullable_correctness
+apply (metis Prf_flat_L nullable_correctness)
+apply(rule ValOrd.intros)
+apply(subst (asm) POSIX_def)
+apply(clarify)
+apply(drule_tac x="v2" in spec)
+by simp
 
 
 section {* Derivatives *}
@@ -346,24 +772,29 @@
   "ders [] r = r"
 | "ders (c # s) r = ders s (der c r)"
 
+
 section {* Injection function *}
 
 fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
 where
-  "injval (CHAR d) c Void = Char d"
+  "injval (EMPTY) c Void = Char c"
+| "injval (CHAR d) c Void = Char d"
+| "injval (CHAR d) c (Char c') = Seq (Char d) (Char c')"
 | "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
 | "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
+| "injval (SEQ r1 r2) c (Char c') = Seq (Char c) (Char c')"
 | "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
 | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
 | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
 
+
 section {* Projection function *}
 
 fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
 where
   "projval (CHAR d) c _ = Void"
-| "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
-| "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
+| "projval (ALT r1 r2) c (Left v1) = Left (projval r1 c v1)"
+| "projval (ALT r1 r2) c (Right v2) = Right (projval r2 c v2)"
 | "projval (SEQ r1 r2) c (Seq v1 v2) = 
      (if flat v1 = [] then Right(projval r2 c v2) 
       else if nullable r1 then Left (Seq (projval r1 c v1) v2)
@@ -374,12 +805,14 @@
 *}
 
 lemma v3:
-  assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
+  assumes "\<turnstile> v : der c r" 
+  shows "\<turnstile> (injval r c v) : r"
 using assms
 apply(induct arbitrary: v rule: der.induct)
 apply(simp)
 apply(erule Prf.cases)
 apply(simp_all)[5]
+apply(simp)
 apply(erule Prf.cases)
 apply(simp_all)[5]
 apply(case_tac "c = c'")
@@ -387,8 +820,10 @@
 apply(erule Prf.cases)
 apply(simp_all)[5]
 apply (metis Prf.intros(5))
+apply(simp)
 apply(erule Prf.cases)
 apply(simp_all)[5]
+apply(simp)
 apply(erule Prf.cases)
 apply(simp_all)[5]
 apply (metis Prf.intros(2))
@@ -413,12 +848,42 @@
 apply(auto)[2]
 done
 
+lemma v3_proj:
+  assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s"
+  shows "\<turnstile> (projval r c v) : der c r"
+using assms
+apply(induct rule: Prf.induct)
+prefer 4
+apply(simp)
+prefer 4
+apply(simp)
+apply (metis Prf.intros(4))
+prefer 2
+apply(simp)
+apply (metis Prf.intros(2))
+prefer 2
+apply(simp)
+apply (metis Prf.intros(3))
+apply(auto)
+apply(rule Prf.intros)
+apply(simp)
+apply (metis Prf_flat_L nullable_correctness)
+apply(rule Prf.intros)
+apply(rule Prf.intros)
+apply (metis Cons_eq_append_conv)
+apply(simp)
+apply(rule Prf.intros)
+apply (metis Cons_eq_append_conv)
+apply(simp)
+done
+
 text {*
-  The string behin the injection value is an added c
+  The string behind the injection value is an added c
 *}
 
 lemma v4:
-  assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
+  assumes "\<turnstile> v : der c r" 
+  shows "flat (injval r c v) = c # (flat v)"
 using assms
 apply(induct arbitrary: v rule: der.induct)
 apply(simp)
@@ -443,10 +908,12 @@
 apply(case_tac "nullable r1")
 apply(simp)
 apply(erule Prf.cases)
-apply(simp_all)[5]
+apply(simp_all (no_asm_use))[5]
 apply(auto)[1]
 apply(erule Prf.cases)
 apply(simp_all)[5]
+apply(clarify)
+apply(simp only: injval.simps flat.simps)
 apply(auto)[1]
 apply (metis mkeps_flat)
 apply(simp)
@@ -454,6 +921,874 @@
 apply(simp_all)[5]
 done
 
+lemma v4_proj:
+  assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s"
+  shows "c # flat (projval r c v) = flat v"
+using assms
+apply(induct rule: Prf.induct)
+prefer 4
+apply(simp)
+prefer 4
+apply(simp)
+prefer 2
+apply(simp)
+prefer 2
+apply(simp)
+apply(auto)
+by (metis Cons_eq_append_conv)
+
+lemma v4_proj2:
+  assumes "\<turnstile> v : r" and "(flat v) = c # s"
+  shows "flat (projval r c v) = s"
+using assms
+by (metis list.inject v4_proj)
+
+lemma injval_inj: "inj_on (injval r c) {v. \<turnstile> v : der c r}"
+apply(induct c r rule: der.induct)
+unfolding inj_on_def
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply (metis list.distinct(1) mkeps_flat v4)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(rotate_tac 6)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis list.distinct(1) mkeps_flat v4)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+done
+
+lemma Values_nullable:
+  assumes "nullable r1"
+  shows "mkeps r1 \<in> Values r1 s"
+using assms
+apply(induct r1 arbitrary: s)
+apply(simp_all)
+apply(simp add: Values_recs)
+apply(simp add: Values_recs)
+apply(simp add: Values_recs)
+apply(auto)[1]
+done
+
+lemma Values_injval:
+  assumes "v \<in> Values (der c r) s"
+  shows "injval r c v \<in> Values r (c#s)"
+using assms
+apply(induct c r arbitrary: v s rule: der.induct)
+apply(simp add: Values_recs)
+apply(simp add: Values_recs)
+apply(case_tac "c = c'")
+apply(simp)
+apply(simp add: Values_recs)
+apply(simp add: prefix_def)
+apply(simp)
+apply(simp add: Values_recs)
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(case_tac "nullable r1")
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(simp add: rest_def)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(simp add: Values_def)
+apply(rule Values_nullable)
+apply(assumption)
+apply(simp add: rest_def)
+apply(subst mkeps_flat)
+apply(assumption)
+apply(simp)
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(simp add: rest_def)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(simp add: Values_def)
+done
+
+lemma Values_projval:
+  assumes "v \<in> Values r (c#s)" "\<exists>s. flat v = c # s"
+  shows "projval r c v \<in> Values (der c r) s"
+using assms
+apply(induct r arbitrary: v s c rule: rexp.induct)
+apply(simp add: Values_recs)
+apply(simp add: Values_recs)
+apply(case_tac "c = char")
+apply(simp)
+apply(simp add: Values_recs)
+apply(simp)
+apply(simp add: Values_recs)
+apply(simp add: prefix_def)
+apply(case_tac "nullable rexp1")
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(simp add: rest_def)
+apply (metis hd_Cons_tl hd_append2 list.sel(1))
+apply(simp add: rest_def)
+apply(simp add: append_eq_Cons_conv)
+apply(auto)[1]
+apply(subst v4_proj2)
+apply(simp add: Values_def)
+apply(assumption)
+apply(simp)
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(auto simp add: Values_def not_nullable_flat)[1]
+apply(simp add: append_eq_Cons_conv)
+apply(auto)[1]
+apply(simp add: append_eq_Cons_conv)
+apply(auto)[1]
+apply(simp add: rest_def)
+apply(subst v4_proj2)
+apply(simp add: Values_def)
+apply(assumption)
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+done
+
+
+definition "MValue v r s \<equiv> (v \<in> Values r s \<and> (\<forall>v' \<in> Values r s. v 2\<succ> v'))"
+
+lemma MValue_ALTE:
+  assumes "MValue v (ALT r1 r2) s"
+  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> 
+         (\<exists>vr. v = Right vr \<and> MValue vr r2 s \<and> (\<forall>vl \<in> Values r1 s. length (flat vl) < length (flat vr)))"
+using assms
+apply(simp add: MValue_def)
+apply(simp add: Values_recs)
+apply(auto)
+apply(drule_tac x="Left x" in bspec)
+apply(simp)
+apply(erule ValOrd2.cases)
+apply(simp_all)
+apply(drule_tac x="Right vr" in bspec)
+apply(simp)
+apply(erule ValOrd2.cases)
+apply(simp_all)
+apply(drule_tac x="Right x" in bspec)
+apply(simp)
+apply(erule ValOrd2.cases)
+apply(simp_all)
+apply(drule_tac x="Left vl" in bspec)
+apply(simp)
+apply(erule ValOrd2.cases)
+apply(simp_all)
+done
+
+lemma MValue_ALTI1:
+  assumes "MValue vl r1 s"  "\<forall>vr \<in> Values r2 s. length (flat vr) \<le> length (flat vl)"
+  shows "MValue (Left vl) (ALT r1 r2) s"
+using assms
+apply(simp add: MValue_def)
+apply(simp add: Values_recs)
+apply(auto)
+apply(rule ValOrd2.intros)
+apply metis
+apply(rule ValOrd2.intros)
+apply metis
+done
+
+lemma MValue_ALTI2:
+  assumes "MValue vr r2 s"  "\<forall>vl \<in> Values r1 s. length (flat vl) < length (flat vr)"
+  shows "MValue (Right vr) (ALT r1 r2) s"
+using assms
+apply(simp add: MValue_def)
+apply(simp add: Values_recs)
+apply(auto)
+apply(rule ValOrd2.intros)
+apply metis
+apply(rule ValOrd2.intros)
+apply metis
+done
+
+lemma t: "(c#xs = c#ys) \<Longrightarrow> xs = ys"
+by (metis list.sel(3))
+
+lemma t2: "(xs = ys) \<Longrightarrow> (c#xs) = (c#ys)"
+by (metis)
+
+lemma "\<not>(nullable r) \<Longrightarrow> \<not>(\<exists>v. \<turnstile> v : r \<and> flat v = [])"
+by (metis Prf_flat_L nullable_correctness)
+
+
+lemma LeftRight:
+  assumes "(Left v1) \<succ>(der c (ALT r1 r2)) (Right v2)"
+  and "\<turnstile> v1 : der c r1" "\<turnstile> v2 : der c r2" 
+  shows "(injval (ALT r1 r2) c (Left v1)) \<succ>(ALT r1 r2) (injval (ALT r1 r2) c (Right v2))"
+using assms
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd.intros)
+apply(clarify)
+apply(subst v4)
+apply(simp)
+apply(subst v4)
+apply(simp)
+apply(simp)
+done
+
+lemma RightLeft:
+  assumes "(Right v1) \<succ>(der c (ALT r1 r2)) (Left v2)"
+  and "\<turnstile> v1 : der c r2" "\<turnstile> v2 : der c r1" 
+  shows "(injval (ALT r1 r2) c (Right v1)) \<succ>(ALT r1 r2) (injval (ALT r1 r2) c (Left v2))"
+using assms
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd.intros)
+apply(clarify)
+apply(subst v4)
+apply(simp)
+apply(subst v4)
+apply(simp)
+apply(simp)
+done
+
+lemma h: 
+  assumes "nullable r1" "\<turnstile> v1 : der c r1"
+  shows "injval r1 c v1 \<succ>r1 mkeps r1"
+using assms
+apply(induct r1 arbitrary: v1 rule: der.induct)
+apply(simp)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(auto)[1]
+apply (metis ValOrd.intros(6))
+apply (metis ValOrd.intros(6))
+apply (metis ValOrd.intros(3) le_add2 list.size(3) mkeps_flat monoid_add_class.add.right_neutral)
+apply(auto)[1]
+apply (metis ValOrd.intros(4) length_greater_0_conv list.distinct(1) list.size(3) mkeps_flat v4)
+apply (metis ValOrd.intros(4) length_greater_0_conv list.distinct(1) list.size(3) mkeps_flat v4)
+apply (metis ValOrd.intros(5))
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply (metis ValOrd.intros(2) list.distinct(1) mkeps_flat v4)
+apply(clarify)
+by (metis ValOrd.intros(1))
+
+lemma LeftRightSeq:
+  assumes "(Left (Seq v1 v2)) \<succ>(der c (SEQ r1 r2)) (Right v3)"
+  and "nullable r1" "\<turnstile> v1 : der c r1"
+  shows "(injval (SEQ r1 r2) c (Seq v1 v2)) \<succ>(SEQ r1 r2) (injval (SEQ r1 r2) c (Right v2))"
+using assms
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply(simp)
+apply(rule ValOrd.intros(2))
+prefer 2
+apply (metis list.distinct(1) mkeps_flat v4)
+by (metis h)
+
+lemma rr1: 
+  assumes "\<turnstile> v : r" "\<not>nullable r" 
+  shows "flat v \<noteq> []"
+using assms
+by (metis Prf_flat_L nullable_correctness)
+
+(* HERE *)
+
+lemma Prf_inj_test:
+  assumes "v1 \<succ>(der c r) v2" 
+          "v1 \<in> Values (der c r) s"
+          "v2 \<in> Values (der c r) s"
+          "injval r c v1 \<in> Values r (c#s)"
+          "injval r c v2 \<in> Values r (c#s)"
+  shows "(injval r c v1) 2\<succ>  (injval r c v2)"
+using assms
+apply(induct c r arbitrary: v1 v2 s rule: der.induct)
+(* NULL case *)
+apply(simp add: Values_recs)
+(* EMPTY case *)
+apply(simp add: Values_recs)
+(* CHAR case *)
+apply(case_tac "c = c'")
+apply(simp)
+apply(simp add: Values_recs)
+apply (metis ValOrd2.intros(8))
+apply(simp add: Values_recs)
+(* ALT case *)
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply (metis ValOrd2.intros(6))
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd2.intros)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd2.intros)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply (metis ValOrd2.intros(5))
+(* SEQ case*)
+apply(simp)
+apply(case_tac "nullable r1")
+apply(simp)
+defer
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply(rule ValOrd2.intros)
+apply(simp)
+apply (metis Ord1)
+apply(clarify)
+apply(rule ValOrd2.intros)
+apply(subgoal_tac "rest v1 (flat v1 @ flat v2) = flat v2")
+apply(simp)
+apply(subgoal_tac "rest (injval r1 c v1) (c # flat v1 @ flat v2) = flat v2")
+apply(simp)
+oops
+
+lemma Prf_inj_test:
+  assumes "v1 \<succ>(der c r) v2" 
+          "v1 \<in> Values (der c r) s"
+          "v2 \<in> Values (der c r) s"
+          "injval r c v1 \<in> Values r (c#s)"
+          "injval r c v2 \<in> Values r (c#s)"
+  shows "(injval r c v1) 2\<succ>  (injval r c v2)"
+using assms
+apply(induct c r arbitrary: v1 v2 s rule: der.induct)
+(* NULL case *)
+apply(simp add: Values_recs)
+(* EMPTY case *)
+apply(simp add: Values_recs)
+(* CHAR case *)
+apply(case_tac "c = c'")
+apply(simp)
+apply(simp add: Values_recs)
+apply (metis ValOrd2.intros(8))
+apply(simp add: Values_recs)
+(* ALT case *)
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply (metis ValOrd2.intros(6))
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd2.intros)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd2.intros)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(subst v4)
+apply(simp add: Values_def)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply (metis ValOrd2.intros(5))
+(* SEQ case*)
+apply(simp)
+apply(case_tac "nullable r1")
+apply(simp)
+defer
+apply(simp)
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply(rule ValOrd2.intros)
+apply(simp)
+apply (metis Ord1)
+apply(clarify)
+apply(rule ValOrd2.intros)
+apply metis
+using injval_inj
+apply(simp add: Values_def inj_on_def)
+apply metis
+apply(simp add: Values_recs)
+apply(auto)[1]
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply (metis Ord1 ValOrd2.intros(1))
+apply(clarify)
+apply(rule ValOrd2.intros(2))
+apply metis
+using injval_inj
+apply(simp add: Values_def inj_on_def)
+apply metis
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd2.intros(2))
+thm h
+apply(rule Ord1)
+apply(rule h)
+apply(simp)
+apply(simp add: Values_def)
+apply(simp add: Values_def)
+apply (metis list.distinct(1) mkeps_flat v4)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply(simp add: Values_def)
+defer
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply(rule ValOrd2.intros(1))
+apply(rotate_tac 1)
+apply(drule_tac x="v2" in meta_spec)
+apply(rotate_tac 8)
+apply(drule_tac x="v2'" in meta_spec)
+apply(rotate_tac 8)
+oops
+
+lemma POSIX_der:
+  assumes "POSIX v (der c r)" "\<turnstile> v : der c r"
+  shows "POSIX (injval r c v) r"
+using assms
+unfolding POSIX_def
+apply(auto)
+thm v3
+apply (erule v3)
+thm v4
+apply(subst (asm) v4)
+apply(assumption)
+apply(drule_tac x="projval r c v'" in spec)
+apply(drule mp)
+apply(rule conjI)
+thm v3_proj
+apply(rule v3_proj)
+apply(simp)
+apply(rule_tac x="flat v" in exI)
+apply(simp)
+thm t
+apply(rule_tac c="c" in  t)
+apply(simp)
+thm v4_proj
+apply(subst v4_proj)
+apply(simp)
+apply(rule_tac x="flat v" in exI)
+apply(simp)
+apply(simp)
+oops
+
+lemma POSIX_der:
+  assumes "POSIX v (der c r)" "\<turnstile> v : der c r"
+  shows "POSIX (injval r c v) r"
+using assms
+apply(induct c r arbitrary: v rule: der.induct)
+(* null case*)
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+(* empty case *)
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+(* char case *)
+apply(simp add: POSIX_def)
+apply(case_tac "c = c'")
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis Prf.intros(5))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis ValOrd.intros(8))
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+(* alt case *)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(simp (no_asm) add: POSIX_def)
+apply(auto)[1]
+apply (metis Prf.intros(2) v3)
+apply(rotate_tac 4)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis POSIX_ALT2 POSIX_def ValOrd.intros(6))
+apply (metis ValOrd.intros(3) order_refl)
+apply(simp (no_asm) add: POSIX_def)
+apply(auto)[1]
+apply (metis Prf.intros(3) v3)
+apply(rotate_tac 4)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+defer
+apply (metis POSIX_ALT1a POSIX_def ValOrd.intros(5))
+prefer 2
+apply(subst (asm) (5) POSIX_def)
+apply(auto)[1]
+apply(rotate_tac 5)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(rule ValOrd.intros)
+apply(subst (asm) v4)
+apply(simp)
+apply(drule_tac x="Left (projval r1a c v1)" in spec)
+apply(clarify)
+apply(drule mp)
+apply(rule conjI)
+apply (metis Prf.intros(2) v3_proj)
+apply(simp)
+apply (metis v4_proj2)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply (metis less_not_refl v4_proj2)
+(* seq case *)
+apply(case_tac "nullable r1")
+defer
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis Prf.intros(1) v3)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(subst (asm) (3) v4)
+apply(simp)
+apply(simp)
+apply(subgoal_tac "flat v1a \<noteq> []")
+prefer 2
+apply (metis Prf_flat_L nullable_correctness)
+apply(subgoal_tac "\<exists>s. flat v1a = c # s")
+prefer 2
+apply (metis append_eq_Cons_conv)
+apply(auto)[1]
+oops
+
+
+lemma POSIX_ex: "\<turnstile> v : r \<Longrightarrow> \<exists>v. POSIX v r"
+apply(induct r arbitrary: v)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(rule_tac x="Void" in exI)
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply (metis Prf.intros(4))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis ValOrd.intros(7))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(rule_tac x="Char c" in exI)
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply (metis Prf.intros(5))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis ValOrd.intros(8))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(drule_tac x="v1" in meta_spec)
+apply(drule_tac x="v2" in meta_spec)
+apply(auto)[1]
+defer
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply (metis POSIX_ALT_I1)
+apply (metis POSIX_ALT_I1 POSIX_ALT_I2)
+apply(case_tac "nullable r1a")
+apply(rule_tac x="Seq (mkeps r1a) va" in exI)
+apply(auto simp add: POSIX_def)[1]
+apply (metis Prf.intros(1) mkeps_nullable)
+apply(simp add: mkeps_flat)
+apply(rotate_tac 7)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(case_tac "mkeps r1 = v1a")
+apply(simp)
+apply (rule ValOrd.intros(1))
+apply (metis append_Nil mkeps_flat)
+apply (rule ValOrd.intros(2))
+apply(drule mkeps_POSIX)
+apply(simp add: POSIX_def)
+oops
+
+lemma POSIX_ex2: "\<turnstile> v : r \<Longrightarrow> \<exists>v. POSIX v r \<and> \<turnstile> v : r"
+apply(induct r arbitrary: v)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(rule_tac x="Void" in exI)
+apply(simp add: POSIX_def)
+apply(auto)[1]
+oops
+
+lemma POSIX_ALT_cases:
+  assumes "\<turnstile> v : (ALT r1 r2)" "POSIX v (ALT r1 r2)"
+  shows "(\<exists>v1. v = Left v1 \<and> POSIX v1 r1) \<or> (\<exists>v2. v = Right v2 \<and> POSIX v2 r2)"
+using assms
+apply(erule_tac Prf.cases)
+apply(simp_all)
+unfolding POSIX_def
+apply(auto)
+apply (metis POSIX_ALT2 POSIX_def assms(2))
+by (metis POSIX_ALT1b assms(2))
+
+lemma POSIX_ALT_cases2:
+  assumes "POSIX v (ALT r1 r2)" "\<turnstile> v : (ALT r1 r2)" 
+  shows "(\<exists>v1. v = Left v1 \<and> POSIX v1 r1) \<or> (\<exists>v2. v = Right v2 \<and> POSIX v2 r2)"
+using assms POSIX_ALT_cases by auto
+
+lemma Prf_flat_empty:
+  assumes "\<turnstile> v : r" "flat v = []"
+  shows "nullable r"
+using assms
+apply(induct)
+apply(auto)
+done
+
+lemma POSIX_proj:
+  assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
+  shows "POSIX (projval r c v) (der c r)"
+using assms
+apply(induct r c v arbitrary: rule: projval.induct)
+defer
+defer
+defer
+defer
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+oops
+
+lemma POSIX_proj:
+  assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
+  shows "POSIX (projval r c v) (der c r)"
+using assms
+apply(induct r arbitrary: c v rule: rexp.induct)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+oops
+
+lemma POSIX_proj:
+  assumes "POSIX v r" "\<turnstile> v : r" "\<exists>s. flat v = c#s"
+  shows "POSIX (projval r c v) (der c r)"
+using assms
+apply(induct r c v arbitrary: rule: projval.induct)
+defer
+defer
+defer
+defer
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+oops
+
+lemma Prf_inj:
+  assumes "v1 \<succ>(der c r) v2" "\<turnstile> v1 : der c r" "\<turnstile> v2 : der c r" "flat v1 = flat v2"
+  shows "(injval r c v1) \<succ>r (injval r c v2)"
+using assms
+apply(induct arbitrary: v1 v2 rule: der.induct)
+(* NULL case *)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+(* EMPTY case *)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+(* CHAR case *)
+apply(case_tac "c = c'")
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd.intros)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+(* ALT case *)
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(rule ValOrd.intros)
+apply(subst v4)
+apply(clarify)
+apply(rotate_tac 3)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(subst v4)
+apply(clarify)
+apply(rotate_tac 2)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(rule ValOrd.intros)
+apply(clarify)
+apply(rotate_tac 3)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(rule ValOrd.intros)
+apply(clarify)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+(* SEQ case*)
+apply(simp)
+apply(case_tac "nullable r1")
+defer
+apply(simp)
+apply(erule ValOrd.cases)
+apply(simp_all)[8]
+apply(clarify)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(clarify)
+apply(rule ValOrd.intros)
+apply(simp)
+oops
+
+
 text {*
   Injection followed by projection is the identity.
 *}
@@ -510,38 +1845,18 @@
   shows "POSIX (injval r c v) r"
 using assms
 apply(induct arbitrary: v rule: der.induct)
+(* NULL case *)
 apply(simp)
 apply(erule Prf.cases)
 apply(simp_all)[5]
+(* EMPTY case *)
 apply(simp)
 apply(erule Prf.cases)
 apply(simp_all)[5]
+(* CHAR case *)
 apply(simp)
 apply(case_tac "c = c'")
 apply(auto simp add: POSIX_def)[1]
 apply(erule Prf.cases)
 apply(simp_all)[5]
-apply(erule Prf.cases)
-apply(simp_all)[5]
-using ValOrd.simps apply blast
-apply(auto)
-apply(erule Prf.cases)
-apply(simp_all)[5]
-(* base cases done *)
-(* ALT case *)
-apply(erule Prf.cases)
-apply(simp_all)[5]
-using POSIX_ALT POSIX_ALT_I1 apply blast
-apply(frule POSIX_ALT1a)
-apply(drule POSIX_ALT1b)
-apply(rule POSIX_ALT_I2)
-apply auto[1]
-apply(subst v4)
-apply(auto)[2]
-apply(rotate_tac 1)
-apply(drule_tac x="v2" in meta_spec)
-apply(simp)
-apply(subst (asm) (4) POSIX_def)
-apply(subst (asm) v4)
-apply(auto)[2]
-(* stuck in the ALT case *)
+oops
\ No newline at end of file