diff -r b04cc5e4e84c -r 7743d2ad71d1 Theories/Regular.thy --- a/Theories/Regular.thy Tue May 31 20:32:49 2011 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,311 +0,0 @@ -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