diff -r 532bb9df225d -r 9613e6ace30f thys/ReStar.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/ReStar.thy Thu Jan 21 12:47:20 2016 +0000 @@ -0,0 +1,2584 @@ + +theory Re + imports "Main" +begin + + +section {* Sequential Composition of Sets *} + +definition + Sequ :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) +where + "A ;; B = {s1 @ s2 | s1 s2. s1 \ A \ s2 \ B}" + +fun spow where + "spow s 0 = []" +| "spow s (Suc n) = s @ spow s n" + +text {* Two Simple Properties about Sequential Composition *} + +lemma seq_empty [simp]: + shows "A ;; {[]} = A" + and "{[]} ;; A = A" +by (simp_all add: Sequ_def) + +lemma seq_null [simp]: + shows "A ;; {} = {}" + and "{} ;; A = {}" +by (simp_all add: Sequ_def) + +lemma seq_image: + assumes "\s1 s2. f (s1 @ s2) = (f s1) @ (f s2)" + shows "f ` (A ;; B) = (f ` A) ;; (f ` B)" +apply(auto simp add: Sequ_def image_def) +apply(rule_tac x="f s1" in exI) +apply(rule_tac x="f s2" in exI) +using assms +apply(auto) +apply(rule_tac x="xa @ xb" in exI) +using assms +apply(auto) +done + +section {* Kleene Star for Sets *} + +inductive_set + Star :: "string set \ string set" ("_\" [101] 102) + for A :: "string set" +where + start[intro]: "[] \ A\" +| step[intro]: "\s1 \ A; s2 \ A\\ \ s1 @ s2 \ A\" + +fun + pow :: "string set \ nat \ string set" ("_ \ _" [100,100] 100) +where + "A \ 0 = {[]}" +| "A \ (Suc n) = A ;; (A \ n)" + +lemma star1: + shows "s \ A\ \ \n. s \ A \ n" + apply(induct rule: Star.induct) + apply (metis Re.pow.simps(1) insertI1) + apply(auto) + apply(rule_tac x="Suc n" in exI) + apply(auto simp add: Sequ_def) + done + +lemma star2: + shows "s \ A \ n \ s \ A\" + apply(induct n arbitrary: s) + apply (metis Re.pow.simps(1) Star.simps empty_iff insertE) + apply(auto simp add: Sequ_def) + done + +lemma star3: + shows "A\ = (\i. A \ i)" +using star1 star2 +apply(auto) +done + +lemma star4: + shows "s \ A \ n \ \ss. s = concat ss \ (\s' \ set ss. s' \ A)" + apply(induct n arbitrary: s) + apply(auto simp add: Sequ_def) + apply(rule_tac x="[]" in exI) + apply(auto) + apply(drule_tac x="s2" in meta_spec) + apply(auto) +by (metis concat.simps(2) insertE set_simps(2)) + +lemma star5: + assumes "f [] = []" + assumes "\s1 s2. f (s1 @ s2) = (f s1) @ (f s2)" + shows "(f ` A) \ n = f ` (A \ n)" +apply(induct n) +apply(simp add: assms) +apply(simp) +apply(subst seq_image[OF assms(2)]) +apply(simp) +done + +lemma star6: + assumes "f [] = []" + assumes "\s1 s2. f (s1 @ s2) = (f s1) @ (f s2)" + shows "(f ` A)\ = f ` (A\)" +apply(simp add: star3) +apply(simp add: image_UN) +apply(subst star5[OF assms]) +apply(simp) +done + + +lemma star_decomp: + assumes a: "c # x \ A\" + shows "\a b. x = a @ b \ c # a \ A \ b \ A\" +using a +by (induct x\"c # x" rule: Star.induct) + (auto simp add: append_eq_Cons_conv) + +section {* Regular Expressions *} + +datatype rexp = + NULL +| EMPTY +| CHAR char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp + +section {* Semantics of Regular Expressions *} + +fun + L :: "rexp \ string set" +where + "L (NULL) = {}" +| "L (EMPTY) = {[]}" +| "L (CHAR c) = {[c]}" +| "L (SEQ r1 r2) = (L r1) ;; (L r2)" +| "L (ALT r1 r2) = (L r1) \ (L r2)" +| "L (STAR r) = (L r)\" + +fun + nullable :: "rexp \ bool" +where + "nullable (NULL) = False" +| "nullable (EMPTY) = True" +| "nullable (CHAR c) = False" +| "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (STAR r1) = True" + +lemma nullable_correctness: + shows "nullable r \ [] \ (L r)" +apply (induct r) +apply(auto simp add: Sequ_def) +done + +section {* Values *} + +datatype val = + Void +| Char char +| Seq val val +| Right val +| Left val +| Star "val list" + +section {* The string behind a value *} + +fun + flat :: "val \ 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)" +| "flat (Star []) = []" +| "flat (Star (v#vs)) = (flat v) @ (flat (Star vs))" + +section {* Relation between values and regular expressions *} + +inductive + Prf :: "val \ rexp \ bool" ("\ _ : _" [100, 100] 100) +where + "\\ v1 : r1; \ v2 : r2\ \ \ Seq v1 v2 : SEQ r1 r2" +| "\ v1 : r1 \ \ Left v1 : ALT r1 r2" +| "\ v2 : r2 \ \ Right v2 : ALT r1 r2" +| "\ Void : EMPTY" +| "\ Char c : CHAR c" +| "\ Star [] : STAR r" +| "\\ v : r; \ Star vs : STAR r\ \ \ Star (v#vs) : STAR r" + +lemma not_nullable_flat: + assumes "\ v : r" "\nullable r" + shows "flat v \ []" +using assms +apply(induct) +apply(auto) +done + +lemma Prf_flat_L: + assumes "\ v : r" shows "flat v \ L r" +using assms +apply(induct v r rule: Prf.induct) +apply(auto simp add: Sequ_def) +done + + +lemma Prf_Star_flat_L: + assumes "\ v : STAR r" shows "flat v \ (L r)\" +using assms +apply(induct v r\"STAR r" arbitrary: r rule: Prf.induct) +apply(auto) +apply(simp add: star3) +apply(auto) +apply(rule_tac x="Suc x" in exI) +apply(auto simp add: Sequ_def) +apply(rule_tac x="flat v" in exI) +apply(rule_tac x="flat (Star vs)" in exI) +apply(auto) +by (metis Prf_flat_L) + +lemma L_flat_Prf: + "L(r) = {flat v | v. \ v : r}" +apply(induct r) +apply(auto dest: Prf_flat_L simp add: Sequ_def) +apply (metis Prf.intros(4) flat.simps(1)) +apply (metis Prf.intros(5) flat.simps(2)) +apply (metis Prf.intros(1) flat.simps(5)) +apply (metis Prf.intros(2) flat.simps(3)) +apply (metis Prf.intros(3) flat.simps(4)) +apply(erule Prf.cases) +apply(auto) +apply(simp add: star3) +apply(auto) +sorry + +section {* Greedy Ordering according to Frisch/Cardelli *} + +inductive + GrOrd :: "val \ val \ bool" ("_ gr\ _") +where + "v1 gr\ v1' \ (Seq v1 v2) gr\ (Seq v1' v2')" +| "v2 gr\ v2' \ (Seq v1 v2) gr\ (Seq v1 v2')" +| "v1 gr\ v2 \ (Left v1) gr\ (Left v2)" +| "v1 gr\ v2 \ (Right v1) gr\ (Right v2)" +| "(Left v2) gr\(Right v1)" +| "(Char c) gr\ (Char c)" +| "(Void) gr\ (Void)" + +lemma Gr_refl: + assumes "\ v : r" + shows "v gr\ v" +using assms +apply(induct) +apply(auto intro: GrOrd.intros) +done + +lemma Gr_total: + assumes "\ v1 : r" "\ v2 : r" + shows "v1 gr\ v2 \ v2 gr\ 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\ v2" "v2 gr\ v3" + and "\ v1 : r" "\ v2 : r" "\ v3 : r" + shows "v1 gr\ 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 \ string \ bool" ("_ \ _" [100, 100] 100) +where + "s1 \ s2 \ \s3. s1 @ s3 = s2" + +definition sprefix :: "string \ string \ bool" ("_ \ _" [100, 100] 100) +where + "s1 \ s2 \ (s1 \ s2 \ s1 \ s2)" + +lemma length_sprefix: + "s1 \ s2 \ length s1 < length s2" +unfolding sprefix_def prefix_def +by (auto) + +definition Prefixes :: "string \ string set" where + "Prefixes s \ {sp. sp \ s}" + +definition Suffixes :: "string \ string set" where + "Suffixes s \ rev ` (Prefixes (rev s))" + +lemma Suffixes_in: + "\s1. s1 @ s2 = s3 \ s2 \ Suffixes s3" +unfolding Suffixes_def Prefixes_def prefix_def image_def +apply(auto) +by (metis rev_rev_ident) + +lemma Prefixes_Cons: + "Prefixes (c # s) = {[]} \ {c # sp | sp. sp \ 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) \ (c # s2)) = (s1 \ s2)" +apply(auto simp add: prefix_def) +done + +lemma prefix_append: + "((s @ s1) \ (s @ s2)) = (s1 \ s2)" +apply(induct s) +apply(simp) +apply(simp add: prefix_Cons) +done + + +definition Values :: "rexp \ string \ val set" where + "Values r s \ {v. \ v : r \ flat v \ s}" + +definition rest :: "val \ string \ string" where + "rest v s \ drop (length (flat v)) s" + +lemma rest_Suffixes: + "rest v s \ 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] \ s then {Char c} else {})" + "Values (ALT r1 r2) s = {Left v | v. v \ Values r1 s} \ {Right v | v. v \ Values r2 s}" + "Values (SEQ r1 r2) s = {Seq v1 v2 | v1 v2. v1 \ Values r1 s \ v2 \ 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="\(x, y). Seq x y" and + A="{(v1, v2) | v1 v2. v1 \ Values r1 s \ v2 \ Values r2 (rest v1 s)}" in finite_surj) +prefer 2 +apply(auto)[1] +apply(rule_tac B="\sp \ Suffixes s. {(v1, v2). v1 \ Values r1 s \ v2 \ 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 \ 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))" + +section {* Derivatives *} + +fun + der :: "char \ rexp \ rexp" +where + "der c (NULL) = NULL" +| "der c (EMPTY) = NULL" +| "der c (CHAR c') = (if c = c' then EMPTY else NULL)" +| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)" +| "der c (SEQ r1 r2) = + (if nullable r1 + then ALT (SEQ (der c r1) r2) (der c r2) + else SEQ (der c r1) r2)" + +fun + ders :: "string \ rexp \ rexp" +where + "ders [] r = r" +| "ders (c # s) r = ders s (der c r)" + + +section {* Injection function *} + +fun injval :: "rexp \ char \ val \ val" +where + "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)" + +fun + lex :: "rexp \ string \ val option" +where + "lex r [] = (if nullable r then Some(mkeps r) else None)" +| "lex r (c#s) = (case (lex (der c r) s) of + None \ None + | Some(v) \ Some(injval r c v))" + +fun + lex2 :: "rexp \ string \ val" +where + "lex2 r [] = mkeps r" +| "lex2 r (c#s) = injval r c (lex2 (der c r) s)" + + +section {* Projection function *} + +fun projval :: "rexp \ char \ val \ 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 (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) + else Seq (projval r1 c v1) v2)" + + + +lemma mkeps_nullable: + assumes "nullable(r)" + shows "\ 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 + +lemma v3: + assumes "\ v : der c r" + shows "\ (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'") +apply(simp) +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)) +apply (metis Prf.intros(3)) +apply(simp) +apply(case_tac "nullable r1") +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto)[1] +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto)[1] +apply (metis Prf.intros(1)) +apply(auto)[1] +apply (metis Prf.intros(1) mkeps_nullable) +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto)[1] +apply(rule Prf.intros) +apply(auto)[2] +done + +lemma v3_proj: + assumes "\ v : r" and "\s. (flat v) = c # s" + shows "\ (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 + +lemma v4: + assumes "\ v : der c r" + shows "flat (injval r c v) = c # (flat v)" +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(simp) +apply(case_tac "c = c'") +apply(simp) +apply(auto)[1] +apply(erule Prf.cases) +apply(simp_all)[5] +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(simp) +apply(case_tac "nullable r1") +apply(simp) +apply(erule Prf.cases) +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) +apply(erule Prf.cases) +apply(simp_all)[5] +done + +lemma v4_proj: + assumes "\ v : r" and "\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 "\ v : r" and "(flat v) = c # s" + shows "flat (projval r c v) = s" +using assms +by (metis list.inject v4_proj) + + +section {* Alternative Posix definition *} + +inductive + PMatch :: "string \ rexp \ val \ bool" ("_ \ _ \ _" [100, 100, 100] 100) +where + "[] \ EMPTY \ Void" +| "[c] \ (CHAR c) \ (Char c)" +| "s \ r1 \ v \ s \ (ALT r1 r2) \ (Left v)" +| "\s \ r2 \ v; s \ L(r1)\ \ s \ (ALT r1 r2) \ (Right v)" +| "\s1 \ r1 \ v1; s2 \ r2 \ v2; + \(\s3 s4. s3 \ [] \ s3 @ s4 = s2 \ (s1 @ s3) \ L r1 \ s4 \ L r2)\ \ + (s1 @ s2) \ (SEQ r1 r2) \ (Seq v1 v2)" + + +lemma PMatch_mkeps: + assumes "nullable r" + shows "[] \ r \ mkeps r" +using assms +apply(induct r) +apply(auto) +apply (metis PMatch.intros(1)) +apply(subst append.simps(1)[symmetric]) +apply (rule PMatch.intros) +apply(simp) +apply(simp) +apply(auto)[1] +apply (rule PMatch.intros) +apply(simp) +apply (rule PMatch.intros) +apply(simp) +apply (rule PMatch.intros) +apply(simp) +by (metis nullable_correctness) + + +lemma PMatch1: + assumes "s \ r \ v" + shows "\ v : r" "flat v = s" +using assms +apply(induct s r v rule: PMatch.induct) +apply(auto) +apply (metis Prf.intros(4)) +apply (metis Prf.intros(5)) +apply (metis Prf.intros(2)) +apply (metis Prf.intros(3)) +apply (metis Prf.intros(1)) +done + +lemma PMatch_Values: + assumes "s \ r \ v" + shows "v \ Values r s" +using assms +apply(simp add: Values_def PMatch1) +by (metis append_Nil2 prefix_def) + +lemma PMatch2: + assumes "s \ (der c r) \ v" + shows "(c#s) \ r \ (injval r c v)" +using assms +apply(induct c r arbitrary: s v rule: der.induct) +apply(auto) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(case_tac "c = c'") +apply(simp) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis PMatch.intros(2)) +apply(simp) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis PMatch.intros(3)) +apply(clarify) +apply(rule PMatch.intros) +apply metis +apply(simp add: L_flat_Prf) +apply(auto)[1] +thm v3_proj +apply(frule_tac c="c" in v3_proj) +apply metis +apply(drule_tac x="projval r1 c v" in spec) +apply(drule mp) +apply (metis v4_proj2) +apply(simp) +apply(case_tac "nullable r1") +apply(simp) +defer +apply(simp) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(clarify) +apply(subst append.simps(2)[symmetric]) +apply(rule PMatch.intros) +apply metis +apply metis +apply(auto)[1] +apply(simp add: L_flat_Prf) +apply(auto)[1] +apply(frule_tac c="c" in v3_proj) +apply metis +apply(drule_tac x="s3" in spec) +apply(drule mp) +apply(rule_tac x="projval r1 c v" in exI) +apply(rule conjI) +apply (metis v4_proj2) +apply(simp) +apply metis +(* nullable case *) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(clarify) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(clarify) +apply(subst append.simps(2)[symmetric]) +apply(rule PMatch.intros) +apply metis +apply metis +apply(auto)[1] +apply(simp add: L_flat_Prf) +apply(auto)[1] +apply(frule_tac c="c" in v3_proj) +apply metis +apply(drule_tac x="s3" in spec) +apply(drule mp) +apply (metis v4_proj2) +apply(simp) +(* interesting case *) +apply(clarify) +apply(simp) +thm L.simps +apply(subst (asm) L.simps(4)[symmetric]) +apply(simp only: L_flat_Prf) +apply(simp) +apply(subst append.simps(1)[symmetric]) +apply(rule PMatch.intros) +apply (metis PMatch_mkeps) +apply metis +apply(auto) +apply(simp only: L_flat_Prf) +apply(simp) +apply(auto) +apply(drule_tac x="Seq (projval r1 c v) vb" in spec) +apply(drule mp) +apply(simp) +apply (metis append_Cons butlast_snoc last_snoc neq_Nil_conv rotate1.simps(2) v4_proj2) +apply(subgoal_tac "\ projval r1 c v : der c r1") +apply (metis Prf.intros(1)) +apply(rule v3_proj) +apply(simp) +by (metis Cons_eq_append_conv) + +lemma lex_correct1: + assumes "s \ L r" + shows "lex r s = None" +using assms +apply(induct s arbitrary: r) +apply(simp) +apply (metis nullable_correctness) +apply(auto) +apply(drule_tac x="der a r" in meta_spec) +apply(drule meta_mp) +apply(auto) +apply(simp add: L_flat_Prf) +by (metis v3 v4) + + +lemma lex_correct2: + assumes "s \ L r" + shows "\v. lex r s = Some(v) \ \ v : r \ flat v = s" +using assms +apply(induct s arbitrary: r) +apply(simp) +apply (metis mkeps_flat mkeps_nullable nullable_correctness) +apply(drule_tac x="der a r" in meta_spec) +apply(drule meta_mp) +apply(simp add: L_flat_Prf) +apply(auto) +apply (metis v3_proj v4_proj2) +apply (metis v3) +apply(rule v4) +by metis + +lemma lex_correct3: + assumes "s \ L r" + shows "\v. lex r s = Some(v) \ s \ r \ v" +using assms +apply(induct s arbitrary: r) +apply(simp) +apply (metis PMatch_mkeps nullable_correctness) +apply(drule_tac x="der a r" in meta_spec) +apply(drule meta_mp) +apply(simp add: L_flat_Prf) +apply(auto) +apply (metis v3_proj v4_proj2) +apply(rule PMatch2) +apply(simp) +done + +lemma lex_correct4: + assumes "s \ L r" + shows "s \ r \ (lex2 r s)" +using assms +apply(induct s arbitrary: r) +apply(simp) +apply (metis PMatch_mkeps nullable_correctness) +apply(simp) +apply(rule PMatch2) +apply(drule_tac x="der a r" in meta_spec) +apply(drule meta_mp) +apply(simp add: L_flat_Prf) +apply(auto) +apply (metis v3_proj v4_proj2) +done + +lemma + "lex2 (ALT (CHAR a) (ALT (CHAR b) (SEQ (CHAR a) (CHAR b)))) [a,b] = Right (Right (Seq (Char a) (Char b)))" +apply(simp) +done + + +section {* Sulzmann's Ordering of values *} + +thm PMatch.intros + +inductive ValOrd :: "string \ val \ rexp \ val \ bool" ("_ \ _ \_ _" [100, 100, 100, 100] 100) +where + "\s2 \ v2 \r2 v2'; flat v1 = s1\ \ (s1 @ s2) \ (Seq v1 v2) \(SEQ r1 r2) (Seq v1 v2')" +| "\s1 \ v1 \r1 v1'; v1 \ v1'\ \ (Seq v1 v2) \(SEQ r1 r2) (Seq v1' v2')" +| "\flat v1 \ s; flat v2 \ flat v1\ \ s \ (Left v1) \(ALT r1 r2) (Right v2)" +| "\flat v2 \ s; flat v1 \ flat v2\ \ s \ (Right v2) \(ALT r1 r2) (Left v1)" +| "s \ v2 \r2 v2' \ s \ (Right v2) \(ALT r1 r2) (Right v2')" +| "s \ v1 \r1 v1' \ s \ (Left v1) \(ALT r1 r2) (Left v1')" +| "s \ Void \EMPTY Void" +| "(c#s) \ (Char c) \(CHAR c) (Char c)" + + +inductive ValOrd :: "val \ rexp \ val \ bool" ("_ \_ _" [100, 100, 100] 100) +where + "v2 \r2 v2' \ (Seq v1 v2) \(SEQ r1 r2) (Seq v1 v2')" +| "\v1 \r1 v1'; v1 \ v1'\ \ (Seq v1 v2) \(SEQ r1 r2) (Seq v1' v2')" +| "length (flat v1) \ length (flat v2) \ (Left v1) \(ALT r1 r2) (Right v2)" +| "length (flat v2) > length (flat v1) \ (Right v2) \(ALT r1 r2) (Left v1)" +| "v2 \r2 v2' \ (Right v2) \(ALT r1 r2) (Right v2')" +| "v1 \r1 v1' \ (Left v1) \(ALT r1 r2) (Left v1')" +| "Void \EMPTY Void" +| "(Char c) \(CHAR c) (Char c)" + +lemma ValOrd_PMatch: + assumes "s \ r \ v1" "\ v2 : r" "flat v2 \ s" + shows "v1 \r v2" +using assms +apply(induct r arbitrary: s v1 v2 rule: rexp.induct) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(erule Prf.cases) +apply(simp_all)[5] +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(7)) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(8)) +defer +apply(erule Prf.cases) +apply(simp_all)[5] +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(6)) +apply (metis PMatch1(2) Prf_flat_L ValOrd.intros(4) length_sprefix sprefix_def) +apply(clarify) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis PMatch1(2) ValOrd.intros(3) length_sprefix less_imp_le_nat order_refl sprefix_def) +apply(clarify) +apply (metis ValOrd.intros(5)) +(* Seq case *) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(auto) +apply(case_tac "v1b = v1a") +apply(auto) +apply(simp add: prefix_def) +apply(auto)[1] +apply (metis PMatch1(2) ValOrd.intros(1) same_append_eq) +apply(simp add: prefix_def) +apply(auto)[1] +apply(simp add: append_eq_append_conv2) +apply(auto) +prefer 2 +apply (metis ValOrd.intros(2)) +prefer 2 +apply (metis ValOrd.intros(2)) +apply(case_tac "us = []") +apply(simp) +apply (metis ValOrd.intros(2) append_Nil2) +apply(drule_tac x="us" in spec) +apply(simp) +apply(drule_tac mp) +apply (metis Prf_flat_L) +apply(drule_tac x="s1 @ us" in meta_spec) +apply(drule_tac x="v1b" in meta_spec) +apply(drule_tac x="v1a" in meta_spec) +apply(drule_tac meta_mp) + +apply(simp) +apply(drule_tac meta_mp) +apply(simp) +apply(simp) +apply(simp) +apply(clarify) +apply (metis ValOrd.intros(6)) +apply(clarify) +apply (metis PMatch1(2) ValOrd.intros(3) length_sprefix less_imp_le_nat order_refl sprefix_def) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(clarify) +apply (metis PMatch1(2) Prf_flat_L ValOrd.intros(4) length_sprefix sprefix_def) +apply (metis ValOrd.intros(5)) +(* Seq case *) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(case_tac "v1 = v1a") +apply(auto) +apply(simp add: prefix_def) +apply(auto)[1] +apply (metis PMatch1(2) ValOrd.intros(1) same_append_eq) +apply(simp add: prefix_def) +apply(auto)[1] +apply(frule PMatch1) +apply(frule PMatch1(2)[symmetric]) +apply(clarify) +apply(simp add: append_eq_append_conv2) +apply(auto) +prefer 2 +apply (metis ValOrd.intros(2)) +prefer 2 +apply (metis ValOrd.intros(2)) +apply(case_tac "us = []") +apply(simp) +apply (metis ValOrd.intros(2) append_Nil2) +apply(drule_tac x="us" in spec) +apply(simp) +apply(drule mp) +apply (metis Prf_flat_L) +apply(drule_tac x="v1a" in meta_spec) +apply(drule_tac meta_mp) +apply(simp) +apply(drule_tac meta_mp) +apply(simp) + +lemma ValOrd_PMatch: + assumes "s \ r \ v1" "\ v2 : r" "flat v2 \ s" + shows "v1 \r v2" +using assms +apply(induct arbitrary: v2 rule: .induct) +apply(erule Prf.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(7)) +apply(erule Prf.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(8)) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(clarify) +apply (metis ValOrd.intros(6)) +apply(clarify) +apply (metis PMatch1(2) ValOrd.intros(3) length_sprefix less_imp_le_nat order_refl sprefix_def) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(clarify) +apply (metis PMatch1(2) Prf_flat_L ValOrd.intros(4) length_sprefix sprefix_def) +apply (metis ValOrd.intros(5)) +(* Seq case *) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(case_tac "v1 = v1a") +apply(auto) +apply(simp add: prefix_def) +apply(auto)[1] +apply (metis PMatch1(2) ValOrd.intros(1) same_append_eq) +apply(simp add: prefix_def) +apply(auto)[1] +apply(frule PMatch1) +apply(frule PMatch1(2)[symmetric]) +apply(clarify) +apply(simp add: append_eq_append_conv2) +apply(auto) +prefer 2 +apply (metis ValOrd.intros(2)) +prefer 2 +apply (metis ValOrd.intros(2)) +apply(case_tac "us = []") +apply(simp) +apply (metis ValOrd.intros(2) append_Nil2) +apply(drule_tac x="us" in spec) +apply(simp) +apply(drule mp) +apply (metis Prf_flat_L) +apply(drule_tac x="v1a" in meta_spec) +apply(drule_tac meta_mp) +apply(simp) +apply(drule_tac meta_mp) +apply(simp) + +apply (metis PMatch1(2) ValOrd.intros(1) same_append_eq) +apply(rule ValOrd.intros(2)) +apply(auto) +apply(drule_tac x="v1a" in meta_spec) +apply(drule_tac meta_mp) +apply(simp) +apply(drule_tac meta_mp) +prefer 2 +apply(simp) +thm append_eq_append_conv +apply(simp add: append_eq_append_conv2) +apply(auto) +apply (metis Prf_flat_L) +apply(case_tac "us = []") +apply(simp) +apply(drule_tac x="us" in spec) +apply(drule mp) + + +inductive ValOrd2 :: "val \ val \ bool" ("_ 2\ _" [100, 100] 100) +where + "v2 2\ v2' \ (Seq v1 v2) 2\ (Seq v1 v2')" +| "\v1 2\ v1'; v1 \ v1'\ \ (Seq v1 v2) 2\ (Seq v1' v2')" +| "length (flat v1) \ length (flat v2) \ (Left v1) 2\ (Right v2)" +| "length (flat v2) > length (flat v1) \ (Right v2) 2\ (Left v1)" +| "v2 2\ v2' \ (Right v2) 2\ (Right v2')" +| "v1 2\ v1' \ (Left v1) 2\ (Left v1')" +| "Void 2\ Void" +| "(Char c) 2\ (Char c)" + +lemma Ord1: + "v1 \r v2 \ v1 2\ v2" +apply(induct rule: ValOrd.induct) +apply(auto intro: ValOrd2.intros) +done + +lemma Ord2: + "v1 2\ v2 \ \r. v1 \r v2" +apply(induct v1 v2 rule: ValOrd2.induct) +apply(auto intro: ValOrd.intros) +done + +lemma Ord3: + "\v1 2\ v2; \ v1 : r\ \ v1 \r v2" +apply(induct v1 v2 arbitrary: r rule: ValOrd2.induct) +apply(auto intro: ValOrd.intros elim: Prf.cases) +done + +section {* Posix definition *} + +definition POSIX :: "val \ rexp \ bool" +where + "POSIX v r \ (\ v : r \ (\v'. (\ v' : r \ flat v' \ flat v) \ v \r v'))" + +lemma ValOrd_refl: + assumes "\ v : r" + shows "v \r v" +using assms +apply(induct) +apply(auto intro: ValOrd.intros) +done + +lemma ValOrd_total: + shows "\\ v1 : r; \ v2 : r\ \ v1 \r v2 \ v2 \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) +by metis + +lemma ValOrd_anti: + shows "\\ v1 : r; \ v2 : r; v1 \r v2; v2 \r v1\ \ 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 POSIX_ALT_I1: + assumes "POSIX v1 r1" + shows "POSIX (Left v1) (ALT r1 r2)" +using assms +unfolding POSIX_def +apply(auto) +apply (metis Prf.intros(2)) +apply(rotate_tac 2) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(rule ValOrd.intros) +apply(auto) +apply(rule ValOrd.intros) +by (metis le_eq_less_or_eq length_sprefix sprefix_def) + +lemma POSIX_ALT_I2: + assumes "POSIX v2 r2" "\v'. \ v' : r1 \ length (flat v2) > length (flat v')" + shows "POSIX (Right v2) (ALT r1 r2)" +using assms +unfolding POSIX_def +apply(auto) +apply (metis Prf.intros) +apply(rotate_tac 3) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(rule ValOrd.intros) +apply metis +apply(rule ValOrd.intros) +apply metis +done + +thm PMatch.intros[no_vars] + +lemma POSIX_PMatch: + assumes "s \ r \ v" "\ v' : r" + shows "length (flat v') \ length (flat v)" +using assms +apply(induct arbitrary: s v 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(erule PMatch.cases) +apply(simp_all)[5] +defer +apply(erule Prf.cases) +apply(simp_all)[5] +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(clarify) +apply(simp add: L_flat_Prf) + +apply(clarify) +apply (metis ValOrd.intros(8)) +apply (metis POSIX_ALT_I1) +apply(rule POSIX_ALT_I2) +apply(simp) +apply(auto)[1] +apply(simp add: POSIX_def) +apply(frule PMatch1(1)) +apply(frule PMatch1(2)) +apply(simp) + + +lemma POSIX_PMatch: + assumes "s \ r \ v" + shows "POSIX v r" +using assms +apply(induct arbitrary: rule: PMatch.induct) +apply(auto) +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(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 (metis POSIX_ALT_I1) +apply(rule POSIX_ALT_I2) +apply(simp) +apply(auto)[1] +apply(simp add: POSIX_def) +apply(frule PMatch1(1)) +apply(frule PMatch1(2)) +apply(simp) + + + +lemma ValOrd_PMatch: + assumes "s \ r \ v1" "\ v2 : r" "flat v2 = s" + shows "v1 \r v2" +using assms +apply(induct arbitrary: v2 rule: PMatch.induct) +apply(erule Prf.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(7)) +apply(erule Prf.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(8)) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(clarify) +apply (metis ValOrd.intros(6)) +apply(clarify) +apply (metis PMatch1(2) ValOrd.intros(3) order_refl) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(clarify) +apply (metis Prf_flat_L) +apply (metis ValOrd.intros(5)) +(* Seq case *) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(case_tac "v1 = v1a") +apply(auto) +apply (metis PMatch1(2) ValOrd.intros(1) same_append_eq) +apply(rule ValOrd.intros(2)) +apply(auto) +apply(drule_tac x="v1a" in meta_spec) +apply(drule_tac meta_mp) +apply(simp) +apply(drule_tac meta_mp) +prefer 2 +apply(simp) +apply(simp add: append_eq_append_conv2) +apply(auto) +apply (metis Prf_flat_L) +apply(case_tac "us = []") +apply(simp) +apply(drule_tac x="us" in spec) +apply(drule mp) + +thm L_flat_Prf +apply(simp add: L_flat_Prf) +thm append_eq_append_conv2 +apply(simp add: append_eq_append_conv2) +apply(auto) +apply(drule_tac x="us" in spec) +apply(drule mp) +apply metis +apply (metis append_Nil2) +apply(case_tac "us = []") +apply(auto) +apply(drule_tac x="s2" in spec) +apply(drule mp) + +apply(auto)[1] +apply(drule_tac x="v1a" in meta_spec) +apply(simp) + +lemma refl_on_ValOrd: + "refl_on (Values r s) {(v1, v2). v1 \r v2 \ v1 \ Values r s \ v2 \ 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 \ rexp \ bool" +where + "POSIX v r \ (\ v : r \ (\v'. (\ v' : r \ flat v = flat v') \ v \r v'))" + +definition POSIX2 :: "val \ rexp \ bool" +where + "POSIX2 v r \ (\ v : r \ (\v'. (\ v' : r \ flat v = flat v') \ v 2\ v'))" + +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)" "\ v1 : r1" "\ 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(erule impE) +apply(rule Prf.intros) +apply(simp) +apply(simp) +apply(erule ValOrd.cases) +apply(simp_all) +apply(clarify) +by (metis ValOrd_refl) + +lemma POSIX_SEQ2: + assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\ v1 : r1" "\ v2 : r2" + shows "POSIX v2 r2" +using assms +unfolding POSIX_def +apply(auto) +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 + +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) +apply(rule Prf.intros) +apply(auto) +apply(erule ValOrd.cases) +apply(simp_all) +done + +lemma POSIX_ALT1a: + assumes "POSIX (Right v2) (ALT r1 r2)" + shows "POSIX v2 r2" +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) +apply(rule Prf.intros) +apply(auto) +apply(erule ValOrd.cases) +apply(simp_all) +done + +lemma POSIX_ALT1b: + assumes "POSIX (Right v2) (ALT r1 r2)" + shows "(\v'. (\ v' : r2 \ flat v' = flat v2) \ v2 \r2 v')" +using assms +apply(drule_tac POSIX_ALT1a) +unfolding POSIX_def +apply(auto) +done + +lemma POSIX_ALT_I1: + assumes "POSIX v1 r1" + shows "POSIX (Left v1) (ALT r1 r2)" +using assms +unfolding POSIX_def +apply(auto) +apply (metis Prf.intros(2)) +apply(rotate_tac 2) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(rule ValOrd.intros) +apply(auto) +apply(rule ValOrd.intros) +by simp + +lemma POSIX_ALT_I2: + assumes "POSIX v2 r2" "\v'. \ v' : r1 \ length (flat v2) > length (flat v')" + shows "POSIX (Right v2) (ALT r1 r2)" +using assms +unfolding POSIX_def +apply(auto) +apply (metis Prf.intros) +apply(rotate_tac 3) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(rule ValOrd.intros) +apply metis +done + +lemma mkeps_POSIX: + assumes "nullable r" + shows "POSIX (mkeps r) r" +using assms +apply(induct r) +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) +apply(auto)[1] +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 (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] +thm mkeps_flat +apply(simp add: mkeps_flat) +apply(auto)[1] +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 + + + +text {* + Injection value is related to r +*} + + + +text {* + The string behind the injection value is an added c +*} + + +lemma injval_inj: "inj_on (injval r c) {v. \ 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 \ 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 \ Values (der c r) s" + shows "injval r c v \ 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 \ Values r (c#s)" "\s. flat v = c # s" + shows "projval r c v \ 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 \ (v \ Values r s \ (\v' \ Values r s. v 2\ v'))" + +lemma MValue_ALTE: + assumes "MValue v (ALT r1 r2) s" + shows "(\vl. v = Left vl \ MValue vl r1 s \ (\vr \ Values r2 s. length (flat vr) \ length (flat vl))) \ + (\vr. v = Right vr \ MValue vr r2 s \ (\vl \ 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" "\vr \ Values r2 s. length (flat vr) \ 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" "\vl \ 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) \ xs = ys" +by (metis list.sel(3)) + +lemma t2: "(xs = ys) \ (c#xs) = (c#ys)" +by (metis) + +lemma "\(nullable r) \ \(\v. \ v : r \ flat v = [])" +by (metis Prf_flat_L nullable_correctness) + + +lemma LeftRight: + assumes "(Left v1) \(der c (ALT r1 r2)) (Right v2)" + and "\ v1 : der c r1" "\ v2 : der c r2" + shows "(injval (ALT r1 r2) c (Left v1)) \(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) \(der c (ALT r1 r2)) (Left v2)" + and "\ v1 : der c r2" "\ v2 : der c r1" + shows "(injval (ALT r1 r2) c (Right v1)) \(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" "\ v1 : der c r1" + shows "injval r1 c v1 \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)) \(der c (SEQ r1 r2)) (Right v3)" + and "nullable r1" "\ v1 : der c r1" + shows "(injval (SEQ r1 r2) c (Seq v1 v2)) \(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 "\ v : r" "\nullable r" + shows "flat v \ []" +using assms +by (metis Prf_flat_L nullable_correctness) + +(* HERE *) + +lemma Prf_inj_test: + assumes "v1 \(der c r) v2" + "v1 \ Values (der c r) s" + "v2 \ Values (der c r) s" + "injval r c v1 \ Values r (c#s)" + "injval r c v2 \ Values r (c#s)" + shows "(injval r c v1) 2\ (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 \(der c r) v2" + "v1 \ Values (der c r) s" + "v2 \ Values (der c r) s" + "injval r c v1 \ Values r (c#s)" + "injval r c v2 \ Values r (c#s)" + shows "(injval r c v1) 2\ (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)" "\ 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)" "\ 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 \ []") +prefer 2 +apply (metis Prf_flat_L nullable_correctness) +apply(subgoal_tac "\s. flat v1a = c # s") +prefer 2 +apply (metis append_eq_Cons_conv) +apply(auto)[1] +oops + + +lemma POSIX_ex: "\ v : r \ \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: "\ v : r \ \v. POSIX v r \ \ 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 "\ v : (ALT r1 r2)" "POSIX v (ALT r1 r2)" + shows "(\v1. v = Left v1 \ POSIX v1 r1) \ (\v2. v = Right v2 \ 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)" "\ v : (ALT r1 r2)" + shows "(\v1. v = Left v1 \ POSIX v1 r1) \ (\v2. v = Right v2 \ POSIX v2 r2)" +using assms POSIX_ALT_cases by auto + +lemma Prf_flat_empty: + assumes "\ v : r" "flat v = []" + shows "nullable r" +using assms +apply(induct) +apply(auto) +done + +lemma POSIX_proj: + assumes "POSIX v r" "\ v : r" "\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" "\ v : r" "\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" "\ v : r" "\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 \(der c r) v2" "\ v1 : der c r" "\ v2 : der c r" "flat v1 = flat v2" + shows "(injval r c v1) \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. +*} + +lemma proj_inj_id: + assumes "\ v : der c r" + shows "projval r c (injval r c v) = v" +using assms +apply(induct r arbitrary: c v rule: rexp.induct) +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(simp) +apply(case_tac "c = char") +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +defer +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(simp) +apply(case_tac "nullable rexp1") +apply(simp) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto)[1] +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto)[1] +apply (metis list.distinct(1) v4) +apply(auto)[1] +apply (metis mkeps_flat) +apply(auto) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto)[1] +apply(simp add: v4) +done + +text {* + + HERE: Crucial lemma that does not go through in the sequence case. + +*} +lemma v5: + assumes "\ v : der c r" "POSIX v (der c r)" + 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] +oops \ No newline at end of file