# HG changeset patch # User urbanc # Date 1311875556 0 # Node ID b755090d0f3d8be8cfff920825221917cf6a7ccb # Parent edacc141060f4a9253382bc741ce95fc4a188c7b added a picture diff -r edacc141060f -r b755090d0f3d Derivatives.thy --- a/Derivatives.thy Thu Jul 28 14:22:10 2011 +0000 +++ b/Derivatives.thy Thu Jul 28 17:52:36 2011 +0000 @@ -66,8 +66,8 @@ apply(auto simp add: Cons_eq_append_conv) done - have "Der c (A\) = Der c ({[]} \ A \ A\)" - by (simp only: star_cases[symmetric]) + have "Der c (A\) = Der c (A \ A\ \ {[]})" + by (simp only: star_unfold_left[symmetric]) also have "... = Der c (A \ A\)" by (simp only: Der_union Der_one) (simp) also have "... = (Der c A) \ A\ \ (Delta A \ Der c (A\))" @@ -157,7 +157,7 @@ pders :: "'a list \ 'a rexp \ ('a rexp) set" where "pders [] r = {r}" -| "pders (c # s) r = (\r'\ (pder c r). (pders s r'))" +| "pders (c # s) r = \ (pders s) ` (pder c r)" abbreviation "pders_set A r \ \s \ A. pders s r" diff -r edacc141060f -r b755090d0f3d Journal/Paper.thy --- a/Journal/Paper.thy Thu Jul 28 14:22:10 2011 +0000 +++ b/Journal/Paper.thy Thu Jul 28 17:52:36 2011 +0000 @@ -67,7 +67,8 @@ tag_str_Star ("tag\<^isub>S\<^isub>T\<^isub>A\<^isub>R _ _" [100, 100] 100) and tag_eq_rel ("\<^raw:$\threesim$>\<^bsub>_\<^esub>") and Delta ("\'(_')") and - nullable ("\'(_')") + nullable ("\'(_')") and + Cons ("_ :: _" [100, 100] 100) lemma meta_eq_app: shows "f \ \x. g x \ f x \ g x" @@ -389,7 +390,7 @@ \begin{prpstn}\label{langprops}\mbox{}\\ \begin{tabular}{@ {}ll} - (i) & @{thm star_cases} \\ + (i) & @{thm star_unfold_left} \\ (ii) & @{thm[mode=IfThen] pow_length}\\ (iii) & @{thm conc_Union_left} \\ \end{tabular} @@ -1098,8 +1099,38 @@ Much more interesting, however, are the inductive cases. They seem hard to solve directly. The reader is invited to try. - Constable et al \cite{Constable00} give the following suggestive picture about - equivalence classes: + In order how to see how our first proof proceeds we use the following suggestive picture + from Constable et al \cite{Constable00}: + + \begin{center} + \begin{tabular}{c@ {\hspace{10mm}}c@ {\hspace{10mm}}c} + \begin{tikzpicture}[scale=1] + %Circle + \draw[thick] (0,0) circle (1.1); + \end{tikzpicture} + & + \begin{tikzpicture}[scale=1] + %Circle + \draw[thick] (0,0) circle (1.1); + %Main rays + \foreach \a in {0, 90,...,359} + \draw[very thick] (0, 0) -- (\a:1.1); + \foreach \a / \l in {45/1, 135/2, 225/3, 315/4} + \draw (\a: 0.65) node {$a_\l$}; + \end{tikzpicture} + & + \begin{tikzpicture}[scale=1] + %Circle + \draw[thick] (0,0) circle (1.1); + %Main rays + \foreach \a in {0, 45,...,359} + \draw[very thick] (0, 0) -- (\a:1.1); + \foreach \a / \l in {22.5/1.1, 67.5/1.2, 112.5/2.1, 157.5/2.2, 202.4/3.1, 247.5/3.2, 292.5/4.1, 337.5/4.2} + \draw (\a: 0.77) node {$a_{\l}$}; + \end{tikzpicture}\\ + @{term UNIV} & @{term "UNIV // (\(lang r))"} & @{term "UNIV // R"} + \end{tabular} + \end{center} Our first proof will rely on some \emph{tagging-functions} defined over strings. Given the inductive hypothesis, it will diff -r edacc141060f -r b755090d0f3d More_Regular_Set.thy --- a/More_Regular_Set.thy Thu Jul 28 14:22:10 2011 +0000 +++ b/More_Regular_Set.thy Thu Jul 28 17:52:36 2011 +0000 @@ -9,25 +9,6 @@ conc (infixr "\" 100) and star ("_\" [101] 102) -lemma conc_add_left: - assumes a: "A = B" - shows "C \ A = C \ B" -using a by simp - -lemma star_cases: - shows "A\ = {[]} \ A \ A\" -proof - { fix x - have "x \ A\ \ x \ {[]} \ A \ A\" - unfolding conc_def - by (induct rule: star_induct) (auto) - } - then show "A\ \ {[]} \ A \ A\" by auto -next - show "{[]} \ A \ A\ \ A\" - unfolding conc_def by auto -qed - lemma star_decom: assumes a: "x \ A\" "x \ []" shows "\a b. x = a @ b \ a \ [] \ a \ A \ b \ A\" @@ -109,9 +90,9 @@ assume eq: "X = B \ A\" have "A\ = {[]} \ A\ \ A" unfolding conc_star_comm[symmetric] - by (rule star_cases) + by(metis Un_commute star_unfold_left) then have "B \ A\ = B \ ({[]} \ A\ \ A)" - by (rule conc_add_left) + by metis also have "\ = B \ B \ (A\ \ A)" unfolding conc_Un_distrib by simp also have "\ = B \ (B \ A\) \ A" diff -r edacc141060f -r b755090d0f3d Myhill_2.thy --- a/Myhill_2.thy Thu Jul 28 14:22:10 2011 +0000 +++ b/Myhill_2.thy Thu Jul 28 17:52:36 2011 +0000 @@ -399,7 +399,7 @@ proof - from h1 have "(x - xa_max) @ z \ []" by (auto simp:strict_prefix_def elim:prefixE) - from star_decom [OF h3 this] + from star_decom [OF h3 this] obtain a b where a_in: "a \ L\<^isub>1" and a_neq: "a \ []" and b_in: "b \ L\<^isub>1\" and ab_max: "(x - xa_max) @ z = a @ b" by blast