diff -r 66c9c06c0f0e -r c11651bbebf5 thys/Re1.thy --- a/thys/Re1.thy Mon Oct 06 14:22:13 2014 +0100 +++ b/thys/Re1.thy Mon Oct 06 15:24:41 2014 +0100 @@ -135,6 +135,10 @@ where "POSIX2 v r \ \ v : r \ (\v'. \ v' : r \ v \r v')" +definition POSIX3 :: "val \ rexp \ bool" +where + "POSIX3 v r \ \ v : r \ (\v'. (\ v' : r \ length (flat v') \ length(flat v)) \ v \r v')" + (* lemma POSIX_SEQ: @@ -192,8 +196,7 @@ using assms unfolding POSIX2_def apply(auto) - -done +oops lemma POSIX_ALT: assumes "POSIX (Left v1) (ALT r1 r2)" @@ -210,6 +213,23 @@ apply(simp_all) done +lemma POSIX2_ALT: + assumes "POSIX2 (Left v1) (ALT r1 r2)" + shows "POSIX2 v1 r1" +using assms +apply(simp add: POSIX2_def) +apply(auto) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(drule_tac x="Left v'" in spec) +apply(drule mp) +apply(rule Prf.intros) +apply(auto) +apply(erule ValOrd.cases) +apply(simp_all) +done + + lemma POSIX_ALT1a: assumes "POSIX (Right v2) (ALT r1 r2)" shows "POSIX v2 r2" @@ -225,6 +245,22 @@ apply(simp_all) done +lemma POSIX2_ALT1a: + assumes "POSIX2 (Right v2) (ALT r1 r2)" + shows "POSIX2 v2 r2" +using assms +unfolding POSIX2_def +apply(auto) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(drule_tac x="Right v'" in spec) +apply(drule mp) +apply(rule Prf.intros) +apply(auto) +apply(erule ValOrd.cases) +apply(simp_all) +done + lemma POSIX_ALT1b: assumes "POSIX (Right v2) (ALT r1 r2)" @@ -250,6 +286,23 @@ apply(rule ValOrd.intros) by simp +lemma POSIX2_ALT_I1: + assumes "POSIX2 v1 r1" + shows "POSIX2 (Left v1) (ALT r1 r2)" +using assms +unfolding POSIX2_def +apply(auto) +apply(rule Prf.intros) +apply(simp) +apply(rotate_tac 2) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(auto) +apply(rule ValOrd.intros) +apply(auto) +apply(rule ValOrd.intros) +oops + lemma POSIX_ALT_I2: assumes "POSIX v2 r2" "\v'. \ v' : r1 \ length (flat v2) > length (flat v')" shows "POSIX (Right v2) (ALT r1 r2)" @@ -325,6 +378,7 @@ apply(induct r) apply(auto)[1] apply(simp add: POSIX2_def) +oops lemma mkeps_POSIX: assumes "nullable r" @@ -368,6 +422,45 @@ by (simp add: ValOrd.intros(5)) +lemma mkeps_POSIX2: + assumes "nullable r" + shows "POSIX2 (mkeps r) r" +using assms +apply(induct r) +apply(simp) +apply(simp) +apply(simp add: POSIX2_def) +apply(rule conjI) +apply(rule Prf.intros) +apply(auto)[1] +apply(erule Prf.cases) +apply(simp_all)[5] +apply(rule ValOrd.intros) +apply(simp) +apply(simp) +apply(simp add: POSIX2_def) +apply(rule conjI) +apply(rule Prf.intros) +apply(simp add: mkeps_nullable) +apply(simp add: mkeps_nullable) +apply(auto)[1] +apply(rotate_tac 6) +apply(erule Prf.cases) +apply(simp_all)[5] +apply(rule ValOrd.intros(2)) +apply(simp) +apply(simp only: nullable.simps) +apply(erule disjE) +apply(simp) +thm POSIX2_ALT1a +apply(rule POSIX2_ALT) +apply(simp add: POSIX2_def) +apply(rule conjI) +apply(rule Prf.intros) +apply(simp add: mkeps_nullable) +oops + + section {* Derivatives *} fun @@ -561,11 +654,48 @@ apply (metis POSIX_ALT_I1) (* maybe it is too early to instantiate this existential quantifier *) (* potentially this is the wrong POSIX value *) -apply(rule_tac x = "Seq v va" in exI ) +apply(case_tac "r1 = NULL") +apply(simp add: POSIX_def) +apply(auto)[1] +apply (metis L.simps(1) L.simps(4) Prf_flat_L mkeps_flat nullable.simps(1) nullable.simps(2) nullable_correctness seq_null(2)) +apply(case_tac "r1 = EMPTY") +apply(rule_tac x = "Seq Void va" in exI ) apply(simp (no_asm) add: POSIX_def) apply(auto) apply(erule Prf.cases) apply(simp_all) +apply(auto)[1] +apply(erule Prf.cases) +apply(simp_all) +apply(rule ValOrd.intros(2)) +apply(rule ValOrd.intros) +apply(case_tac "\c. r1 = CHAR c") +apply(auto) +apply(rule_tac x = "Seq (Char c) va" in exI ) +apply(simp (no_asm) add: POSIX_def) +apply(auto) +apply(erule Prf.cases) +apply(simp_all) +apply(auto)[1] +apply(erule Prf.cases) +apply(simp_all) +apply(auto)[1] +apply(rule ValOrd.intros(2)) +apply(rule ValOrd.intros) +apply(case_tac "\r1a r1b. r1 = ALT r1a r1b") +apply(auto) +apply(rule_tac x = "Seq () va" in exI ) +apply(simp (no_asm) add: POSIX_def) +apply(auto) +apply(erule Prf.cases) +apply(simp_all) +apply(auto)[1] +apply(erule Prf.cases) +apply(simp_all) +apply(auto)[1] +apply(rule ValOrd.intros(2)) +apply(rule ValOrd.intros) + apply(case_tac "v \r1a v1") apply (metis ValOrd.intros(2)) apply(simp add: POSIX_def)