# HG changeset patch # User urbanc # Date 1297158687 0 # Node ID 1589bf5c1ad81399904a2d0561549a40315befd4 # Parent d63baacbdb165d1f38f0fb4ce679e0f3a60382a5 added an abbreviation for folds ALT NULL diff -r d63baacbdb16 -r 1589bf5c1ad8 Myhill_1.thy --- a/Myhill_1.thy Mon Feb 07 20:30:10 2011 +0000 +++ b/Myhill_1.thy Tue Feb 08 09:51:27 2011 +0000 @@ -320,11 +320,17 @@ 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 @@ -333,7 +339,7 @@ lemma folds_alt_simp [simp]: assumes a: "finite rs" - shows "L (folds ALT NULL rs) = \ (L ` rs)" + shows "L (\rs) = \ (L ` rs)" apply(rule set_eqI) apply(simp add: folds_def) apply(rule someI2_ex) @@ -382,6 +388,12 @@ apply(auto) done +lemma finals_included_in_UNIV: + shows "finals A \ UNIV // \A" +unfolding finals_def +unfolding quotient_def +by auto + section {* Direction @{text "finite partition \ regular language"}*} @@ -503,7 +515,7 @@ text {* The following @{text "items_of rhs X"} returns all @{text "X"}-items in @{text "rhs"}. - *} +*} definition "items_of rhs X \ {Trn X r | r. (Trn X r) \ rhs}" @@ -515,7 +527,7 @@ *} definition - "rexp_of rhs X \ folds ALT NULL ((snd o the_Trn) ` items_of rhs X)" + "rexp_of rhs X \ \((snd o the_Trn) ` items_of rhs X)" text {* The following @{text "lam_of rhs"} returns all pure regular expression items in @{text "rhs"}. @@ -531,8 +543,9 @@ is used to compute compute the regular expression corresponds to @{text "rhs"}. *} + definition - "rexp_of_lam rhs \ folds ALT NULL (the_r ` lam_of rhs)" + "rexp_of_lam rhs \ \(the_r ` lam_of rhs)" text {* The following @{text "attach_rexp rexp' itm"} attach @@ -1282,7 +1295,7 @@ by (auto dest: bchoice) def rs \ "f ` (finals A)" have "A = \ (finals A)" using lang_is_union_of_finals by auto - also have "\ = L (folds ALT NULL rs)" + also have "\ = L (\rs)" proof - have "finite rs" proof -