diff -r 030939b7d475 -r 532bb9df225d thys/Re.thy --- a/thys/Re.thy Wed Jan 06 12:24:29 2016 +0000 +++ b/thys/Re.thy Thu Jan 14 17:19:40 2016 +0000 @@ -324,6 +324,14 @@ definition rest :: "val \ string \ string" where "rest v s \ drop (length (flat v)) s" +lemma rest_flat: + assumes "flat v1 \ s" + shows "flat v1 @ rest v1 s = s" +using assms +apply(auto simp add: rest_def prefix_def) +done + + lemma rest_Suffixes: "rest v s \ Suffixes s" unfolding rest_def @@ -836,6 +844,51 @@ thm PMatch.intros +inductive ValOrds :: "string \ val \ rexp \ val \ bool" ("_ \ _ \_ _" [100, 100, 100] 100) +where + "\s2 \ v2 \r2 v2'; flat v1 = s1\ \ (s1 @ s2) \ (Seq v1 v2) \(SEQ r1 r2) (Seq v1 v2')" +| "\s1 \ v1 \r1 v1'; v1 \ v1'; flat v2 = s2; (flat v1' @ flat v2') \ (s1 @ s2)\ + \ s1 @ s2 \ (Seq v1 v2) \(SEQ r1 r2) (Seq v1' v2')" +| "\(flat v2) \ (flat v1); flat v1 = s\ \ s \ (Left v1) \(ALT r1 r2) (Right v2)" +| "\(flat v1) \ (flat v2); flat v2 = s\ \ s \ (Right v2) \(ALT r1 r2) (Left v1)" +| "s \ v2 \r2 v2' \ s \ (Right v2) \(ALT r1 r2) (Right v2')" +| "s \ v1 \r1 v1' \ s \ (Left v1) \(ALT r1 r2) (Left v1')" +| "[] \ Void \EMPTY Void" +| "[c] \ (Char c) \(CHAR c) (Char c)" + +lemma valord_prefix: + assumes "s \ v1 \r v2" + shows "flat v1 \ s" and "flat v2 \ s" +using assms +apply(induct) +apply(auto) +apply(simp add: prefix_def rest_def) +apply(auto)[1] +apply(simp add: prefix_def rest_def) +apply(auto)[1] +apply(auto simp add: prefix_def rest_def) +done + +lemma valord_prefix2: + assumes "s \ v1 \r v2" + shows "flat v2 \ flat v1" +using assms +apply(induct) +apply(auto) +apply(simp add: prefix_def rest_def) +apply(simp add: prefix_def rest_def) +apply(auto)[1] +defer +apply(simp add: prefix_def rest_def) +apply(auto)[1] +apply (metis append_eq_append_conv_if append_eq_conv_conj) +apply(simp add: prefix_def rest_def) +apply(auto)[1] +apply (metis append_eq_append_conv_if append_take_drop_id le_eq_less_or_eq) +apply(simp add: prefix_def rest_def) +apply(simp add: prefix_def rest_def) + + inductive ValOrd :: "val \ rexp \ val \ bool" ("_ \_ _" [100, 100, 100] 100) where "v2 \r2 v2' \ (Seq v1 v2) \(SEQ r1 r2) (Seq v1 v2')" @@ -1023,10 +1076,69 @@ thm PMatch.intros[no_vars] lemma POSIX_PMatch: - assumes "s \ r \ v" "\ v' : r" - shows "length (flat v') \ length (flat v)" + assumes "s \ r \ v" "v' \ Values r s" + shows "v \r v' \ length (flat v') < length (flat v)" using assms -apply(induct arbitrary: s v v' rule: rexp.induct) +apply(induct r arbitrary: s v v' rule: rexp.induct) +apply(simp_all add: Values_recs) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(7)) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(simp add: prefix_def) +apply (metis ValOrd.intros(8)) +apply(auto)[1] +defer +apply(auto)[1] +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis ValOrd.intros(6)) +apply (metis (no_types, lifting) PMatch1(2) Prf_flat_L Values_def length_sprefix mem_Collect_eq sprefix_def) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply (metis (no_types, lifting) PMatch1(2) ValOrd.intros(3) Values_def length_sprefix mem_Collect_eq order_refl sprefix_def) +apply (metis ValOrd.intros(5)) +apply(erule PMatch.cases) +apply(simp_all)[5] +apply(auto) +apply(case_tac "v1a = v1") +apply(simp) +apply(rule ValOrd.intros(1)) +apply (metis PMatch1(2) append_Nil comm_monoid_diff_class.diff_cancel drop_0 drop_all drop_append order_refl rest_def) +apply(rule ValOrd.intros(2)) +apply(auto) +apply(drule_tac x="s1" in meta_spec) +apply(drule_tac x="v1a" in meta_spec) +apply(drule_tac x="v1" in meta_spec) +apply(auto) +apply(drule meta_mp) +defer +apply(auto)[1] +apply(frule PMatch1) +apply(drule PMatch1(2)) +apply(frule PMatch1) +apply(drule PMatch1(2)) +apply(auto)[1] +apply(simp add: Values_def) +apply(simp add: prefix_def) +apply(auto)[1] +apply(simp add: append_eq_append_conv2) +apply(auto)[1] +apply(rotate_tac 10) +apply(drule sym) +apply(simp) +apply(simp add: rest_def) +apply(case_tac "s3a = []") +apply(auto)[1] + + +apply (metis ValOrd.intros(6)) +apply (metis (no_types, lifting) PMatch1(2) ValOrd.intros(3) Values_def length_sprefix mem_Collect_eq order_refl sprefix_def) +apply(auto)[1] +apply (metis (no_types, lifting) PMatch1(2) Prf_flat_L Values_def length_sprefix mem_Collect_eq sprefix_def) +apply (metis ValOrd.intros(5)) +apply(auto)[1] apply(erule Prf.cases) apply(simp_all)[5] apply(erule Prf.cases)