# HG changeset patch # User krauss # Date 1278688503 -7200 # Node ID 8380686be5cd6a85ab2357e38ea9b7ee9b152ec8 # Parent 00ff97087ab54925821f8ceb5f0c81c578a5ff4d moved example to its own file in HOL/ex diff -r 00ff97087ab5 -r 8380686be5cd src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Jul 09 17:00:42 2010 +0200 +++ b/src/HOL/IsaMakefile Fri Jul 09 17:15:03 2010 +0200 @@ -1006,7 +1006,8 @@ ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy \ ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy ex/Sudoku.thy \ ex/Tarski.thy ex/Termination.thy ex/Transfer_Ex.thy ex/Tree23.thy \ - ex/Unification.thy ex/document/root.bib ex/document/root.tex \ + ex/Unification.thy ex/While_Combinator_Example.thy ex/document/root.bib \ + ex/document/root.tex \ ex/set.thy ex/svc_funcs.ML ex/svc_test.thy @$(ISABELLE_TOOL) usedir $(OUT)/HOL ex diff -r 00ff97087ab5 -r 8380686be5cd src/HOL/Library/While_Combinator.thy --- a/src/HOL/Library/While_Combinator.thy Fri Jul 09 17:00:42 2010 +0200 +++ b/src/HOL/Library/While_Combinator.thy Fri Jul 09 17:15:03 2010 +0200 @@ -10,7 +10,7 @@ imports Main begin -subsection {* Option result *} +subsection {* Partial version *} definition while_option :: "('a \ bool) \ ('a \ 'a) \ 'a \ 'a option" where "while_option b c s = (if (\k. ~ b ((c ^^ k) s)) @@ -81,7 +81,7 @@ qed -subsection {* Totalized version *} +subsection {* Total version *} definition while :: "('a \ bool) \ ('a \ 'a) \ 'a \ 'a" where "while b c s = the (while_option b c s)" @@ -127,54 +127,5 @@ apply blast done -text {* - \medskip An application: computation of the @{term lfp} on finite - sets via iteration. -*} - -theorem lfp_conv_while: - "[| mono f; finite U; f U = U |] ==> - lfp f = fst (while (\(A, fA). A \ fA) (\(A, fA). (fA, f fA)) ({}, f {}))" -apply (rule_tac P = "\(A, B). (A \ U \ B = f A \ A \ B \ B \ lfp f)" and - r = "((Pow U \ UNIV) \ (Pow U \ UNIV)) \ - inv_image finite_psubset (op - U o fst)" in while_rule) - apply (subst lfp_unfold) - apply assumption - apply (simp add: monoD) - apply (subst lfp_unfold) - apply assumption - apply clarsimp - apply (blast dest: monoD) - apply (fastsimp intro!: lfp_lowerbound) - apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset]) -apply (clarsimp simp add: finite_psubset_def order_less_le) -apply (blast intro!: finite_Diff dest: monoD) -done - - -subsection {* Example *} - -text{* Cannot use @{thm[source]set_eq_subset} because it leads to -looping because the antisymmetry simproc turns the subset relationship -back into equality. *} - -theorem "P (lfp (\N::int set. {0} \ {(n + 2) mod 6 | n. n \ N})) = - P {0, 4, 2}" -proof - - have seteq: "!!A B. (A = B) = ((!a : A. a:B) & (!b:B. b:A))" - by blast - have aux: "!!f A B. {f n | n. A n \ B n} = {f n | n. A n} \ {f n | n. B n}" - apply blast - done - show ?thesis - apply (subst lfp_conv_while [where ?U = "{0, 1, 2, 3, 4, 5}"]) - apply (rule monoI) - apply blast - apply simp - apply (simp add: aux set_eq_subset) - txt {* The fixpoint computation is performed purely by rewriting: *} - apply (simp add: while_unfold aux seteq del: subset_empty) - done -qed end diff -r 00ff97087ab5 -r 8380686be5cd src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Fri Jul 09 17:00:42 2010 +0200 +++ b/src/HOL/ex/ROOT.ML Fri Jul 09 17:15:03 2010 +0200 @@ -23,6 +23,7 @@ "InductiveInvariant_examples", "LocaleTest2", "Records", + "While_Combinator_Example", "MonoidGroup", "BinEx", "Hex_Bin_Examples", diff -r 00ff97087ab5 -r 8380686be5cd src/HOL/ex/While_Combinator_Example.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/While_Combinator_Example.thy Fri Jul 09 17:15:03 2010 +0200 @@ -0,0 +1,60 @@ +(* Title: HOL/Library/While_Combinator.thy + Author: Tobias Nipkow + Copyright 2000 TU Muenchen +*) + +header {* An application of the While combinator *} + +theory While_Combinator_Example +imports While_Combinator +begin + +text {* Computation of the @{term lfp} on finite sets via + iteration. *} + +theorem lfp_conv_while: + "[| mono f; finite U; f U = U |] ==> + lfp f = fst (while (\(A, fA). A \ fA) (\(A, fA). (fA, f fA)) ({}, f {}))" +apply (rule_tac P = "\(A, B). (A \ U \ B = f A \ A \ B \ B \ lfp f)" and + r = "((Pow U \ UNIV) \ (Pow U \ UNIV)) \ + inv_image finite_psubset (op - U o fst)" in while_rule) + apply (subst lfp_unfold) + apply assumption + apply (simp add: monoD) + apply (subst lfp_unfold) + apply assumption + apply clarsimp + apply (blast dest: monoD) + apply (fastsimp intro!: lfp_lowerbound) + apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset]) +apply (clarsimp simp add: finite_psubset_def order_less_le) +apply (blast intro!: finite_Diff dest: monoD) +done + + +subsection {* Example *} + +text{* Cannot use @{thm[source]set_eq_subset} because it leads to +looping because the antisymmetry simproc turns the subset relationship +back into equality. *} + +theorem "P (lfp (\N::int set. {0} \ {(n + 2) mod 6 | n. n \ N})) = + P {0, 4, 2}" +proof - + have seteq: "!!A B. (A = B) = ((!a : A. a:B) & (!b:B. b:A))" + by blast + have aux: "!!f A B. {f n | n. A n \ B n} = {f n | n. A n} \ {f n | n. B n}" + apply blast + done + show ?thesis + apply (subst lfp_conv_while [where ?U = "{0, 1, 2, 3, 4, 5}"]) + apply (rule monoI) + apply blast + apply simp + apply (simp add: aux set_eq_subset) + txt {* The fixpoint computation is performed purely by rewriting: *} + apply (simp add: while_unfold aux seteq del: subset_empty) + done +qed + +end \ No newline at end of file