diff -r fff2e1b40dfc -r 32b222d77fa0 thys/Positions.thy --- a/thys/Positions.thy Wed Jul 19 14:55:46 2017 +0100 +++ b/thys/Positions.thy Fri Aug 11 20:29:01 2017 +0100 @@ -31,10 +31,8 @@ lemma Pos_stars: "Pos (Stars vs) = {[]} \ (\n < length vs. {n#ps | ps. ps \ Pos (vs ! n)})" apply(induct vs) -apply(simp) -apply(simp add: insert_ident) -apply(rule subset_antisym) -using less_Suc_eq_0_disj by auto +apply(auto simp add: insert_ident less_Suc_eq_0_disj) +done lemma Pos_empty: shows "[] \ Pos v" @@ -45,31 +43,25 @@ "intlen [] = 0" | "intlen (x # xs) = 1 + intlen xs" +lemma intlen_int: + shows "intlen xs = int (length xs)" +by (induct xs)(simp_all) + lemma intlen_bigger: shows "0 \ intlen xs" by (induct xs)(auto) lemma intlen_append: shows "intlen (xs @ ys) = intlen xs + intlen ys" -by (induct xs arbitrary: ys) (auto) +by (simp add: intlen_int) lemma intlen_length: shows "intlen xs < intlen ys \ length xs < length ys" -apply(induct xs arbitrary: ys) -apply (auto simp add: intlen_bigger not_less) -apply (metis intlen.elims intlen_bigger le_imp_0_less) -apply (smt Suc_lessI intlen.simps(2) length_Suc_conv nat_neq_iff) -by (smt Suc_lessE intlen.simps(2) length_Suc_conv) +by (simp add: intlen_int) lemma intlen_length_eq: shows "intlen xs = intlen ys \ length xs = length ys" -apply(induct xs arbitrary: ys) -apply (auto simp add: intlen_bigger not_less) -apply(case_tac ys) -apply(simp_all) -apply (smt intlen_bigger) -apply (smt intlen.elims intlen_bigger length_Suc_conv) -by (metis intlen.simps(2) length_Suc_conv) +by (simp add: intlen_int) definition pflat_len :: "val \ nat list => int" where @@ -90,7 +82,7 @@ lemma pflat_len_Stars_simps: assumes "n < length vs" shows "pflat_len (Stars vs) (n#p) = pflat_len (vs!n) p" -using assms +using assms apply(induct vs arbitrary: n p) apply(auto simp add: less_Suc_eq_0_disj pflat_len_simps) done @@ -98,7 +90,8 @@ lemma pflat_len_outside: assumes "p \ Pos v1" shows "pflat_len v1 p = -1 " -using assms by (auto simp add: pflat_len_def) +using assms by (simp add: pflat_len_def) + section {* Orderings *} @@ -175,15 +168,10 @@ lemma PosOrd_shorterE: assumes "v1 :\val v2" shows "length (flat v2) \ length (flat v1)" -using assms -apply(auto simp add: pflat_len_simps PosOrd_ex_def PosOrd_def) -apply(case_tac p) -apply(simp add: pflat_len_simps intlen_length) -apply(simp) -apply(drule_tac x="[]" in bspec) -apply(simp add: Pos_empty) -apply(simp add: pflat_len_simps le_less intlen_length_eq) -done +using assms unfolding PosOrd_ex_def PosOrd_def +apply(auto simp add: pflat_len_def intlen_int split: if_splits) +apply (metis Pos_empty Un_iff at.simps(1) eq_iff lex_simps(1) nat_less_le) +by (metis Pos_empty UnI2 at.simps(1) lex_simps(2) lex_trichotomous linear) lemma PosOrd_shorterI: assumes "length (flat v2) < length (flat v1)" @@ -206,8 +194,7 @@ unfolding PosOrd_ex_def apply(rule_tac x="[0]" in exI) using assms -apply(auto simp add: PosOrd_def pflat_len_simps) -apply(smt intlen_bigger) +apply(auto simp add: PosOrd_def pflat_len_simps intlen_int) done lemma PosOrd_Left_eq: @@ -547,34 +534,35 @@ by (metis PosOrd_SeqI1 PosOrd_shorterI WW1 antisym_conv3 append_eq_append_conv assms(2)) + section {* The Posix Value is smaller than any other Value *} lemma Posix_PosOrd: - assumes "s \ r \ v1" "v2 \ CPT r s" + assumes "s \ r \ v1" "v2 \ CV r s" shows "v1 :\val v2" using assms proof (induct arbitrary: v2 rule: Posix.induct) case (Posix_ONE v) - have "v \ CPT ONE []" by fact + have "v \ CV ONE []" by fact then have "v = Void" - by (simp add: CPT_simps) + by (simp add: CV_simps) then show "Void :\val v" by (simp add: PosOrd_ex_eq_def) next case (Posix_CHAR c v) - have "v \ CPT (CHAR c) [c]" by fact + have "v \ CV (CHAR c) [c]" by fact then have "v = Char c" - by (simp add: CPT_simps) + by (simp add: CV_simps) then show "Char c :\val v" by (simp add: PosOrd_ex_eq_def) next case (Posix_ALT1 s r1 v r2 v2) have as1: "s \ r1 \ v" by fact - have IH: "\v2. v2 \ CPT r1 s \ v :\val v2" by fact - have "v2 \ CPT (ALT r1 r2) s" by fact + have IH: "\v2. v2 \ CV r1 s \ v :\val v2" by fact + have "v2 \ CV (ALT r1 r2) s" by fact then have "\ v2 : ALT r1 r2" "flat v2 = s" - by(auto simp add: CPT_def prefix_list_def) + by(auto simp add: CV_def prefix_list_def) then consider (Left) v3 where "v2 = Left v3" "\ v3 : r1" "flat v3 = s" | (Right) v3 where "v2 = Right v3" "\ v3 : r2" "flat v3 = s" @@ -582,8 +570,8 @@ then show "Left v :\val v2" proof(cases) case (Left v3) - have "v3 \ CPT r1 s" using Left(2,3) - by (auto simp add: CPT_def prefix_list_def) + have "v3 \ CV r1 s" using Left(2,3) + by (auto simp add: CV_def prefix_list_def) with IH have "v :\val v3" by simp moreover have "flat v3 = flat v" using as1 Left(3) @@ -603,10 +591,10 @@ case (Posix_ALT2 s r2 v r1 v2) have as1: "s \ r2 \ v" by fact have as2: "s \ L r1" by fact - have IH: "\v2. v2 \ CPT r2 s \ v :\val v2" by fact - have "v2 \ CPT (ALT r1 r2) s" by fact + have IH: "\v2. v2 \ CV r2 s \ v :\val v2" by fact + have "v2 \ CV (ALT r1 r2) s" by fact then have "\ v2 : ALT r1 r2" "flat v2 = s" - by(auto simp add: CPT_def prefix_list_def) + by(auto simp add: CV_def prefix_list_def) then consider (Left) v3 where "v2 = Left v3" "\ v3 : r1" "flat v3 = s" | (Right) v3 where "v2 = Right v3" "\ v3 : r2" "flat v3 = s" @@ -614,8 +602,8 @@ then show "Right v :\val v2" proof (cases) case (Right v3) - have "v3 \ CPT r2 s" using Right(2,3) - by (auto simp add: CPT_def prefix_list_def) + have "v3 \ CV r2 s" using Right(2,3) + by (auto simp add: CV_def prefix_list_def) with IH have "v :\val v3" by simp moreover have "flat v3 = flat v" using as1 Right(3) @@ -625,10 +613,10 @@ then show "Right v :\val v2" unfolding Right . next case (Left v3) - have "v3 \ CPT r1 s" using Left(2,3) as2 - by (auto simp add: CPT_def prefix_list_def) + have "v3 \ CV r1 s" using Left(2,3) as2 + by (auto simp add: CV_def prefix_list_def) then have "flat v3 = flat v \ \ v3 : r1" using as1 Left(3) - by (simp add: Posix1(2) CPT_def) + by (simp add: Posix1(2) CV_def) then have "False" using as1 as2 Left by (auto simp add: Posix1(2) L_flat_Prf1 Prf_CPrf) then show "Right v :\val v2" by simp @@ -637,22 +625,22 @@ case (Posix_SEQ s1 r1 v1 s2 r2 v2 v3) have "s1 \ r1 \ v1" "s2 \ r2 \ v2" by fact+ then have as1: "s1 = flat v1" "s2 = flat v2" by (simp_all add: Posix1(2)) - have IH1: "\v3. v3 \ CPT r1 s1 \ v1 :\val v3" by fact - have IH2: "\v3. v3 \ CPT r2 s2 \ v2 :\val v3" by fact + have IH1: "\v3. v3 \ CV r1 s1 \ v1 :\val v3" by fact + have IH2: "\v3. v3 \ CV r2 s2 \ v2 :\val v3" by fact have cond: "\ (\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 - have "v3 \ CPT (SEQ r1 r2) (s1 @ s2)" by fact + have "v3 \ CV (SEQ r1 r2) (s1 @ s2)" by fact then obtain v3a v3b where eqs: "v3 = Seq v3a v3b" "\ v3a : r1" "\ v3b : r2" "flat v3a @ flat v3b = s1 @ s2" - by (force simp add: prefix_list_def CPT_def elim: CPrf.cases) + by (force simp add: prefix_list_def CV_def elim: CPrf.cases) with cond have "flat v3a \pre s1" unfolding prefix_list_def by (smt L_flat_Prf1 Prf_CPrf append_eq_append_conv2 append_self_conv) then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat v3b = s2)" using eqs by (simp add: sprefix_list_def append_eq_conv_conj) then have q2: "v1 :\val v3a \ (flat v3a = s1 \ flat v3b = s2)" using PosOrd_spreI as1(1) eqs by blast - then have "v1 :\val v3a \ (v3a \ CPT r1 s1 \ v3b \ CPT r2 s2)" using eqs(2,3) - by (auto simp add: CPT_def) + then have "v1 :\val v3a \ (v3a \ CV r1 s1 \ v3b \ CV r2 s2)" using eqs(2,3) + by (auto simp add: CV_def) then have "v1 :\val v3a \ (v1 :\val v3a \ v2 :\val v3b)" using IH1 IH2 by blast then have "Seq v1 v2 :\val Seq v3a v3b" using eqs q2 as1 unfolding PosOrd_ex_eq_def by (auto simp add: PosOrd_SeqI1 PosOrd_SeqI2) @@ -661,17 +649,17 @@ case (Posix_STAR1 s1 r v s2 vs v3) have "s1 \ r \ v" "s2 \ STAR r \ Stars vs" by fact+ then have as1: "s1 = flat v" "s2 = flat (Stars vs)" by (auto dest: Posix1(2)) - have IH1: "\v3. v3 \ CPT r s1 \ v :\val v3" by fact - have IH2: "\v3. v3 \ CPT (STAR r) s2 \ Stars vs :\val v3" by fact + have IH1: "\v3. v3 \ CV r s1 \ v :\val v3" by fact + have IH2: "\v3. v3 \ CV (STAR r) s2 \ Stars vs :\val v3" by fact have cond: "\ (\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 have cond2: "flat v \ []" by fact - have "v3 \ CPT (STAR r) (s1 @ s2)" by fact + have "v3 \ CV (STAR r) (s1 @ s2)" by fact then consider (NonEmpty) v3a vs3 where "v3 = Stars (v3a # vs3)" "\ v3a : r" "\ Stars vs3 : STAR r" "flat (Stars (v3a # vs3)) = s1 @ s2" | (Empty) "v3 = Stars []" - unfolding CPT_def + unfolding CV_def apply(auto) apply(erule CPrf.cases) apply(simp_all) @@ -690,8 +678,8 @@ by (simp add: sprefix_list_def append_eq_conv_conj) then have q2: "v :\val v3a \ (flat v3a = s1 \ flat (Stars vs3) = s2)" using PosOrd_spreI as1(1) NonEmpty(4) by blast - then have "v :\val v3a \ (v3a \ CPT r s1 \ Stars vs3 \ CPT (STAR r) s2)" - using NonEmpty(2,3) by (auto simp add: CPT_def) + then have "v :\val v3a \ (v3a \ CV r s1 \ Stars vs3 \ CV (STAR r) s2)" + using NonEmpty(2,3) by (auto simp add: CV_def) then have "v :\val v3a \ (v :\val v3a \ Stars vs :\val Stars vs3)" using IH1 IH2 by blast then have "v :\val v3a \ (v = v3a \ Stars vs :\val Stars vs3)" unfolding PosOrd_ex_eq_def by auto @@ -708,58 +696,26 @@ qed next case (Posix_STAR2 r v2) - have "v2 \ CPT (STAR r) []" by fact + have "v2 \ CV (STAR r) []" by fact then have "v2 = Stars []" - unfolding CPT_def by (auto elim: CPrf.cases) + unfolding CV_def by (auto elim: CPrf.cases) then show "Stars [] :\val v2" by (simp add: PosOrd_ex_eq_def) qed -lemma Posix_PosOrd_stronger: - assumes "s \ r \ v1" "v2 \ CPTpre r s" - shows "v1 :\val v2" -proof - - from assms(2) have "v2 \ CPT r s \ flat v2 \spre s" - unfolding CPTpre_def CPT_def sprefix_list_def prefix_list_def by auto - moreover - { assume "v2 \ CPT r s" - with assms(1) - have "v1 :\val v2" by (rule Posix_PosOrd) - } - moreover - { assume "flat v2 \spre s" - then have "flat v2 \spre flat v1" using assms(1) - using Posix1(2) by blast - then have "v1 :\val v2" - by (simp add: PosOrd_ex_eq_def PosOrd_spreI) - } - ultimately show "v1 :\val v2" by blast -qed lemma Posix_PosOrd_reverse: assumes "s \ r \ v1" - shows "\(\v2 \ CPTpre r s. v2 :\val v1)" + shows "\(\v2 \ CV r s. v2 :\val v1)" using assms -by (metis Posix_PosOrd_stronger less_irrefl PosOrd_def +by (metis Posix_PosOrd less_irrefl PosOrd_def PosOrd_ex_eq_def PosOrd_ex_def PosOrd_trans) -lemma test2: - assumes "\v \ set vs. flat v \ r \ v \ flat v \ []" - shows "(Stars vs) \ CPT (STAR r) (flat (Stars vs))" -using assms -apply(induct vs) -apply(auto simp add: CPT_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) - lemma PosOrd_Posix_Stars: assumes "\v \ set vs. flat v \ r \ v \ flat v \ []" - and "\(\vs2 \ PT (STAR r) (flat (Stars vs)). vs2 :\val (Stars vs))" + and "\(\vs2 \ LV (STAR r) (flat (Stars vs)). vs2 :\val (Stars vs))" shows "(flat (Stars vs)) \ (STAR r) \ Stars vs" using assms proof(induct vs) @@ -769,24 +725,24 @@ next case (Cons v vs) have IH: "\\v\set vs. flat v \ r \ v \ flat v \ []; - \ (\vs2\PT (STAR r) (flat (Stars vs)). vs2 :\val Stars vs)\ + \ (\vs2\LV (STAR r) (flat (Stars vs)). vs2 :\val Stars vs)\ \ flat (Stars vs) \ STAR r \ Stars vs" by fact have as2: "\v\set (v # vs). flat v \ r \ v \ flat v \ []" by fact - have as3: "\ (\vs2\PT (STAR r) (flat (Stars (v # vs))). vs2 :\val Stars (v # vs))" by fact + have as3: "\ (\vs2\LV (STAR r) (flat (Stars (v # vs))). vs2 :\val Stars (v # vs))" by fact have "flat v \ r \ v" using as2 by simp moreover have "flat (Stars vs) \ STAR r \ Stars vs" proof (rule IH) show "\v\set vs. flat v \ r \ v \ flat v \ []" using as2 by simp next - show "\ (\vs2\PT (STAR r) (flat (Stars vs)). vs2 :\val Stars vs)" using as3 + show "\ (\vs2\LV (STAR r) (flat (Stars vs)). vs2 :\val Stars vs)" using as3 apply(auto) - apply(subst (asm) (2) PT_def) + apply(subst (asm) (2) LV_def) apply(auto) apply(erule Prf.cases) apply(simp_all) apply(drule_tac x="Stars (v # vs)" in bspec) - apply(simp add: PT_def CPT_def) + apply(simp add: LV_def CV_def) using Posix_Prf Prf.intros(6) calculation apply(rule_tac Prf.intros) apply(simp add:) @@ -810,7 +766,7 @@ apply(simp_all) apply(clarify) apply(drule_tac x="Stars (va#vs)" in bspec) - apply(auto simp add: PT_def)[1] + apply(auto simp add: LV_def)[1] apply(rule Prf.intros) apply(simp) by (simp add: PosOrd_StarsI PosOrd_shorterI) @@ -823,69 +779,69 @@ section {* The Smallest Value is indeed the Posix Value *} text {* - The next lemma seems to require PT instead of CPT in the Star-case. + The next lemma seems to require LV instead of CV in the Star-case. *} lemma PosOrd_Posix: - assumes "v1 \ CPT r s" "\v\<^sub>2 \ PT r s. \ v\<^sub>2 :\val v1" + assumes "v1 \ CV r s" "\v\<^sub>2 \ LV r s. \ v\<^sub>2 :\val v1" shows "s \ r \ v1" using assms proof(induct r arbitrary: s v1) case (ZERO s v1) - have "v1 \ CPT ZERO s" by fact - then show "s \ ZERO \ v1" unfolding CPT_def + have "v1 \ CV ZERO s" by fact + then show "s \ ZERO \ v1" unfolding CV_def by (auto elim: CPrf.cases) next case (ONE s v1) - have "v1 \ CPT ONE s" by fact - then show "s \ ONE \ v1" unfolding CPT_def + have "v1 \ CV ONE s" by fact + then show "s \ ONE \ v1" unfolding CV_def by(auto elim!: CPrf.cases intro: Posix.intros) next case (CHAR c s v1) - have "v1 \ CPT (CHAR c) s" by fact - then show "s \ CHAR c \ v1" unfolding CPT_def + have "v1 \ CV (CHAR c) s" by fact + then show "s \ CHAR c \ v1" unfolding CV_def by (auto elim!: CPrf.cases intro: Posix.intros) next case (ALT r1 r2 s v1) - have IH1: "\s v1. \v1 \ CPT r1 s; \v2 \ PT r1 s. \ v2 :\val v1\ \ s \ r1 \ v1" by fact - have IH2: "\s v1. \v1 \ CPT r2 s; \v2 \ PT r2 s. \ v2 :\val v1\ \ s \ r2 \ v1" by fact - have as1: "\v2\PT (ALT r1 r2) s. \ v2 :\val v1" by fact - have as2: "v1 \ CPT (ALT r1 r2) s" by fact + have IH1: "\s v1. \v1 \ CV r1 s; \v2 \ LV r1 s. \ v2 :\val v1\ \ s \ r1 \ v1" by fact + have IH2: "\s v1. \v1 \ CV r2 s; \v2 \ LV r2 s. \ v2 :\val v1\ \ s \ r2 \ v1" by fact + have as1: "\v2\LV (ALT r1 r2) s. \ v2 :\val v1" by fact + have as2: "v1 \ CV (ALT r1 r2) s" by fact then consider (Left) v1' where "v1 = Left v1'" "s = flat v1'" - "v1' \ CPT r1 s" + "v1' \ CV r1 s" | (Right) v1' where "v1 = Right v1'" "s = flat v1'" - "v1' \ CPT r2 s" - unfolding CPT_def by (auto elim: CPrf.cases) + "v1' \ CV r2 s" + unfolding CV_def by (auto elim: CPrf.cases) then show "s \ ALT r1 r2 \ v1" proof (cases) case (Left v1') - have "v1' \ CPT r1 s" using as2 - unfolding CPT_def Left by (auto elim: CPrf.cases) + have "v1' \ CV r1 s" using as2 + unfolding CV_def Left by (auto elim: CPrf.cases) moreover - have "\v2 \ PT r1 s. \ v2 :\val v1'" using as1 - unfolding PT_def Left using Prf.intros(2) PosOrd_Left_eq by force + have "\v2 \ LV r1 s. \ v2 :\val v1'" using as1 + unfolding LV_def Left using Prf.intros(2) PosOrd_Left_eq by force ultimately have "s \ r1 \ v1'" using IH1 by simp then have "s \ ALT r1 r2 \ Left v1'" by (rule Posix.intros) then show "s \ ALT r1 r2 \ v1" using Left by simp next case (Right v1') - have "v1' \ CPT r2 s" using as2 - unfolding CPT_def Right by (auto elim: CPrf.cases) + have "v1' \ CV r2 s" using as2 + unfolding CV_def Right by (auto elim: CPrf.cases) moreover - have "\v2 \ PT r2 s. \ v2 :\val v1'" using as1 - unfolding PT_def Right using Prf.intros(3) PosOrd_RightI by force + have "\v2 \ LV r2 s. \ v2 :\val v1'" using as1 + unfolding LV_def Right using Prf.intros(3) PosOrd_RightI by force ultimately have "s \ r2 \ v1'" using IH2 by simp moreover { assume "s \ L r1" - then obtain v' where "v' \ PT r1 s" - unfolding PT_def using L_flat_Prf2 by blast - then have "Left v' \ PT (ALT r1 r2) s" - unfolding PT_def by (auto intro: Prf.intros) + then obtain v' where "v' \ LV r1 s" + unfolding LV_def using L_flat_Prf2 by blast + then have "Left v' \ LV (ALT r1 r2) s" + unfolding LV_def by (auto intro: Prf.intros) with as1 have "\ (Left v' :\val Right v1') \ (flat v' = s)" - unfolding PT_def Right by (auto) + unfolding LV_def Right by (auto) then have False using PosOrd_Left_Right Right by blast } then have "s \ L r1" by rule @@ -894,36 +850,36 @@ qed next case (SEQ r1 r2 s v1) - have IH1: "\s v1. \v1 \ CPT r1 s; \v2 \ PT r1 s. \ v2 :\val v1\ \ s \ r1 \ v1" by fact - have IH2: "\s v1. \v1 \ CPT r2 s; \v2 \ PT r2 s. \ v2 :\val v1\ \ s \ r2 \ v1" by fact - have as1: "\v2\PT (SEQ r1 r2) s. \ v2 :\val v1" by fact - have as2: "v1 \ CPT (SEQ r1 r2) s" by fact + have IH1: "\s v1. \v1 \ CV r1 s; \v2 \ LV r1 s. \ v2 :\val v1\ \ s \ r1 \ v1" by fact + have IH2: "\s v1. \v1 \ CV r2 s; \v2 \ LV r2 s. \ v2 :\val v1\ \ s \ r2 \ v1" by fact + have as1: "\v2\LV (SEQ r1 r2) s. \ v2 :\val v1" by fact + have as2: "v1 \ CV (SEQ r1 r2) s" by fact then obtain v1a v1b where eqs: "v1 = Seq v1a v1b" "s = flat v1a @ flat v1b" - "v1a \ CPT r1 (flat v1a)" "v1b \ CPT r2 (flat v1b)" - unfolding CPT_def by(auto elim: CPrf.cases) - have "\v2 \ PT r1 (flat v1a). \ v2 :\val v1a" + "v1a \ CV r1 (flat v1a)" "v1b \ CV r2 (flat v1b)" + unfolding CV_def by(auto elim: CPrf.cases) + have "\v2 \ LV r1 (flat v1a). \ v2 :\val v1a" proof fix v2 - assume "v2 \ PT r1 (flat v1a)" - with eqs(2,4) have "Seq v2 v1b \ PT (SEQ r1 r2) s" - by (simp add: CPT_def PT_def Prf.intros(1) Prf_CPrf) + assume "v2 \ LV r1 (flat v1a)" + with eqs(2,4) have "Seq v2 v1b \ LV (SEQ r1 r2) s" + by (simp add: CV_def LV_def Prf.intros(1) Prf_CPrf) with as1 have "\ Seq v2 v1b :\val Seq v1a v1b \ flat (Seq v2 v1b) = flat (Seq v1a v1b)" - using eqs by (simp add: PT_def) + using eqs by (simp add: LV_def) then show "\ v2 :\val v1a" using PosOrd_SeqI1 by blast qed then have "flat v1a \ r1 \ v1a" using IH1 eqs by simp moreover - have "\v2 \ PT r2 (flat v1b). \ v2 :\val v1b" + have "\v2 \ LV r2 (flat v1b). \ v2 :\val v1b" proof fix v2 - assume "v2 \ PT r2 (flat v1b)" - with eqs(2,3,4) have "Seq v1a v2 \ PT (SEQ r1 r2) s" - by (simp add: CPT_def PT_def Prf.intros Prf_CPrf) + assume "v2 \ LV r2 (flat v1b)" + with eqs(2,3,4) have "Seq v1a v2 \ LV (SEQ r1 r2) s" + by (simp add: CV_def LV_def Prf.intros Prf_CPrf) with as1 have "\ Seq v1a v2 :\val Seq v1a v1b \ flat v2 = flat v1b" - using eqs by (simp add: PT_def) + using eqs by (simp add: LV_def) then show "\ v2 :\val v1b" using PosOrd_SeqI2 by auto qed @@ -935,8 +891,8 @@ then obtain s3 s4 where q1: "s3 \ [] \ s3 @ s4 = flat v1b \ flat v1a @ s3 \ L r1 \ s4 \ L r2" by blast then obtain vA vB where q2: "flat vA = flat v1a @ s3" "\ vA : r1" "flat vB = s4" "\ vB : r2" using L_flat_Prf2 by blast - then have "Seq vA vB \ PT (SEQ r1 r2) s" unfolding eqs using q1 - by (auto simp add: PT_def intro: Prf.intros) + then have "Seq vA vB \ LV (SEQ r1 r2) s" unfolding eqs using q1 + by (auto simp add: LV_def intro: Prf.intros) with as1 have "\ Seq vA vB :\val Seq v1a v1b" unfolding eqs by auto then have "\ vA :\val v1a \ length (flat vA) > length (flat v1a)" using q1 q2 PosOrd_SeqI1 by auto then show "False" @@ -947,57 +903,57 @@ by (rule Posix.intros) next case (STAR r s v1) - have IH: "\s v1. \v1 \ CPT r s; \v2\PT r s. \ v2 :\val v1\ \ s \ r \ v1" by fact - have as1: "\v2\PT (STAR r) s. \ v2 :\val v1" by fact - have as2: "v1 \ CPT (STAR r) s" by fact + have IH: "\s v1. \v1 \ CV r s; \v2\LV r s. \ v2 :\val v1\ \ s \ r \ v1" by fact + have as1: "\v2\LV (STAR r) s. \ v2 :\val v1" by fact + have as2: "v1 \ CV (STAR r) s" by fact then obtain vs where eqs: "v1 = Stars vs" "s = flat (Stars vs)" - "\v \ set vs. v \ CPT r (flat v)" - unfolding CPT_def by (auto elim: CPrf.cases dest!: CPrf_stars) + "\v \ set vs. v \ CV r (flat v)" + unfolding CV_def by (auto elim: CPrf.cases) have "\v\set vs. flat v \ r \ v \ flat v \ []" proof fix v assume a: "v \ set vs" then obtain pre post where e: "vs = pre @ [v] @ post" by (metis append_Cons append_Nil in_set_conv_decomp_first) - then have q: "\v2\PT (STAR r) s. \ v2 :\val Stars (pre @ [v] @ post)" + then have q: "\v2\LV (STAR r) s. \ v2 :\val Stars (pre @ [v] @ post)" using as1 unfolding eqs by simp - have "\v2\PT r (flat v). \ v2 :\val v" unfolding eqs + have "\v2\LV r (flat v). \ v2 :\val v" unfolding eqs proof (rule ballI, rule notI) fix v2 assume w: "v2 :\val v" - assume "v2 \ PT r (flat v)" - then have "Stars (pre @ [v2] @ post) \ PT (STAR r) s" + assume "v2 \ LV r (flat v)" + then have "Stars (pre @ [v2] @ post) \ LV (STAR r) s" using as2 unfolding e eqs - apply(auto simp add: CPT_def PT_def intro!: Prf.intros)[1] - using CPrf_Stars_appendE CPrf_stars Prf_CPrf apply blast - by (meson CPrf_Stars_appendE CPrf_stars Prf_CPrf list.set_intros(2)) + apply(auto simp add: CV_def LV_def intro!: Prf.intros)[1] + using CPrf_Stars_appendE Prf_CPrf Prf_elims(6) list.set_intros apply blast + by (metis CPrf_Stars_appendE Prf_CPrf Prf_elims(6) list.set_intros(2) val.inject(5)) then have "\ Stars (pre @ [v2] @ post) :\val Stars (pre @ [v] @ post)" using q by simp with w show "False" - using PT_def \v2 \ PT r (flat v)\ append_Cons flat.simps(7) mem_Collect_eq + using LV_def \v2 \ LV r (flat v)\ append_Cons flat.simps(7) mem_Collect_eq PosOrd_StarsI PosOrd_Stars_appendI by auto qed with IH - show "flat v \ r \ v \ flat v \ []" using a as2 unfolding eqs - using eqs(3) by (smt CPT_def CPrf_stars mem_Collect_eq) + show "flat v \ r \ v \ flat v \ []" using a as2 unfolding eqs CV_def + by (auto elim: CPrf.cases) qed moreover - have "\ (\vs2\PT (STAR r) (flat (Stars vs)). vs2 :\val Stars vs)" + have "\ (\vs2\LV (STAR r) (flat (Stars vs)). vs2 :\val Stars vs)" proof - assume "\vs2 \ PT (STAR r) (flat (Stars vs)). vs2 :\val Stars vs" + assume "\vs2 \ LV (STAR r) (flat (Stars vs)). vs2 :\val Stars vs" then obtain vs2 where "\ Stars vs2 : STAR r" "flat (Stars vs2) = flat (Stars vs)" "Stars vs2 :\val Stars vs" - unfolding PT_def - apply(auto elim: Prf.cases) + unfolding LV_def + apply(auto) apply(erule Prf.cases) apply(auto intro: Prf.intros) done then show "False" using as1 unfolding eqs apply - apply(drule_tac x="Stars vs2" in bspec) - apply(auto simp add: PT_def) + apply(auto simp add: LV_def) done qed ultimately have "flat (Stars vs) \ STAR r \ Stars vs"