--- a/thys/Spec.thy Wed Jul 19 14:55:46 2017 +0100
+++ b/thys/Spec.thy Fri Aug 11 20:29:01 2017 +0100
@@ -1,9 +1,8 @@
theory Spec
- imports Main
+ imports Main "~~/src/HOL/Library/Sublist"
begin
-
section {* Sequential Composition of Languages *}
definition
@@ -172,13 +171,15 @@
lemma ders_correctness:
shows "L (ders s r) = Ders s (L r)"
-apply(induct s arbitrary: r)
-apply(simp_all add: Ders_def der_correctness Der_def)
-done
+by (induct s arbitrary: r)
+ (simp_all add: Ders_def der_correctness Der_def)
+
section {* Lemmas about ders *}
+(* not really needed *)
+
lemma ders_ZERO:
shows "ders s (ZERO) = ZERO"
apply(induct s)
@@ -201,9 +202,8 @@
lemma ders_ALT:
shows "ders s (ALT r1 r2) = ALT (ders s r1) (ders s r2)"
-apply(induct s arbitrary: r1 r2)
-apply(simp_all)
-done
+by (induct s arbitrary: r1 r2)(simp_all)
+
section {* Values *}
@@ -229,8 +229,11 @@
| "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) = concat (map flat vs)"
+ "flat (Stars vs) = flats vs"
by (induct vs) (auto)
@@ -273,7 +276,7 @@
lemma Star_val:
assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<turnstile> v : r"
- shows "\<exists>vs. concat (map flat vs) = concat ss \<and> (\<forall>v\<in>set vs. \<turnstile> v : r)"
+ shows "\<exists>vs. flats vs = concat ss \<and> (\<forall>v\<in>set vs. \<turnstile> v : r)"
using assms
apply(induct ss)
apply(auto)
@@ -313,7 +316,7 @@
have "s \<in> L (STAR r)" by fact
then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r"
using Star_string by auto
- then obtain vs where "concat (map flat vs) = s" "\<forall>v\<in>set vs. \<turnstile> v : r"
+ then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<turnstile> v : r"
using IH Star_val by blast
then show "\<exists>v. \<turnstile> v : STAR r \<and> flat v = s"
using Prf.intros(6) flat_Stars by blast
@@ -331,8 +334,8 @@
shows "L(r) = {flat v | v. \<turnstile> v : r}"
using L_flat_Prf1 L_flat_Prf2 by blast
-section {* CPT and CPTpre *}
+section {* Canonical Values *}
inductive
CPrf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile> _ : _" [100, 100] 100)
@@ -350,71 +353,153 @@
using assms
by (induct)(auto intro: Prf.intros)
-lemma CPrf_stars:
- assumes "\<Turnstile> Stars vs : STAR r"
- shows "\<forall>v \<in> set vs. flat v \<noteq> [] \<and> \<Turnstile> v : r"
-using assms
-apply(erule_tac CPrf.cases)
-apply(simp_all)
-done
-
lemma CPrf_Stars_appendE:
assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r"
shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r"
using assms
apply(erule_tac CPrf.cases)
-apply(auto intro: CPrf.intros elim: Prf.cases)
+apply(auto intro: CPrf.intros)
done
-definition PT :: "rexp \<Rightarrow> string \<Rightarrow> val set"
-where "PT r s \<equiv> {v. flat v = s \<and> \<turnstile> v : r}"
+
+section {* Sets of Lexical and Canonical Values *}
-definition
- "CPT r s = {v. flat v = s \<and> \<Turnstile> v : r}"
+definition
+ LV :: "rexp \<Rightarrow> string \<Rightarrow> val set"
+where "LV r s \<equiv> {v. \<turnstile> v : r \<and> flat v = s}"
definition
- "CPTpre r s = {v. \<exists>s'. flat v @ s' = s \<and> \<Turnstile> v : r}"
+ CV :: "rexp \<Rightarrow> string \<Rightarrow> val set"
+where "CV r s \<equiv> {v. \<Turnstile> v : r \<and> flat v = s}"
+
+lemma LV_CV_subset:
+ shows "CV r s \<subseteq> LV r s"
+unfolding CV_def LV_def by(auto simp add: Prf_CPrf)
+
+abbreviation
+ "Prefixes s \<equiv> {s'. prefixeq s' s}"
+
+abbreviation
+ "Suffixes s \<equiv> {s'. suffixeq s' s}"
+
+abbreviation
+ "SSuffixes s \<equiv> {s'. suffix s' s}"
-lemma CPT_CPTpre_subset:
- shows "CPT r s \<subseteq> CPTpre r s"
-by(auto simp add: CPT_def CPTpre_def)
+lemma Suffixes_cons [simp]:
+ shows "Suffixes (c # s) = Suffixes s \<union> {c # s}"
+by (auto simp add: suffixeq_def Cons_eq_append_conv)
+
+lemma CV_simps:
+ shows "CV ZERO s = {}"
+ and "CV ONE s = (if s = [] then {Void} else {})"
+ and "CV (CHAR c) s = (if s = [c] then {Char c} else {})"
+ and "CV (ALT r1 r2) s = Left ` CV r1 s \<union> Right ` CV r2 s"
+unfolding CV_def
+by (auto intro: CPrf.intros elim: CPrf.cases)
+
+lemma finite_Suffixes:
+ shows "finite (Suffixes s)"
+by (induct s) (simp_all)
-lemma CPT_simps:
- shows "CPT ZERO s = {}"
- and "CPT ONE s = (if s = [] then {Void} else {})"
- and "CPT (CHAR c) s = (if s = [c] then {Char c} else {})"
- and "CPT (ALT r1 r2) s = Left ` CPT r1 s \<union> Right ` CPT r2 s"
- and "CPT (SEQ r1 r2) s =
- {Seq v1 v2 | v1 v2. flat v1 @ flat v2 = s \<and> v1 \<in> CPT r1 (flat v1) \<and> v2 \<in> CPT r2 (flat v2)}"
- and "CPT (STAR r) s =
- Stars ` {vs. concat (map flat vs) = s \<and> (\<forall>v \<in> set vs. v \<in> CPT r (flat v) \<and> flat v \<noteq> [])}"
-apply -
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
+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 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 CV_SEQ_subset:
+ "CV (SEQ r1 r2) s \<subseteq> (\<lambda>(v1,v2). Seq v1 v2) ` ((\<Union>s' \<in> Prefixes s. CV r1 s') \<times> (\<Union>s' \<in> Suffixes s. CV r2 s'))"
+unfolding image_def CV_def prefixeq_def suffixeq_def
+by (auto elim: CPrf.cases)
+
+lemma CV_STAR_subset:
+ "CV (STAR r) s \<subseteq> {Stars []} \<union>
+ (\<lambda>(v,vs). Stars (v#vs)) ` ((\<Union>s' \<in> Prefixes s. CV r s') \<times> (\<Union>s2 \<in> SSuffixes s. Stars -` (CV (STAR r) s2)))"
+unfolding image_def CV_def prefixeq_def suffix_def
+apply(auto)
apply(erule CPrf.cases)
-apply(simp_all)[6]
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[6]
-apply(erule CPrf.cases)
-apply(simp_all)[6]
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[6]
-apply(erule CPrf.cases)
-apply(simp_all)[6]
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[6]
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[6]
-(* STAR case *)
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[6]
+apply(auto)
+apply(case_tac vs)
+apply(auto intro: CPrf.intros)
done
+lemma CV_STAR_finite:
+ assumes "\<forall>s. finite (CV r s)"
+ shows "finite (CV (STAR r) s)"
+proof(induct s rule: length_induct)
+ fix s::"char list"
+ assume "\<forall>s'. length s' < length s \<longrightarrow> finite (CV (STAR r) s')"
+ then have IH: "\<forall>s' \<in> SSuffixes s. finite (CV (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. CV r s'"
+ def S2 \<equiv> "\<Union>s2 \<in> SSuffixes s. Stars -` (CV (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 "CV (STAR r) s \<subseteq> {Stars []} \<union> f ` (S1 \<times> S2)" unfolding S1_def S2_def f_def
+ by (rule CV_STAR_subset)
+ ultimately
+ show "finite (CV (STAR r) s)" by (simp add: finite_subset)
+qed
+
+
+lemma CV_finite:
+ shows "finite (CV r s)"
+proof(induct r arbitrary: s)
+ case (ZERO s)
+ show "finite (CV ZERO s)" by (simp add: CV_simps)
+next
+ case (ONE s)
+ show "finite (CV ONE s)" by (simp add: CV_simps)
+next
+ case (CHAR c s)
+ show "finite (CV (CHAR c) s)" by (simp add: CV_simps)
+next
+ case (ALT r1 r2 s)
+ then show "finite (CV (ALT r1 r2) s)" by (simp add: CV_simps)
+next
+ case (SEQ r1 r2 s)
+ def f \<equiv> "\<lambda>(v1, v2). Seq v1 v2"
+ def S1 \<equiv> "\<Union>s' \<in> Prefixes s. CV r1 s'"
+ def S2 \<equiv> "\<Union>s' \<in> Suffixes s. CV r2 s'"
+ have IHs: "\<And>s. finite (CV r1 s)" "\<And>s. finite (CV 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 "CV (SEQ r1 r2) s \<subseteq> f ` (S1 \<times> S2)"
+ unfolding f_def S1_def S2_def by (auto simp add: CV_SEQ_subset)
+ ultimately
+ show "finite (CV (SEQ r1 r2) s)"
+ by (simp add: finite_subset)
+next
+ case (STAR r s)
+ then show "finite (CV (STAR r) s)" by (simp add: CV_STAR_finite)
+qed
+
+
section {* Our POSIX Definition *}
@@ -531,12 +616,12 @@
Our POSIX value is a canonical value.
*}
-lemma Posix_CPT:
+lemma Posix_CV:
assumes "s \<in> r \<rightarrow> v"
- shows "v \<in> CPT r s"
+ shows "v \<in> CV r s"
using assms
apply(induct rule: Posix.induct)
-apply(auto simp add: CPT_def intro: CPrf.intros elim: CPrf.cases)
+apply(auto simp add: CV_def intro: CPrf.intros elim: CPrf.cases)
apply(rotate_tac 5)
apply(erule CPrf.cases)
apply(simp_all)
@@ -544,203 +629,17 @@
apply(simp_all)
done
-
-
-(*
-lemma CPTpre_STAR_finite:
- assumes "\<And>s. finite (CPT r s)"
- shows "finite (CPT (STAR r) s)"
-apply(induct s rule: length_induct)
-apply(case_tac xs)
-apply(simp)
-apply(simp add: CPT_simps)
-apply(auto)
-apply(rule finite_imageI)
-using assms
-thm finite_Un
-prefer 2
-apply(simp add: CPT_simps)
-apply(rule finite_imageI)
-apply(rule finite_subset)
-apply(rule CPTpre_subsets)
-apply(simp)
-apply(rule_tac B="(\<lambda>(v, vs). Stars (v#vs)) ` {(v, vs). v \<in> CPTpre r (a#list) \<and> flat v \<noteq> [] \<and> Stars vs \<in> CPTpre (STAR r) (drop (length (flat v)) (a#list))}" in finite_subset)
-apply(auto)[1]
-apply(rule finite_imageI)
-apply(simp add: Collect_case_prod_Sigma)
-apply(rule finite_SigmaI)
-apply(rule assms)
-apply(case_tac "flat v = []")
-apply(simp)
-apply(drule_tac x="drop (length (flat v)) (a # list)" in spec)
-apply(simp)
-apply(auto)[1]
-apply(rule test)
-apply(simp)
-done
-
-lemma CPTpre_subsets:
- "CPTpre ZERO s = {}"
- "CPTpre ONE s \<subseteq> {Void}"
- "CPTpre (CHAR c) s \<subseteq> {Char c}"
- "CPTpre (ALT r1 r2) s \<subseteq> Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s"
- "CPTpre (SEQ r1 r2) s \<subseteq> {Seq v1 v2 | v1 v2. v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}"
- "CPTpre (STAR r) s \<subseteq> {Stars []} \<union>
- {Stars (v#vs) | v vs. v \<in> CPTpre r s \<and> flat v \<noteq> [] \<and> Stars vs \<in> CPTpre (STAR r) (drop (length (flat v)) s)}"
- "CPTpre (STAR r) [] = {Stars []}"
-apply(auto simp add: CPTpre_def)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(erule CPrf.cases)
-apply(simp_all)
-
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(rule CPrf.intros)
-done
-
-
-lemma CPTpre_simps:
- shows "CPTpre ONE s = {Void}"
- and "CPTpre (CHAR c) (d#s) = (if c = d then {Char c} else {})"
- and "CPTpre (ALT r1 r2) s = Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s"
- and "CPTpre (SEQ r1 r2) s =
- {Seq v1 v2 | v1 v2. v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}"
-apply -
-apply(rule subset_antisym)
-apply(rule CPTpre_subsets)
-apply(auto simp add: CPTpre_def intro: "CPrf.intros")[1]
-apply(case_tac "c = d")
-apply(simp)
-apply(rule subset_antisym)
-apply(rule CPTpre_subsets)
-apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
-apply(simp)
-apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(rule subset_antisym)
-apply(rule CPTpre_subsets)
-apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
-apply(rule subset_antisym)
-apply(rule CPTpre_subsets)
-apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
-done
-
-lemma CPT_simps:
- shows "CPT ONE s = (if s = [] then {Void} else {})"
- and "CPT (CHAR c) [d] = (if c = d then {Char c} else {})"
- and "CPT (ALT r1 r2) s = Left ` CPT r1 s \<union> Right ` CPT r2 s"
- and "CPT (SEQ r1 r2) s =
- {Seq v1 v2 | v1 v2 s1 s2. s1 @ s2 = s \<and> v1 \<in> CPT r1 s1 \<and> v2 \<in> CPT r2 s2}"
-apply -
-apply(rule subset_antisym)
-apply(auto simp add: CPT_def)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[7]
-apply(erule CPrf.cases)
-apply(simp_all)[7]
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
-apply(auto simp add: CPT_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[7]
-apply(erule CPrf.cases)
-apply(simp_all)[7]
-apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[7]
-apply(clarify)
-apply blast
-apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1]
-apply(erule CPrf.cases)
-apply(simp_all)[7]
-done
-
-lemma test:
- assumes "finite A"
- shows "finite {vs. Stars vs \<in> A}"
-using assms
-apply(induct A)
-apply(simp)
-apply(auto)
-apply(case_tac x)
-apply(simp_all)
-done
-
-lemma CPTpre_STAR_finite:
- assumes "\<And>s. finite (CPTpre r s)"
- shows "finite (CPTpre (STAR r) s)"
-apply(induct s rule: length_induct)
-apply(case_tac xs)
-apply(simp)
-apply(simp add: CPTpre_subsets)
-apply(rule finite_subset)
-apply(rule CPTpre_subsets)
-apply(simp)
-apply(rule_tac B="(\<lambda>(v, vs). Stars (v#vs)) ` {(v, vs). v \<in> CPTpre r (a#list) \<and> flat v \<noteq> [] \<and> Stars vs \<in> CPTpre (STAR r) (drop (length (flat v)) (a#list))}" in finite_subset)
-apply(auto)[1]
-apply(rule finite_imageI)
-apply(simp add: Collect_case_prod_Sigma)
-apply(rule finite_SigmaI)
-apply(rule assms)
-apply(case_tac "flat v = []")
-apply(simp)
-apply(drule_tac x="drop (length (flat v)) (a # list)" in spec)
-apply(simp)
-apply(auto)[1]
-apply(rule test)
-apply(simp)
-done
-
-lemma CPTpre_finite:
- shows "finite (CPTpre r s)"
-apply(induct r arbitrary: s)
-apply(simp add: CPTpre_subsets)
-apply(rule finite_subset)
-apply(rule CPTpre_subsets)
-apply(simp)
-apply(rule finite_subset)
-apply(rule CPTpre_subsets)
-apply(simp)
-apply(rule finite_subset)
-apply(rule CPTpre_subsets)
-apply(rule_tac B="(\<lambda>(v1, v2). Seq v1 v2) ` {(v1, v2). v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}" in finite_subset)
-apply(auto)[1]
-apply(rule finite_imageI)
-apply(simp add: Collect_case_prod_Sigma)
-apply(rule finite_subset)
-apply(rule CPTpre_subsets)
-apply(simp)
-by (simp add: CPTpre_STAR_finite)
-
-
-lemma CPT_finite:
- shows "finite (CPT r s)"
-apply(rule finite_subset)
-apply(rule CPT_CPTpre_subset)
-apply(rule CPTpre_finite)
-done
-*)
-
lemma test2:
assumes "\<forall>v \<in> set vs. flat v \<in> r \<rightarrow> v \<and> flat v \<noteq> []"
- shows "(Stars vs) \<in> CPT (STAR r) (flat (Stars vs))"
+ shows "(Stars vs) \<in> CV (STAR r) (flat (Stars vs))"
using assms
apply(induct vs)
-apply(auto simp add: CPT_def)
+apply(auto simp add: CV_def)
apply(rule CPrf.intros)
apply(simp)
apply(rule CPrf.intros)
apply(simp_all)
-by (metis (no_types, lifting) CPT_def Posix_CPT mem_Collect_eq)
+by (metis (no_types, lifting) CV_def Posix_CV mem_Collect_eq)
end
\ No newline at end of file