diff -r f16019b11179 -r e85099ac4c6c thys/SpecExt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/SpecExt.thy Wed Sep 06 00:52:08 2017 +0100 @@ -0,0 +1,978 @@ + +theory SpecExt + imports Main "~~/src/HOL/Library/Sublist" +begin + +section {* Sequential Composition of Languages *} + +definition + Sequ :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) +where + "A ;; B = {s1 @ s2 | s1 s2. s1 \ A \ s2 \ B}" + +text {* Two Simple Properties about Sequential Composition *} + +lemma Sequ_empty_string [simp]: + shows "A ;; {[]} = A" + and "{[]} ;; A = A" +by (simp_all add: Sequ_def) + +lemma Sequ_empty [simp]: + shows "A ;; {} = {}" + and "{} ;; A = {}" +by (simp_all add: Sequ_def) + +lemma Sequ_assoc: + shows "(A ;; B) ;; C = A ;; (B ;; C)" +apply(auto simp add: Sequ_def) +apply blast +by (metis append_assoc) + +lemma Sequ_Union_in: + shows "(A ;; (\x\ B. C x)) = (\x\ B. A ;; C x)" +by (auto simp add: Sequ_def) + +section {* Semantic Derivative (Left Quotient) of Languages *} + +definition + Der :: "char \ string set \ string set" +where + "Der c A \ {s. c # s \ A}" + +definition + Ders :: "string \ string set \ string set" +where + "Ders s A \ {s'. s @ s' \ A}" + +lemma Der_null [simp]: + shows "Der c {} = {}" +unfolding Der_def +by auto + +lemma Der_empty [simp]: + shows "Der c {[]} = {}" +unfolding Der_def +by auto + +lemma Der_char [simp]: + shows "Der c {[d]} = (if c = d then {[]} else {})" +unfolding Der_def +by auto + +lemma Der_union [simp]: + shows "Der c (A \ B) = Der c A \ Der c B" +unfolding Der_def +by auto + +lemma Der_UNION [simp]: + shows "Der c (\x\A. B x) = (\x\A. Der c (B x))" +by (auto simp add: Der_def) + +lemma Der_Sequ [simp]: + shows "Der c (A ;; B) = (Der c A) ;; B \ (if [] \ A then Der c B else {})" +unfolding Der_def Sequ_def +by (auto simp add: Cons_eq_append_conv) + + +section {* Kleene Star for Languages *} + +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\" + +(* Arden's lemma *) + +lemma Star_cases: + shows "A\ = {[]} \ A ;; A\" +unfolding Sequ_def +by (auto) (metis Star.simps) + +lemma Star_decomp: + assumes "c # x \ A\" + shows "\s1 s2. x = s1 @ s2 \ c # s1 \ A \ s2 \ A\" +using assms +by (induct x\"c # x" rule: Star.induct) + (auto simp add: append_eq_Cons_conv) + +lemma Star_Der_Sequ: + shows "Der c (A\) \ (Der c A) ;; A\" +unfolding Der_def Sequ_def +by(auto simp add: Star_decomp) + + +lemma Der_star [simp]: + shows "Der c (A\) = (Der c A) ;; A\" +proof - + have "Der c (A\) = Der c ({[]} \ A ;; A\)" + by (simp only: Star_cases[symmetric]) + also have "... = Der c (A ;; A\)" + by (simp only: Der_union Der_empty) (simp) + also have "... = (Der c A) ;; A\ \ (if [] \ A then Der c (A\) else {})" + by simp + also have "... = (Der c A) ;; A\" + using Star_Der_Sequ by auto + finally show "Der c (A\) = (Der c A) ;; A\" . +qed + +section {* Power operation for Sets *} + +fun + Pow :: "string set \ nat \ string set" ("_ \ _" [101, 102] 101) +where + "A \ 0 = {[]}" +| "A \ (Suc n) = A ;; (A \ n)" + +lemma Pow_empty [simp]: + shows "[] \ A \ n \ (n = 0 \ [] \ A)" +by(induct n) (auto simp add: Sequ_def) + +lemma Pow_Suc_rev: + "A \ (Suc n) = (A \ n) ;; A" +apply(induct n arbitrary: A) +apply(simp_all) +by (metis Sequ_assoc) + + +lemma Pow_decomp: + assumes "c # x \ A \ n" + shows "\s1 s2. x = s1 @ s2 \ c # s1 \ A \ s2 \ A \ (n - 1)" +using assms +apply(induct n) +apply(auto simp add: Cons_eq_append_conv Sequ_def) +apply(case_tac n) +apply(auto simp add: Sequ_def) +apply(blast) +done + +lemma Star_Pow: + assumes "s \ A\" + shows "\n. s \ A \ n" +using assms +apply(induct) +apply(auto) +apply(rule_tac x="Suc n" in exI) +apply(auto simp add: Sequ_def) +done + +lemma Pow_Star: + assumes "s \ A \ n" + shows "s \ A\" +using assms +apply(induct n arbitrary: s) +apply(auto simp add: Sequ_def) +done + +lemma Der_Pow_0: + shows "Der c (A \ 0) = {}" +by(simp add: Der_def) + +lemma Der_Pow_Suc: + shows "Der c (A \ (Suc n)) = (Der c A) ;; (A \ n)" +unfolding Der_def Sequ_def +apply(auto simp add: Cons_eq_append_conv Sequ_def dest!: Pow_decomp) +apply(case_tac n) +apply(force simp add: Sequ_def)+ +done + +lemma Der_Pow [simp]: + shows "Der c (A \ n) = (if n = 0 then {} else (Der c A) ;; (A \ (n - 1)))" +apply(case_tac n) +apply(simp_all del: Pow.simps add: Der_Pow_0 Der_Pow_Suc) +done + +lemma Der_Pow_Sequ [simp]: + shows "Der c (A ;; A \ n) = (Der c A) ;; (A \ n)" +by (simp only: Pow.simps[symmetric] Der_Pow) (simp) + + +lemma Pow_Sequ_Un: + assumes "0 < x" + shows "(\n \ {..x}. (A \ n)) = ({[]} \ (\n \ {..x - Suc 0}. A ;; (A \ n)))" +using assms +apply(auto simp add: Sequ_def) +apply(smt Pow.elims Sequ_def Suc_le_mono Suc_pred atMost_iff empty_iff insert_iff mem_Collect_eq) +apply(rule_tac x="Suc xa" in bexI) +apply(auto simp add: Sequ_def) +done + +lemma Pow_Sequ_Un2: + assumes "0 < x" + shows "(\n \ {x..}. (A \ n)) = (\n \ {x - Suc 0..}. A ;; (A \ n))" +using assms +apply(auto simp add: Sequ_def) +apply(case_tac n) +apply(auto simp add: Sequ_def) +apply fastforce +apply(case_tac x) +apply(auto) +apply(rule_tac x="Suc xa" in bexI) +apply(auto simp add: Sequ_def) +done + +section {* Regular Expressions *} + +datatype rexp = + ZERO +| ONE +| CHAR char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp +| UPNTIMES rexp nat +| NTIMES rexp nat +| FROMNTIMES rexp nat +| NMTIMES rexp nat nat + +section {* Semantics of Regular Expressions *} + +fun + L :: "rexp \ string set" +where + "L (ZERO) = {}" +| "L (ONE) = {[]}" +| "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)\" +| "L (UPNTIMES r n) = (\i\ {..n} . (L r) \ i)" +| "L (NTIMES r n) = (L r) \ n" +| "L (FROMNTIMES r n) = (\i\ {n..} . (L r) \ i)" +| "L (NMTIMES r n m) = (\i\{n..m} . (L r) \ i)" + +section {* Nullable, Derivatives *} + +fun + nullable :: "rexp \ bool" +where + "nullable (ZERO) = False" +| "nullable (ONE) = True" +| "nullable (CHAR c) = False" +| "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (STAR r) = True" +| "nullable (UPNTIMES r n) = True" +| "nullable (NTIMES r n) = (if n = 0 then True else nullable r)" +| "nullable (FROMNTIMES r n) = (if n = 0 then True else nullable r)" +| "nullable (NMTIMES r n m) = (if m < n then False else (if n = 0 then True else nullable r))" + +fun + der :: "char \ rexp \ rexp" +where + "der c (ZERO) = ZERO" +| "der c (ONE) = ZERO" +| "der c (CHAR d) = (if c = d then ONE else ZERO)" +| "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)" +| "der c (STAR r) = SEQ (der c r) (STAR r)" +| "der c (UPNTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (UPNTIMES r (n - 1)))" +| "der c (NTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (NTIMES r (n - 1)))" +| "der c (FROMNTIMES r n) = SEQ (der c r) (FROMNTIMES r (n - 1))" +| "der c (NMTIMES r n m) = + (if m < n then ZERO + else (if n = 0 then (if m = 0 then ZERO else + SEQ (der c r) (UPNTIMES r (m - 1))) else + SEQ (der c r) (NMTIMES r (n - 1) (m - 1))))" + +fun + ders :: "string \ rexp \ rexp" +where + "ders [] r = r" +| "ders (c # s) r = ders s (der c r)" + + +lemma nullable_correctness: + shows "nullable r \ [] \ (L r)" +by(induct r) (auto simp add: Sequ_def) + + +lemma der_correctness: + shows "L (der c r) = Der c (L r)" +apply(induct r) +apply(simp add: nullable_correctness del: Der_UNION) +apply(simp add: nullable_correctness del: Der_UNION) +apply(simp add: nullable_correctness del: Der_UNION) +apply(simp add: nullable_correctness del: Der_UNION) +apply(simp add: nullable_correctness del: Der_UNION) +apply(simp add: nullable_correctness del: Der_UNION) +prefer 2 +apply(simp add: nullable_correctness del: Der_UNION) +apply(simp add: nullable_correctness del: Der_UNION) +apply(rule impI) +apply(subst Sequ_Union_in) +apply(subst Der_Pow_Sequ[symmetric]) +apply(subst Pow.simps[symmetric]) +apply(subst Der_UNION[symmetric]) +apply(subst Pow_Sequ_Un) +apply(simp) +apply(simp only: Der_union Der_empty) +apply(simp) +apply(simp add: nullable_correctness del: Der_UNION) +apply(subst Sequ_Union_in) +apply(subst Der_Pow_Sequ[symmetric]) +apply(subst Pow.simps[symmetric]) +apply(case_tac x2) +prefer 2 +apply(subst Pow_Sequ_Un2) +apply(simp) +apply(simp) +apply(auto simp add: Sequ_def Der_def)[1] +apply(rule_tac x="Suc xa" in exI) +apply(auto simp add: Sequ_def)[1] +apply(drule Pow_decomp) +apply(auto)[1] +apply (metis append_Cons) +apply(simp add: nullable_correctness del: Der_UNION) +apply(rule impI) +apply(rule conjI) +apply(rule impI) +apply(subst Sequ_Union_in) +apply(subst Der_Pow_Sequ[symmetric]) +apply(subst Pow.simps[symmetric]) +apply(subst Der_UNION[symmetric]) +apply(case_tac x3a) +apply(simp) +apply(clarify) +apply(auto simp add: Sequ_def Der_def Cons_eq_append_conv)[1] +apply(rule_tac x="Suc xa" in bexI) +apply(auto simp add: Sequ_def)[2] +apply (metis append_Cons) +apply (metis (no_types, hide_lams) Pow_decomp atMost_iff diff_Suc_eq_diff_pred diff_is_0_eq) +apply(rule impI)+ +apply(subst Sequ_Union_in) +apply(subst Der_Pow_Sequ[symmetric]) +apply(subst Pow.simps[symmetric]) +apply(subst Der_UNION[symmetric]) +apply(case_tac x2) +apply(simp) +apply(simp del: Pow.simps) +apply(auto simp add: Sequ_def Der_def) +apply (metis One_nat_def Suc_le_D Suc_le_mono atLeastAtMost_iff diff_Suc_1 not_le) +by fastforce + + + +lemma ders_correctness: + shows "L (ders s r) = Ders s (L r)" +by (induct s arbitrary: r) + (simp_all add: Ders_def der_correctness Der_def) + + + +section {* Values *} + +datatype val = + Void +| Char char +| Seq val val +| Right val +| Left val +| Stars "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 (Stars []) = []" +| "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" + +abbreviation + "flats vs \ concat (map flat vs)" + +lemma flat_Stars [simp]: + "flat (Stars vs) = flats vs" +by (induct vs) (auto) + +lemma Star_concat: + assumes "\s \ set ss. s \ A" + shows "concat ss \ A\" +using assms by (induct ss) (auto) + +lemma Star_cstring: + assumes "s \ A\" + shows "\ss. concat ss = s \ (\s \ set ss. s \ A \ s \ [])" +using assms +apply(induct rule: Star.induct) +apply(auto)[1] +apply(rule_tac x="[]" in exI) +apply(simp) +apply(erule exE) +apply(clarify) +apply(case_tac "s1 = []") +apply(rule_tac x="ss" in exI) +apply(simp) +apply(rule_tac x="s1#ss" in exI) +apply(simp) +done + +lemma Aux: + assumes "\s\set ss. s = []" + shows "concat ss = []" +using assms +by (induct ss) (auto) + +lemma Pow_cstring_nonempty: + assumes "s \ A \ n" + shows "\ss. concat ss = s \ length ss \ n \ (\s \ set ss. s \ A \ s \ [])" +using assms +apply(induct n arbitrary: s) +apply(auto) +apply(simp add: Sequ_def) +apply(erule exE)+ +apply(clarify) +apply(drule_tac x="s2" in meta_spec) +apply(simp) +apply(clarify) +apply(case_tac "s1 = []") +apply(simp) +apply(rule_tac x="ss" in exI) +apply(simp) +apply(rule_tac x="s1 # ss" in exI) +apply(simp) +done + +lemma Pow_cstring: + assumes "s \ A \ n" + shows "\ss1 ss2. concat (ss1 @ ss2) = s \ length (ss1 @ ss2) = n \ + (\s \ set ss1. s \ A \ s \ []) \ (\s \ set ss2. s \ A \ s = [])" +using assms +apply(induct n arbitrary: s) +apply(auto)[1] +apply(simp only: Pow_Suc_rev) +apply(simp add: Sequ_def) +apply(erule exE)+ +apply(clarify) +apply(drule_tac x="s1" in meta_spec) +apply(simp) +apply(erule exE)+ +apply(clarify) +apply(case_tac "s2 = []") +apply(simp) +apply(rule_tac x="ss1" in exI) +apply(rule_tac x="s2#ss2" in exI) +apply(simp) +apply(rule_tac x="ss1 @ [s2]" in exI) +apply(rule_tac x="ss2" in exI) +apply(simp) +apply(subst Aux) +apply(auto)[1] +apply(subst Aux) +apply(auto)[1] +apply(simp) +done + + +section {* Lexical Values *} + + + +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 : ONE" +| "\ Char c : CHAR c" +| "\\v \ set vs. \ v : r \ flat v \ []\ \ \ Stars vs : STAR r" +| "\\v \ set vs. \ v : r \ flat v \ []; length vs \ n\ \ \ Stars vs : UPNTIMES r n" +| "\\v \ set vs1. \ v : r \ flat v \ []; + \v \ set vs2. \ v : r \ flat v = []; + length (vs1 @ vs2) = n\ \ \ Stars (vs1 @ vs2) : NTIMES r n" +| "\\v \ set vs1. \ v : r \ flat v \ []; + \v \ set vs2. \ v : r \ flat v = []; + length (vs1 @ vs2) \ n\ \ \ Stars (vs1 @ vs2) : FROMNTIMES r n" +| "\\v \ set vs1. \ v : r \ flat v \ []; + \v \ set vs2. \ v : r \ flat v = []; + length (vs1 @ vs2) \ n; length (vs1 @ vs2) \ m\ \ \ Stars (vs1 @ vs2) : NMTIMES r n m" + +inductive_cases Prf_elims: + "\ v : ZERO" + "\ v : SEQ r1 r2" + "\ v : ALT r1 r2" + "\ v : ONE" + "\ v : CHAR c" + "\ vs : STAR r" + "\ vs : UPNTIMES r n" + "\ vs : NTIMES r n" + "\ vs : FROMNTIMES r n" + "\ vs : NMTIMES r n m" + +lemma Prf_Stars_appendE: + assumes "\ Stars (vs1 @ vs2) : STAR r" + shows "\ Stars vs1 : STAR r \ \ Stars vs2 : STAR r" +using assms +by (auto intro: Prf.intros elim!: Prf_elims) + +lemma flats_empty: + assumes "(\v\set vs. flat v = [])" + shows "flats vs = []" +using assms +by(induct vs) (simp_all) + +lemma Star_cval: + assumes "\s\set ss. \v. s = flat v \ \ v : r" + shows "\vs. flats vs = concat ss \ (\v\set vs. \ v : r \ flat v \ [])" +using assms +apply(induct ss) +apply(auto) +apply(rule_tac x="[]" in exI) +apply(simp) +apply(case_tac "flat v = []") +apply(rule_tac x="vs" in exI) +apply(simp) +apply(rule_tac x="v#vs" in exI) +apply(simp) +done + + +lemma flats_cval: + assumes "\s\set ss. \v. s = flat v \ \ v : r" + shows "\vs1 vs2. flats (vs1 @ vs2) = concat ss \ length (vs1 @ vs2) = length ss \ + (\v\set vs1. \ v : r \ flat v \ []) \ + (\v\set vs2. \ v : r \ flat v = [])" +using assms +apply(induct ss rule: rev_induct) +apply(rule_tac x="[]" in exI)+ +apply(simp) +apply(simp) +apply(clarify) +apply(case_tac "flat v = []") +apply(rule_tac x="vs1" in exI) +apply(rule_tac x="v#vs2" in exI) +apply(simp) +apply(rule_tac x="vs1 @ [v]" in exI) +apply(rule_tac x="vs2" in exI) +apply(simp) +apply(subst (asm) (2) flats_empty) +apply(simp) +apply(simp) +done + +lemma flats_cval_nonempty: + assumes "\s\set ss. \v. s = flat v \ \ v : r" + shows "\vs. flats vs = concat ss \ length vs \ length ss \ + (\v\set vs. \ v : r \ flat v \ [])" +using assms +apply(induct ss) +apply(rule_tac x="[]" in exI) +apply(simp) +apply(simp) +apply(clarify) +apply(case_tac "flat v = []") +apply(rule_tac x="vs" in exI) +apply(simp) +apply(rule_tac x="v # vs" in exI) +apply(simp) +done + +lemma Pow_flats: + assumes "\v \ set vs. flat v \ A" + shows "flats vs \ A \ length vs" +using assms +by(induct vs)(auto simp add: Sequ_def) + +lemma Pow_flats_appends: + assumes "\v \ set vs1. flat v \ A" "\v \ set vs2. flat v \ A" + shows "flats vs1 @ flats vs2 \ A \ (length vs1 + length vs2)" +using assms +apply(induct vs1) +apply(auto simp add: Sequ_def Pow_flats) +done + +lemma L_flat_Prf1: + assumes "\ v : r" + shows "flat v \ L r" +using assms +apply(induct) +apply(auto simp add: Sequ_def Star_concat Pow_flats) +apply(meson Pow_flats atMost_iff) +using Pow_flats_appends apply blast +apply(meson Pow_flats_appends atLeast_iff) +apply(meson Pow_flats_appends atLeastAtMost_iff) +done + +lemma L_flat_Prf2: + assumes "s \ L r" + shows "\v. \ v : r \ flat v = s" +using assms +proof(induct r arbitrary: s) + case (STAR r s) + have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact + have "s \ L (STAR r)" by fact + then obtain ss where "concat ss = s" "\s \ set ss. s \ L r \ s \ []" + using Star_cstring by auto + then obtain vs where "flats vs = s" "\v\set vs. \ v : r \ flat v \ []" + using IH Star_cval by metis + then show "\v. \ v : STAR r \ flat v = s" + using Prf.intros(6) flat_Stars by blast +next + case (SEQ r1 r2 s) + then show "\v. \ v : SEQ r1 r2 \ flat v = s" + unfolding Sequ_def L.simps by (fastforce intro: Prf.intros) +next + case (ALT r1 r2 s) + then show "\v. \ v : ALT r1 r2 \ flat v = s" + unfolding L.simps by (fastforce intro: Prf.intros) +next + case (NTIMES r n) + have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact + have "s \ L (NTIMES r n)" by fact + then obtain ss1 ss2 where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = n" + "\s \ set ss1. s \ L r \ s \ []" "\s \ set ss2. s \ L r \ s = []" + using Pow_cstring by force + then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = n" + "\v\set vs1. \ v : r \ flat v \ []" "\v\set vs2. \ v : r \ flat v = []" + using IH flats_cval + apply - + apply(drule_tac x="ss1 @ ss2" in meta_spec) + apply(drule_tac x="r" in meta_spec) + apply(drule meta_mp) + apply(simp) + apply (metis Un_iff) + apply(clarify) + apply(drule_tac x="vs1" in meta_spec) + apply(drule_tac x="vs2" in meta_spec) + apply(simp) + done + then show "\v. \ v : NTIMES r n \ flat v = s" + using Prf.intros(8) flat_Stars by blast +next + case (FROMNTIMES r n) + have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact + have "s \ L (FROMNTIMES r n)" by fact + then obtain ss1 ss2 m where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = m" "n \ m" + "\s \ set ss1. s \ L r \ s \ []" "\s \ set ss2. s \ L r \ s = []" + using Pow_cstring by auto blast + then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = m" "n \ m" + "\v\set vs1. \ v : r \ flat v \ []" "\v\set vs2. \ v : r \ flat v = []" + using IH flats_cval + apply - + apply(drule_tac x="ss1 @ ss2" in meta_spec) + apply(drule_tac x="r" in meta_spec) + apply(drule meta_mp) + apply(simp) + apply (metis Un_iff) + apply(clarify) + apply(drule_tac x="vs1" in meta_spec) + apply(drule_tac x="vs2" in meta_spec) + apply(simp) + done + then show "\v. \ v : FROMNTIMES r n \ flat v = s" + using Prf.intros(9) flat_Stars by blast +next + case (NMTIMES r n m) + have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact + have "s \ L (NMTIMES r n m)" by fact + then obtain ss1 ss2 k where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = k" "n \ k" "k \ m" + "\s \ set ss1. s \ L r \ s \ []" "\s \ set ss2. s \ L r \ s = []" + using Pow_cstring by (auto, blast) + then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = k" "n \ k" "k \ m" + "\v\set vs1. \ v : r \ flat v \ []" "\v\set vs2. \ v : r \ flat v = []" + using IH flats_cval + apply - + apply(drule_tac x="ss1 @ ss2" in meta_spec) + apply(drule_tac x="r" in meta_spec) + apply(drule meta_mp) + apply(simp) + apply (metis Un_iff) + apply(clarify) + apply(drule_tac x="vs1" in meta_spec) + apply(drule_tac x="vs2" in meta_spec) + apply(simp) + done + then show "\v. \ v : NMTIMES r n m \ flat v = s" + apply(rule_tac x="Stars (vs1 @ vs2)" in exI) + apply(simp) + apply(rule Prf.intros) + apply(auto) + done +next + case (UPNTIMES r n s) + have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact + have "s \ L (UPNTIMES r n)" by fact + then obtain ss where "concat ss = s" "\s \ set ss. s \ L r \ s \ []" "length ss \ n" + using Pow_cstring_nonempty by force + then obtain vs where "flats vs = s" "\v\set vs. \ v : r \ flat v \ []" "length vs \ n" + using IH flats_cval_nonempty by (smt order.trans) + then show "\v. \ v : UPNTIMES r n \ flat v = s" + using Prf.intros(7) flat_Stars by blast +qed (auto intro: Prf.intros) + + +lemma L_flat_Prf: + shows "L(r) = {flat v | v. \ v : r}" +using L_flat_Prf1 L_flat_Prf2 by blast + + + +section {* Sets of Lexical Values *} + +text {* + Shows that lexical values are finite for a given regex and string. +*} + +definition + LV :: "rexp \ string \ val set" +where "LV r s \ {v. \ v : r \ flat v = s}" + +lemma LV_simps: + shows "LV ZERO s = {}" + and "LV ONE s = (if s = [] then {Void} else {})" + and "LV (CHAR c) s = (if s = [c] then {Char c} else {})" + and "LV (ALT r1 r2) s = Left ` LV r1 s \ Right ` LV r2 s" +unfolding LV_def +by (auto intro: Prf.intros elim: Prf.cases) + + +abbreviation + "Prefixes s \ {s'. prefixeq s' s}" + +abbreviation + "Suffixes s \ {s'. suffixeq s' s}" + +abbreviation + "SSuffixes s \ {s'. suffix s' s}" + +lemma Suffixes_cons [simp]: + shows "Suffixes (c # s) = Suffixes s \ {c # s}" +by (auto simp add: suffixeq_def Cons_eq_append_conv) + + +lemma finite_Suffixes: + shows "finite (Suffixes s)" +by (induct s) (simp_all) + +lemma finite_SSuffixes: + shows "finite (SSuffixes s)" +proof - + have "SSuffixes s \ Suffixes s" + unfolding suffix_def suffixeq_def by auto + then show "finite (SSuffixes s)" + using finite_Suffixes finite_subset by blast +qed + +lemma finite_Prefixes: + shows "finite (Prefixes s)" +proof - + have "finite (Suffixes (rev s))" + by (rule finite_Suffixes) + then have "finite (rev ` Suffixes (rev s))" by simp + moreover + have "rev ` (Suffixes (rev s)) = Prefixes s" + unfolding suffixeq_def prefixeq_def image_def + by (auto)(metis rev_append rev_rev_ident)+ + ultimately show "finite (Prefixes s)" by simp +qed + +lemma LV_STAR_finite: + assumes "\s. finite (LV r s)" + shows "finite (LV (STAR r) s)" +proof(induct s rule: length_induct) + fix s::"char list" + assume "\s'. length s' < length s \ finite (LV (STAR r) s')" + then have IH: "\s' \ SSuffixes s. finite (LV (STAR r) s')" + by (auto simp add: suffix_def) + def f \ "\(v, vs). Stars (v # vs)" + def S1 \ "\s' \ Prefixes s. LV r s'" + def S2 \ "\s2 \ SSuffixes s. Stars -` (LV (STAR r) s2)" + have "finite S1" using assms + unfolding S1_def by (simp_all add: finite_Prefixes) + moreover + with IH have "finite S2" unfolding S2_def + by (auto simp add: finite_SSuffixes inj_on_def finite_vimageI) + ultimately + have "finite ({Stars []} \ f ` (S1 \ S2))" by simp + moreover + have "LV (STAR r) s \ {Stars []} \ f ` (S1 \ S2)" + unfolding S1_def S2_def f_def + unfolding LV_def image_def prefixeq_def suffix_def + apply(auto elim: Prf_elims) + apply(erule Prf_elims) + apply(auto) + apply(case_tac vs) + apply(auto intro: Prf.intros) + done + ultimately + show "finite (LV (STAR r) s)" by (simp add: finite_subset) +qed + +lemma LV_UPNTIMES_STAR: + "LV (UPNTIMES r n) s \ LV (STAR r) s" +by(auto simp add: LV_def intro: Prf.intros elim: Prf_elims) + +lemma LV_FROMNTIMES_STAR: + "LV (FROMNTIMES r n) s \ LV (STAR r) s" +apply(auto simp add: LV_def intro: Prf.intros elim!: Prf_elims) +oops + +lemma LV_finite: + shows "finite (LV r s)" +proof(induct r arbitrary: s) + case (ZERO s) + show "finite (LV ZERO s)" by (simp add: LV_simps) +next + case (ONE s) + show "finite (LV ONE s)" by (simp add: LV_simps) +next + case (CHAR c s) + show "finite (LV (CHAR c) s)" by (simp add: LV_simps) +next + case (ALT r1 r2 s) + then show "finite (LV (ALT r1 r2) s)" by (simp add: LV_simps) +next + case (SEQ r1 r2 s) + def f \ "\(v1, v2). Seq v1 v2" + def S1 \ "\s' \ Prefixes s. LV r1 s'" + def S2 \ "\s' \ Suffixes s. LV r2 s'" + have IHs: "\s. finite (LV r1 s)" "\s. finite (LV r2 s)" by fact+ + then have "finite S1" "finite S2" unfolding S1_def S2_def + by (simp_all add: finite_Prefixes finite_Suffixes) + moreover + have "LV (SEQ r1 r2) s \ f ` (S1 \ S2)" + unfolding f_def S1_def S2_def + unfolding LV_def image_def prefixeq_def suffixeq_def + by (auto elim: Prf.cases) + ultimately + show "finite (LV (SEQ r1 r2) s)" + by (simp add: finite_subset) +next + case (STAR r s) + then show "finite (LV (STAR r) s)" by (simp add: LV_STAR_finite) +next + case (NTIMES r n s) + have "\s. finite (LV r s)" by fact + then show "finite (LV (NTIMES r n) s)" + apply(simp add: LV_def) +qed + + + +section {* Our POSIX Definition *} + +inductive + Posix :: "string \ rexp \ val \ bool" ("_ \ _ \ _" [100, 100, 100] 100) +where + Posix_ONE: "[] \ ONE \ Void" +| Posix_CHAR: "[c] \ (CHAR c) \ (Char c)" +| Posix_ALT1: "s \ r1 \ v \ s \ (ALT r1 r2) \ (Left v)" +| Posix_ALT2: "\s \ r2 \ v; s \ L(r1)\ \ s \ (ALT r1 r2) \ (Right v)" +| Posix_SEQ: "\s1 \ r1 \ v1; s2 \ r2 \ v2; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r1 \ s\<^sub>4 \ L r2)\ \ + (s1 @ s2) \ (SEQ r1 r2) \ (Seq v1 v2)" +| Posix_STAR1: "\s1 \ r \ v; s2 \ STAR r \ Stars vs; flat v \ []; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (STAR r))\ + \ (s1 @ s2) \ STAR r \ Stars (v # vs)" +| Posix_STAR2: "[] \ STAR r \ Stars []" + +inductive_cases Posix_elims: + "s \ ZERO \ v" + "s \ ONE \ v" + "s \ CHAR c \ v" + "s \ ALT r1 r2 \ v" + "s \ SEQ r1 r2 \ v" + "s \ STAR r \ v" + +lemma Posix1: + assumes "s \ r \ v" + shows "s \ L r" "flat v = s" +using assms +by (induct s r v rule: Posix.induct) + (auto simp add: Sequ_def) + +text {* + Our Posix definition determines a unique value. +*} + +lemma Posix_determ: + assumes "s \ r \ v1" "s \ r \ v2" + shows "v1 = v2" +using assms +proof (induct s r v1 arbitrary: v2 rule: Posix.induct) + case (Posix_ONE v2) + have "[] \ ONE \ v2" by fact + then show "Void = v2" by cases auto +next + case (Posix_CHAR c v2) + have "[c] \ CHAR c \ v2" by fact + then show "Char c = v2" by cases auto +next + case (Posix_ALT1 s r1 v r2 v2) + have "s \ ALT r1 r2 \ v2" by fact + moreover + have "s \ r1 \ v" by fact + then have "s \ L r1" by (simp add: Posix1) + ultimately obtain v' where eq: "v2 = Left v'" "s \ r1 \ v'" by cases auto + moreover + have IH: "\v2. s \ r1 \ v2 \ v = v2" by fact + ultimately have "v = v'" by simp + then show "Left v = v2" using eq by simp +next + case (Posix_ALT2 s r2 v r1 v2) + have "s \ ALT r1 r2 \ v2" by fact + moreover + have "s \ L r1" by fact + ultimately obtain v' where eq: "v2 = Right v'" "s \ r2 \ v'" + by cases (auto simp add: Posix1) + moreover + have IH: "\v2. s \ r2 \ v2 \ v = v2" by fact + ultimately have "v = v'" by simp + then show "Right v = v2" using eq by simp +next + case (Posix_SEQ s1 r1 v1 s2 r2 v2 v') + have "(s1 @ s2) \ SEQ r1 r2 \ v'" + "s1 \ r1 \ v1" "s2 \ r2 \ v2" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" by fact+ + then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \ r1 \ v1'" "s2 \ r2 \ v2'" + apply(cases) apply (auto simp add: append_eq_append_conv2) + using Posix1(1) by fastforce+ + moreover + have IHs: "\v1'. s1 \ r1 \ v1' \ v1 = v1'" + "\v2'. s2 \ r2 \ v2' \ v2 = v2'" by fact+ + ultimately show "Seq v1 v2 = v'" by simp +next + case (Posix_STAR1 s1 r v s2 vs v2) + have "(s1 @ s2) \ STAR r \ v2" + "s1 \ r \ v" "s2 \ STAR r \ Stars vs" "flat v \ []" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (STAR r))" by fact+ + then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (STAR r) \ (Stars vs')" + apply(cases) apply (auto simp add: append_eq_append_conv2) + using Posix1(1) apply fastforce + apply (metis Posix1(1) Posix_STAR1.hyps(6) append_Nil append_Nil2) + using Posix1(2) by blast + moreover + have IHs: "\v2. s1 \ r \ v2 \ v = v2" + "\v2. s2 \ STAR r \ v2 \ Stars vs = v2" by fact+ + ultimately show "Stars (v # vs) = v2" by auto +next + case (Posix_STAR2 r v2) + have "[] \ STAR r \ v2" by fact + then show "Stars [] = v2" by cases (auto simp add: Posix1) +qed + + +text {* + Our POSIX value is a lexical value. +*} + +lemma Posix_LV: + assumes "s \ r \ v" + shows "v \ LV r s" +using assms unfolding LV_def +apply(induct rule: Posix.induct) +apply(auto simp add: intro!: Prf.intros elim!: Prf_elims) +done + +end \ No newline at end of file