theory SpecExt imports Main "~~/src/HOL/Library/Sublist"beginsection {* Sequential Composition of Languages *}definition Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)where "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> 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 blastby (metis append_assoc)lemma Sequ_Union_in: shows "(A ;; (\<Union>x\<in> B. C x)) = (\<Union>x\<in> B. A ;; C x)" by (auto simp add: Sequ_def)section {* Semantic Derivative (Left Quotient) of Languages *}definition Der :: "char \<Rightarrow> string set \<Rightarrow> string set"where "Der c A \<equiv> {s. c # s \<in> A}"definition Ders :: "string \<Rightarrow> string set \<Rightarrow> string set"where "Ders s A \<equiv> {s'. s @ s' \<in> A}"lemma Der_null [simp]: shows "Der c {} = {}"unfolding Der_defby autolemma Der_empty [simp]: shows "Der c {[]} = {}"unfolding Der_defby autolemma Der_char [simp]: shows "Der c {[d]} = (if c = d then {[]} else {})"unfolding Der_defby autolemma Der_union [simp]: shows "Der c (A \<union> B) = Der c A \<union> Der c B"unfolding Der_defby autolemma Der_UNION [simp]: shows "Der c (\<Union>x\<in>A. B x) = (\<Union>x\<in>A. Der c (B x))"by (auto simp add: Der_def)lemma Der_Sequ [simp]: shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})"unfolding Der_def Sequ_defby (auto simp add: Cons_eq_append_conv)section {* Kleene Star for Languages *}inductive_set Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102) for A :: "string set"where start[intro]: "[] \<in> A\<star>"| step[intro]: "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>"(* Arden's lemma *)lemma Star_cases: shows "A\<star> = {[]} \<union> A ;; A\<star>"unfolding Sequ_defby (auto) (metis Star.simps)lemma Star_decomp: assumes "c # x \<in> A\<star>" shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A\<star>"using assmsby (induct x\<equiv>"c # x" rule: Star.induct) (auto simp add: append_eq_Cons_conv)lemma Star_Der_Sequ: shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"unfolding Der_def Sequ_defby(auto simp add: Star_decomp)lemma Der_star [simp]: shows "Der c (A\<star>) = (Der c A) ;; A\<star>"proof - have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)" by (simp only: Star_cases[symmetric]) also have "... = Der c (A ;; A\<star>)" by (simp only: Der_union Der_empty) (simp) also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})" by simp also have "... = (Der c A) ;; A\<star>" using Star_Der_Sequ by auto finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .qedsection {* Power operation for Sets *}fun Pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _" [101, 102] 101)where "A \<up> 0 = {[]}"| "A \<up> (Suc n) = A ;; (A \<up> n)"lemma Pow_empty [simp]: shows "[] \<in> A \<up> n \<longleftrightarrow> (n = 0 \<or> [] \<in> A)"by(induct n) (auto simp add: Sequ_def)lemma Pow_Suc_rev: "A \<up> (Suc n) = (A \<up> n) ;; A"apply(induct n arbitrary: A)apply(simp_all)by (metis Sequ_assoc)lemma Pow_decomp: assumes "c # x \<in> A \<up> n" shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A \<up> (n - 1)"using assmsapply(induct n) apply(auto simp add: Cons_eq_append_conv Sequ_def)apply(case_tac n)apply(auto simp add: Sequ_def)apply(blast)donelemma Star_Pow: assumes "s \<in> A\<star>" shows "\<exists>n. s \<in> A \<up> n"using assmsapply(induct)apply(auto)apply(rule_tac x="Suc n" in exI)apply(auto simp add: Sequ_def)donelemma Pow_Star: assumes "s \<in> A \<up> n" shows "s \<in> A\<star>"using assmsapply(induct n arbitrary: s)apply(auto simp add: Sequ_def)donelemma Der_Pow_0: shows "Der c (A \<up> 0) = {}"by(simp add: Der_def)lemma Der_Pow_Suc: shows "Der c (A \<up> (Suc n)) = (Der c A) ;; (A \<up> 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)+donelemma Der_Pow [simp]: shows "Der c (A \<up> n) = (if n = 0 then {} else (Der c A) ;; (A \<up> (n - 1)))"apply(case_tac n)apply(simp_all del: Pow.simps add: Der_Pow_0 Der_Pow_Suc)donelemma Der_Pow_Sequ [simp]: shows "Der c (A ;; A \<up> n) = (Der c A) ;; (A \<up> n)"by (simp only: Pow.simps[symmetric] Der_Pow) (simp)lemma Pow_Sequ_Un: assumes "0 < x" shows "(\<Union>n \<in> {..x}. (A \<up> n)) = ({[]} \<union> (\<Union>n \<in> {..x - Suc 0}. A ;; (A \<up> n)))"using assmsapply(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)donelemma Pow_Sequ_Un2: assumes "0 < x" shows "(\<Union>n \<in> {x..}. (A \<up> n)) = (\<Union>n \<in> {x - Suc 0..}. A ;; (A \<up> n))"using assmsapply(auto simp add: Sequ_def)apply(case_tac n)apply(auto simp add: Sequ_def)apply fastforceapply(case_tac x)apply(auto)apply(rule_tac x="Suc xa" in bexI)apply(auto simp add: Sequ_def)donesection {* 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 natsection {* Semantics of Regular Expressions *}fun L :: "rexp \<Rightarrow> 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) \<union> (L r2)"| "L (STAR r) = (L r)\<star>"| "L (UPNTIMES r n) = (\<Union>i\<in> {..n} . (L r) \<up> i)"| "L (NTIMES r n) = (L r) \<up> n"| "L (FROMNTIMES r n) = (\<Union>i\<in> {n..} . (L r) \<up> i)"| "L (NMTIMES r n m) = (\<Union>i\<in>{n..m} . (L r) \<up> i)" section {* Nullable, Derivatives *}fun nullable :: "rexp \<Rightarrow> bool"where "nullable (ZERO) = False"| "nullable (ONE) = True"| "nullable (CHAR c) = False"| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"| "nullable (SEQ r1 r2) = (nullable r1 \<and> 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 \<Rightarrow> rexp \<Rightarrow> 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 \<Rightarrow> rexp \<Rightarrow> rexp"where "ders [] r = r"| "ders (c # s) r = ders s (der c r)"lemma nullable_correctness: shows "nullable r \<longleftrightarrow> [] \<in> (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 2apply(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 2apply(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 fastforcelemma 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 \<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)"| "flat (Stars []) = []"| "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" abbreviation "flats vs \<equiv> concat (map flat vs)"lemma flat_Stars [simp]: "flat (Stars vs) = flats vs"by (induct vs) (auto)lemma Star_concat: assumes "\<forall>s \<in> set ss. s \<in> A" shows "concat ss \<in> A\<star>"using assms by (induct ss) (auto)lemma Star_cstring: assumes "s \<in> A\<star>" shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])"using assmsapply(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)donelemma Aux: assumes "\<forall>s\<in>set ss. s = []" shows "concat ss = []"using assmsby (induct ss) (auto)lemma Pow_cstring_nonempty: assumes "s \<in> A \<up> n" shows "\<exists>ss. concat ss = s \<and> length ss \<le> n \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])"using assmsapply(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)donelemma Pow_cstring: assumes "s \<in> A \<up> n" shows "\<exists>ss1 ss2. concat (ss1 @ ss2) = s \<and> length (ss1 @ ss2) = n \<and> (\<forall>s \<in> set ss1. s \<in> A \<and> s \<noteq> []) \<and> (\<forall>s \<in> set ss2. s \<in> A \<and> s = [])"using assmsapply(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)donesection {* Lexical Values *}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"| "\<Turnstile> v2 : r2 \<Longrightarrow> \<Turnstile> Right v2 : ALT r1 r2"| "\<Turnstile> Void : ONE"| "\<Turnstile> Char c : CHAR c"| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : STAR r"| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []; length vs \<le> n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : UPNTIMES r n"| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; \<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; length (vs1 @ vs2) = n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : NTIMES r n"| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; \<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; length (vs1 @ vs2) \<ge> n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : FROMNTIMES r n"| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; \<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; length (vs1 @ vs2) \<ge> n; length (vs1 @ vs2) \<le> m\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : NMTIMES r n m"inductive_cases Prf_elims: "\<Turnstile> v : ZERO" "\<Turnstile> v : SEQ r1 r2" "\<Turnstile> v : ALT r1 r2" "\<Turnstile> v : ONE" "\<Turnstile> v : CHAR c" "\<Turnstile> vs : STAR r" "\<Turnstile> vs : UPNTIMES r n" "\<Turnstile> vs : NTIMES r n" "\<Turnstile> vs : FROMNTIMES r n" "\<Turnstile> vs : NMTIMES r n m"lemma Prf_Stars_appendE: assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r" shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r" using assmsby (auto intro: Prf.intros elim!: Prf_elims)lemma flats_empty: assumes "(\<forall>v\<in>set vs. flat v = [])" shows "flats vs = []"using assmsby(induct vs) (simp_all)lemma Star_cval: assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r" shows "\<exists>vs. flats vs = concat ss \<and> (\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])"using assmsapply(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)donelemma flats_cval: assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r" shows "\<exists>vs1 vs2. flats (vs1 @ vs2) = concat ss \<and> length (vs1 @ vs2) = length ss \<and> (\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []) \<and> (\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = [])"using assmsapply(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)donelemma flats_cval_nonempty: assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r" shows "\<exists>vs. flats vs = concat ss \<and> length vs \<le> length ss \<and> (\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])" using assmsapply(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)donelemma Pow_flats: assumes "\<forall>v \<in> set vs. flat v \<in> A" shows "flats vs \<in> A \<up> length vs"using assmsby(induct vs)(auto simp add: Sequ_def)lemma Pow_flats_appends: assumes "\<forall>v \<in> set vs1. flat v \<in> A" "\<forall>v \<in> set vs2. flat v \<in> A" shows "flats vs1 @ flats vs2 \<in> A \<up> (length vs1 + length vs2)"using assmsapply(induct vs1)apply(auto simp add: Sequ_def Pow_flats)donelemma L_flat_Prf1: assumes "\<Turnstile> v : r" shows "flat v \<in> L r"using assmsapply(induct) apply(auto simp add: Sequ_def Star_concat Pow_flats)apply(meson Pow_flats atMost_iff)using Pow_flats_appends apply blastapply(meson Pow_flats_appends atLeast_iff)apply(meson Pow_flats_appends atLeastAtMost_iff)donelemma L_flat_Prf2: assumes "s \<in> L r" shows "\<exists>v. \<Turnstile> v : r \<and> flat v = s"using assmsproof(induct r arbitrary: s) case (STAR r s) have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact have "s \<in> L (STAR r)" by fact then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []" using Star_cstring by auto then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []" using IH Star_cval by metis then show "\<exists>v. \<Turnstile> v : STAR r \<and> flat v = s" using Prf.intros(6) flat_Stars by blastnext case (SEQ r1 r2 s) then show "\<exists>v. \<Turnstile> v : SEQ r1 r2 \<and> flat v = s" unfolding Sequ_def L.simps by (fastforce intro: Prf.intros)next case (ALT r1 r2 s) then show "\<exists>v. \<Turnstile> v : ALT r1 r2 \<and> flat v = s" unfolding L.simps by (fastforce intro: Prf.intros)next case (NTIMES r n) have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact have "s \<in> L (NTIMES r n)" by fact then obtain ss1 ss2 where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = n" "\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []" using Pow_cstring by force then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = n" "\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> 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 "\<exists>v. \<Turnstile> v : NTIMES r n \<and> flat v = s" using Prf.intros(8) flat_Stars by blastnext case (FROMNTIMES r n) have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact have "s \<in> L (FROMNTIMES r n)" by fact then obtain ss1 ss2 m where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = m" "n \<le> m" "\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []" using Pow_cstring by auto blast then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = m" "n \<le> m" "\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> 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 "\<exists>v. \<Turnstile> v : FROMNTIMES r n \<and> flat v = s" using Prf.intros(9) flat_Stars by blastnext case (NMTIMES r n m) have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact have "s \<in> L (NMTIMES r n m)" by fact then obtain ss1 ss2 k where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = k" "n \<le> k" "k \<le> m" "\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []" using Pow_cstring by (auto, blast) then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = k" "n \<le> k" "k \<le> m" "\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> 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 "\<exists>v. \<Turnstile> v : NMTIMES r n m \<and> flat v = s" apply(rule_tac x="Stars (vs1 @ vs2)" in exI) apply(simp) apply(rule Prf.intros) apply(auto) donenext case (UPNTIMES r n s) have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact have "s \<in> L (UPNTIMES r n)" by fact then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []" "length ss \<le> n" using Pow_cstring_nonempty by force then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []" "length vs \<le> n" using IH flats_cval_nonempty by (smt order.trans) then show "\<exists>v. \<Turnstile> v : UPNTIMES r n \<and> flat v = s" using Prf.intros(7) flat_Stars by blastqed (auto intro: Prf.intros)lemma L_flat_Prf: shows "L(r) = {flat v | v. \<Turnstile> v : r}"using L_flat_Prf1 L_flat_Prf2 by blastsection {* Sets of Lexical Values *}text {* Shows that lexical values are finite for a given regex and string.*}definition LV :: "rexp \<Rightarrow> string \<Rightarrow> val set"where "LV r s \<equiv> {v. \<Turnstile> v : r \<and> 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 \<union> Right ` LV r2 s"unfolding LV_defapply(auto intro: Prf.intros elim: Prf.cases)doneabbreviation "Prefixes s \<equiv> {s'. prefixeq s' s}"abbreviation "Suffixes s \<equiv> {s'. suffixeq s' s}"abbreviation "SSuffixes s \<equiv> {s'. suffix s' s}"lemma Suffixes_cons [simp]: shows "Suffixes (c # s) = Suffixes s \<union> {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 \<subseteq> Suffixes s" unfolding suffix_def suffixeq_def by auto then show "finite (SSuffixes s)" using finite_Suffixes finite_subset by blastqedlemma 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 simpqedlemma LV_STAR_finite: assumes "\<forall>s. finite (LV r s)" shows "finite (LV (STAR r) s)"proof(induct s rule: length_induct) fix s::"char list" assume "\<forall>s'. length s' < length s \<longrightarrow> finite (LV (STAR r) s')" then have IH: "\<forall>s' \<in> SSuffixes s. finite (LV (STAR r) s')" by (auto simp add: suffix_def) def f \<equiv> "\<lambda>(v, vs). Stars (v # vs)" def S1 \<equiv> "\<Union>s' \<in> Prefixes s. LV r s'" def S2 \<equiv> "\<Union>s2 \<in> 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 []} \<union> f ` (S1 \<times> S2))" by simp moreover have "LV (STAR r) s \<subseteq> {Stars []} \<union> f ` (S1 \<times> 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 \<subseteq> LV (STAR r) s"by(auto simp add: LV_def intro: Prf.intros elim: Prf_elims)(*lemma LV_NTIMES_finite: assumes "\<forall>s. finite (LV r s)" shows "finite (LV (NTIMES r n) s)"proof(induct s rule: length_induct) fix s::"char list" assume "\<forall>s'. length s' < length s \<longrightarrow> finite (LV (NTIMES r n) s')" then have IH: "\<forall>s' \<in> SSuffixes s. finite (LV (NTIMES r n) s')" by (auto simp add: suffix_def) def f \<equiv> "\<lambda>(v, vs). Stars (v # vs)" def S1 \<equiv> "\<Union>s' \<in> Prefixes s. LV r s'" def S2 \<equiv> "\<Union>s2 \<in> SSuffixes s. Stars -` (LV (NTIMES r n) 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 []} \<union> f ` (S1 \<times> S2))" by simp moreover have "LV (NTIMES r n) s \<subseteq> {Stars []} \<union> f ` (S1 \<times> 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 vs1) apply(auto intro: Prf.intros) done ultimately show "finite (LV (STAR r) s)" by (simp add: finite_subset)qed *)lemma LV_NTIMES_0: shows "LV (NTIMES r 0) s \<subseteq> {Stars []}"unfolding LV_defapply(auto elim: Prf_elims)donelemma LV_NTIMES_1: shows "LV (NTIMES r 1) s \<subseteq> (\<lambda>v. Stars [v]) ` (LV r s)"unfolding LV_defapply(auto elim!: Prf_elims)apply(case_tac vs1)apply(simp)apply(case_tac vs2)apply(simp)apply(simp)apply(simp)donelemma LV_NTIMES_2: shows "LV (NTIMES r 2) [] \<subseteq> (\<lambda>(v1,v2). Stars [v1,v2]) ` (LV r [] \<times> LV r [])"unfolding LV_defapply(auto elim!: Prf_elims simp add: image_def)apply(case_tac vs1)apply(auto)apply(case_tac vs2)apply(auto)apply(case_tac list)apply(auto)donelemma LV_NTIMES_3: shows "LV (NTIMES r (Suc n)) [] = (\<lambda>(v,vs). Stars (v#vs)) ` (LV r [] \<times> (Stars -` (LV (NTIMES r n) [])))"unfolding LV_defapply(auto elim!: Prf_elims simp add: image_def)apply(case_tac vs1)apply(auto)apply(case_tac vs2)apply(auto)apply(subst append.simps(1)[symmetric])apply(rule Prf.intros)apply(auto)apply(subst append.simps(1)[symmetric])apply(rule Prf.intros)apply(auto)donethm card_cartesian_productlemma LV_empty_finite: shows "card (LV (NTIMES r n) []) \<le> ((card (LV r [])) ^ n)"apply(induct n arbitrary:)using LV_NTIMES_0apply (metis card_empty card_insert_disjoint card_mono empty_iff finite.emptyI finite.insertI nat_power_eq_Suc_0_iff)apply(simp add: LV_NTIMES_3)apply(subst card_image)apply(simp add: inj_on_def)apply(subst card_cartesian_product)apply(subst card_vimage_inj)apply(simp add: inj_on_def)apply(auto simp add: LV_def elim: Prf_elims)[1]using nat_mult_le_cancel_disj by blastlemma LV_NTIMES_STAR: "LV (NTIMES r n) s \<subseteq> LV (STAR r) s"apply(auto simp add: LV_def intro: Prf.intros elim!: Prf_elims)apply(rule Prf.intros)oopslemma LV_FROMNTIMES_STAR: "LV (FROMNTIMES r n) s \<subseteq> LV (STAR r) s"apply(auto simp add: LV_def intro: Prf.intros elim!: Prf_elims)oopslemma 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 \<equiv> "\<lambda>(v1, v2). Seq v1 v2" def S1 \<equiv> "\<Union>s' \<in> Prefixes s. LV r1 s'" def S2 \<equiv> "\<Union>s' \<in> Suffixes s. LV r2 s'" have IHs: "\<And>s. finite (LV r1 s)" "\<And>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 \<subseteq> f ` (S1 \<times> 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 (UPNTIMES r n s) have "\<And>s. finite (LV r s)" by fact then show "finite (LV (UPNTIMES r n) s)" by (meson LV_STAR_finite LV_UPNTIMES_STAR rev_finite_subset)next case (FROMNTIMES r n s) have "\<And>s. finite (LV r s)" by fact then show "finite (LV (FROMNTIMES r n) s)"qedsection {* Our POSIX Definition *}inductive Posix :: "string \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<in> _ \<rightarrow> _" [100, 100, 100] 100)where Posix_ONE: "[] \<in> ONE \<rightarrow> Void"| Posix_CHAR: "[c] \<in> (CHAR c) \<rightarrow> (Char c)"| Posix_ALT1: "s \<in> r1 \<rightarrow> v \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Left v)"| Posix_ALT2: "\<lbrakk>s \<in> r2 \<rightarrow> v; s \<notin> L(r1)\<rbrakk> \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Right v)"| Posix_SEQ: "\<lbrakk>s1 \<in> r1 \<rightarrow> v1; s2 \<in> r2 \<rightarrow> v2; \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r1 \<and> s\<^sub>4 \<in> L r2)\<rbrakk> \<Longrightarrow> (s1 @ s2) \<in> (SEQ r1 r2) \<rightarrow> (Seq v1 v2)"| Posix_STAR1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v \<noteq> []; \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk> \<Longrightarrow> (s1 @ s2) \<in> STAR r \<rightarrow> Stars (v # vs)"| Posix_STAR2: "[] \<in> STAR r \<rightarrow> Stars []"inductive_cases Posix_elims: "s \<in> ZERO \<rightarrow> v" "s \<in> ONE \<rightarrow> v" "s \<in> CHAR c \<rightarrow> v" "s \<in> ALT r1 r2 \<rightarrow> v" "s \<in> SEQ r1 r2 \<rightarrow> v" "s \<in> STAR r \<rightarrow> v"lemma Posix1: assumes "s \<in> r \<rightarrow> v" shows "s \<in> L r" "flat v = s"using assmsby (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 \<in> r \<rightarrow> v1" "s \<in> r \<rightarrow> v2" shows "v1 = v2"using assmsproof (induct s r v1 arbitrary: v2 rule: Posix.induct) case (Posix_ONE v2) have "[] \<in> ONE \<rightarrow> v2" by fact then show "Void = v2" by cases autonext case (Posix_CHAR c v2) have "[c] \<in> CHAR c \<rightarrow> v2" by fact then show "Char c = v2" by cases autonext case (Posix_ALT1 s r1 v r2 v2) have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact moreover have "s \<in> r1 \<rightarrow> v" by fact then have "s \<in> L r1" by (simp add: Posix1) ultimately obtain v' where eq: "v2 = Left v'" "s \<in> r1 \<rightarrow> v'" by cases auto moreover have IH: "\<And>v2. s \<in> r1 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact ultimately have "v = v'" by simp then show "Left v = v2" using eq by simpnext case (Posix_ALT2 s r2 v r1 v2) have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact moreover have "s \<notin> L r1" by fact ultimately obtain v' where eq: "v2 = Right v'" "s \<in> r2 \<rightarrow> v'" by cases (auto simp add: Posix1) moreover have IH: "\<And>v2. s \<in> r2 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact ultimately have "v = v'" by simp then show "Right v = v2" using eq by simpnext case (Posix_SEQ s1 r1 v1 s2 r2 v2 v') have "(s1 @ s2) \<in> SEQ r1 r2 \<rightarrow> v'" "s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by fact+ then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \<in> r1 \<rightarrow> v1'" "s2 \<in> r2 \<rightarrow> v2'" apply(cases) apply (auto simp add: append_eq_append_conv2) using Posix1(1) by fastforce+ moreover have IHs: "\<And>v1'. s1 \<in> r1 \<rightarrow> v1' \<Longrightarrow> v1 = v1'" "\<And>v2'. s2 \<in> r2 \<rightarrow> v2' \<Longrightarrow> v2 = v2'" by fact+ ultimately show "Seq v1 v2 = v'" by simpnext case (Posix_STAR1 s1 r v s2 vs v2) have "(s1 @ s2) \<in> STAR r \<rightarrow> v2" "s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" "flat v \<noteq> []" "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact+ then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (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: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" "\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ ultimately show "Stars (v # vs) = v2" by autonext case (Posix_STAR2 r v2) have "[] \<in> STAR r \<rightarrow> v2" by fact then show "Stars [] = v2" by cases (auto simp add: Posix1)qedtext {* Our POSIX value is a lexical value.*}lemma Posix_LV: assumes "s \<in> r \<rightarrow> v" shows "v \<in> LV r s"using assms unfolding LV_defapply(induct rule: Posix.induct)apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)doneend