# HG changeset patch # User urbanc # Date 1300882650 0 # Node ID e122cb146ecc5588fd27f76e32b1756e4e5e1330 # Parent 3b7477db346207bce9ff0bdebe93d2d4c91c6e1b added the most current versions of the theories. diff -r 3b7477db3462 -r e122cb146ecc Myhill_1.thy --- a/Myhill_1.thy Tue Mar 15 15:53:22 2011 +0000 +++ b/Myhill_1.thy Wed Mar 23 12:17:30 2011 +0000 @@ -1,5 +1,6 @@ theory Myhill_1 -imports Main Folds While_Combinator +imports Main Folds + "~~/src/HOL/Library/While_Combinator" begin section {* Preliminary definitions *} diff -r 3b7477db3462 -r e122cb146ecc Myhill_2.thy --- a/Myhill_2.thy Tue Mar 15 15:53:22 2011 +0000 +++ b/Myhill_2.thy Wed Mar 23 12:17:30 2011 +0000 @@ -1,5 +1,7 @@ theory Myhill_2 - imports Myhill_1 List_Prefix Prefix_subtract + imports Myhill_1 + Prefix_subtract + "~~/src/HOL/Library/List_Prefix" begin section {* Direction @{text "regular language \finite partition"} *} diff -r 3b7477db3462 -r e122cb146ecc Paper/Paper.thy --- a/Paper/Paper.thy Tue Mar 15 15:53:22 2011 +0000 +++ b/Paper/Paper.thy Wed Mar 23 12:17:30 2011 +0000 @@ -1,6 +1,6 @@ (*<*) theory Paper -imports "../Myhill" "LaTeXsugar" +imports "../Myhill_2" "~~/src/HOL/Library/LaTeXsugar" begin declare [[show_question_marks = false]] @@ -878,7 +878,7 @@ \begin{proof}[Base Cases] The cases for @{const NULL}, @{const EMPTY} and @{const CHAR} are routine, because - we can easily establish + we can easily establish that \begin{center} \begin{tabular}{l} @@ -937,7 +937,7 @@ finite by assumption. Now @{term "f (UNIV // =tag=)"} is a subset of @{text "range f"}, and so also finite. Injectivity amounts to showing that @{text "X = Y"} under the assumptions that @{text "X, Y \ "}~@{term "UNIV // =tag="} and @{text "f X = f Y"}. - From the assumption we can obtain @{text "x \ X"} and @{text "y \ Y"} with + From the assumptions we can obtain @{text "x \ X"} and @{text "y \ Y"} with @{text "tag x = tag y"}. Since @{text x} and @{text y} are tag-related, this in turn means that the equivalence classes @{text X} and @{text Y} must be equal.\qed @@ -1318,7 +1318,7 @@ Our formalisation consists of 780 lines of Isabelle/Isar code for the first - direction and 475 for the second, plus around 300 lines of standard material about + direction and 460 for the second, plus around 300 lines of standard material about regular languages. While this might be seen as too large to count as a concise proof pearl, this should be seen in the context of the work done by Constable at al \cite{Constable00} who formalised the Myhill-Nerode theorem diff -r 3b7477db3462 -r e122cb146ecc Paper/ROOT.ML --- a/Paper/ROOT.ML Tue Mar 15 15:53:22 2011 +0000 +++ b/Paper/ROOT.ML Wed Mar 23 12:17:30 2011 +0000 @@ -1,4 +1,4 @@ no_document use_thy "../Myhill"; -no_document use_thy "LaTeXsugar"; +no_document use_thy "~~/src/HOL/Library/LaTeXsugar"; use_thy "Paper" \ No newline at end of file diff -r 3b7477db3462 -r e122cb146ecc Prefix_subtract.thy --- a/Prefix_subtract.thy Tue Mar 15 15:53:22 2011 +0000 +++ b/Prefix_subtract.thy Wed Mar 23 12:17:30 2011 +0000 @@ -1,5 +1,6 @@ theory Prefix_subtract - imports Main List_Prefix + imports Main + "~~/src/HOL/Library/List_Prefix" begin section {* A small theory of prefix subtraction *} diff -r 3b7477db3462 -r e122cb146ecc Theories/Closure.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Theories/Closure.thy Wed Mar 23 12:17:30 2011 +0000 @@ -0,0 +1,140 @@ +theory Closure +imports Myhill_2 +begin + +section {* Closure properties of regular languages *} + +abbreviation + regular :: "lang \ bool" +where + "regular A \ \r::rexp. A = L r" + + +lemma closure_union[intro]: + assumes "regular A" "regular B" + shows "regular (A \ B)" +proof - + from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto + then have "A \ B = L (ALT r1 r2)" by simp + then show "regular (A \ B)" by blast +qed + +lemma closure_seq[intro]: + assumes "regular A" "regular B" + shows "regular (A ;; B)" +proof - + from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto + then have "A ;; B = L (SEQ r1 r2)" by simp + then show "regular (A ;; B)" by blast +qed + +lemma closure_star[intro]: + assumes "regular A" + shows "regular (A\)" +proof - + from assms obtain r::rexp where "L r = A" by auto + then have "A\ = L (STAR r)" by simp + then show "regular (A\)" by blast +qed + +lemma closure_complement[intro]: + assumes "regular A" + shows "regular (- A)" +proof - + from assms have "finite (UNIV // \A)" by (simp add: Myhill_Nerode) + then have "finite (UNIV // \(-A))" by (simp add: str_eq_rel_def) + then show "regular (- A)" by (simp add: Myhill_Nerode) +qed + +lemma closure_difference[intro]: + assumes "regular A" "regular B" + shows "regular (A - B)" +proof - + have "A - B = - (- A \ B)" by blast + moreover + have "regular (- (- A \ B))" + using assms by blast + ultimately show "regular (A - B)" by simp +qed + +lemma closure_intersection[intro]: + assumes "regular A" "regular B" + shows "regular (A \ B)" +proof - + have "A \ B = - (- A \ - B)" by blast + moreover + have "regular (- (- A \ - B))" + using assms by blast + ultimately show "regular (A \ B)" by simp +qed + + +text {* closure under string reversal *} + +fun + Rev :: "rexp \ rexp" +where + "Rev NULL = NULL" +| "Rev EMPTY = EMPTY" +| "Rev (CHAR c) = CHAR c" +| "Rev (ALT r1 r2) = ALT (Rev r1) (Rev r2)" +| "Rev (SEQ r1 r2) = SEQ (Rev r2) (Rev r1)" +| "Rev (STAR r) = STAR (Rev r)" + +lemma rev_Seq: + "(rev ` A) ;; (rev ` B) = rev ` (B ;; A)" +unfolding Seq_def image_def +apply(auto) +apply(rule_tac x="xb @ xa" in exI) +apply(auto) +done + +lemma rev_Star1: + assumes a: "s \ (rev ` A)\" + shows "s \ rev ` (A\)" +using a +proof(induct rule: star_induct) + case (step s1 s2) + have inj: "inj (rev::string \ string)" unfolding inj_on_def by auto + have "s1 \ rev ` A" "s2 \ rev ` (A\)" by fact+ + then obtain x1 x2 where "x1 \ A" "x2 \ A\" and eqs: "s1 = rev x1" "s2 = rev x2" by auto + then have "x1 \ A\" "x2 \ A\" by (auto intro: star_intro2) + then have "x2 @ x1 \ A\" by (auto intro: star_intro1) + then have "rev (x2 @ x1) \ rev ` A\" using inj by (simp only: inj_image_mem_iff) + then show "s1 @ s2 \ rev ` A\" using eqs by simp +qed (auto) + +lemma rev_Star2: + assumes a: "s \ A\" + shows "rev s \ (rev ` A)\" +using a +proof(induct rule: star_induct) + case (step s1 s2) + have inj: "inj (rev::string \ string)" unfolding inj_on_def by auto + have "s1 \ A"by fact + then have "rev s1 \ rev ` A" using inj by (simp only: inj_image_mem_iff) + then have "rev s1 \ (rev ` A)\" by (auto intro: star_intro2) + moreover + have "rev s2 \ (rev ` A)\" by fact + ultimately show "rev (s1 @ s2) \ (rev ` A)\" by (auto intro: star_intro1) +qed (auto) + +lemma rev_Star: + "(rev ` A)\ = rev ` (A\)" +using rev_Star1 rev_Star2 by auto + +lemma rev_lang: + "L (Rev r) = rev ` (L r)" +by (induct r) (simp_all add: rev_Star rev_Seq image_Un) + +lemma closure_reversal[intro]: + assumes "regular A" + shows "regular (rev ` A)" +proof - + from assms obtain r::rexp where "L r = A" by auto + then have "L (Rev r) = rev ` A" by (simp add: rev_lang) + then show "regular (rev` A)" by blast +qed + + +end \ No newline at end of file diff -r 3b7477db3462 -r e122cb146ecc Theories/Derivs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Theories/Derivs.thy Wed Mar 23 12:17:30 2011 +0000 @@ -0,0 +1,448 @@ +theory Derivs +imports Closure +begin + +section {* Experiments with Derivatives -- independent of Myhill-Nerode *} + +subsection {* Left-Quotients *} + +definition + Delta :: "lang \ lang" +where + "Delta A = (if [] \ A then {[]} else {})" + +definition + Der :: "char \ lang \ lang" +where + "Der c A \ {s. [c] @ s \ A}" + +definition + Ders :: "string \ lang \ lang" +where + "Ders s A \ {s'. s @ s' \ A}" + +definition + Ders_set :: "lang \ lang \ lang" +where + "Ders_set A B \ {s' | s s'. s @ s' \ B \ s \ A}" + +lemma Ders_set_Ders: + shows "Ders_set A B = (\s \ A. Ders s B)" +unfolding Ders_set_def Ders_def +by auto + +lemma Der_null [simp]: + shows "Der c {} = {}" +unfolding Der_def +by auto + +lemma Der_empty [simp]: + shows "Der c {[]} = {}" +unfolding Der_def +by auto + +lemma Der_char [simp]: + shows "Der c {[d]} = (if c = d then {[]} else {})" +unfolding Der_def +by auto + +lemma Der_union [simp]: + shows "Der c (A \ B) = Der c A \ Der c B" +unfolding Der_def +by auto + +lemma Der_seq [simp]: + shows "Der c (A ;; B) = (Der c A) ;; B \ (Delta A ;; Der c B)" +unfolding Der_def Delta_def +unfolding Seq_def +by (auto simp add: Cons_eq_append_conv) + +lemma Der_star [simp]: + shows "Der c (A\) = (Der c A) ;; A\" +apply(subst star_cases) +apply(simp only: Delta_def Der_union Der_seq Der_empty) +apply(simp add: Der_def Seq_def) +apply(auto) +apply(drule star_decom) +apply(auto simp add: Cons_eq_append_conv) +done + +lemma Ders_singleton: + shows "Ders [c] A = Der c A" +unfolding Der_def Ders_def +by simp + +lemma Ders_append: + shows "Ders (s1 @ s2) A = Ders s2 (Ders s1 A)" +unfolding Ders_def by simp + +lemma MN_Rel_Ders: + shows "x \A y \ Ders x A = Ders y A" +unfolding Ders_def str_eq_def str_eq_rel_def +by auto + +subsection {* Brozowsky's derivatives of regular expressions *} + +fun + nullable :: "rexp \ bool" +where + "nullable (NULL) = False" +| "nullable (EMPTY) = True" +| "nullable (CHAR c) = False" +| "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (STAR r) = True" + +fun + der :: "char \ rexp \ 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) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)" +| "der c (STAR r) = SEQ (der c r) (STAR r)" + +function + ders :: "string \ rexp \ rexp" +where + "ders [] r = r" +| "ders (s @ [c]) r = der c (ders s r)" +by (auto) (metis rev_cases) + +termination + by (relation "measure (length o fst)") (auto) + +lemma Delta_nullable: + shows "Delta (L r) = (if nullable r then {[]} else {})" +unfolding Delta_def +by (induct r) (auto simp add: Seq_def split: if_splits) + +lemma Der_der: + fixes r::rexp + shows "Der c (L r) = L (der c r)" +by (induct r) (simp_all add: Delta_nullable) + +lemma Ders_ders: + fixes r::rexp + shows "Ders s (L r) = L (ders s r)" +apply(induct s rule: rev_induct) +apply(simp add: Ders_def) +apply(simp only: ders.simps) +apply(simp only: Ders_append) +apply(simp only: Ders_singleton) +apply(simp only: Der_der) +done + + +subsection {* Antimirov's Partial Derivatives *} + +abbreviation + "SEQS R r \ {SEQ r' r | r'. r' \ R}" + +fun + pder :: "char \ rexp \ rexp set" +where + "pder c NULL = {NULL}" +| "pder c EMPTY = {NULL}" +| "pder c (CHAR c') = (if c = c' then {EMPTY} else {NULL})" +| "pder c (ALT r1 r2) = (pder c r1) \ (pder c r2)" +| "pder c (SEQ r1 r2) = SEQS (pder c r1) r2 \ (if nullable r1 then pder c r2 else {})" +| "pder c (STAR r) = SEQS (pder c r) (STAR r)" + +abbreviation + "pder_set c R \ \r \ R. pder c r" + +function + pders :: "string \ rexp \ rexp set" +where + "pders [] r = {r}" +| "pders (s @ [c]) r = pder_set c (pders s r)" +by (auto) (metis rev_cases) + +termination + by (relation "measure (length o fst)") (auto) + +abbreviation + "pders_set A r \ \s \ A. pders s r" + +lemma pders_append: + "pders (s1 @ s2) r = \ (pders s2) ` (pders s1 r)" +apply(induct s2 arbitrary: s1 r rule: rev_induct) +apply(simp) +apply(subst append_assoc[symmetric]) +apply(simp only: pders.simps) +apply(auto) +done + +lemma pders_singleton: + "pders [c] r = pder c r" +apply(subst append_Nil[symmetric]) +apply(simp only: pders.simps) +apply(simp) +done + +lemma pder_set_lang: + shows "(\ (L ` pder_set c R)) = (\r \ R. (\L ` (pder c r)))" +unfolding image_def +by auto + +lemma + shows seq_UNION_left: "B ;; (\n\C. A n) = (\n\C. B ;; A n)" + and seq_UNION_right: "(\n\C. A n) ;; B = (\n\C. A n ;; B)" +unfolding Seq_def by auto + +lemma Der_pder: + fixes r::rexp + shows "Der c (L r) = \ L ` (pder c r)" +by (induct r) (auto simp add: Delta_nullable seq_UNION_right) + +lemma Ders_pders: + fixes r::rexp + shows "Ders s (L r) = \ L ` (pders s r)" +proof (induct s rule: rev_induct) + case (snoc c s) + have ih: "Ders s (L r) = \ L ` (pders s r)" by fact + have "Ders (s @ [c]) (L r) = Ders [c] (Ders s (L r))" + by (simp add: Ders_append) + also have "\ = Der c (\ L ` (pders s r))" using ih + by (simp add: Ders_singleton) + also have "\ = (\r\pders s r. Der c (L r))" + unfolding Der_def image_def by auto + also have "\ = (\r\pders s r. (\ L ` (pder c r)))" + by (simp add: Der_pder) + also have "\ = (\L ` (pder_set c (pders s r)))" + by (simp add: pder_set_lang) + also have "\ = (\L ` (pders (s @ [c]) r))" + by simp + finally show "Ders (s @ [c]) (L r) = \L ` pders (s @ [c]) r" . +qed (simp add: Ders_def) + +lemma Ders_set_pders_set: + fixes r::rexp + shows "Ders_set A (L r) = (\ L ` (pders_set A r))" +by (simp add: Ders_set_Ders Ders_pders) + +lemma pders_NULL [simp]: + shows "pders s NULL = {NULL}" +by (induct s rule: rev_induct) (simp_all) + +lemma pders_EMPTY [simp]: + shows "pders s EMPTY = (if s = [] then {EMPTY} else {NULL})" +by (induct s rule: rev_induct) (auto) + +lemma pders_CHAR [simp]: + shows "pders s (CHAR c) = (if s = [] then {CHAR c} else (if s = [c] then {EMPTY} else {NULL}))" +by (induct s rule: rev_induct) (auto) + +lemma pders_ALT [simp]: + shows "pders s (ALT r1 r2) = (if s = [] then {ALT r1 r2} else (pders s r1) \ (pders s r2))" +by (induct s rule: rev_induct) (auto) + +definition + "Suf s \ {v. v \ [] \ (\u. u @ v = s)}" + +lemma Suf: + shows "Suf (s @ [c]) = (Suf s) ;; {[c]} \ {[c]}" +unfolding Suf_def Seq_def +by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv) + +lemma Suf_Union: + shows "(\v \ Suf s ;; {[c]}. P v) = (\v \ Suf s. P (v @ [c]))" +by (auto simp add: Seq_def) + +lemma inclusion1: + shows "pder_set c (SEQS R r2) \ SEQS (pder_set c R) r2 \ (pder c r2)" +apply(auto simp add: if_splits) +apply(blast) +done + +lemma pders_SEQ: + shows "pders s (SEQ r1 r2) \ SEQS (pders s r1) r2 \ (\v \ Suf s. pders v r2)" +proof (induct s rule: rev_induct) + case (snoc c s) + have ih: "pders s (SEQ r1 r2) \ SEQS (pders s r1) r2 \ (\v \ Suf s. pders v r2)" + by fact + have "pders (s @ [c]) (SEQ r1 r2) = pder_set c (pders s (SEQ r1 r2))" by simp + also have "\ \ pder_set c (SEQS (pders s r1) r2 \ (\v \ Suf s. pders v r2))" + using ih by auto + also have "\ = pder_set c (SEQS (pders s r1) r2) \ pder_set c (\v \ Suf s. pders v r2)" + by (simp) + also have "\ = pder_set c (SEQS (pders s r1) r2) \ (\v \ Suf s. pder_set c (pders v r2))" + by (simp) + also have "\ \ pder_set c (SEQS (pders s r1) r2) \ (pder c r2) \ (\v \ Suf s. pders (v @ [c]) r2)" + by (auto) + also have "\ \ SEQS (pder_set c (pders s r1)) r2 \ (pder c r2) \ (\v \ Suf s. pders (v @ [c]) r2)" + using inclusion1 by blast + also have "\ = SEQS (pders (s @ [c]) r1) r2 \ (\v \ Suf (s @ [c]). pders v r2)" + apply(subst (2) pders.simps) + apply(simp only: Suf) + apply(simp add: Suf_Union pders_singleton) + apply(auto) + done + finally show ?case . +qed (simp) + +lemma pders_STAR: + assumes a: "s \ []" + shows "pders s (STAR r) \ (\v \ Suf s. SEQS (pders v r) (STAR r))" +using a +proof (induct s rule: rev_induct) + case (snoc c s) + have ih: "s \ [] \ pders s (STAR r) \ (\v\Suf s. SEQS (pders v r) (STAR r))" by fact + { assume asm: "s \ []" + have "pders (s @ [c]) (STAR r) = pder_set c (pders s (STAR r))" by simp + also have "\ \ (pder_set c (\v\Suf s. SEQS (pders v r) (STAR r)))" + using ih[OF asm] by blast + also have "\ = (\v\Suf s. pder_set c (SEQS (pders v r) (STAR r)))" + by simp + also have "\ \ (\v\Suf s. (SEQS (pder_set c (pders v r)) (STAR r) \ pder c (STAR r)))" + using inclusion1 by blast + also have "\ = (\v\Suf s. (SEQS (pder_set c (pders v r)) (STAR r))) \ pder c (STAR r)" + using asm by (auto simp add: Suf_def) + also have "\ = (\v\Suf s. (SEQS (pders (v @ [c]) r) (STAR r))) \ (SEQS (pder c r) (STAR r))" + by simp + also have "\ = (\v\Suf (s @ [c]). (SEQS (pders v r) (STAR r)))" + apply(simp only: Suf) + apply(simp add: Suf_Union pders_singleton) + apply(auto) + done + finally have ?case . + } + moreover + { assume asm: "s = []" + then have ?case + apply(simp add: pders_singleton Suf_def) + apply(auto) + apply(rule_tac x="[c]" in exI) + apply(simp add: pders_singleton) + done + } + ultimately show ?case by blast +qed (simp) + +abbreviation + "UNIV1 \ UNIV - {[]}" + +lemma pders_set_NULL: + shows "pders_set UNIV1 NULL = {NULL}" +by auto + +lemma pders_set_EMPTY: + shows "pders_set UNIV1 EMPTY = {NULL}" +by (auto split: if_splits) + +lemma pders_set_CHAR: + shows "pders_set UNIV1 (CHAR c) \ {EMPTY, NULL}" +by (auto split: if_splits) + +lemma pders_set_ALT: + shows "pders_set UNIV1 (ALT r1 r2) = pders_set UNIV1 r1 \ pders_set UNIV1 r2" +by auto + +lemma pders_set_SEQ_aux: + assumes a: "s \ UNIV1" + shows "pders_set (Suf s) r2 \ pders_set UNIV1 r2" +using a by (auto simp add: Suf_def) + +lemma pders_set_SEQ: + shows "pders_set UNIV1 (SEQ r1 r2) \ SEQS (pders_set UNIV1 r1) r2 \ pders_set UNIV1 r2" +apply(rule UN_least) +apply(rule subset_trans) +apply(rule pders_SEQ) +apply(simp) +apply(rule conjI) +apply(auto)[1] +apply(rule subset_trans) +apply(rule pders_set_SEQ_aux) +apply(auto) +done + +lemma pders_set_STAR: + shows "pders_set UNIV1 (STAR r) \ SEQS (pders_set UNIV1 r) (STAR r)" +apply(rule UN_least) +apply(rule subset_trans) +apply(rule pders_STAR) +apply(simp) +apply(simp add: Suf_def) +apply(auto) +done + +lemma finite_SEQS: + assumes a: "finite A" + shows "finite (SEQS A r)" +using a by (auto) + +lemma finite_pders_set_UNIV1: + shows "finite (pders_set UNIV1 r)" +apply(induct r) +apply(simp) +apply(simp only: pders_set_EMPTY) +apply(simp) +apply(rule finite_subset) +apply(rule pders_set_CHAR) +apply(simp) +apply(rule finite_subset) +apply(rule pders_set_SEQ) +apply(simp only: finite_SEQS finite_Un) +apply(simp) +apply(simp only: pders_set_ALT) +apply(simp) +apply(rule finite_subset) +apply(rule pders_set_STAR) +apply(simp only: finite_SEQS) +done + +lemma pders_set_UNIV_UNIV1: + shows "pders_set UNIV r = pders [] r \ pders_set UNIV1 r" +apply(auto) +apply(rule_tac x="[]" in exI) +apply(simp) +done + +lemma finite_pders_set_UNIV: + shows "finite (pders_set UNIV r)" +unfolding pders_set_UNIV_UNIV1 +by (simp add: finite_pders_set_UNIV1) + +lemma finite_pders_set: + shows "finite (pders_set A r)" +apply(rule rev_finite_subset) +apply(rule_tac r="r" in finite_pders_set_UNIV) +apply(auto) +done + +lemma finite_pders: + shows "finite (pders s r)" +using finite_pders_set[where A="{s}" and r="r"] +by simp + +lemma closure_left_quotient: + assumes "regular A" + shows "regular (Ders_set B A)" +proof - + from assms obtain r::rexp where eq: "L r = A" by auto + have fin: "finite (pders_set B r)" by (rule finite_pders_set) + + have "Ders_set B (L r) = (\ L ` (pders_set B r))" + by (simp add: Ders_set_pders_set) + also have "\ = L (\(pders_set B r))" using fin by simp + finally have "Ders_set B A = L (\(pders_set B r))" using eq + by simp + then show "regular (Ders_set B A)" by auto +qed + + +fun + width :: "rexp \ nat" +where + "width (NULL) = 0" +| "width (EMPTY) = 0" +| "width (CHAR c) = 1" +| "width (ALT r1 r2) = width r1 + width r2" +| "width (SEQ r1 r2) = width r1 + width r2" +| "width (STAR r) = width r" + + + +end \ No newline at end of file diff -r 3b7477db3462 -r e122cb146ecc Theories/Folds.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Theories/Folds.thy Wed Mar 23 12:17:30 2011 +0000 @@ -0,0 +1,22 @@ +theory Folds +imports Main +begin + +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" + + +end \ No newline at end of file diff -r 3b7477db3462 -r e122cb146ecc Theories/Myhill_1.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Theories/Myhill_1.thy Wed Mar 23 12:17:30 2011 +0000 @@ -0,0 +1,783 @@ +theory Myhill_1 +imports Main Folds Regular + "~~/src/HOL/Library/While_Combinator" +begin + +section {* Direction @{text "finite partition \ regular language"} *} + +lemma Pair_Collect[simp]: + shows "(x, y) \ {(x, y). P x y} \ P x y" +by simp + +text {* Myhill-Nerode relation *} + +definition + str_eq_rel :: "lang \ (string \ string) set" ("\_" [100] 100) +where + "\A \ {(x, y). (\z. x @ z \ A \ y @ z \ A)}" + +definition + finals :: "lang \ lang set" +where + "finals A \ {\A `` {s} | s . s \ A}" + +lemma lang_is_union_of_finals: + shows "A = \ finals A" +unfolding finals_def +unfolding Image_def +unfolding str_eq_rel_def +by (auto) (metis append_Nil2) + +lemma finals_in_partitions: + shows "finals A \ (UNIV // \A)" +unfolding finals_def quotient_def +by auto + +section {* Equational systems *} + +text {* The two kinds of terms in the rhs of equations. *} + +datatype rhs_trm = + Lam "rexp" (* Lambda-marker *) + | Trn "lang" "rexp" (* Transition *) + + +overloading L_rhs_trm \ "L:: rhs_trm \ lang" +begin + fun L_rhs_trm:: "rhs_trm \ lang" + where + "L_rhs_trm (Lam r) = L r" + | "L_rhs_trm (Trn X r) = X ;; L r" +end + +overloading L_rhs \ "L:: rhs_trm set \ lang" +begin + fun L_rhs:: "rhs_trm set \ lang" + where + "L_rhs rhs = \ (L ` rhs)" +end + +lemma L_rhs_set: + shows "L {Trn X r | r. P r} = \{L (Trn X r) | r. P r}" +by (auto simp del: L_rhs_trm.simps) + +lemma L_rhs_union_distrib: + fixes A B::"rhs_trm set" + shows "L A \ L B = L (A \ B)" +by simp + + + +text {* Transitions between equivalence classes *} + +definition + transition :: "lang \ char \ lang \ bool" ("_ \_\_" [100,100,100] 100) +where + "Y \c\ X \ Y ;; {[c]} \ X" + +text {* Initial equational system *} + +definition + "Init_rhs CS X \ + if ([] \ X) then + {Lam EMPTY} \ {Trn Y (CHAR c) | Y c. Y \ CS \ Y \c\ X} + else + {Trn Y (CHAR c)| Y c. Y \ CS \ Y \c\ X}" + +definition + "Init CS \ {(X, Init_rhs CS X) | X. X \ CS}" + + +section {* Arden Operation on equations *} + +fun + Append_rexp :: "rexp \ rhs_trm \ rhs_trm" +where + "Append_rexp r (Lam rexp) = Lam (SEQ rexp r)" +| "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)" + + +definition + "Append_rexp_rhs rhs rexp \ (Append_rexp rexp) ` rhs" + +definition + "Arden X rhs \ + Append_rexp_rhs (rhs - {Trn X r | r. Trn X r \ rhs}) (STAR (\ {r. Trn X r \ rhs}))" + + +section {* Substitution Operation on equations *} + +definition + "Subst rhs X xrhs \ + (rhs - {Trn X r | r. Trn X r \ rhs}) \ (Append_rexp_rhs xrhs (\ {r. Trn X r \ rhs}))" + +definition + Subst_all :: "(lang \ rhs_trm set) set \ lang \ rhs_trm set \ (lang \ rhs_trm set) set" +where + "Subst_all ES X xrhs \ {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \ ES}" + +definition + "Remove ES X xrhs \ + Subst_all (ES - {(X, xrhs)}) X (Arden X xrhs)" + + +section {* While-combinator *} + +definition + "Iter X ES \ (let (Y, yrhs) = SOME (Y, yrhs). (Y, yrhs) \ ES \ X \ Y + in Remove ES Y yrhs)" + +lemma IterI2: + assumes "(Y, yrhs) \ ES" + and "X \ Y" + and "\Y yrhs. \(Y, yrhs) \ ES; X \ Y\ \ Q (Remove ES Y yrhs)" + shows "Q (Iter X ES)" +unfolding Iter_def using assms +by (rule_tac a="(Y, yrhs)" in someI2) (auto) + +abbreviation + "Cond ES \ card ES \ 1" + +definition + "Solve X ES \ while Cond (Iter X) ES" + + +section {* Invariants *} + +definition + "distinctness ES \ + \ X rhs rhs'. (X, rhs) \ ES \ (X, rhs') \ ES \ rhs = rhs'" + +definition + "soundness ES \ \(X, rhs) \ ES. X = L rhs" + +definition + "ardenable rhs \ (\ Y r. Trn Y r \ rhs \ [] \ L r)" + +definition + "ardenable_all ES \ \(X, rhs) \ ES. ardenable rhs" + +definition + "finite_rhs ES \ \(X, rhs) \ ES. finite rhs" + +lemma finite_rhs_def2: + "finite_rhs ES = (\ X rhs. (X, rhs) \ ES \ finite rhs)" +unfolding finite_rhs_def by auto + +definition + "rhss rhs \ {X | X r. Trn X r \ rhs}" + +definition + "lhss ES \ {Y | Y yrhs. (Y, yrhs) \ ES}" + +definition + "validity ES \ \(X, rhs) \ ES. rhss rhs \ lhss ES" + +lemma rhss_union_distrib: + shows "rhss (A \ B) = rhss A \ rhss B" +by (auto simp add: rhss_def) + +lemma lhss_union_distrib: + shows "lhss (A \ B) = lhss A \ lhss B" +by (auto simp add: lhss_def) + + +definition + "invariant ES \ finite ES + \ finite_rhs ES + \ soundness ES + \ distinctness ES + \ ardenable_all ES + \ validity ES" + + +lemma invariantI: + assumes "soundness ES" "finite ES" "distinctness ES" "ardenable_all ES" + "finite_rhs ES" "validity ES" + shows "invariant ES" +using assms by (simp add: invariant_def) + + +subsection {* The proof of this direction *} + +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 rhs_trm_soundness: + 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 "L ({Trn X r| r. Trn X r \ rhs}) = X ;; (L (\{r. Trn X r \ rhs}))" + by (simp only: L_rhs_set L_rhs_trm.simps) (auto simp add: Seq_def) +qed + +lemma lang_of_append_rexp: + "L (Append_rexp r rhs_trm) = L rhs_trm ;; L r" +by (induct rule: Append_rexp.induct) + (auto simp add: seq_assoc) + +lemma lang_of_append_rexp_rhs: + "L (Append_rexp_rhs rhs r) = L rhs ;; L r" +unfolding Append_rexp_rhs_def +by (auto simp add: Seq_def lang_of_append_rexp) + + + +subsubsection {* Intialization *} + +lemma defined_by_str: + assumes "s \ X" "X \ UNIV // \A" + shows "X = \A `` {s}" +using assms +unfolding quotient_def Image_def str_eq_rel_def +by auto + +lemma every_eqclass_has_transition: + assumes has_str: "s @ [c] \ X" + and in_CS: "X \ UNIV // \A" + obtains Y where "Y \ UNIV // \A" and "Y ;; {[c]} \ X" and "s \ Y" +proof - + def Y \ "\A `` {s}" + have "Y \ UNIV // \A" + unfolding Y_def quotient_def by auto + moreover + have "X = \A `` {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 using that by blast +qed + +lemma l_eq_r_in_eqs: + assumes X_in_eqs: "(X, rhs) \ Init (UNIV // \A)" + shows "X = L rhs" +proof + show "X \ L rhs" + proof + fix x + assume in_X: "x \ X" + { assume empty: "x = []" + then have "x \ L rhs" using X_in_eqs in_X + unfolding Init_def Init_rhs_def + by auto + } + moreover + { assume not_empty: "x \ []" + then obtain s c where decom: "x = s @ [c]" + using rev_cases by blast + have "X \ UNIV // \A" using X_in_eqs unfolding Init_def by auto + then obtain Y where "Y \ UNIV // \A" "Y ;; {[c]} \ X" "s \ Y" + using decom in_X every_eqclass_has_transition by blast + then have "x \ L {Trn Y (CHAR c)| Y c. Y \ UNIV // \A \ Y \c\ X}" + unfolding transition_def + using decom by (force simp add: Seq_def) + then have "x \ L rhs" using X_in_eqs in_X + unfolding Init_def Init_rhs_def by simp + } + ultimately show "x \ L rhs" by blast + qed +next + show "L rhs \ X" using X_in_eqs + unfolding Init_def Init_rhs_def transition_def + by auto +qed + +lemma test: + assumes X_in_eqs: "(X, rhs) \ Init (UNIV // \A)" + shows "X = \ (L ` rhs)" +using assms l_eq_r_in_eqs by (simp) + +lemma finite_Init_rhs: + assumes finite: "finite CS" + shows "finite (Init_rhs CS X)" +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 + then have "finite S" using S_def + by (rule_tac B = "CS \ UNIV" in finite_subset) (auto) + moreover have "{Trn Y (CHAR c) |Y c. Y \ CS \ Y ;; {[c]} \ X} = h ` S" + unfolding S_def h_def image_def by auto + ultimately + have "finite {Trn Y (CHAR c) |Y c. Y \ CS \ Y ;; {[c]} \ X}" by auto + then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp +qed + +lemma Init_ES_satisfies_invariant: + assumes finite_CS: "finite (UNIV // \A)" + shows "invariant (Init (UNIV // \A))" +proof (rule invariantI) + show "soundness (Init (UNIV // \A))" + unfolding soundness_def + using l_eq_r_in_eqs by auto + show "finite (Init (UNIV // \A))" using finite_CS + unfolding Init_def by simp + show "distinctness (Init (UNIV // \A))" + unfolding distinctness_def Init_def by simp + show "ardenable_all (Init (UNIV // \A))" + unfolding ardenable_all_def Init_def Init_rhs_def ardenable_def + by auto + show "finite_rhs (Init (UNIV // \A))" + using finite_Init_rhs[OF finite_CS] + unfolding finite_rhs_def Init_def by auto + show "validity (Init (UNIV // \A))" + unfolding validity_def Init_def Init_rhs_def rhss_def lhss_def + by auto +qed + +subsubsection {* Interation step *} + +lemma Arden_keeps_eq: + assumes l_eq_r: "X = L rhs" + and not_empty: "ardenable rhs" + and finite: "finite rhs" + shows "X = L (Arden X rhs)" +proof - + def A \ "L (\{r. Trn X r \ rhs})" + def b \ "{Trn X r | r. Trn X r \ rhs}" + def B \ "L (rhs - b)" + have not_empty2: "[] \ A" + using finite_Trn[OF finite] not_empty + unfolding A_def ardenable_def by simp + have "X = L rhs" using l_eq_r by simp + also have "\ = L (b \ (rhs - b))" unfolding b_def by auto + also have "\ = L b \ B" unfolding B_def by (simp only: L_rhs_union_distrib) + also have "\ = X ;; A \ B" + unfolding b_def + unfolding rhs_trm_soundness[OF finite] + unfolding A_def + by blast + finally have "X = X ;; A \ B" . + then have "X = B ;; A\" + by (simp add: arden[OF not_empty2]) + also have "\ = L (Arden X rhs)" + unfolding Arden_def A_def B_def b_def + by (simp only: lang_of_append_rexp_rhs L_rexp.simps) + finally show "X = L (Arden X rhs)" by simp +qed + +lemma Append_keeps_finite: + "finite rhs \ finite (Append_rexp_rhs rhs r)" +by (auto simp:Append_rexp_rhs_def) + +lemma Arden_keeps_finite: + "finite rhs \ finite (Arden X rhs)" +by (auto simp:Arden_def Append_keeps_finite) + +lemma Append_keeps_nonempty: + "ardenable rhs \ ardenable (Append_rexp_rhs rhs r)" +apply (auto simp:ardenable_def Append_rexp_rhs_def) +by (case_tac x, auto simp:Seq_def) + +lemma nonempty_set_sub: + "ardenable rhs \ ardenable (rhs - A)" +by (auto simp:ardenable_def) + +lemma nonempty_set_union: + "\ardenable rhs; ardenable rhs'\ \ ardenable (rhs \ rhs')" +by (auto simp:ardenable_def) + +lemma Arden_keeps_nonempty: + "ardenable rhs \ ardenable (Arden X rhs)" +by (simp only:Arden_def Append_keeps_nonempty nonempty_set_sub) + + +lemma Subst_keeps_nonempty: + "\ardenable rhs; ardenable xrhs\ \ ardenable (Subst rhs X xrhs)" +by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub) + +lemma Subst_keeps_eq: + assumes substor: "X = L xrhs" + and finite: "finite rhs" + shows "L (Subst rhs X xrhs) = L rhs" (is "?Left = ?Right") +proof- + def A \ "L (rhs - {Trn X r | r. Trn X r \ rhs})" + have "?Left = A \ L (Append_rexp_rhs xrhs (\{r. Trn X r \ rhs}))" + unfolding 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 - {Trn X r | r. Trn X r \ rhs}) \ ({Trn X r | r. Trn X r \ rhs})" by auto + thus ?thesis + unfolding A_def + unfolding L_rhs_union_distrib + by simp + qed + moreover have "L (Append_rexp_rhs xrhs (\{r. Trn X r \ rhs})) = L ({Trn X r | r. Trn X r \ rhs})" + using finite substor by (simp only: lang_of_append_rexp_rhs rhs_trm_soundness) + ultimately show ?thesis by simp +qed + +lemma Subst_keeps_finite_rhs: + "\finite rhs; finite yrhs\ \ finite (Subst rhs Y yrhs)" +by (auto simp: Subst_def Append_keeps_finite) + +lemma Subst_all_keeps_finite: + assumes finite: "finite ES" + shows "finite (Subst_all ES Y yrhs)" +proof - + def eqns \ "{(X::lang, rhs) |X rhs. (X, rhs) \ ES}" + def h \ "\(X::lang, rhs). (X, Subst rhs Y yrhs)" + have "finite (h ` eqns)" using finite h_def eqns_def by auto + moreover + have "Subst_all ES Y yrhs = h ` eqns" unfolding h_def eqns_def Subst_all_def by auto + ultimately + show "finite (Subst_all ES Y yrhs)" by simp +qed + +lemma Subst_all_keeps_finite_rhs: + "\finite_rhs ES; finite yrhs\ \ finite_rhs (Subst_all ES Y yrhs)" +by (auto intro:Subst_keeps_finite_rhs simp add:Subst_all_def finite_rhs_def) + +lemma append_rhs_keeps_cls: + "rhss (Append_rexp_rhs rhs r) = rhss rhs" +apply (auto simp:rhss_def Append_rexp_rhs_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_removes_cl: + "rhss (Arden Y yrhs) = rhss yrhs - {Y}" +apply (simp add:Arden_def append_rhs_keeps_cls) +by (auto simp:rhss_def) + +lemma lhss_keeps_cls: + "lhss (Subst_all ES Y yrhs) = lhss ES" +by (auto simp:lhss_def Subst_all_def) + +lemma Subst_updates_cls: + "X \ rhss xrhs \ + rhss (Subst rhs X xrhs) = rhss rhs \ rhss xrhs - {X}" +apply (simp only:Subst_def append_rhs_keeps_cls rhss_union_distrib) +by (auto simp:rhss_def) + +lemma Subst_all_keeps_validity: + assumes sc: "validity (ES \ {(Y, yrhs)})" (is "validity ?A") + shows "validity (Subst_all ES Y (Arden Y yrhs))" (is "validity ?B") +proof - + { fix X xrhs' + assume "(X, xrhs') \ ?B" + then obtain xrhs + where xrhs_xrhs': "xrhs' = Subst xrhs Y (Arden Y yrhs)" + and X_in: "(X, xrhs) \ ES" by (simp add:Subst_all_def, blast) + have "rhss xrhs' \ lhss ?B" + proof- + have "lhss ?B = lhss ES" by (auto simp add:lhss_def Subst_all_def) + moreover have "rhss xrhs' \ lhss ES" + proof- + have "rhss xrhs' \ rhss xrhs \ rhss (Arden Y yrhs) - {Y}" + proof- + have "Y \ rhss (Arden Y yrhs)" + using Arden_removes_cl by simp + thus ?thesis using xrhs_xrhs' by (auto simp:Subst_updates_cls) + qed + moreover have "rhss xrhs \ lhss ES \ {Y}" using X_in sc + apply (simp only:validity_def lhss_union_distrib) + by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lhss_def) + moreover have "rhss (Arden Y yrhs) \ lhss ES \ {Y}" + using sc + by (auto simp add:Arden_removes_cl validity_def lhss_def) + ultimately show ?thesis by auto + qed + ultimately show ?thesis by simp + qed + } thus ?thesis by (auto simp only:Subst_all_def validity_def) +qed + +lemma Subst_all_satisfies_invariant: + assumes invariant_ES: "invariant (ES \ {(Y, yrhs)})" + shows "invariant (Subst_all ES Y (Arden Y yrhs))" +proof (rule invariantI) + have Y_eq_yrhs: "Y = L yrhs" + using invariant_ES by (simp only:invariant_def soundness_def, blast) + have finite_yrhs: "finite yrhs" + using invariant_ES by (auto simp:invariant_def finite_rhs_def) + have nonempty_yrhs: "ardenable yrhs" + using invariant_ES by (auto simp:invariant_def ardenable_all_def) + show "soundness (Subst_all ES Y (Arden Y yrhs))" + proof - + have "Y = L (Arden Y yrhs)" + using Y_eq_yrhs invariant_ES finite_yrhs + using finite_Trn[OF finite_yrhs] + apply(rule_tac Arden_keeps_eq) + apply(simp_all) + unfolding invariant_def ardenable_all_def ardenable_def + apply(auto) + done + thus ?thesis using invariant_ES + unfolding invariant_def finite_rhs_def2 soundness_def Subst_all_def + by (auto simp add: Subst_keeps_eq simp del: L_rhs.simps) + qed + show "finite (Subst_all ES Y (Arden Y yrhs))" + using invariant_ES by (simp add:invariant_def Subst_all_keeps_finite) + show "distinctness (Subst_all ES Y (Arden Y yrhs))" + using invariant_ES + unfolding distinctness_def Subst_all_def invariant_def by auto + show "ardenable_all (Subst_all ES Y (Arden Y yrhs))" + proof - + { fix X rhs + assume "(X, rhs) \ ES" + hence "ardenable rhs" using invariant_ES + by (auto simp add:invariant_def ardenable_all_def) + with nonempty_yrhs + have "ardenable (Subst rhs Y (Arden Y yrhs))" + by (simp add:nonempty_yrhs + Subst_keeps_nonempty Arden_keeps_nonempty) + } thus ?thesis by (auto simp add:ardenable_all_def Subst_all_def) + qed + show "finite_rhs (Subst_all ES Y (Arden Y yrhs))" + proof- + have "finite_rhs ES" using invariant_ES + by (simp add:invariant_def finite_rhs_def) + moreover have "finite (Arden Y yrhs)" + proof - + have "finite yrhs" using invariant_ES + by (auto simp:invariant_def finite_rhs_def) + thus ?thesis using Arden_keeps_finite by simp + qed + ultimately show ?thesis + by (simp add:Subst_all_keeps_finite_rhs) + qed + show "validity (Subst_all ES Y (Arden Y yrhs))" + using invariant_ES Subst_all_keeps_validity by (simp add:invariant_def) +qed + +lemma Remove_in_card_measure: + assumes finite: "finite ES" + and in_ES: "(X, rhs) \ ES" + shows "(Remove ES X rhs, ES) \ measure card" +proof - + def f \ "\ x. ((fst x)::lang, Subst (snd x) X (Arden X rhs))" + def ES' \ "ES - {(X, rhs)}" + have "Subst_all ES' X (Arden X rhs) = f ` ES'" + apply (auto simp: Subst_all_def f_def image_def) + by (rule_tac x = "(Y, yrhs)" in bexI, simp+) + then have "card (Subst_all ES' X (Arden X rhs)) \ card ES'" + unfolding ES'_def using finite by (auto intro: card_image_le) + also have "\ < card ES" unfolding ES'_def + using in_ES finite by (rule_tac card_Diff1_less) + finally show "(Remove ES X rhs, ES) \ measure card" + unfolding Remove_def ES'_def by simp +qed + + +lemma Subst_all_cls_remains: + "(X, xrhs) \ ES \ \ xrhs'. (X, xrhs') \ (Subst_all ES Y yrhs)" +by (auto simp: Subst_all_def) + +lemma card_noteq_1_has_more: + assumes card:"Cond ES" + and e_in: "(X, xrhs) \ ES" + and finite: "finite ES" + shows "\(Y, yrhs) \ ES. (X, xrhs) \ (Y, yrhs)" +proof- + have "card ES > 1" using card e_in finite + by (cases "card ES") (auto) + then have "card (ES - {(X, xrhs)}) > 0" + using finite e_in by auto + then have "(ES - {(X, xrhs)}) \ {}" using finite by (rule_tac notI, simp) + then show "\(Y, yrhs) \ ES. (X, xrhs) \ (Y, yrhs)" + by auto +qed + +lemma iteration_step_measure: + assumes Inv_ES: "invariant ES" + and X_in_ES: "(X, xrhs) \ ES" + and Cnd: "Cond ES " + shows "(Iter X ES, ES) \ measure card" +proof - + have fin: "finite ES" using Inv_ES unfolding invariant_def by simp + then obtain Y yrhs + where Y_in_ES: "(Y, yrhs) \ ES" and not_eq: "(X, xrhs) \ (Y, yrhs)" + using Cnd X_in_ES by (drule_tac card_noteq_1_has_more) (auto) + then have "(Y, yrhs) \ ES " "X \ Y" + using X_in_ES Inv_ES unfolding invariant_def distinctness_def + by auto + then show "(Iter X ES, ES) \ measure card" + apply(rule IterI2) + apply(rule Remove_in_card_measure) + apply(simp_all add: fin) + done +qed + +lemma iteration_step_invariant: + assumes Inv_ES: "invariant ES" + and X_in_ES: "(X, xrhs) \ ES" + and Cnd: "Cond ES" + shows "invariant (Iter X ES)" +proof - + have finite_ES: "finite ES" using Inv_ES by (simp add: invariant_def) + then obtain Y yrhs + where Y_in_ES: "(Y, yrhs) \ ES" and not_eq: "(X, xrhs) \ (Y, yrhs)" + using Cnd X_in_ES by (drule_tac card_noteq_1_has_more) (auto) + then have "(Y, yrhs) \ ES" "X \ Y" + using X_in_ES Inv_ES unfolding invariant_def distinctness_def + by auto + then show "invariant (Iter X ES)" + proof(rule IterI2) + fix Y yrhs + assume h: "(Y, yrhs) \ ES" "X \ Y" + then have "ES - {(Y, yrhs)} \ {(Y, yrhs)} = ES" by auto + then show "invariant (Remove ES Y yrhs)" unfolding Remove_def + using Inv_ES + by (rule_tac Subst_all_satisfies_invariant) (simp) + qed +qed + +lemma iteration_step_ex: + assumes Inv_ES: "invariant ES" + and X_in_ES: "(X, xrhs) \ ES" + and Cnd: "Cond ES" + shows "\xrhs'. (X, xrhs') \ (Iter X ES)" +proof - + have finite_ES: "finite ES" using Inv_ES by (simp add: invariant_def) + then obtain Y yrhs + where "(Y, yrhs) \ ES" "(X, xrhs) \ (Y, yrhs)" + using Cnd X_in_ES by (drule_tac card_noteq_1_has_more) (auto) + then have "(Y, yrhs) \ ES " "X \ Y" + using X_in_ES Inv_ES unfolding invariant_def distinctness_def + by auto + then show "\xrhs'. (X, xrhs') \ (Iter X ES)" + apply(rule IterI2) + unfolding Remove_def + apply(rule Subst_all_cls_remains) + using X_in_ES + apply(auto) + done +qed + + +subsubsection {* Conclusion of the proof *} + +lemma Solve: + assumes fin: "finite (UNIV // \A)" + and X_in: "X \ (UNIV // \A)" + shows "\rhs. Solve X (Init (UNIV // \A)) = {(X, rhs)} \ invariant {(X, rhs)}" +proof - + def Inv \ "\ES. invariant ES \ (\rhs. (X, rhs) \ ES)" + have "Inv (Init (UNIV // \A))" unfolding Inv_def + using fin X_in by (simp add: Init_ES_satisfies_invariant, simp add: Init_def) + moreover + { fix ES + assume inv: "Inv ES" and crd: "Cond ES" + then have "Inv (Iter X ES)" + unfolding Inv_def + by (auto simp add: iteration_step_invariant iteration_step_ex) } + moreover + { fix ES + assume inv: "Inv ES" and not_crd: "\Cond ES" + from inv obtain rhs where "(X, rhs) \ ES" unfolding Inv_def by auto + moreover + from not_crd have "card ES = 1" by simp + ultimately + have "ES = {(X, rhs)}" by (auto simp add: card_Suc_eq) + then have "\rhs'. ES = {(X, rhs')} \ invariant {(X, rhs')}" using inv + unfolding Inv_def by auto } + moreover + have "wf (measure card)" by simp + moreover + { fix ES + assume inv: "Inv ES" and crd: "Cond ES" + then have "(Iter X ES, ES) \ measure card" + unfolding Inv_def + apply(clarify) + apply(rule_tac iteration_step_measure) + apply(auto) + done } + ultimately + show "\rhs. Solve X (Init (UNIV // \A)) = {(X, rhs)} \ invariant {(X, rhs)}" + unfolding Solve_def by (rule while_rule) +qed + +lemma every_eqcl_has_reg: + assumes finite_CS: "finite (UNIV // \A)" + and X_in_CS: "X \ (UNIV // \A)" + shows "\r::rexp. X = L r" +proof - + from finite_CS X_in_CS + obtain xrhs where Inv_ES: "invariant {(X, xrhs)}" + using Solve by metis + + def A \ "Arden X xrhs" + have "rhss xrhs \ {X}" using Inv_ES + unfolding validity_def invariant_def rhss_def lhss_def + by auto + then have "rhss A = {}" unfolding A_def + by (simp add: Arden_removes_cl) + then have eq: "{Lam r | r. Lam r \ A} = A" unfolding rhss_def + by (auto, case_tac x, auto) + + have "finite A" using Inv_ES unfolding A_def invariant_def finite_rhs_def + using Arden_keeps_finite by auto + then have fin: "finite {r. Lam r \ A}" by (rule finite_Lam) + + have "X = L xrhs" using Inv_ES unfolding invariant_def soundness_def + by simp + then have "X = L A" using Inv_ES + unfolding A_def invariant_def ardenable_all_def finite_rhs_def + by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn) + then have "X = L {Lam r | r. Lam r \ A}" using eq by simp + then have "X = L (\{r. Lam r \ A})" using fin by auto + then show "\r::rexp. X = L r" by blast +qed + +lemma bchoice_finite_set: + assumes a: "\x \ S. \y. x = f y" + and b: "finite S" + shows "\ys. (\ S) = \(f ` ys) \ finite ys" +using bchoice[OF a] b +apply(erule_tac exE) +apply(rule_tac x="fa ` S" in exI) +apply(auto) +done + +theorem Myhill_Nerode1: + assumes finite_CS: "finite (UNIV // \A)" + shows "\r::rexp. A = L r" +proof - + have fin: "finite (finals A)" + using finals_in_partitions finite_CS by (rule finite_subset) + have "\X \ (UNIV // \A). \r::rexp. X = L r" + using finite_CS every_eqcl_has_reg by blast + then have a: "\X \ finals A. \r::rexp. X = L r" + using finals_in_partitions by auto + then obtain rs::"rexp set" where "\ (finals A) = \(L ` rs)" "finite rs" + using fin by (auto dest: bchoice_finite_set) + then have "A = L (\rs)" + unfolding lang_is_union_of_finals[symmetric] by simp + then show "\r::rexp. A = L r" by blast +qed + + +end \ No newline at end of file diff -r 3b7477db3462 -r e122cb146ecc Theories/Myhill_2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Theories/Myhill_2.thy Wed Mar 23 12:17:30 2011 +0000 @@ -0,0 +1,470 @@ +theory Myhill_2 + imports Myhill_1 Prefix_subtract + "~~/src/HOL/Library/List_Prefix" +begin + +section {* Direction @{text "regular language \finite partition"} *} + +definition + str_eq :: "string \ lang \ string \ bool" ("_ \_ _") +where + "x \A y \ (x, y) \ (\A)" + +definition + tag_eq_rel :: "(string \ 'b) \ (string \ string) set" ("=_=") +where + "=tag= \ {(x, y) | x y. tag x = tag y}" + +lemma finite_eq_tag_rel: + assumes rng_fnt: "finite (range tag)" + shows "finite (UNIV // =tag=)" +proof - + let "?f" = "\X. tag ` X" and ?A = "(UNIV // =tag=)" + have "finite (?f ` ?A)" + proof - + have "range ?f \ (Pow (range tag))" unfolding Pow_def by auto + moreover + have "finite (Pow (range tag))" using rng_fnt by simp + ultimately + have "finite (range ?f)" unfolding image_def by (blast intro: finite_subset) + moreover + have "?f ` ?A \ range ?f" by auto + ultimately show "finite (?f ` ?A)" by (rule rev_finite_subset) + qed + moreover + have "inj_on ?f ?A" + proof - + { fix X Y + assume X_in: "X \ ?A" + and Y_in: "Y \ ?A" + and tag_eq: "?f X = ?f Y" + then obtain x y + where "x \ X" "y \ Y" "tag x = tag y" + unfolding quotient_def Image_def image_def tag_eq_rel_def + by (simp) (blast) + with X_in Y_in + have "X = Y" + unfolding quotient_def tag_eq_rel_def by auto + } + then show "inj_on ?f ?A" unfolding inj_on_def by auto + qed + ultimately show "finite (UNIV // =tag=)" by (rule finite_imageD) +qed + +lemma refined_partition_finite: + assumes fnt: "finite (UNIV // R1)" + and refined: "R1 \ R2" + and eq1: "equiv UNIV R1" and eq2: "equiv UNIV R2" + shows "finite (UNIV // R2)" +proof - + let ?f = "\X. {R1 `` {x} | x. x \ X}" + and ?A = "UNIV // R2" and ?B = "UNIV // R1" + have "?f ` ?A \ Pow ?B" + unfolding image_def Pow_def quotient_def by auto + moreover + have "finite (Pow ?B)" using fnt by simp + ultimately + have "finite (?f ` ?A)" by (rule finite_subset) + moreover + have "inj_on ?f ?A" + proof - + { fix X Y + assume X_in: "X \ ?A" and Y_in: "Y \ ?A" and eq_f: "?f X = ?f Y" + from quotientE [OF X_in] + obtain x where "X = R2 `` {x}" by blast + with equiv_class_self[OF eq2] have x_in: "x \ X" by simp + then have "R1 ``{x} \ ?f X" by auto + with eq_f have "R1 `` {x} \ ?f Y" by simp + then obtain y + where y_in: "y \ Y" and eq_r1_xy: "R1 `` {x} = R1 `` {y}" by auto + with eq_equiv_class[OF _ eq1] + have "(x, y) \ R1" by blast + with refined have "(x, y) \ R2" by auto + with quotient_eqI [OF eq2 X_in Y_in x_in y_in] + have "X = Y" . + } + then show "inj_on ?f ?A" unfolding inj_on_def by blast + qed + ultimately show "finite (UNIV // R2)" by (rule finite_imageD) +qed + +lemma tag_finite_imageD: + assumes rng_fnt: "finite (range tag)" + and same_tag_eqvt: "\m n. tag m = tag n \ m \A n" + shows "finite (UNIV // \A)" +proof (rule_tac refined_partition_finite [of "=tag="]) + show "finite (UNIV // =tag=)" by (rule finite_eq_tag_rel[OF rng_fnt]) +next + from same_tag_eqvt + show "=tag= \ \A" unfolding tag_eq_rel_def str_eq_def + by auto +next + show "equiv UNIV =tag=" + unfolding equiv_def tag_eq_rel_def refl_on_def sym_def trans_def + by auto +next + show "equiv UNIV (\A)" + unfolding equiv_def str_eq_rel_def sym_def refl_on_def trans_def + by blast +qed + + +subsection {* The proof *} + +subsubsection {* The base case for @{const "NULL"} *} + +lemma quot_null_eq: + shows "UNIV // \{} = {UNIV}" +unfolding quotient_def Image_def str_eq_rel_def by auto + +lemma quot_null_finiteI [intro]: + shows "finite (UNIV // \{})" +unfolding quot_null_eq by simp + + +subsubsection {* The base case for @{const "EMPTY"} *} + +lemma quot_empty_subset: + shows "UNIV // \{[]} \ {{[]}, UNIV - {[]}}" +proof + fix x + assume "x \ UNIV // \{[]}" + then obtain y where h: "x = {z. (y, z) \ \{[]}}" + unfolding quotient_def Image_def by blast + show "x \ {{[]}, UNIV - {[]}}" + proof (cases "y = []") + case True with h + have "x = {[]}" by (auto simp: str_eq_rel_def) + thus ?thesis by simp + next + case False with h + have "x = UNIV - {[]}" by (auto simp: str_eq_rel_def) + thus ?thesis by simp + qed +qed + +lemma quot_empty_finiteI [intro]: + shows "finite (UNIV // \{[]})" +by (rule finite_subset[OF quot_empty_subset]) (simp) + + +subsubsection {* The base case for @{const "CHAR"} *} + +lemma quot_char_subset: + "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" +proof + fix x + assume "x \ UNIV // \{[c]}" + then obtain y where h: "x = {z. (y, z) \ \{[c]}}" + unfolding quotient_def Image_def by blast + show "x \ {{[]},{[c]}, UNIV - {[], [c]}}" + proof - + { assume "y = []" hence "x = {[]}" using h + by (auto simp:str_eq_rel_def) } + moreover + { assume "y = [c]" hence "x = {[c]}" using h + by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def) } + moreover + { assume "y \ []" and "y \ [c]" + hence "\ z. (y @ z) \ [c]" by (case_tac y, auto) + moreover have "\ p. (p \ [] \ p \ [c]) = (\ q. p @ q \ [c])" + by (case_tac p, auto) + ultimately have "x = UNIV - {[],[c]}" using h + by (auto simp add:str_eq_rel_def) + } + ultimately show ?thesis by blast + qed +qed + +lemma quot_char_finiteI [intro]: + shows "finite (UNIV // \{[c]})" +by (rule finite_subset[OF quot_char_subset]) (simp) + + +subsubsection {* The inductive case for @{const ALT} *} + +definition + tag_str_ALT :: "lang \ lang \ string \ (lang \ lang)" +where + "tag_str_ALT A B \ (\x. (\A `` {x}, \B `` {x}))" + +lemma quot_union_finiteI [intro]: + fixes L1 L2::"lang" + assumes finite1: "finite (UNIV // \A)" + and finite2: "finite (UNIV // \B)" + shows "finite (UNIV // \(A \ B))" +proof (rule_tac tag = "tag_str_ALT A B" in tag_finite_imageD) + have "finite ((UNIV // \A) \ (UNIV // \B))" + using finite1 finite2 by auto + then show "finite (range (tag_str_ALT A B))" + unfolding tag_str_ALT_def quotient_def + by (rule rev_finite_subset) (auto) +next + show "\x y. tag_str_ALT A B x = tag_str_ALT A B y \ x \(A \ B) y" + unfolding tag_str_ALT_def + unfolding str_eq_def + unfolding str_eq_rel_def + by auto +qed + + +subsubsection {* The inductive case for @{text "SEQ"}*} + +definition + tag_str_SEQ :: "lang \ lang \ string \ (lang \ lang set)" +where + "tag_str_SEQ L1 L2 \ + (\x. (\L1 `` {x}, {(\L2 `` {x - xa}) | xa. xa \ x \ xa \ L1}))" + +lemma Seq_in_cases: + assumes "x @ z \ A ;; B" + shows "(\ x' \ x. x' \ A \ (x - x') @ z \ B) \ + (\ z' \ z. (x @ z') \ A \ (z - z') \ B)" +using assms +unfolding Seq_def prefix_def +by (auto simp add: append_eq_append_conv2) + +lemma tag_str_SEQ_injI: + assumes eq_tag: "tag_str_SEQ A B x = tag_str_SEQ A B y" + shows "x \(A ;; B) y" +proof - + { fix x y z + assume xz_in_seq: "x @ z \ A ;; B" + and tag_xy: "tag_str_SEQ A B x = tag_str_SEQ A B y" + have"y @ z \ A ;; B" + proof - + { (* first case with x' in A and (x - x') @ z in B *) + fix x' + assume h1: "x' \ x" and h2: "x' \ A" and h3: "(x - x') @ z \ B" + obtain y' + where "y' \ y" + and "y' \ A" + and "(y - y') @ z \ B" + proof - + have "{\B `` {x - x'} |x'. x' \ x \ x' \ A} = + {\B `` {y - y'} |y'. y' \ y \ y' \ A}" (is "?Left = ?Right") + using tag_xy unfolding tag_str_SEQ_def by simp + moreover + have "\B `` {x - x'} \ ?Left" using h1 h2 by auto + ultimately + have "\B `` {x - x'} \ ?Right" by simp + then obtain y' + where eq_xy': "\B `` {x - x'} = \B `` {y - y'}" + and pref_y': "y' \ y" and y'_in: "y' \ A" + by simp blast + + have "(x - x') \B (y - y')" using eq_xy' + unfolding Image_def str_eq_rel_def str_eq_def by auto + with h3 have "(y - y') @ z \ B" + unfolding str_eq_rel_def str_eq_def by simp + with pref_y' y'_in + show ?thesis using that by blast + qed + then have "y @ z \ A ;; B" by (erule_tac prefixE) (auto simp: Seq_def) + } + moreover + { (* second case with x @ z' in A and z - z' in B *) + fix z' + assume h1: "z' \ z" and h2: "(x @ z') \ A" and h3: "z - z' \ B" + have "\A `` {x} = \A `` {y}" + using tag_xy unfolding tag_str_SEQ_def by simp + with h2 have "y @ z' \ A" + unfolding Image_def str_eq_rel_def str_eq_def by auto + with h1 h3 have "y @ z \ A ;; B" + unfolding prefix_def Seq_def + by (auto) (metis append_assoc) + } + ultimately show "y @ z \ A ;; B" + using Seq_in_cases [OF xz_in_seq] by blast + qed + } + from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]] + show "x \(A ;; B) y" unfolding str_eq_def str_eq_rel_def by blast +qed + +lemma quot_seq_finiteI [intro]: + fixes L1 L2::"lang" + assumes fin1: "finite (UNIV // \L1)" + and fin2: "finite (UNIV // \L2)" + shows "finite (UNIV // \(L1 ;; L2))" +proof (rule_tac tag = "tag_str_SEQ L1 L2" in tag_finite_imageD) + show "\x y. tag_str_SEQ L1 L2 x = tag_str_SEQ L1 L2 y \ x \(L1 ;; L2) y" + by (rule tag_str_SEQ_injI) +next + have *: "finite ((UNIV // \L1) \ (Pow (UNIV // \L2)))" + using fin1 fin2 by auto + show "finite (range (tag_str_SEQ L1 L2))" + unfolding tag_str_SEQ_def + apply(rule finite_subset[OF _ *]) + unfolding quotient_def + by auto +qed + + +subsubsection {* The inductive case for @{const "STAR"} *} + +definition + tag_str_STAR :: "lang \ string \ lang set" +where + "tag_str_STAR L1 \ (\x. {\L1 `` {x - xa} | xa. xa < x \ xa \ L1\})" + +text {* A technical lemma. *} +lemma finite_set_has_max: "\finite A; A \ {}\ \ + (\ max \ A. \ a \ A. f a <= (f max :: nat))" +proof (induct rule:finite.induct) + case emptyI thus ?case by simp +next + case (insertI A a) + show ?case + proof (cases "A = {}") + case True thus ?thesis by (rule_tac x = a in bexI, auto) + next + case False + with insertI.hyps and False + obtain max + where h1: "max \ A" + and h2: "\a\A. f a \ f max" by blast + show ?thesis + proof (cases "f a \ f max") + assume "f a \ f max" + with h1 h2 show ?thesis by (rule_tac x = max in bexI, auto) + next + assume "\ (f a \ f max)" + thus ?thesis using h2 by (rule_tac x = a in bexI, auto) + qed + qed +qed + + +text {* The following is a technical lemma, which helps to show the range finiteness of tag function. *} + +lemma finite_strict_prefix_set: "finite {xa. xa < (x::string)}" +apply (induct x rule:rev_induct, simp) +apply (subgoal_tac "{xa. xa < xs @ [x]} = {xa. xa < xs} \ {xs}") +by (auto simp:strict_prefix_def) + + +lemma tag_str_STAR_injI: + assumes eq_tag: "tag_str_STAR L\<^isub>1 v = tag_str_STAR L\<^isub>1 w" + shows "v \(L\<^isub>1\) w" +proof- + { fix x y z + assume xz_in_star: "x @ z \ L\<^isub>1\" + and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y" + have "y @ z \ L\<^isub>1\" + proof(cases "x = []") + case True + with tag_xy have "y = []" + by (auto simp add: tag_str_STAR_def strict_prefix_def) + thus ?thesis using xz_in_star True by simp + next + case False + let ?S = "{xa. xa < x \ xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\}" + have "finite ?S" + by (rule_tac B = "{xa. xa < x}" in finite_subset, + auto simp:finite_strict_prefix_set) + moreover have "?S \ {}" using False xz_in_star + by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def) + ultimately have "\ xa_max \ ?S. \ xa \ ?S. length xa \ length xa_max" + using finite_set_has_max by blast + then obtain xa_max + where h1: "xa_max < x" + and h2: "xa_max \ L\<^isub>1\" + and h3: "(x - xa_max) @ z \ L\<^isub>1\" + and h4:"\ xa < x. xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\ + \ length xa \ length xa_max" + by blast + obtain ya + where h5: "ya < y" and h6: "ya \ L\<^isub>1\" + and eq_xya: "(x - xa_max) \L\<^isub>1 (y - ya)" + proof- + from tag_xy have "{\L\<^isub>1 `` {x - xa} |xa. xa < x \ xa \ L\<^isub>1\} = + {\L\<^isub>1 `` {y - xa} |xa. xa < y \ xa \ L\<^isub>1\}" (is "?left = ?right") + by (auto simp:tag_str_STAR_def) + moreover have "\L\<^isub>1 `` {x - xa_max} \ ?left" using h1 h2 by auto + ultimately have "\L\<^isub>1 `` {x - xa_max} \ ?right" by simp + thus ?thesis using that + apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast + qed + have "(y - ya) @ z \ L\<^isub>1\" + proof- + obtain za zb where eq_zab: "z = za @ zb" + and l_za: "(y - ya)@za \ L\<^isub>1" and ls_zb: "zb \ L\<^isub>1\" + proof - + from h1 have "(x - xa_max) @ z \ []" + by (auto simp:strict_prefix_def elim:prefixE) + 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 + let ?za = "a - (x - xa_max)" and ?zb = "b" + have pfx: "(x - xa_max) \ a" (is "?P1") + and eq_z: "z = ?za @ ?zb" (is "?P2") + proof - + have "((x - xa_max) \ a \ (a - (x - xa_max)) @ b = z) \ + (a < (x - xa_max) \ ((x - xa_max) - a) @ z = b)" + using append_eq_dest[OF ab_max] by (auto simp:strict_prefix_def) + moreover { + assume np: "a < (x - xa_max)" + and b_eqs: "((x - xa_max) - a) @ z = b" + have "False" + proof - + let ?xa_max' = "xa_max @ a" + have "?xa_max' < x" + using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) + moreover have "?xa_max' \ L\<^isub>1\" + using a_in h2 by (simp add:star_intro3) + moreover have "(x - ?xa_max') @ z \ L\<^isub>1\" + using b_eqs b_in np h1 by (simp add:diff_diff_append) + moreover have "\ (length ?xa_max' \ length xa_max)" + using a_neq by simp + ultimately show ?thesis using h4 by blast + qed } + ultimately show ?P1 and ?P2 by auto + qed + hence "(x - xa_max)@?za \ L\<^isub>1" using a_in by (auto elim:prefixE) + with eq_xya have "(y - ya) @ ?za \ L\<^isub>1" + by (auto simp:str_eq_def str_eq_rel_def) + with eq_z and b_in + show ?thesis using that by blast + qed + have "((y - ya) @ za) @ zb \ L\<^isub>1\" using l_za ls_zb by blast + with eq_zab show ?thesis by simp + qed + with h5 h6 show ?thesis + by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE) + qed + } + from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]] + show ?thesis unfolding str_eq_def str_eq_rel_def by blast +qed + +lemma quot_star_finiteI [intro]: + assumes finite1: "finite (UNIV // \L1)" + shows "finite (UNIV // \(L1\))" +proof (rule_tac tag = "tag_str_STAR L1" in tag_finite_imageD) + show "\x y. tag_str_STAR L1 x = tag_str_STAR L1 y \ x \(L1\) y" + by (rule tag_str_STAR_injI) +next + have *: "finite (Pow (UNIV // \L1))" + using finite1 by auto + show "finite (range (tag_str_STAR L1))" + unfolding tag_str_STAR_def + apply(rule finite_subset[OF _ *]) + unfolding quotient_def + by auto +qed + +subsubsection{* The conclusion *} + +lemma Myhill_Nerode2: + fixes r::"rexp" + shows "finite (UNIV // \(L r))" +by (induct r) (auto) + + +theorem Myhill_Nerode: + shows "(\r::rexp. A = L r) \ finite (UNIV // \A)" +using Myhill_Nerode1 Myhill_Nerode2 by auto + +end diff -r 3b7477db3462 -r e122cb146ecc Theories/Prefix_subtract.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Theories/Prefix_subtract.thy Wed Mar 23 12:17:30 2011 +0000 @@ -0,0 +1,60 @@ +theory Prefix_subtract + imports Main "~~/src/HOL/Library/List_Prefix" +begin + + +section {* A small theory of prefix subtraction *} + +text {* + The notion of @{text "prefix_subtract"} makes + the second direction of the Myhill-Nerode theorem + more readable. +*} + +instantiation list :: (type) minus +begin + +fun minus_list :: "'a list \ 'a list \ 'a list" +where + "minus_list [] xs = []" +| "minus_list (x#xs) [] = x#xs" +| "minus_list (x#xs) (y#ys) = (if x = y then minus_list xs ys else (x#xs))" + +instance by default + +end + +lemma [simp]: "x - [] = x" +by (induct x) (auto) + +lemma [simp]: "(x @ y) - x = y" +by (induct x) (auto) + +lemma [simp]: "x - x = []" +by (induct x) (auto) + +lemma [simp]: "x = z @ y \ x - z = y " +by (induct x) (auto) + +lemma diff_prefix: + "\c \ a - b; b \ a\ \ b @ c \ a" +by (auto elim: prefixE) + +lemma diff_diff_append: + "\c < a - b; b < a\ \ (a - b) - c = a - (b @ c)" +apply (clarsimp simp:strict_prefix_def) +by (drule diff_prefix, auto elim:prefixE) + +lemma append_eq_cases: + assumes a: "x @ y = m @ n" + shows "x \ m \ m \ x" +unfolding prefix_def using a +by (auto simp add: append_eq_append_conv2) + +lemma append_eq_dest: + assumes a: "x @ y = m @ n" + shows "(x \ m \ (m - x) @ n = y) \ (m \ x \ (x - m) @ y = n)" +using append_eq_cases[OF a] a +by (auto elim: prefixE) + +end diff -r 3b7477db3462 -r e122cb146ecc Theories/Regular.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Theories/Regular.thy Wed Mar 23 12:17:30 2011 +0000 @@ -0,0 +1,311 @@ +theory Regular +imports Main Folds +begin + +section {* Preliminary definitions *} + +type_synonym 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) + +lemma seq_null [simp]: + shows "A ;; {} = {}" + and "{} ;; 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 +by (induct rule: star_induct) (blast)+ + +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 seq_Union_left +unfolding seq_pow_comm 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 modified version of Arden's lemma *} + +text {* A helper lemma for Arden *} + +lemma arden_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 arden: + 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 arden_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 arden_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 function @{text L} is overloaded, with the idea that @{text "L x"} + evaluates to the language represented by the object @{text x}. +*} + +consts L:: "'a \ lang" + +overloading L_rexp \ "L:: rexp \ lang" +begin +fun + L_rexp :: "rexp \ lang" +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 + + +text {* ALT-combination for a set of regular expressions *} + +abbreviation + Setalt ("\_" [1000] 999) +where + "\A \ folds ALT NULL A" + +text {* + For finite sets, @{term Setalt} is preserved under @{term L}. +*} + +lemma folds_alt_simp [simp]: + fixes rs::"rexp set" + assumes a: "finite rs" + shows "L (\rs) = \ (L ` rs)" +unfolding folds_def +apply(rule set_eqI) +apply(rule someI2_ex) +apply(rule_tac finite_imp_fold_graph[OF a]) +apply(erule fold_graph.induct) +apply(auto) +done + +end \ No newline at end of file