diff -r 12772d537b71 -r 462d893ecb3d thys/Positions.thy --- a/thys/Positions.thy Fri Aug 25 15:05:20 2017 +0200 +++ b/thys/Positions.thy Fri Aug 25 23:52:49 2017 +0200 @@ -671,6 +671,7 @@ by (auto simp add: LV_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 + thm PosOrd_SeqI1 PosOrd_SeqI2 unfolding PosOrd_ex_eq_def by (auto simp add: PosOrd_SeqI1 PosOrd_SeqI2) then show "Seq v1 v2 :\val v3" unfolding eqs by blast next @@ -690,11 +691,10 @@ unfolding LV_def apply(auto) apply(erule Prf.cases) - apply(simp_all) - apply(auto)[1] + apply(auto) apply(case_tac vs) - apply(auto) - using Prf.intros(6) by blast + apply(auto intro: Prf.intros) + done then show "Stars (v # vs) :\val v3" proof (cases) case (NonEmpty v3a vs3) @@ -821,8 +821,10 @@ next case (ONE s v1) have "v1 \ LV ONE s" by fact - then show "s \ ONE \ v1" unfolding LV_def - by(auto elim!: Prf.cases intro: Posix.intros) + then have "v1 = Void" "s = []" unfolding LV_def + by(auto elim: Prf.cases) + then show "s \ ONE \ v1" + by(auto intro: Posix.intros) next case (CHAR c s v1) have "v1 \ LV (CHAR c) s" by fact @@ -832,7 +834,7 @@ case (ALT r1 r2 s v1) have IH1: "\s v1. \v1 \ LV r1 s; \v2 \ LV r1 s. \ v2 :\val v1\ \ s \ r1 \ v1" by fact have IH2: "\s v1. \v1 \ LV 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 as1: "\v2 \ LV (ALT r1 r2) s. \ v2 :\val v1" by fact have as2: "v1 \ LV (ALT r1 r2) s" by fact then consider (Left) v1' where @@ -845,18 +847,17 @@ then show "s \ ALT r1 r2 \ v1" proof (cases) case (Left v1') - have "v1' \ LV r1 s" using as2 - unfolding LV_def Left by (auto elim: Prf.cases) + have "v1' \ LV r1 s" using Left(3) . moreover have "\v2 \ LV r1 s. \ v2 :\val v1'" using as1 - unfolding LV_def Left using Prf.intros(2) PosOrd_Left_eq by force + unfolding Left(1,2) unfolding LV_def + 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' \ LV r2 s" using as2 - unfolding LV_def Right by (auto elim: Prf.cases) + have "v1' \ LV r2 s" using Right(3) . moreover have "\v2 \ LV r2 s. \ v2 :\val v1'" using as1 unfolding LV_def Right using Prf.intros(3) PosOrd_RightI by force