diff -r f95a405c3180 -r 39cef7b9212a thys/Re1.thy --- a/thys/Re1.thy Fri Apr 10 22:38:36 2015 +0100 +++ b/thys/Re1.thy Sat Apr 25 10:58:48 2015 +0100 @@ -140,6 +140,18 @@ value "flats(Seq(Char c)(Char b))" +inductive StrOrd :: "string list \ string list \ bool" ("_ \ _" [100, 100] 100) +where + "[] \ []" +| "xs \ ys \ (x#xs) \ (x#ys)" +| "length x \ length y \ (x#xs) \ (y#xs)" + +lemma StrOrd_append: + "xs \ ys \ (zs @ xs) \ (zs @ ys)" +apply(induct zs) +apply(auto intro: StrOrd.intros) +done + section {* Relation between values and regular expressions *} inductive Prf :: "val \ rexp \ bool" ("\ _ : _" [100, 100] 100) @@ -233,35 +245,35 @@ | "Void \EMPTY Void" | "(Char c) \(CHAR c) (Char c)" -inductive ValOrd2 :: "val \ rexp \ nat \ val \ bool" ("_ \_,_ _" [100, 100, 100, 100] 100) +inductive ValOrd2 :: "val \ val \ bool" ("_ 2\ _" [100, 100] 100) where - "v2 \r2,n v2' \ (Seq v1 v2) \(SEQ r1 r2),(length (flat v1) + n) (Seq v1 v2')" -| "\v1 \r1,n v1'; v1 \ v1'\ \ (Seq v1 v2) \(SEQ r1 r2),(n + length (flat v2)) (Seq v1' v2')" -| "length (flat v1) \ length (flat v2) \ (Left v1) \(ALT r1 r2),(length (flat v1)) (Right v2)" -| "length (flat v2) > length (flat v1) \ (Right v2) \(ALT r1 r2),(length (flat v2)) (Left v1)" -| "v2 \r2,n v2' \ (Right v2) \(ALT r1 r2),n (Right v2')" -| "v1 \r1,n v1' \ (Left v1) \(ALT r1 r2),n (Left v1')" -| "Void \EMPTY,0 Void" -| "(Char c) \(CHAR c),1 (Char c)" + "v2 2\ v2' \ (Seq v1 v2) 2\ (Seq v1 v2')" +| "\v1 2\ v1'; v1 \ v1'\ \ (Seq v1 v2) 2\ (Seq v1' v2')" +| "length (flat v1) \ length (flat v2) \ (Left v1) 2\ (Right v2)" +| "length (flat v2) > length (flat v1) \ (Right v2) 2\ (Left v1)" +| "v2 2\ v2' \ (Right v2) 2\ (Right v2')" +| "v1 2\ v1' \ (Left v1) 2\ (Left v1')" +| "Void 2\ Void" +| "(Char c) 2\ (Char c)" -lemma - assumes "v1 \r,n v2" - shows "length(flat v1) = n" -using assms -apply(induct) -apply(auto) +lemma Ord1: + "v1 \r v2 \ v1 2\ v2" +apply(induct rule: ValOrd.induct) +apply(auto intro: ValOrd2.intros) done -lemma - assumes "v1 \r,n v2" - shows "length(flat v2) <= n" -using assms -apply(induct) -apply(auto) -oops +lemma Ord2: + "v1 2\ v2 \ \r. v1 \r v2" +apply(induct v1 v2 rule: ValOrd2.induct) +apply(auto intro: ValOrd.intros) +done +lemma Ord3: + "\v1 2\ v2; \ v1 : r\ \ v1 \r v2" +apply(induct v1 v2 arbitrary: r rule: ValOrd2.induct) +apply(auto intro: ValOrd.intros elim: Prf.cases) +done -section {* The ordering is reflexive *} lemma ValOrd_refl: assumes "\ v : r" @@ -277,6 +289,19 @@ where "POSIX v r \ (\ v : r \ (\v'. (\ v' : r \ flat v = flat v') \ v \r v'))" +definition POSIX2 :: "val \ rexp \ bool" +where + "POSIX2 v r \ (\ v : r \ (\v'. (\ v' : r \ flat v = flat v') \ v 2\ v'))" + +lemma "POSIX v r = POSIX2 v r" +unfolding POSIX_def POSIX2_def +apply(auto) +apply(rule Ord1) +apply(auto) +apply(rule Ord3) +apply(auto) +done + (* an alternative definition: might cause problems with theorem mkeps_POSIX @@ -885,6 +910,37 @@ apply (metis list.distinct(1) mkeps_flat v4) by (metis h) +lemma POSIX_der: + assumes "POSIX2 v (der c r)" "\ v : der c r" + shows "POSIX2 (injval r c v) r" +using assms +unfolding POSIX2_def +apply(auto) +thm v3 +apply (erule v3) +thm v4 +apply(subst (asm) v4) +apply(assumption) +apply(drule_tac x="projval r c v'" in spec) +apply(drule mp) +apply(rule conjI) +thm v3_proj +apply(rule v3_proj) +apply(simp) +apply(rule_tac x="flat v" in exI) +apply(simp) +thm t +apply(rule_tac c="c" in t) +apply(simp) +thm v4_proj +apply(subst v4_proj) +apply(simp) +apply(rule_tac x="flat v" in exI) +apply(simp) +apply(simp) +oops + + lemma Prf_inj_test: assumes "v1 \(der c r) v2" "\ v1 : der c r" "\ v2 : der c r" "\ vs : rs" "flat v1 = flat v2" shows "Seqs (injval r c v1) vs \(SEQS r rs) Seqs (injval r c v2) vs"