thys/Positions.thy
changeset 266 fff2e1b40dfc
parent 265 d36be1e356c0
child 267 32b222d77fa0
--- a/thys/Positions.thy	Tue Jul 18 18:39:20 2017 +0100
+++ b/thys/Positions.thy	Wed Jul 19 14:55:46 2017 +0100
@@ -1,6 +1,6 @@
    
 theory Positions
-  imports "Lexer" 
+  imports "Spec" 
 begin
 
 section {* Positions in Values *}
@@ -163,26 +163,6 @@
 
 
 
-definition pflat_len2 :: "val \<Rightarrow> nat list => (bool * nat)"
-where
-  "pflat_len2 v p \<equiv> (if p \<in> Pos v then (True, length (flat (at v p))) else (False, 0))"
-
-instance prod :: (ord, ord) ord
-  by (rule Orderings.class.Orderings.ord.of_class.intro)
-
-
-lemma "(0, 0) < (3::nat, 2::nat)"
-
-
-
-
-definition PosOrd2:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>val2 _ _" [60, 60, 59] 60)
-where
-  "v1 \<sqsubset>val2 p v2 \<equiv> (fst (pflat_len2 v1 p) > fst (pflat_len2 v2 p) \<or>
-                     snd (pflat_len2 v1 p) > fst (pflat_len2 v2 p)) \<and>
-                    (\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q)"
-
-
 definition PosOrd_ex:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubset>val _" [60, 59] 60)
 where
   "v1 :\<sqsubset>val v2 \<equiv> \<exists>p. v1 \<sqsubset>val p v2"
@@ -567,287 +547,6 @@
 by (metis PosOrd_SeqI1 PosOrd_shorterI WW1 antisym_conv3 append_eq_append_conv assms(2))
 
 
-section {* CPT and CPTpre *}
-
-
-inductive 
-  CPrf :: "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"
-| "\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> [] \<Longrightarrow> \<Turnstile> Stars vs : STAR r"
-
-lemma Prf_CPrf:
-  assumes "\<Turnstile> v : r"
-  shows "\<turnstile> v : r"
-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)
-done
-
-definition PT :: "rexp \<Rightarrow> string \<Rightarrow> val set"
-where "PT r s \<equiv> {v. flat v = s \<and> \<turnstile> v : r}"
-
-definition
-  "CPT r s = {v. flat v = s \<and> \<Turnstile> v : r}"
-
-definition
-  "CPTpre r s = {v. \<exists>s'. flat v @ s' = s \<and> \<Turnstile> v : r}"
-
-lemma CPT_CPTpre_subset:
-  shows "CPT r s \<subseteq> CPTpre r s"
-by(auto simp add: CPT_def CPTpre_def)
-
-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]
-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]
-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 Posix_CPT:
-  assumes "s \<in> r \<rightarrow> v"
-  shows "v \<in> CPT r s"
-using assms
-apply(induct rule: Posix.induct)
-apply(auto simp add: CPT_def intro: CPrf.intros elim: CPrf.cases)
-apply(rotate_tac 5)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(rule CPrf.intros)
-apply(simp_all)
-done
-
-
 section {* The Posix Value is smaller than any other Value *}
 
 
@@ -1044,81 +743,6 @@
 by (metis Posix_PosOrd_stronger less_irrefl PosOrd_def 
     PosOrd_ex_eq_def PosOrd_ex_def PosOrd_trans)
 
-(*
-lemma PosOrd_Posix_Stars:
-  assumes "(Stars vs) \<in> CPT (STAR r) (flat (Stars vs))" "\<forall>v \<in> set vs. flat v \<in> r \<rightarrow> v"
-  and "\<not>(\<exists>vs2 \<in> PT (STAR r) (flat (Stars vs)). vs2 :\<sqsubset>val (Stars vs))"
-  shows "(flat (Stars vs)) \<in> (STAR r) \<rightarrow> Stars vs" 
-using assms
-apply(induct vs)
-apply(simp)
-apply(rule Posix.intros)
-apply(simp (no_asm))
-apply(rule Posix.intros)
-apply(auto)[1]
-apply(auto simp add: CPT_def PT_def)[1]
-defer
-apply(simp)
-apply(drule meta_mp)
-apply(auto simp add: CPT_def PT_def)[1]
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(drule meta_mp)
-apply(auto simp add: CPT_def PT_def)[1]
-apply(erule Prf.cases)
-apply(simp_all)
-using CPrf_stars PosOrd_irrefl apply fastforce
-apply(clarify)
-apply(drule_tac x="Stars (a#v#vsa)" in spec)
-apply(simp)
-apply(drule mp)
-apply (meson CPrf_stars Prf.intros(7) Prf_CPrf list.set_intros(1))
-apply(subst (asm) (2) PosOrd_ex_def)
-apply(simp)
-apply (metis flat.simps(7) flat_Stars PosOrd_StarsI2 PosOrd_ex_def)
-apply(auto simp add: CPT_def PT_def)[1]
-using CPrf_stars apply auto[1]
-apply(auto)[1]
-apply(auto simp add: CPT_def PT_def)[1]
-apply(subgoal_tac "\<exists>vA. flat vA = flat a @ s\<^sub>3 \<and> \<turnstile> vA : r")
-prefer 2
-apply (meson L_flat_Prf2)
-apply(subgoal_tac "\<exists>vB. flat (Stars vB) = s\<^sub>4 \<and> \<turnstile> (Stars vB) : (STAR r)")
-apply(clarify)
-apply(drule_tac x="Stars (vA # vB)" in spec)
-apply(simp)
-apply(drule mp)
-using Prf.intros(7) apply blast
-apply(subst (asm) (2) PosOrd_ex_def)
-apply(simp)
-prefer 2
-apply(simp)
-using Star_values_exists apply blast
-prefer 2
-apply(drule meta_mp)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(drule meta_mp)
-apply(auto)[1]
-prefer 2
-apply(simp)
-apply(erule CPrf.cases)
-apply(simp_all)
-apply(clarify)
-apply(rotate_tac 3)
-apply(erule Prf.cases)
-apply(simp_all)
-apply (metis CPrf_stars intlen.cases less_irrefl list.set_intros(1) PosOrd_def PosOrd_ex_def)
-apply(drule_tac x="Stars (v#va#vsb)" in spec)
-apply(drule mp)
-apply (simp add: Posix1a Prf.intros(7))
-apply(simp)
-apply(subst (asm) (2) PosOrd_ex_def)
-apply(simp)
-apply (metis flat.simps(7) flat_Stars PosOrd_StarsI2 PosOrd_ex_def)
-by (metis PosOrd_StarsI PosOrd_ex_def PosOrd_spreI append_assoc append_self_conv flat.simps(7) flat_Stars prefix_list_def sprefix_list_def)
-*)
-
 
 lemma test2: 
   assumes "\<forall>v \<in> set vs. flat v \<in> r \<rightarrow> v \<and> flat v \<noteq> []"
@@ -1163,7 +787,7 @@
         apply(simp_all)
         apply(drule_tac x="Stars (v # vs)" in bspec)
         apply(simp add: PT_def CPT_def)
-        using Posix1a Prf.intros(6) calculation
+        using Posix_Prf Prf.intros(6) calculation
         apply(rule_tac Prf.intros)
         apply(simp add:)
         apply (simp add: PosOrd_StarsI2)