diff -r a8749366ad5d -r 8b432cef3805 thys/LexerExt.thy --- a/thys/LexerExt.thy Thu Mar 02 00:56:40 2017 +0000 +++ b/thys/LexerExt.thy Thu Mar 02 12:39:27 2017 +0000 @@ -13,23 +13,28 @@ text {* Two Simple Properties about Sequential Composition *} -lemma seq_empty [simp]: +lemma Sequ_empty [simp]: shows "A ;; {[]} = A" and "{[]} ;; A = A" by (simp_all add: Sequ_def) -lemma seq_null [simp]: +lemma Sequ_null [simp]: shows "A ;; {} = {}" and "{} ;; A = {}" by (simp_all add: Sequ_def) -lemma seq_assoc: +lemma Sequ_assoc: shows "A ;; (B ;; C) = (A ;; B) ;; C" apply(auto simp add: Sequ_def) apply(metis append_assoc) apply(metis) done +lemma Sequ_UNION: + shows "(\x\A. B ;; C x) = B ;; (\x\A. C x)" +by (auto simp add: Sequ_def) + + section {* Semantic Derivative (Left Quotient) of Languages *} definition @@ -84,26 +89,104 @@ shows "[] \ A \ n \ (n = 0 \ [] \ A)" by(induct n) (auto simp add: Sequ_def) +lemma Der_Pow [simp]: + shows "Der c (A \ (Suc n)) = + (Der c A) ;; (A \ n) \ (if [] \ A then Der c (A \ n) else {})" +unfolding Der_def Sequ_def +by(auto simp add: Cons_eq_append_conv Sequ_def) + +lemma Der_Pow_subseteq: + assumes "[] \ A" + shows "Der c (A \ n) \ (Der c A) ;; (A \ n)" +using assms +apply(induct n) +apply(simp add: Sequ_def Der_def) +apply(simp) +apply(rule conjI) +apply (smt Sequ_def append_Nil2 mem_Collect_eq Sequ_assoc subsetI) +apply(subgoal_tac "((Der c A) ;; (A \ n)) \ ((Der c A) ;; (A ;; (A \ n)))") +apply(auto)[1] +by (smt Sequ_def append_Nil2 mem_Collect_eq Sequ_assoc subsetI) + +lemma test: + shows "(\x\Suc n. Der c (A \ x)) = (\x\n. Der c A ;; A \ x)" +apply(induct n) +apply(simp) +apply(auto)[1] +apply(case_tac xa) +apply(simp) +apply(simp) +apply(auto)[1] +apply(case_tac "[] \ A") +apply(simp) +apply(simp) +by (smt Der_Pow Der_Pow_subseteq UN_insert atMost_Suc sup.orderE sup_bot.right_neutral) + +lemma test2: + shows "(\x\(Suc ` A). Der c (B \ x)) = (\x\A. Der c B ;; B \ x)" +apply(auto)[1] +apply(case_tac "[] \ B") +apply(simp) +using Der_Pow_subseteq apply blast +apply(simp) +done + + section {* Kleene Star for Languages *} -inductive_set - Star :: "string set \ string set" ("_\" [101] 102) - for A :: "string set" -where - start[intro]: "[] \ A\" -| step[intro]: "\s1 \ A; s2 \ A\\ \ s1 @ s2 \ A\" +definition + Star :: "string set \ string set" ("_\" [101] 102) where + "A\ = (\n. A \ n)" + +lemma Star_empty [intro]: + shows "[] \ A\" +unfolding Star_def +by auto + +lemma Star_step [intro]: + assumes "s1 \ A" "s2 \ A\" + shows "s1 @ s2 \ A\" +proof - + from assms obtain n where "s1 \A" "s2 \ A \ n" + unfolding Star_def by auto + then have "s1 @ s2 \ A ;; (A \ n)" + by (auto simp add: Sequ_def) + then have "s1 @ s2 \ A \ (Suc n)" + by simp + then show "s1 @ s2 \ A\" + unfolding Star_def + by (auto simp del: Pow.simps) +qed lemma star_cases: shows "A\ = {[]} \ A ;; A\" -unfolding Sequ_def -by (auto) (metis Star.simps) +unfolding Star_def +apply(simp add: Sequ_def) +apply(auto) +apply(case_tac xa) +apply(auto simp add: Sequ_def) +apply(rule_tac x="Suc xa" in exI) +apply(auto simp add: Sequ_def) +done -lemma star_decomp: - assumes a: "c # x \ A\" - shows "\a b. x = a @ b \ c # a \ A \ b \ A\" -using a -by (induct x\"c # x" rule: Star.induct) - (auto simp add: append_eq_Cons_conv) +lemma Der_Star1: + shows "Der c (A ;; A\) = (Der c A) ;; A\" +proof - + have "(Der c A) ;; A\ = (Der c A) ;; (\n. A \ n)" + unfolding Star_def by simp + also have "... = (\n. Der c A ;; A \ n)" + unfolding Sequ_UNION by simp + also have "... = (\x\(Suc ` UNIV). Der c (A \ x))" + unfolding test2 by simp + also have "... = (\n. Der c (A \ (Suc n)))" + by (simp) + also have "... = Der c (\n. A \ (Suc n))" + unfolding Der_UNION by simp + also have "... = Der c (A ;; (\n. A \ n))" + by (simp only: Pow.simps Sequ_UNION) + finally show "Der c (A ;; A\) = (Der c A) ;; A\" + unfolding Star_def[symmetric] by blast +qed lemma Der_star [simp]: shows "Der c (A\) = (Der c A) ;; A\" @@ -112,35 +195,12 @@ by (simp only: star_cases[symmetric]) also have "... = Der c (A ;; A\)" by (simp only: Der_union Der_empty) (simp) - also have "... = (Der c A) ;; A\ \ (if [] \ A then Der c (A\) else {})" - by simp also have "... = (Der c A) ;; A\" - unfolding Sequ_def Der_def - by (auto dest: star_decomp) + using Der_Star1 by simp finally show "Der c (A\) = (Der c A) ;; A\" . qed -lemma Star_in_Pow: - assumes a: "s \ A\" - shows "\n. s \ A \ n" -using a -apply(induct) -apply(auto) -apply(rule_tac x="Suc n" in exI) -apply(auto simp add: Sequ_def) -done -lemma Pow_in_Star: - assumes a: "s \ A \ n" - shows "s \ A\" -using a -by (induct n arbitrary: s) (auto simp add: Sequ_def) - - -lemma Star_def2: - shows "A\ = (\n. A \ n)" -using Star_in_Pow Pow_in_Star -by (auto) section {* Regular Expressions *} @@ -208,8 +268,10 @@ | "der c (FROMNTIMES r 0) = SEQ (der c r) (FROMNTIMES r 0)" | "der c (FROMNTIMES r (Suc n)) = SEQ (der c r) (FROMNTIMES r n)" | "der c (NMTIMES r n m) = - (if m < n then ZERO else (if n = 0 then (if m = 0 then ZERO else SEQ (der c r) (UPNTIMES r (m - 1))) else - SEQ (der c r) (NMTIMES r (n - 1) (m - 1))))" + (if m < n then ZERO + else (if n = 0 then (if m = 0 then ZERO else + SEQ (der c r) (UPNTIMES r (m - 1))) else + SEQ (der c r) (NMTIMES r (n - 1) (m - 1))))" fun ders :: "string \ rexp \ rexp" @@ -224,67 +286,6 @@ apply(auto simp add: Sequ_def) done - -lemma Suc_reduce_Union2: - "(\x\{Suc n..}. B x) = (\x\{n..}. B (Suc x))" -apply(auto) -apply(rule_tac x="xa - 1" in bexI) -apply(simp) -apply(simp) -done - -lemma Seq_UNION: - shows "(\x\A. B ;; C x) = B ;; (\x\A. C x)" -by (auto simp add: Sequ_def) - -lemma Der_Pow [simp]: - shows "Der c (A \ (Suc n)) = - (Der c A) ;; (A \ n) \ (if [] \ A then Der c (A \ n) else {})" -unfolding Der_def Sequ_def -by(auto simp add: Cons_eq_append_conv Sequ_def) - -lemma Suc_Union: - "(\x\Suc m. B x) = (B (Suc m) \ (\x\m. B x))" -by (metis UN_insert atMost_Suc) - - -lemma Der_Pow_subseteq: - assumes "[] \ A" - shows "Der c (A \ n) \ (Der c A) ;; (A \ n)" -using assms -apply(induct n) -apply(simp add: Sequ_def Der_def) -apply(simp) -apply(rule conjI) -apply (smt Sequ_def append_Nil2 mem_Collect_eq seq_assoc subsetI) -apply(subgoal_tac "((Der c A) ;; (A \ n)) \ ((Der c A) ;; (A ;; (A \ n)))") -apply(auto)[1] -by (smt Sequ_def append_Nil2 mem_Collect_eq seq_assoc subsetI) - -lemma test: - shows "(\x\Suc n. Der c (L r \ x)) = (\x\n. Der c (L r) ;; L r \ x)" -apply(induct n) -apply(simp) -apply(auto)[1] -apply(case_tac xa) -apply(simp) -apply(simp) -apply(auto)[1] -apply(case_tac "[] \ L r") -apply(simp) -apply(simp) -by (smt Der_Pow Suc_Union inf_sup_aci(5) inf_sup_aci(7) sup_idem) - -lemma test2: - shows "(\x\(Suc ` A). Der c (L r \ x)) = (\x\A. Der c (L r) ;; L r \ x)" -apply(auto)[1] -apply(case_tac "[] \ L r") -apply(simp) -defer -apply(simp) -using Der_Pow_subseteq by blast - - lemma Der_Pow_notin: assumes "[] \ A" shows "Der c (A \ (Suc n)) = (Der c A) ;; (A \ n)" @@ -297,7 +298,7 @@ apply(simp_all add: nullable_correctness)[7] apply(simp only: der.simps L.simps) apply(simp only: Der_UNION) -apply(simp only: Seq_UNION[symmetric]) +apply(simp only: Sequ_UNION[symmetric]) apply(simp add: test) apply(simp) (* NTIMES *) @@ -309,11 +310,16 @@ apply(simp only: der.simps) apply(simp only: L.simps) apply(simp) -using Der_star Star_def2 apply auto[1] +using Der_star Star_def apply auto[1] apply(simp only: der.simps) apply(simp only: L.simps) apply(simp add: Der_UNION) -apply(smt Der_Pow Der_Pow_notin Der_Pow_subseteq SUP_cong Seq_UNION Suc_reduce_Union2 sup.absorb_iff1) +apply(subst Sequ_UNION[symmetric]) +apply(subst test2[symmetric]) +apply(subgoal_tac "(Suc ` {n..}) = {Suc n..}") +apply simp +apply(auto simp add: image_def)[1] +using Suc_le_D apply blast (* NMTIMES *) apply(case_tac "n \ m") prefer 2 @@ -321,17 +327,17 @@ apply(simp only: L.simps) apply(simp) apply(auto) -apply(subst (asm) Seq_UNION[symmetric]) +apply(subst (asm) Sequ_UNION[symmetric]) apply(subst (asm) test[symmetric]) apply(auto simp add: Der_UNION)[1] apply(auto simp add: Der_UNION)[1] -apply(subst Seq_UNION[symmetric]) +apply(subst Sequ_UNION[symmetric]) apply(subst test[symmetric]) apply(auto)[1] -apply(subst (asm) Seq_UNION[symmetric]) +apply(subst (asm) Sequ_UNION[symmetric]) apply(subst (asm) test2[symmetric]) apply(auto simp add: Der_UNION)[1] -apply(subst Seq_UNION[symmetric]) +apply(subst Sequ_UNION[symmetric]) apply(subst test2[symmetric]) apply(auto simp add: Der_UNION)[1] done @@ -459,17 +465,25 @@ apply(simp) done +lemma Star_Pow: + assumes "s \ A \ n" + shows "\ss. concat ss = s \ (\s \ set ss. s \ A) \ (length ss = n)" +using assms +apply(induct n arbitrary: s) +apply(auto simp add: Sequ_def) +apply(drule_tac x="s2" in meta_spec) +apply(auto) +apply(rule_tac x="s1#ss" in exI) +apply(simp) +done + lemma Star_string: assumes "s \ A\" shows "\ss. concat ss = s \ (\s \ set ss. s \ A)" using assms -apply(induct rule: Star.induct) -apply(auto) -apply(rule_tac x="[]" in exI) -apply(simp) -apply(rule_tac x="s1#ss" in exI) -apply(simp) -done +apply(auto simp add: Star_def) +using Star_Pow by blast + lemma Star_val: assumes "\s\set ss. \v. s = flat v \ \ v : r" @@ -489,17 +503,7 @@ by (metis List.bind_def bind_simps(2) length_Suc_conv set_ConsD) -lemma Star_Pow: - assumes "s \ A \ n" - shows "\ss. concat ss = s \ (\s \ set ss. s \ A) \ (length ss = n)" -using assms -apply(induct n arbitrary: s) -apply(auto simp add: Sequ_def) -apply(drule_tac x="s2" in meta_spec) -apply(auto) -apply(rule_tac x="s1#ss" in exI) -apply(simp) -done + lemma L_flat_Prf2: @@ -764,8 +768,8 @@ apply(rule_tac x="Suc x" in bexI) apply(auto simp add: Sequ_def) apply(rule_tac x="Suc x" in bexI) -apply(auto simp add: Sequ_def) -apply(simp add: Star_in_Pow) +using Sequ_def apply auto[2] +apply(simp add: Star_def) apply(rule_tac x="Suc x" in bexI) apply(auto simp add: Sequ_def) done @@ -1393,7 +1397,7 @@ apply(erule Posix.cases) apply(simp_all) apply (simp add: Posix1a Prf_injval_flat) - apply(simp only: Star_def2) + apply(simp only: Star_def) apply(simp only: der_correctness Der_def) apply(simp) done