# HG changeset patch # User urbanc # Date 1297222410 0 # Node ID f41351709800d2be02e4cb745e41ddb2f0d5080c # Parent f438f4dbaada30d2b73a956bbd95b79706700cfa saved a copy of the current Myhill for reference diff -r f438f4dbaada -r f41351709800 Attic/Myhill_1.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Attic/Myhill_1.thy Wed Feb 09 03:33:30 2011 +0000 @@ -0,0 +1,1293 @@ +theory Myhill_1 + imports Main +begin + +(* +text {* + \begin{figure} + \centering + \scalebox{0.95}{ + \begin{tikzpicture}[->,>=latex,shorten >=1pt,auto,node distance=1.2cm, semithick] + \node[state,initial] (n1) {$1$}; + \node[state,accepting] (n2) [right = 10em of n1] {$2$}; + + \path (n1) edge [bend left] node {$0$} (n2) + (n1) edge [loop above] node{$1$} (n1) + (n2) edge [loop above] node{$0$} (n2) + (n2) edge [bend left] node {$1$} (n1) + ; + \end{tikzpicture}} + \caption{An example automaton (or partition)}\label{fig:example_automata} + \end{figure} +*} + +*) + + +section {* Preliminary definitions *} + +types lang = "string set" + +text {* Sequential composition of two languages *} + +definition + Seq :: "lang \ lang \ lang" (infixr ";;" 100) +where + "A ;; B = {s\<^isub>1 @ s\<^isub>2 | s\<^isub>1 s\<^isub>2. s\<^isub>1 \ A \ s\<^isub>2 \ B}" + + +text {* Some properties of operator @{text ";;"}. *} + +lemma seq_add_left: + assumes a: "A = B" + shows "C ;; A = C ;; B" +using a by simp + +lemma seq_union_distrib_right: + shows "(A \ B) ;; C = (A ;; C) \ (B ;; C)" +unfolding Seq_def by auto + +lemma seq_union_distrib_left: + shows "C ;; (A \ B) = (C ;; A) \ (C ;; B)" +unfolding Seq_def by auto + +lemma seq_intro: + assumes a: "x \ A" "y \ B" + shows "x @ y \ A ;; B " +using a by (auto simp: Seq_def) + +lemma seq_assoc: + shows "(A ;; B) ;; C = A ;; (B ;; C)" +unfolding Seq_def +apply(auto) +apply(blast) +by (metis append_assoc) + +lemma seq_empty [simp]: + shows "A ;; {[]} = A" + and "{[]} ;; A = A" +by (simp_all add: Seq_def) + + +text {* Power and Star of a language *} + +fun + pow :: "lang \ nat \ lang" (infixl "\" 100) +where + "A \ 0 = {[]}" +| "A \ (Suc n) = A ;; (A \ n)" + +definition + Star :: "lang \ lang" ("_\" [101] 102) +where + "A\ \ (\n. A \ n)" + + +lemma star_start[intro]: + shows "[] \ A\" +proof - + have "[] \ A \ 0" by auto + then show "[] \ A\" unfolding Star_def by blast +qed + +lemma star_step [intro]: + assumes a: "s1 \ A" + and b: "s2 \ A\" + shows "s1 @ s2 \ A\" +proof - + from b obtain n where "s2 \ A \ n" unfolding Star_def by auto + then have "s1 @ s2 \ A \ (Suc n)" using a by (auto simp add: Seq_def) + then show "s1 @ s2 \ A\" unfolding Star_def by blast +qed + +lemma star_induct[consumes 1, case_names start step]: + assumes a: "x \ A\" + and b: "P []" + and c: "\s1 s2. \s1 \ A; s2 \ A\; P s2\ \ P (s1 @ s2)" + shows "P x" +proof - + from a obtain n where "x \ A \ n" unfolding Star_def by auto + then show "P x" + by (induct n arbitrary: x) + (auto intro!: b c simp add: Seq_def Star_def) +qed + +lemma star_intro1: + assumes a: "x \ A\" + and b: "y \ A\" + shows "x @ y \ A\" +using a b +by (induct rule: star_induct) (auto) + +lemma star_intro2: + assumes a: "y \ A" + shows "y \ A\" +proof - + from a have "y @ [] \ A\" by blast + then show "y \ A\" by simp +qed + +lemma star_intro3: + assumes a: "x \ A\" + and b: "y \ A" + shows "x @ y \ A\" +using a b by (blast intro: star_intro1 star_intro2) + +lemma star_cases: + shows "A\ = {[]} \ A ;; A\" +proof + { fix x + have "x \ A\ \ x \ {[]} \ A ;; A\" + unfolding Seq_def + by (induct rule: star_induct) (auto) + } + then show "A\ \ {[]} \ A ;; A\" by auto +next + show "{[]} \ A ;; A\ \ A\" + unfolding Seq_def by auto +qed + +lemma star_decom: + assumes a: "x \ A\" "x \ []" + shows "\a b. x = a @ b \ a \ [] \ a \ A \ b \ A\" +using a +apply(induct rule: star_induct) +apply(simp) +apply(blast) +done + +lemma + shows seq_Union_left: "B ;; (\n. A \ n) = (\n. B ;; (A \ n))" + and seq_Union_right: "(\n. A \ n) ;; B = (\n. (A \ n) ;; B)" +unfolding Seq_def by auto + +lemma seq_pow_comm: + shows "A ;; (A \ n) = (A \ n) ;; A" +by (induct n) (simp_all add: seq_assoc[symmetric]) + +lemma seq_star_comm: + shows "A ;; A\ = A\ ;; A" +unfolding Star_def +unfolding seq_Union_left +unfolding seq_pow_comm +unfolding seq_Union_right +by simp + +text {* Two lemmas about the length of strings in @{text "A \ n"} *} + +lemma pow_length: + assumes a: "[] \ A" + and b: "s \ A \ Suc n" + shows "n < length s" +using b +proof (induct n arbitrary: s) + case 0 + have "s \ A \ Suc 0" by fact + with a have "s \ []" by auto + then show "0 < length s" by auto +next + case (Suc n) + have ih: "\s. s \ A \ Suc n \ n < length s" by fact + have "s \ A \ Suc (Suc n)" by fact + then obtain s1 s2 where eq: "s = s1 @ s2" and *: "s1 \ A" and **: "s2 \ A \ Suc n" + by (auto simp add: Seq_def) + from ih ** have "n < length s2" by simp + moreover have "0 < length s1" using * a by auto + ultimately show "Suc n < length s" unfolding eq + by (simp only: length_append) +qed + +lemma seq_pow_length: + assumes a: "[] \ A" + and b: "s \ B ;; (A \ Suc n)" + shows "n < length s" +proof - + from b obtain s1 s2 where eq: "s = s1 @ s2" and *: "s2 \ A \ Suc n" + unfolding Seq_def by auto + from * have " n < length s2" by (rule pow_length[OF a]) + then show "n < length s" using eq by simp +qed + + +section {* A slightly modified version of Arden's lemma *} + + +text {* A helper lemma for Arden *} + +lemma ardens_helper: + assumes eq: "X = X ;; A \ B" + shows "X = X ;; (A \ Suc n) \ (\m\{0..n}. B ;; (A \ m))" +proof (induct n) + case 0 + show "X = X ;; (A \ Suc 0) \ (\(m::nat)\{0..0}. B ;; (A \ m))" + using eq by simp +next + case (Suc n) + have ih: "X = X ;; (A \ Suc n) \ (\m\{0..n}. B ;; (A \ m))" by fact + also have "\ = (X ;; A \ B) ;; (A \ Suc n) \ (\m\{0..n}. B ;; (A \ m))" using eq by simp + also have "\ = X ;; (A \ Suc (Suc n)) \ (B ;; (A \ Suc n)) \ (\m\{0..n}. B ;; (A \ m))" + by (simp add: seq_union_distrib_right seq_assoc) + also have "\ = X ;; (A \ Suc (Suc n)) \ (\m\{0..Suc n}. B ;; (A \ m))" + by (auto simp add: le_Suc_eq) + finally show "X = X ;; (A \ Suc (Suc n)) \ (\m\{0..Suc n}. B ;; (A \ m))" . +qed + +theorem ardens_revised: + assumes nemp: "[] \ A" + shows "X = X ;; A \ B \ X = B ;; A\" +proof + assume eq: "X = B ;; A\" + have "A\ = {[]} \ A\ ;; A" + unfolding seq_star_comm[symmetric] + by (rule star_cases) + then have "B ;; A\ = B ;; ({[]} \ A\ ;; A)" + by (rule seq_add_left) + also have "\ = B \ B ;; (A\ ;; A)" + unfolding seq_union_distrib_left by simp + also have "\ = B \ (B ;; A\) ;; A" + by (simp only: seq_assoc) + finally show "X = X ;; A \ B" + using eq by blast +next + assume eq: "X = X ;; A \ B" + { fix n::nat + have "B ;; (A \ n) \ X" using ardens_helper[OF eq, of "n"] by auto } + then have "B ;; A\ \ X" + unfolding Seq_def Star_def UNION_def + by auto + moreover + { fix s::string + obtain k where "k = length s" by auto + then have not_in: "s \ X ;; (A \ Suc k)" + using seq_pow_length[OF nemp] by blast + assume "s \ X" + then have "s \ X ;; (A \ Suc k) \ (\m\{0..k}. B ;; (A \ m))" + using ardens_helper[OF eq, of "k"] by auto + then have "s \ (\m\{0..k}. B ;; (A \ m))" using not_in by auto + moreover + have "(\m\{0..k}. B ;; (A \ m)) \ (\n. B ;; (A \ n))" by auto + ultimately + have "s \ B ;; A\" + unfolding seq_Union_left Star_def + by auto } + then have "X \ B ;; A\" by auto + ultimately + show "X = B ;; A\" by simp +qed + + +section {* Regular Expressions *} + +datatype rexp = + NULL +| EMPTY +| CHAR char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp + + +text {* + The following @{text "L"} is an overloaded operator, where @{text "L(x)"} evaluates to + the language represented by the syntactic object @{text "x"}. +*} + +consts L:: "'a \ lang" + +text {* The @{text "L (rexp)"} for regular expressions. *} + +overloading L_rexp \ "L:: rexp \ lang" +begin +fun + L_rexp :: "rexp \ string set" +where + "L_rexp (NULL) = {}" + | "L_rexp (EMPTY) = {[]}" + | "L_rexp (CHAR c) = {[c]}" + | "L_rexp (SEQ r1 r2) = (L_rexp r1) ;; (L_rexp r2)" + | "L_rexp (ALT r1 r2) = (L_rexp r1) \ (L_rexp r2)" + | "L_rexp (STAR r) = (L_rexp r)\" +end + + +section {* Folds for Sets *} + +text {* + To obtain equational system out of finite set of equivalence classes, a fold operation + on finite sets @{text "folds"} is defined. The use of @{text "SOME"} makes @{text "folds"} + more robust than the @{text "fold"} in the Isabelle library. The expression @{text "folds f"} + makes sense when @{text "f"} is not @{text "associative"} and @{text "commutitive"}, + while @{text "fold f"} does not. +*} + + +definition + folds :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" +where + "folds f z S \ SOME x. fold_graph f z S x" + +abbreviation + Setalt ("\_" [1000] 999) +where + "\A == folds ALT NULL A" + +text {* + The following lemma ensures that the arbitrary choice made by the + @{text "SOME"} in @{text "folds"} does not affect the @{text "L"}-value + of the resultant regular expression. +*} + +lemma folds_alt_simp [simp]: + assumes a: "finite rs" + shows "L (\rs) = \ (L ` rs)" +apply(rule set_eqI) +apply(simp add: folds_def) +apply(rule someI2_ex) +apply(rule_tac finite_imp_fold_graph[OF a]) +apply(erule fold_graph.induct) +apply(auto) +done + + +text {* Just a technical lemma for collections and pairs *} + +lemma Pair_Collect[simp]: + shows "(x, y) \ {(x, y). P x y} \ P x y" +by simp + +text {* + @{text "\A"} is an equivalence class defined by language @{text "A"}. +*} +definition + str_eq_rel :: "lang \ (string \ string) set" ("\_" [100] 100) +where + "\A \ {(x, y). (\z. x @ z \ A \ y @ z \ A)}" + +text {* + Among the equivalence clases of @{text "\A"}, the set @{text "finals A"} singles out + those which contains the strings from @{text "A"}. +*} + +definition + finals :: "lang \ lang set" +where + "finals A \ {\A `` {x} | x . x \ A}" + +text {* + The following lemma establishes the relationshipt between + @{text "finals A"} and @{text "A"}. +*} + +lemma lang_is_union_of_finals: + shows "A = \ finals A" +unfolding finals_def +unfolding Image_def +unfolding str_eq_rel_def +apply(auto) +apply(drule_tac x = "[]" in spec) +apply(auto) +done + +lemma finals_in_partitions: + shows "finals A \ (UNIV // \A)" +unfolding finals_def +unfolding quotient_def +by auto + +section {* Direction @{text "finite partition \ regular language"}*} + +text {* + The relationship between equivalent classes can be described by an + equational system. For example, in equational system \eqref{example_eqns}, + $X_0, X_1$ are equivalent classes. The first equation says every string in + $X_0$ is obtained either by appending one $b$ to a string in $X_0$ or by + appending one $a$ to a string in $X_1$ or just be an empty string + (represented by the regular expression $\lambda$). Similary, the second + equation tells how the strings inside $X_1$ are composed. + + \begin{equation}\label{example_eqns} + \begin{aligned} + X_0 & = X_0 b + X_1 a + \lambda \\ + X_1 & = X_0 a + X_1 b + \end{aligned} + \end{equation} + + \noindent + The summands on the right hand side is represented by the following data + type @{text "rhs_item"}, mnemonic for 'right hand side item'. Generally, + there are two kinds of right hand side items, one kind corresponds to pure + regular expressions, like the $\lambda$ in \eqref{example_eqns}, the other + kind corresponds to transitions from one one equivalent class to another, + like the $X_0 b, X_1 a$ etc. + +*} + +datatype rhs_item = + Lam "rexp" (* Lambda *) + | Trn "lang" "rexp" (* Transition *) + + +text {* + In this formalization, pure regular expressions like $\lambda$ is + repsented by @{text "Lam(EMPTY)"}, while transitions like $X_0 a$ is + represented by @{term "Trn X\<^isub>0 (CHAR a)"}. +*} + +text {* + Every right-hand side item @{text "itm"} defines a language given + by @{text "L(itm)"}, defined as: +*} + +overloading L_rhs_e \ "L:: rhs_item \ lang" +begin + fun L_rhs_e:: "rhs_item \ lang" + where + "L_rhs_e (Lam r) = L r" + | "L_rhs_e (Trn X r) = X ;; L r" +end + +text {* + The right hand side of every equation is represented by a set of + items. The string set defined by such a set @{text "itms"} is given + by @{text "L(itms)"}, defined as: +*} + +overloading L_rhs \ "L:: rhs_item set \ lang" +begin + fun L_rhs:: "rhs_item set \ lang" + where + "L_rhs rhs = \ (L ` rhs)" +end + +text {* + Given a set of equivalence classes @{text "CS"} and one equivalence class @{text "X"} among + @{text "CS"}, the term @{text "init_rhs CS X"} is used to extract the right hand side of + the equation describing the formation of @{text "X"}. The definition of @{text "init_rhs"} + is: +*} + +definition + transition :: "lang \ rexp \ lang \ bool" ("_ \_\_" [100,100,100] 100) +where + "Y \r\ X \ Y ;; (L r) \ X" + +definition + "init_rhs CS X \ + if ([] \ X) then + {Lam EMPTY} \ {Trn Y (CHAR c) | Y c. Y \ CS \ Y \(CHAR c)\ X} + else + {Trn Y (CHAR c)| Y c. Y \ CS \ Y \(CHAR c)\ X}" + +text {* + In the definition of @{text "init_rhs"}, the term + @{text "{Trn Y (CHAR c)| Y c. Y \ CS \ Y ;; {[c]} \ X}"} appearing on both branches + describes the formation of strings in @{text "X"} out of transitions, while + the term @{text "{Lam(EMPTY)}"} describes the empty string which is intrinsically contained in + @{text "X"} rather than by transition. This @{text "{Lam(EMPTY)}"} corresponds to + the $\lambda$ in \eqref{example_eqns}. + + With the help of @{text "init_rhs"}, the equitional system descrbing the formation of every + equivalent class inside @{text "CS"} is given by the following @{text "eqs(CS)"}. +*} + + +definition "eqs CS \ {(X, init_rhs CS X) | X. X \ CS}" + + + +(************ arden's lemma variation ********************) + +text {* + The following @{text "trns_of rhs X"} returns all @{text "X"}-items in @{text "rhs"}. +*} + +definition + "trns_of rhs X \ {Trn X r | r. Trn X r \ rhs}" + +text {* + The following @{text "attach_rexp rexp' itm"} attach + the regular expression @{text "rexp'"} to + the right of right hand side item @{text "itm"}. +*} + +fun + attach_rexp :: "rexp \ rhs_item \ rhs_item" +where + "attach_rexp rexp' (Lam rexp) = Lam (SEQ rexp rexp')" +| "attach_rexp rexp' (Trn X rexp) = Trn X (SEQ rexp rexp')" + +text {* + The following @{text "append_rhs_rexp rhs rexp"} attaches + @{text "rexp"} to every item in @{text "rhs"}. +*} + +definition + "append_rhs_rexp rhs rexp \ (attach_rexp rexp) ` rhs" + +text {* + With the help of the two functions immediately above, Ardens' + transformation on right hand side @{text "rhs"} is implemented + by the following function @{text "arden_variate X rhs"}. + After this transformation, the recursive occurence of @{text "X"} + in @{text "rhs"} will be eliminated, while the string set defined + by @{text "rhs"} is kept unchanged. +*} + +definition + "arden_variate X rhs \ + append_rhs_rexp (rhs - trns_of rhs X) (STAR (\ {r. Trn X r \ rhs}))" + + +(*********** substitution of ES *************) + +text {* + Suppose the equation defining @{text "X"} is $X = xrhs$, + the purpose of @{text "rhs_subst"} is to substitute all occurences of @{text "X"} in + @{text "rhs"} by @{text "xrhs"}. + A litte thought may reveal that the final result + should be: first append $(a_1 | a_2 | \ldots | a_n)$ to every item of @{text "xrhs"} and then + union the result with all non-@{text "X"}-items of @{text "rhs"}. + *} + +definition + "rhs_subst rhs X xrhs \ + (rhs - (trns_of rhs X)) \ (append_rhs_rexp xrhs (\ {r. Trn X r \ rhs}))" + +text {* + Suppose the equation defining @{text "X"} is $X = xrhs$, the follwing + @{text "eqs_subst ES X xrhs"} substitute @{text "xrhs"} into every equation + of the equational system @{text "ES"}. + *} + +definition + "eqs_subst ES X xrhs \ {(Y, rhs_subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \ ES}" + +text {* + The computation of regular expressions for equivalence classes is accomplished + using a iteration principle given by the following lemma. +*} + +lemma wf_iter [rule_format]: + fixes f + assumes step: "\ e. \P e; \ Q e\ \ (\ e'. P e' \ (f(e'), f(e)) \ less_than)" + shows pe: "P e \ (\ e'. P e' \ Q e')" +proof(induct e rule: wf_induct + [OF wf_inv_image[OF wf_less_than, where f = "f"]], clarify) + fix x + assume h [rule_format]: + "\y. (y, x) \ inv_image less_than f \ P y \ (\e'. P e' \ Q e')" + and px: "P x" + show "\e'. P e' \ Q e'" + proof(cases "Q x") + assume "Q x" with px show ?thesis by blast + next + assume nq: "\ Q x" + from step [OF px nq] + obtain e' where pe': "P e'" and ltf: "(f e', f x) \ less_than" by auto + show ?thesis + proof(rule h) + from ltf show "(e', x) \ inv_image less_than f" + by (simp add:inv_image_def) + next + from pe' show "P e'" . + qed + qed +qed + +text {* + The @{text "P"} in lemma @{text "wf_iter"} is an invariant kept throughout the iteration procedure. + The particular invariant used to solve our problem is defined by function @{text "Inv(ES)"}, + an invariant over equal system @{text "ES"}. + Every definition starting next till @{text "Inv"} stipulates a property to be satisfied by @{text "ES"}. +*} + +text {* + Every variable is defined at most onece in @{text "ES"}. + *} + +definition + "distinct_equas ES \ + \ X rhs rhs'. (X, rhs) \ ES \ (X, rhs') \ ES \ rhs = rhs'" + +text {* + Every equation in @{text "ES"} (represented by @{text "(X, rhs)"}) is valid, i.e. @{text "(X = L rhs)"}. + *} +definition + "valid_eqns ES \ \ X rhs. (X, rhs) \ ES \ (X = L rhs)" + +text {* + The following @{text "rhs_nonempty rhs"} requires regular expressions occuring in transitional + items of @{text "rhs"} does not contain empty string. This is necessary for + the application of Arden's transformation to @{text "rhs"}. + *} + +definition + "rhs_nonempty rhs \ (\ Y r. Trn Y r \ rhs \ [] \ L r)" + +text {* + The following @{text "ardenable ES"} requires that Arden's transformation is applicable + to every equation of equational system @{text "ES"}. + *} + +definition + "ardenable ES \ \ X rhs. (X, rhs) \ ES \ rhs_nonempty rhs" + +(* The following non_empty seems useless. *) +definition + "non_empty ES \ \ X rhs. (X, rhs) \ ES \ X \ {}" + +text {* + The following @{text "finite_rhs ES"} requires every equation in @{text "rhs"} be finite. + *} +definition + "finite_rhs ES \ \ X rhs. (X, rhs) \ ES \ finite rhs" + +text {* + The following @{text "classes_of rhs"} returns all variables (or equivalent classes) + occuring in @{text "rhs"}. + *} +definition + "classes_of rhs \ {X. \ r. Trn X r \ rhs}" + +text {* + The following @{text "lefts_of ES"} returns all variables + defined by equational system @{text "ES"}. + *} +definition + "lefts_of ES \ {Y | Y yrhs. (Y, yrhs) \ ES}" + +text {* + The following @{text "self_contained ES"} requires that every + variable occuring on the right hand side of equations is already defined by some + equation in @{text "ES"}. + *} +definition + "self_contained ES \ \ (X, xrhs) \ ES. classes_of xrhs \ lefts_of ES" + + +text {* + The invariant @{text "Inv(ES)"} is a conjunction of all the previously defined constaints. + *} +definition + "Inv ES \ valid_eqns ES \ finite ES \ distinct_equas ES \ ardenable ES \ + non_empty ES \ finite_rhs ES \ self_contained ES" + +subsection {* The proof of this direction *} + +subsubsection {* Basic properties *} + +text {* + The following are some basic properties of the above definitions. +*} + +lemma L_rhs_union_distrib: + fixes A B::"rhs_item set" + shows "L A \ L B = L (A \ B)" +by simp + +lemma finite_Trn: + assumes fin: "finite rhs" + shows "finite {r. Trn Y r \ rhs}" +proof - + have "finite {Trn Y r | Y r. Trn Y r \ rhs}" + by (rule rev_finite_subset[OF fin]) (auto) + then have "finite ((\(Y, r). Trn Y r) ` {(Y, r) | Y r. Trn Y r \ rhs})" + by (simp add: image_Collect) + then have "finite {(Y, r) | Y r. Trn Y r \ rhs}" + by (erule_tac finite_imageD) (simp add: inj_on_def) + then show "finite {r. Trn Y r \ rhs}" + by (erule_tac f="snd" in finite_surj) (auto simp add: image_def) +qed + +lemma finite_Lam: + assumes fin:"finite rhs" + shows "finite {r. Lam r \ rhs}" +proof - + have "finite {Lam r | r. Lam r \ rhs}" + by (rule rev_finite_subset[OF fin]) (auto) + then show "finite {r. Lam r \ rhs}" + apply(simp add: image_Collect[symmetric]) + apply(erule finite_imageD) + apply(auto simp add: inj_on_def) + done +qed + +lemma rexp_of_empty: + assumes finite:"finite rhs" + and nonempty:"rhs_nonempty rhs" + shows "[] \ L (\ {r. Trn X r \ rhs})" +using finite nonempty rhs_nonempty_def +using finite_Trn[OF finite] +by (auto) + +lemma [intro!]: + "P (Trn X r) \ (\a. (\r. a = Trn X r \ P a))" by auto + +lemma lang_of_rexp_of: + assumes finite:"finite rhs" + shows "L ({Trn X r| r. Trn X r \ rhs}) = X ;; (L (\{r. Trn X r \ rhs}))" +proof - + have "finite {r. Trn X r \ rhs}" + by (rule finite_Trn[OF finite]) + then show ?thesis + apply(auto simp add: Seq_def) + apply(rule_tac x = "s\<^isub>1" in exI, rule_tac x = "s\<^isub>2" in exI, auto) + apply(rule_tac x= "Trn X xa" in exI) + apply(auto simp: Seq_def) + done +qed + +lemma rexp_of_lam_eq_lam_set: + assumes fin: "finite rhs" + shows "L (\{r. Lam r \ rhs}) = L ({Lam r | r. Lam r \ rhs})" +proof - + have "finite ({r. Lam r \ rhs})" using fin by (rule finite_Lam) + then show ?thesis by auto +qed + +lemma [simp]: + "L (attach_rexp r xb) = L xb ;; L r" +apply (cases xb, auto simp: Seq_def) +apply(rule_tac x = "s\<^isub>1 @ s\<^isub>1'" in exI, rule_tac x = "s\<^isub>2'" in exI) +apply(auto simp: Seq_def) +done + +lemma lang_of_append_rhs: + "L (append_rhs_rexp rhs r) = L rhs ;; L r" +apply (auto simp:append_rhs_rexp_def image_def) +apply (auto simp:Seq_def) +apply (rule_tac x = "L xb ;; L r" in exI, auto simp add:Seq_def) +by (rule_tac x = "attach_rexp r xb" in exI, auto simp:Seq_def) + +lemma classes_of_union_distrib: + "classes_of A \ classes_of B = classes_of (A \ B)" +by (auto simp add:classes_of_def) + +lemma lefts_of_union_distrib: + "lefts_of A \ lefts_of B = lefts_of (A \ B)" +by (auto simp:lefts_of_def) + + +subsubsection {* Intialization *} + +text {* + The following several lemmas until @{text "init_ES_satisfy_Inv"} shows that + the initial equational system satisfies invariant @{text "Inv"}. +*} + +lemma defined_by_str: + "\s \ X; X \ UNIV // (\Lang)\ \ X = (\Lang) `` {s}" +by (auto simp:quotient_def Image_def str_eq_rel_def) + +lemma every_eqclass_has_transition: + assumes has_str: "s @ [c] \ X" + and in_CS: "X \ UNIV // (\Lang)" + obtains Y where "Y \ UNIV // (\Lang)" and "Y ;; {[c]} \ X" and "s \ Y" +proof - + def Y \ "(\Lang) `` {s}" + have "Y \ UNIV // (\Lang)" + unfolding Y_def quotient_def by auto + moreover + have "X = (\Lang) `` {s @ [c]}" + using has_str in_CS defined_by_str by blast + then have "Y ;; {[c]} \ X" + unfolding Y_def Image_def Seq_def + unfolding str_eq_rel_def + by clarsimp + moreover + have "s \ Y" unfolding Y_def + unfolding Image_def str_eq_rel_def by simp + ultimately show thesis by (blast intro: that) +qed + +lemma l_eq_r_in_eqs: + assumes X_in_eqs: "(X, xrhs) \ (eqs (UNIV // (\Lang)))" + shows "X = L xrhs" +proof + show "X \ L xrhs" + proof + fix x + assume "(1)": "x \ X" + show "x \ L xrhs" + proof (cases "x = []") + assume empty: "x = []" + thus ?thesis using X_in_eqs "(1)" + by (auto simp:eqs_def init_rhs_def) + next + assume not_empty: "x \ []" + then obtain clist c where decom: "x = clist @ [c]" + by (case_tac x rule:rev_cases, auto) + have "X \ UNIV // (\Lang)" using X_in_eqs by (auto simp:eqs_def) + then obtain Y + where "Y \ UNIV // (\Lang)" + and "Y ;; {[c]} \ X" + and "clist \ Y" + using decom "(1)" every_eqclass_has_transition by blast + hence + "x \ L {Trn Y (CHAR c)| Y c. Y \ UNIV // (\Lang) \ Y \(CHAR c)\ X}" + unfolding transition_def + using "(1)" decom + by (simp, rule_tac x = "Trn Y (CHAR c)" in exI, simp add:Seq_def) + thus ?thesis using X_in_eqs "(1)" + by (simp add: eqs_def init_rhs_def) + qed + qed +next + show "L xrhs \ X" using X_in_eqs + by (auto simp:eqs_def init_rhs_def transition_def) +qed + +lemma finite_init_rhs: + assumes finite: "finite CS" + shows "finite (init_rhs CS X)" +proof- + have "finite {Trn Y (CHAR c) |Y c. Y \ CS \ Y ;; {[c]} \ X}" (is "finite ?A") + proof - + def S \ "{(Y, c)| Y c. Y \ CS \ Y ;; {[c]} \ X}" + def h \ "\ (Y, c). Trn Y (CHAR c)" + have "finite (CS \ (UNIV::char set))" using finite by auto + hence "finite S" using S_def + by (rule_tac B = "CS \ UNIV" in finite_subset, auto) + moreover have "?A = h ` S" by (auto simp: S_def h_def image_def) + ultimately show ?thesis + by auto + qed + thus ?thesis by (simp add:init_rhs_def transition_def) +qed + +lemma init_ES_satisfy_Inv: + assumes finite_CS: "finite (UNIV // (\Lang))" + shows "Inv (eqs (UNIV // (\Lang)))" +proof - + have "finite (eqs (UNIV // (\Lang)))" using finite_CS + by (simp add:eqs_def) + moreover have "distinct_equas (eqs (UNIV // (\Lang)))" + by (simp add:distinct_equas_def eqs_def) + moreover have "ardenable (eqs (UNIV // (\Lang)))" + by (auto simp add:ardenable_def eqs_def init_rhs_def rhs_nonempty_def del:L_rhs.simps) + moreover have "valid_eqns (eqs (UNIV // (\Lang)))" + using l_eq_r_in_eqs by (simp add:valid_eqns_def) + moreover have "non_empty (eqs (UNIV // (\Lang)))" + by (auto simp:non_empty_def eqs_def quotient_def Image_def str_eq_rel_def) + moreover have "finite_rhs (eqs (UNIV // (\Lang)))" + using finite_init_rhs[OF finite_CS] + by (auto simp:finite_rhs_def eqs_def) + moreover have "self_contained (eqs (UNIV // (\Lang)))" + by (auto simp:self_contained_def eqs_def init_rhs_def classes_of_def lefts_of_def) + ultimately show ?thesis by (simp add:Inv_def) +qed + +subsubsection {* + Interation step + *} + +text {* + From this point until @{text "iteration_step"}, it is proved + that there exists iteration steps which keep @{text "Inv(ES)"} while + decreasing the size of @{text "ES"}. +*} + +lemma arden_variate_keeps_eq: + assumes l_eq_r: "X = L rhs" + and not_empty: "[] \ L (\{r. Trn X r \ rhs})" + and finite: "finite rhs" + shows "X = L (arden_variate X rhs)" +proof - + def A \ "L (\{r. Trn X r \ rhs})" + def b \ "rhs - trns_of rhs X" + def B \ "L b" + have "X = B ;; A\" + proof- + have "L rhs = L(trns_of rhs X \ b)" by (auto simp: b_def trns_of_def) + also have "\ = X ;; A \ B" + unfolding trns_of_def + unfolding L_rhs_union_distrib[symmetric] + by (simp only: lang_of_rexp_of finite B_def A_def) + finally show ?thesis + using l_eq_r not_empty + apply(rule_tac ardens_revised[THEN iffD1]) + apply(simp add: A_def) + apply(simp) + done + qed + moreover have "L (arden_variate X rhs) = (B ;; A\)" + by (simp only:arden_variate_def L_rhs_union_distrib lang_of_append_rhs + B_def A_def b_def L_rexp.simps seq_union_distrib_left) + ultimately show ?thesis by simp +qed + +lemma append_keeps_finite: + "finite rhs \ finite (append_rhs_rexp rhs r)" +by (auto simp:append_rhs_rexp_def) + +lemma arden_variate_keeps_finite: + "finite rhs \ finite (arden_variate X rhs)" +by (auto simp:arden_variate_def append_keeps_finite) + +lemma append_keeps_nonempty: + "rhs_nonempty rhs \ rhs_nonempty (append_rhs_rexp rhs r)" +apply (auto simp:rhs_nonempty_def append_rhs_rexp_def) +by (case_tac x, auto simp:Seq_def) + +lemma nonempty_set_sub: + "rhs_nonempty rhs \ rhs_nonempty (rhs - A)" +by (auto simp:rhs_nonempty_def) + +lemma nonempty_set_union: + "\rhs_nonempty rhs; rhs_nonempty rhs'\ \ rhs_nonempty (rhs \ rhs')" +by (auto simp:rhs_nonempty_def) + +lemma arden_variate_keeps_nonempty: + "rhs_nonempty rhs \ rhs_nonempty (arden_variate X rhs)" +by (simp only:arden_variate_def append_keeps_nonempty nonempty_set_sub) + + +lemma rhs_subst_keeps_nonempty: + "\rhs_nonempty rhs; rhs_nonempty xrhs\ \ rhs_nonempty (rhs_subst rhs X xrhs)" +by (simp only:rhs_subst_def append_keeps_nonempty nonempty_set_union nonempty_set_sub) + +lemma rhs_subst_keeps_eq: + assumes substor: "X = L xrhs" + and finite: "finite rhs" + shows "L (rhs_subst rhs X xrhs) = L rhs" (is "?Left = ?Right") +proof- + def A \ "L (rhs - trns_of rhs X)" + have "?Left = A \ L (append_rhs_rexp xrhs (\{r. Trn X r \ rhs}))" + unfolding rhs_subst_def + unfolding L_rhs_union_distrib[symmetric] + by (simp add: A_def) + moreover have "?Right = A \ L ({Trn X r | r. Trn X r \ rhs})" + proof- + have "rhs = (rhs - trns_of rhs X) \ (trns_of rhs X)" by (auto simp add: trns_of_def) + thus ?thesis + unfolding A_def + unfolding L_rhs_union_distrib + unfolding trns_of_def + by simp + qed + moreover have "L (append_rhs_rexp xrhs (\{r. Trn X r \ rhs})) = L ({Trn X r | r. Trn X r \ rhs})" + using finite substor by (simp only:lang_of_append_rhs lang_of_rexp_of) + ultimately show ?thesis by simp +qed + +lemma rhs_subst_keeps_finite_rhs: + "\finite rhs; finite yrhs\ \ finite (rhs_subst rhs Y yrhs)" +by (auto simp:rhs_subst_def append_keeps_finite) + +lemma eqs_subst_keeps_finite: + assumes finite:"finite (ES:: (string set \ rhs_item set) set)" + shows "finite (eqs_subst ES Y yrhs)" +proof - + have "finite {(Ya, rhs_subst yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \ ES}" + (is "finite ?A") + proof- + def eqns' \ "{((Ya::string set), yrhsa)| Ya yrhsa. (Ya, yrhsa) \ ES}" + def h \ "\ ((Ya::string set), yrhsa). (Ya, rhs_subst yrhsa Y yrhs)" + have "finite (h ` eqns')" using finite h_def eqns'_def by auto + moreover have "?A = h ` eqns'" by (auto simp:h_def eqns'_def) + ultimately show ?thesis by auto + qed + thus ?thesis by (simp add:eqs_subst_def) +qed + +lemma eqs_subst_keeps_finite_rhs: + "\finite_rhs ES; finite yrhs\ \ finite_rhs (eqs_subst ES Y yrhs)" +by (auto intro:rhs_subst_keeps_finite_rhs simp add:eqs_subst_def finite_rhs_def) + +lemma append_rhs_keeps_cls: + "classes_of (append_rhs_rexp rhs r) = classes_of rhs" +apply (auto simp:classes_of_def append_rhs_rexp_def) +apply (case_tac xa, auto simp:image_def) +by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "Trn x ra" in bexI, simp+) + +lemma arden_variate_removes_cl: + "classes_of (arden_variate Y yrhs) = classes_of yrhs - {Y}" +apply (simp add:arden_variate_def append_rhs_keeps_cls trns_of_def) +by (auto simp:classes_of_def) + +lemma lefts_of_keeps_cls: + "lefts_of (eqs_subst ES Y yrhs) = lefts_of ES" +by (auto simp:lefts_of_def eqs_subst_def) + +lemma rhs_subst_updates_cls: + "X \ classes_of xrhs \ + classes_of (rhs_subst rhs X xrhs) = classes_of rhs \ classes_of xrhs - {X}" +apply (simp only:rhs_subst_def append_rhs_keeps_cls + classes_of_union_distrib[THEN sym]) +by (auto simp:classes_of_def trns_of_def) + +lemma eqs_subst_keeps_self_contained: + fixes Y + assumes sc: "self_contained (ES \ {(Y, yrhs)})" (is "self_contained ?A") + shows "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" + (is "self_contained ?B") +proof- + { fix X xrhs' + assume "(X, xrhs') \ ?B" + then obtain xrhs + where xrhs_xrhs': "xrhs' = rhs_subst xrhs Y (arden_variate Y yrhs)" + and X_in: "(X, xrhs) \ ES" by (simp add:eqs_subst_def, blast) + have "classes_of xrhs' \ lefts_of ?B" + proof- + have "lefts_of ?B = lefts_of ES" by (auto simp add:lefts_of_def eqs_subst_def) + moreover have "classes_of xrhs' \ lefts_of ES" + proof- + have "classes_of xrhs' \ + classes_of xrhs \ classes_of (arden_variate Y yrhs) - {Y}" + proof- + have "Y \ classes_of (arden_variate Y yrhs)" + using arden_variate_removes_cl by simp + thus ?thesis using xrhs_xrhs' by (auto simp:rhs_subst_updates_cls) + qed + moreover have "classes_of xrhs \ lefts_of ES \ {Y}" using X_in sc + apply (simp only:self_contained_def lefts_of_union_distrib[THEN sym]) + by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lefts_of_def) + moreover have "classes_of (arden_variate Y yrhs) \ lefts_of ES \ {Y}" + using sc + by (auto simp add:arden_variate_removes_cl self_contained_def lefts_of_def) + ultimately show ?thesis by auto + qed + ultimately show ?thesis by simp + qed + } thus ?thesis by (auto simp only:eqs_subst_def self_contained_def) +qed + +lemma eqs_subst_satisfy_Inv: + assumes Inv_ES: "Inv (ES \ {(Y, yrhs)})" + shows "Inv (eqs_subst ES Y (arden_variate Y yrhs))" +proof - + have finite_yrhs: "finite yrhs" + using Inv_ES by (auto simp:Inv_def finite_rhs_def) + have nonempty_yrhs: "rhs_nonempty yrhs" + using Inv_ES by (auto simp:Inv_def ardenable_def) + have Y_eq_yrhs: "Y = L yrhs" + using Inv_ES by (simp only:Inv_def valid_eqns_def, blast) + have "distinct_equas (eqs_subst ES Y (arden_variate Y yrhs))" + using Inv_ES + by (auto simp:distinct_equas_def eqs_subst_def Inv_def) + moreover have "finite (eqs_subst ES Y (arden_variate Y yrhs))" + using Inv_ES by (simp add:Inv_def eqs_subst_keeps_finite) + moreover have "finite_rhs (eqs_subst ES Y (arden_variate Y yrhs))" + proof- + have "finite_rhs ES" using Inv_ES + by (simp add:Inv_def finite_rhs_def) + moreover have "finite (arden_variate Y yrhs)" + proof - + have "finite yrhs" using Inv_ES + by (auto simp:Inv_def finite_rhs_def) + thus ?thesis using arden_variate_keeps_finite by simp + qed + ultimately show ?thesis + by (simp add:eqs_subst_keeps_finite_rhs) + qed + moreover have "ardenable (eqs_subst ES Y (arden_variate Y yrhs))" + proof - + { fix X rhs + assume "(X, rhs) \ ES" + hence "rhs_nonempty rhs" using prems Inv_ES + by (simp add:Inv_def ardenable_def) + with nonempty_yrhs + have "rhs_nonempty (rhs_subst rhs Y (arden_variate Y yrhs))" + by (simp add:nonempty_yrhs + rhs_subst_keeps_nonempty arden_variate_keeps_nonempty) + } thus ?thesis by (auto simp add:ardenable_def eqs_subst_def) + qed + moreover have "valid_eqns (eqs_subst ES Y (arden_variate Y yrhs))" + proof- + have "Y = L (arden_variate Y yrhs)" + using Y_eq_yrhs Inv_ES finite_yrhs nonempty_yrhs + by (rule_tac arden_variate_keeps_eq, (simp add:rexp_of_empty)+) + thus ?thesis using Inv_ES + by (clarsimp simp add:valid_eqns_def + eqs_subst_def rhs_subst_keeps_eq Inv_def finite_rhs_def + simp del:L_rhs.simps) + qed + moreover have + non_empty_subst: "non_empty (eqs_subst ES Y (arden_variate Y yrhs))" + using Inv_ES by (auto simp:Inv_def non_empty_def eqs_subst_def) + moreover + have self_subst: "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" + using Inv_ES eqs_subst_keeps_self_contained by (simp add:Inv_def) + ultimately show ?thesis using Inv_ES by (simp add:Inv_def) +qed + +lemma eqs_subst_card_le: + assumes finite: "finite (ES::(string set \ rhs_item set) set)" + shows "card (eqs_subst ES Y yrhs) <= card ES" +proof- + def f \ "\ x. ((fst x)::string set, rhs_subst (snd x) Y yrhs)" + have "eqs_subst ES Y yrhs = f ` ES" + apply (auto simp:eqs_subst_def f_def image_def) + by (rule_tac x = "(Ya, yrhsa)" in bexI, simp+) + thus ?thesis using finite by (auto intro:card_image_le) +qed + +lemma eqs_subst_cls_remains: + "(X, xrhs) \ ES \ \ xrhs'. (X, xrhs') \ (eqs_subst ES Y yrhs)" +by (auto simp:eqs_subst_def) + +lemma card_noteq_1_has_more: + assumes card:"card S \ 1" + and e_in: "e \ S" + and finite: "finite S" + obtains e' where "e' \ S \ e \ e'" +proof- + have "card (S - {e}) > 0" + proof - + have "card S > 1" using card e_in finite + by (case_tac "card S", auto) + thus ?thesis using finite e_in by auto + qed + hence "S - {e} \ {}" using finite by (rule_tac notI, simp) + thus "(\e'. e' \ S \ e \ e' \ thesis) \ thesis" by auto +qed + +lemma iteration_step: + assumes Inv_ES: "Inv ES" + and X_in_ES: "(X, xrhs) \ ES" + and not_T: "card ES \ 1" + shows "\ ES'. (Inv ES' \ (\ xrhs'.(X, xrhs') \ ES')) \ + (card ES', card ES) \ less_than" (is "\ ES'. ?P ES'") +proof - + have finite_ES: "finite ES" using Inv_ES by (simp add:Inv_def) + then obtain Y yrhs + where Y_in_ES: "(Y, yrhs) \ ES" and not_eq: "(X, xrhs) \ (Y, yrhs)" + using not_T X_in_ES by (drule_tac card_noteq_1_has_more, auto) + def ES' == "ES - {(Y, yrhs)}" + let ?ES'' = "eqs_subst ES' Y (arden_variate Y yrhs)" + have "?P ?ES''" + proof - + have "Inv ?ES''" using Y_in_ES Inv_ES + by (rule_tac eqs_subst_satisfy_Inv, simp add:ES'_def insert_absorb) + moreover have "\xrhs'. (X, xrhs') \ ?ES''" using not_eq X_in_ES + by (rule_tac ES = ES' in eqs_subst_cls_remains, auto simp add:ES'_def) + moreover have "(card ?ES'', card ES) \ less_than" + proof - + have "finite ES'" using finite_ES ES'_def by auto + moreover have "card ES' < card ES" using finite_ES Y_in_ES + by (auto simp:ES'_def card_gt_0_iff intro:diff_Suc_less) + ultimately show ?thesis + by (auto dest:eqs_subst_card_le elim:le_less_trans) + qed + ultimately show ?thesis by simp + qed + thus ?thesis by blast +qed + +subsubsection {* + Conclusion of the proof + *} + +text {* + From this point until @{text "hard_direction"}, the hard direction is proved + through a simple application of the iteration principle. +*} + +lemma iteration_conc: + assumes history: "Inv ES" + and X_in_ES: "\ xrhs. (X, xrhs) \ ES" + shows + "\ ES'. (Inv ES' \ (\ xrhs'. (X, xrhs') \ ES')) \ card ES' = 1" + (is "\ ES'. ?P ES'") +proof (cases "card ES = 1") + case True + thus ?thesis using history X_in_ES + by blast +next + case False + thus ?thesis using history iteration_step X_in_ES + by (rule_tac f = card in wf_iter, auto) +qed + +lemma last_cl_exists_rexp: + assumes ES_single: "ES = {(X, xrhs)}" + and Inv_ES: "Inv ES" + shows "\ (r::rexp). L r = X" (is "\ r. ?P r") +proof- + def A \ "arden_variate X xrhs" + have "?P (\{r. Lam r \ A})" + proof - + have "L (\{r. Lam r \ A}) = L ({Lam r | r. Lam r \ A})" + proof(rule rexp_of_lam_eq_lam_set) + show "finite A" + unfolding A_def + using Inv_ES ES_single + by (rule_tac arden_variate_keeps_finite) + (auto simp add: Inv_def finite_rhs_def) + qed + also have "\ = L A" + proof- + have "{Lam r | r. Lam r \ A} = A" + proof- + have "classes_of A = {}" using Inv_ES ES_single + unfolding A_def + by (simp add:arden_variate_removes_cl + self_contained_def Inv_def lefts_of_def) + thus ?thesis + unfolding A_def + by (auto simp only: classes_of_def, case_tac x, auto) + qed + thus ?thesis by simp + qed + also have "\ = X" + unfolding A_def + proof(rule arden_variate_keeps_eq [THEN sym]) + show "X = L xrhs" using Inv_ES ES_single + by (auto simp only:Inv_def valid_eqns_def) + next + from Inv_ES ES_single show "[] \ L (\{r. Trn X r \ xrhs})" + by(simp add:Inv_def ardenable_def rexp_of_empty finite_rhs_def) + next + from Inv_ES ES_single show "finite xrhs" + by (simp add:Inv_def finite_rhs_def) + qed + finally show ?thesis by simp + qed + thus ?thesis by auto +qed + +lemma every_eqcl_has_reg: + assumes finite_CS: "finite (UNIV // (\Lang))" + and X_in_CS: "X \ (UNIV // (\Lang))" + shows "\ (reg::rexp). L reg = X" (is "\ r. ?E r") +proof - + from X_in_CS have "\ xrhs. (X, xrhs) \ (eqs (UNIV // (\Lang)))" + by (auto simp:eqs_def init_rhs_def) + then obtain ES xrhs where Inv_ES: "Inv ES" + and X_in_ES: "(X, xrhs) \ ES" + and card_ES: "card ES = 1" + using finite_CS X_in_CS init_ES_satisfy_Inv iteration_conc + by blast + hence ES_single_equa: "ES = {(X, xrhs)}" + by (auto simp:Inv_def dest!:card_Suc_Diff1 simp:card_eq_0_iff) + thus ?thesis using Inv_ES + by (rule last_cl_exists_rexp) +qed + +theorem hard_direction: + assumes finite_CS: "finite (UNIV // \A)" + shows "\r::rexp. A = L r" +proof - + have "\ X \ (UNIV // \A). \reg::rexp. X = L reg" + using finite_CS every_eqcl_has_reg by blast + then obtain f + where f_prop: "\ X \ (UNIV // \A). X = L ((f X)::rexp)" + by (auto dest: bchoice) + def rs \ "f ` (finals A)" + have "A = \ (finals A)" using lang_is_union_of_finals by auto + also have "\ = L (\rs)" + proof - + have "finite rs" + proof - + have "finite (finals A)" + using finite_CS finals_in_partitions[of "A"] + by (erule_tac finite_subset, simp) + thus ?thesis using rs_def by auto + qed + thus ?thesis + using f_prop rs_def finals_in_partitions[of "A"] by auto + qed + finally show ?thesis by blast +qed + +end \ No newline at end of file