--- a/thys/Re1.thy Fri Sep 19 14:01:01 2014 +0100
+++ b/thys/Re1.thy Mon Sep 22 14:57:02 2014 +0100
@@ -168,6 +168,33 @@
done
*)
+
+
+
+lemma POSIX_ALT2:
+ assumes "POSIX (Left v1) (ALT r1 r2)"
+ shows "POSIX v1 r1"
+using assms
+unfolding POSIX_def
+apply(auto)
+apply(drule_tac x="Left v'" in spec)
+apply(simp)
+apply(drule mp)
+apply(rule Prf.intros)
+apply(auto)
+apply(erule ValOrd.cases)
+apply(simp_all)
+done
+
+lemma POSIX2_ALT:
+ assumes "POSIX2 (Left v1) (ALT r1 r2)"
+ shows "POSIX2 v1 r1"
+using assms
+unfolding POSIX2_def
+apply(auto)
+
+done
+
lemma POSIX_ALT:
assumes "POSIX (Left v1) (ALT r1 r2)"
shows "POSIX v1 r1"
@@ -291,6 +318,14 @@
value.
*}
+lemma mkeps_POSIX2:
+ assumes "nullable r"
+ shows "POSIX2 (mkeps r) r"
+using assms
+apply(induct r)
+apply(auto)[1]
+apply(simp add: POSIX2_def)
+
lemma mkeps_POSIX:
assumes "nullable r"
shows "POSIX (mkeps r) r"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Re1.thy~ Mon Sep 22 14:57:02 2014 +0100
@@ -0,0 +1,605 @@
+
+theory Re1
+ imports "Main"
+begin
+
+section {* Sequential Composition of Sets *}
+
+definition
+ Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
+where
+ "A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}"
+
+text {* Two Simple Properties about Sequential Composition *}
+
+lemma seq_empty [simp]:
+ shows "A ;; {[]} = A"
+ and "{[]} ;; A = A"
+by (simp_all add: Sequ_def)
+
+lemma seq_null [simp]:
+ shows "A ;; {} = {}"
+ and "{} ;; A = {}"
+by (simp_all add: Sequ_def)
+
+section {* Regular Expressions *}
+
+datatype rexp =
+ NULL
+| EMPTY
+| CHAR char
+| SEQ rexp rexp
+| ALT rexp rexp
+
+section {* Semantics of Regular Expressions *}
+
+fun
+ L :: "rexp \<Rightarrow> string set"
+where
+ "L (NULL) = {}"
+| "L (EMPTY) = {[]}"
+| "L (CHAR c) = {[c]}"
+| "L (SEQ r1 r2) = (L r1) ;; (L r2)"
+| "L (ALT r1 r2) = (L r1) \<union> (L r2)"
+
+
+section {* Values *}
+
+datatype val =
+ Void
+| Char char
+| Seq val val
+| Right val
+| Left val
+
+section {* Relation between values and regular expressions *}
+
+inductive Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100)
+where
+ "\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2"
+| "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2"
+| "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2"
+| "\<turnstile> Void : EMPTY"
+| "\<turnstile> Char c : CHAR c"
+
+section {* The string behind a value *}
+
+fun flat :: "val \<Rightarrow> string"
+where
+ "flat(Void) = []"
+| "flat(Char c) = [c]"
+| "flat(Left v) = flat(v)"
+| "flat(Right v) = flat(v)"
+| "flat(Seq v1 v2) = flat(v1) @ flat(v2)"
+
+
+lemma Prf_flat_L:
+ assumes "\<turnstile> v : r" shows "flat v \<in> L r"
+using assms
+apply(induct)
+apply(auto simp add: Sequ_def)
+done
+
+lemma L_flat_Prf:
+ "L(r) = {flat v | v. \<turnstile> v : r}"
+apply(induct r)
+apply(auto dest: Prf_flat_L simp add: Sequ_def)
+apply (metis Prf.intros(4) flat.simps(1))
+apply (metis Prf.intros(5) flat.simps(2))
+apply (metis Prf.intros(1) flat.simps(5))
+apply (metis Prf.intros(2) flat.simps(3))
+apply (metis Prf.intros(3) flat.simps(4))
+apply(erule Prf.cases)
+apply(auto)
+done
+
+section {* Ordering of values *}
+
+inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100)
+where
+ "\<lbrakk>v1 = v1'; v2 \<succ>r2 v2'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')"
+| "v1 \<succ>r1 v1' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')"
+| "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)"
+| "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)"
+| "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')"
+| "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')"
+| "Void \<succ>EMPTY Void"
+| "(Char c) \<succ>(CHAR c) (Char c)"
+
+(*
+lemma
+ assumes "r = SEQ (ALT EMPTY EMPTY) (ALT EMPTY (CHAR c))"
+ shows "(Seq (Left Void) (Right (Char c))) \<succ>r (Seq (Left Void) (Left Void))"
+using assms
+apply(simp)
+apply(rule ValOrd.intros)
+apply(rule ValOrd.intros)
+apply(rule ValOrd.intros)
+apply(rule ValOrd.intros)
+apply(simp)
+done
+*)
+
+section {* Posix definition *}
+
+definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool"
+where
+ "POSIX v r \<equiv> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v')"
+
+(*
+an alternative definition: might cause problems
+with theorem mkeps_POSIX
+*)
+
+definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool"
+where
+ "POSIX2 v r \<equiv> \<turnstile> v : r \<and> (\<forall>v'. \<turnstile> v' : r \<longrightarrow> v \<succ>r v')"
+
+
+(*
+lemma POSIX_SEQ:
+ assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2"
+ shows "POSIX v1 r1 \<and> POSIX v2 r2"
+using assms
+unfolding POSIX_def
+apply(auto)
+apply(drule_tac x="Seq v' v2" in spec)
+apply(simp)
+apply (smt Prf.intros(1) ValOrd.simps assms(3) rexp.inject(2) val.distinct(15) val.distinct(17) val.distinct(3) val.distinct(9) val.inject(2))
+apply(drule_tac x="Seq v1 v'" in spec)
+apply(simp)
+by (smt Prf.intros(1) ValOrd.simps rexp.inject(2) val.distinct(15) val.distinct(17) val.distinct(3) val.distinct(9) val.inject(2))
+*)
+
+(*
+lemma POSIX_SEQ_I:
+ assumes "POSIX v1 r1" "POSIX v2 r2"
+ shows "POSIX (Seq v1 v2) (SEQ r1 r2)"
+using assms
+unfolding POSIX_def
+apply(auto)
+apply(rotate_tac 2)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(rule ValOrd.intros)
+
+apply(auto)
+done
+*)
+
+lemma POSIX_ALT:
+ assumes "POSIX (Left v1) (ALT r1 r2)"
+ shows "POSIX v1 r1"
+using assms
+unfolding POSIX_def
+apply(auto)
+apply(drule_tac x="Left v'" in spec)
+apply(simp)
+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"
+using assms
+unfolding POSIX_def
+apply(auto)
+apply(drule_tac x="Right v'" in spec)
+apply(simp)
+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)"
+ shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')"
+using assms
+apply(drule_tac POSIX_ALT1a)
+unfolding POSIX_def
+apply(auto)
+done
+
+lemma POSIX_ALT_I1:
+ assumes "POSIX v1 r1"
+ shows "POSIX (Left v1) (ALT r1 r2)"
+using assms
+unfolding POSIX_def
+apply(auto)
+apply(rotate_tac 3)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)
+apply(rule ValOrd.intros)
+apply(auto)
+apply(rule ValOrd.intros)
+by simp
+
+lemma POSIX_ALT_I2:
+ assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')"
+ shows "POSIX (Right v2) (ALT r1 r2)"
+using assms
+unfolding POSIX_def
+apply(auto)
+apply(rotate_tac 3)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)
+apply(rule ValOrd.intros)
+apply metis
+done
+
+
+section {* The ordering is reflexive *}
+
+lemma ValOrd_refl:
+ assumes "\<turnstile> v : r"
+ shows "v \<succ>r v"
+using assms
+apply(induct)
+apply(auto intro: ValOrd.intros)
+done
+
+
+section {* The Matcher *}
+
+fun
+ nullable :: "rexp \<Rightarrow> bool"
+where
+ "nullable (NULL) = False"
+| "nullable (EMPTY) = True"
+| "nullable (CHAR c) = False"
+| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
+| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
+
+lemma nullable_correctness:
+ shows "nullable r \<longleftrightarrow> [] \<in> (L r)"
+apply (induct r)
+apply(auto simp add: Sequ_def)
+done
+
+fun mkeps :: "rexp \<Rightarrow> val"
+where
+ "mkeps(EMPTY) = Void"
+| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
+| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
+
+lemma mkeps_nullable:
+ assumes "nullable(r)" shows "\<turnstile> mkeps r : r"
+using assms
+apply(induct rule: nullable.induct)
+apply(auto intro: Prf.intros)
+done
+
+lemma mkeps_flat:
+ assumes "nullable(r)" shows "flat (mkeps r) = []"
+using assms
+apply(induct rule: nullable.induct)
+apply(auto)
+done
+
+text {*
+ The value mkeps returns is always the correct POSIX
+ value.
+*}
+
+lemma mkeps_POSIX:
+ assumes "nullable r"
+ shows "POSIX (mkeps r) r"
+using assms
+apply(induct r)
+apply(auto)[1]
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis ValOrd.intros)
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)
+apply (simp add: ValOrd.intros(2) mkeps_flat)
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)
+apply (simp add: ValOrd.intros(6))
+apply (simp add: ValOrd.intros(3))
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)
+apply (simp add: ValOrd.intros(6))
+apply (simp add: ValOrd.intros(3))
+apply(simp add: POSIX_def)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)
+apply (metis Prf_flat_L mkeps_flat nullable_correctness)
+by (simp add: ValOrd.intros(5))
+
+
+section {* Derivatives *}
+
+fun
+ der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
+where
+ "der c (NULL) = NULL"
+| "der c (EMPTY) = NULL"
+| "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
+| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
+| "der c (SEQ r1 r2) =
+ (if nullable r1
+ then ALT (SEQ (der c r1) r2) (der c r2)
+ else SEQ (der c r1) r2)"
+
+fun
+ ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
+where
+ "ders [] r = r"
+| "ders (c # s) r = ders s (der c r)"
+
+section {* Injection function *}
+
+fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
+where
+ "injval (CHAR d) c Void = Char d"
+| "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)"
+| "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)"
+| "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2"
+| "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
+| "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
+
+section {* Projection function *}
+
+fun projval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
+where
+ "projval (CHAR d) c _ = Void"
+| "projval (ALT r1 r2) c (Left v1) = Left(projval r1 c v1)"
+| "projval (ALT r1 r2) c (Right v2) = Right(projval r2 c v2)"
+| "projval (SEQ r1 r2) c (Seq v1 v2) =
+ (if flat v1 = [] then Right(projval r2 c v2)
+ else if nullable r1 then Left (Seq (projval r1 c v1) v2)
+ else Seq (projval r1 c v1) v2)"
+
+text {*
+ Injection value is related to r
+*}
+
+lemma v3:
+ assumes "\<turnstile> v : der c r" shows "\<turnstile> (injval r c v) : r"
+using assms
+apply(induct arbitrary: v rule: der.induct)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(case_tac "c = c'")
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis Prf.intros(5))
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply (metis Prf.intros(2))
+apply (metis Prf.intros(3))
+apply(simp)
+apply(case_tac "nullable r1")
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply (metis Prf.intros(1))
+apply(auto)[1]
+apply (metis Prf.intros(1) mkeps_nullable)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(rule Prf.intros)
+apply(auto)[2]
+done
+
+text {*
+ The string behin the injection value is an added c
+*}
+
+lemma v4:
+ assumes "\<turnstile> v : der c r" shows "flat (injval r c v) = c # (flat v)"
+using assms
+apply(induct arbitrary: v rule: der.induct)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(case_tac "c = c'")
+apply(simp)
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(case_tac "nullable r1")
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply (metis mkeps_flat)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+done
+
+text {*
+ Injection followed by projection is the identity.
+*}
+
+lemma proj_inj_id:
+ assumes "\<turnstile> v : der c r"
+ shows "projval r c (injval r c v) = v"
+using assms
+apply(induct r arbitrary: c v rule: rexp.induct)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(case_tac "c = char")
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+defer
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(case_tac "nullable rexp1")
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply (metis list.distinct(1) v4)
+apply(auto)[1]
+apply (metis mkeps_flat)
+apply(auto)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(auto)[1]
+apply(simp add: v4)
+done
+
+lemma "\<exists>v. POSIX v r"
+apply(induct r)
+apply(rule exI)
+apply(simp add: POSIX_def)
+apply (metis (full_types) Prf_flat_L der.simps(1) der.simps(2) der.simps(3) flat.simps(1) nullable.simps(1) nullable_correctness proj_inj_id projval.simps(1) v3 v4)
+apply(rule_tac x = "Void" in exI)
+apply(simp add: POSIX_def)
+apply (metis POSIX_def flat.simps(1) mkeps.simps(1) mkeps_POSIX nullable.simps(2))
+apply(rule_tac x = "Char char" in exI)
+apply(simp add: POSIX_def)
+apply(auto) [1]
+apply(erule Prf.cases)
+apply(simp_all) [5]
+apply (metis ValOrd.intros(8))
+defer
+apply(auto)
+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(simp (no_asm) add: POSIX_def)
+apply(auto)
+apply(erule Prf.cases)
+apply(simp_all)
+apply(case_tac "v \<succ>r1a v1")
+apply (metis ValOrd.intros(2))
+apply(simp add: POSIX_def)
+apply(case_tac "flat v = flat v1")
+apply(auto)[1]
+apply(simp only: append_eq_append_conv2)
+apply(auto)
+thm append_eq_append_conv2
+
+text {*
+
+ HERE: Crucial lemma that does not go through in the sequence case.
+
+*}
+lemma v5:
+ assumes "\<turnstile> v : der c r" "POSIX v (der c r)"
+ shows "POSIX (injval r c v) r"
+using assms
+apply(induct arbitrary: v rule: der.induct)
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(simp)
+apply(case_tac "c = c'")
+apply(auto simp add: POSIX_def)[1]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+apply(erule Prf.cases)
+apply(simp_all)[5]
+using ValOrd.simps apply blast
+apply(auto)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+(* base cases done *)
+(* ALT case *)
+apply(erule Prf.cases)
+apply(simp_all)[5]
+using POSIX_ALT POSIX_ALT_I1 apply blast
+apply(clarify)
+apply(subgoal_tac "POSIX v2 (der c r2)")
+prefer 2
+apply(auto simp add: POSIX_def)[1]
+apply (metis POSIX_ALT1a POSIX_def flat.simps(4))
+apply(rotate_tac 1)
+apply(drule_tac x="v2" in meta_spec)
+apply(simp)
+apply(subgoal_tac "\<turnstile> Right (injval r2 c v2) : (ALT r1 r2)")
+prefer 2
+apply (metis Prf.intros(3) v3)
+apply(rule ccontr)
+apply(auto simp add: POSIX_def)[1]
+
+apply(rule allI)
+apply(rule impI)
+apply(erule conjE)
+thm POSIX_ALT_I2
+apply(frule POSIX_ALT1a)
+apply(drule POSIX_ALT1b)
+apply(rule POSIX_ALT_I2)
+apply auto[1]
+apply(subst v4)
+apply(auto)[2]
+apply(rotate_tac 1)
+apply(drule_tac x="v2" in meta_spec)
+apply(simp)
+apply(subst (asm) (4) POSIX_def)
+apply(subst (asm) v4)
+apply(auto)[2]
+(* stuck in the ALT case *)