wenzelm@10251: (* Title: HOL/Library/While.thy wenzelm@10251: ID: $Id$ wenzelm@10251: Author: Tobias Nipkow wenzelm@10251: Copyright 2000 TU Muenchen wenzelm@10251: *) wenzelm@10251: wenzelm@10251: header {* wenzelm@10251: \title{A general ``while'' combinator} wenzelm@10251: \author{Tobias Nipkow} wenzelm@10251: *} wenzelm@10251: wenzelm@10251: theory While_Combinator = Main: wenzelm@10251: wenzelm@10251: text {* wenzelm@10251: We define a while-combinator @{term while} and prove: (a) an wenzelm@10251: unrestricted unfolding law (even if while diverges!) (I got this wenzelm@10251: idea from Wolfgang Goerigk), and (b) the invariant rule for reasoning wenzelm@10251: about @{term while}. wenzelm@10251: *} wenzelm@10251: wenzelm@10251: consts while_aux :: "('a => bool) \ ('a => 'a) \ 'a => 'a" wenzelm@10251: recdef while_aux wenzelm@10251: "same_fst (\b. True) (\b. same_fst (\c. True) (\c. wenzelm@10251: {(t, s). b s \ c s = t \ wenzelm@10251: \ (\f. f 0 = s \ (\i. b (f i) \ c (f i) = f (i + 1)))}))" wenzelm@10251: "while_aux (b, c, s) = wenzelm@10251: (if (\f. f 0 = s \ (\i. b (f i) \ c (f i) = f (i + 1))) wenzelm@10251: then arbitrary wenzelm@10251: else if b s then while_aux (b, c, c s) wenzelm@10251: else s)" wenzelm@10251: wenzelm@10774: recdef_tc while_aux_tc: while_aux wenzelm@10774: apply (rule wf_same_fst) wenzelm@10774: apply (rule wf_same_fst) wenzelm@10774: apply (simp add: wf_iff_no_infinite_down_chain) wenzelm@10774: apply blast wenzelm@10774: done wenzelm@10774: wenzelm@10251: constdefs wenzelm@10251: while :: "('a => bool) => ('a => 'a) => 'a => 'a" wenzelm@10251: "while b c s == while_aux (b, c, s)" wenzelm@10251: wenzelm@10251: lemma while_aux_unfold: wenzelm@10251: "while_aux (b, c, s) = wenzelm@10251: (if \f. f 0 = s \ (\i. b (f i) \ c (f i) = f (i + 1)) wenzelm@10251: then arbitrary wenzelm@10251: else if b s then while_aux (b, c, c s) wenzelm@10251: else s)" wenzelm@10251: apply (rule while_aux_tc [THEN while_aux.simps [THEN trans]]) wenzelm@10251: apply (simp add: same_fst_def) wenzelm@10251: apply (rule refl) wenzelm@10251: done wenzelm@10251: wenzelm@10251: text {* wenzelm@10251: The recursion equation for @{term while}: directly executable! wenzelm@10251: *} wenzelm@10251: wenzelm@10251: theorem while_unfold: wenzelm@10251: "while b c s = (if b s then while b c (c s) else s)" wenzelm@10251: apply (unfold while_def) wenzelm@10251: apply (rule while_aux_unfold [THEN trans]) wenzelm@10251: apply auto wenzelm@10251: apply (subst while_aux_unfold) wenzelm@10251: apply simp wenzelm@10251: apply clarify wenzelm@10251: apply (erule_tac x = "\i. f (Suc i)" in allE) wenzelm@10251: apply blast wenzelm@10251: done wenzelm@10251: nipkow@10984: hide const while_aux nipkow@10984: wenzelm@10251: text {* wenzelm@10251: The proof rule for @{term while}, where @{term P} is the invariant. wenzelm@10251: *} wenzelm@10251: nipkow@10653: theorem while_rule_lemma[rule_format]: nipkow@10984: "[| !!s. P s ==> b s ==> P (c s); nipkow@10984: !!s. P s ==> \ b s ==> Q s; nipkow@10984: wf {(t, s). P s \ b s \ t = c s} |] ==> wenzelm@10251: P s --> Q (while b c s)" wenzelm@10251: proof - wenzelm@10251: case antecedent wenzelm@10251: assume wf: "wf {(t, s). P s \ b s \ t = c s}" wenzelm@10251: show ?thesis wenzelm@10251: apply (induct s rule: wf [THEN wf_induct]) wenzelm@10251: apply simp wenzelm@10251: apply clarify wenzelm@10251: apply (subst while_unfold) wenzelm@10251: apply (simp add: antecedent) wenzelm@10251: done wenzelm@10251: qed wenzelm@10251: nipkow@10653: theorem while_rule: nipkow@10984: "[| P s; nipkow@10984: !!s. [| P s; b s |] ==> P (c s); nipkow@10984: !!s. [| P s; \ b s |] ==> Q s; wenzelm@10997: wf r; nipkow@10984: !!s. [| P s; b s |] ==> (c s, s) \ r |] ==> nipkow@10984: Q (while b c s)" nipkow@10653: apply (rule while_rule_lemma) nipkow@10653: prefer 4 apply assumption nipkow@10653: apply blast nipkow@10653: apply blast nipkow@10653: apply(erule wf_subset) nipkow@10653: apply blast nipkow@10653: done nipkow@10653: nipkow@10984: text {* nipkow@10984: \medskip An application: computation of the @{term lfp} on finite nipkow@10984: sets via iteration. nipkow@10984: *} nipkow@10984: nipkow@10984: theorem lfp_conv_while: nipkow@10984: "[| mono f; finite U; f U = U |] ==> nipkow@10984: lfp f = fst (while (\(A, fA). A \ fA) (\(A, fA). (fA, f fA)) ({}, f {}))" nipkow@10984: apply (rule_tac P = "\(A, B). (A \ U \ B = f A \ A \ B \ B \ lfp f)" and nipkow@10984: r = "((Pow U <*> UNIV) <*> (Pow U <*> UNIV)) \ nipkow@10984: inv_image finite_psubset (op - U o fst)" in while_rule) nipkow@10984: apply (subst lfp_unfold) nipkow@10984: apply assumption nipkow@10984: apply (simp add: monoD) nipkow@10984: apply (subst lfp_unfold) nipkow@10984: apply assumption nipkow@10984: apply clarsimp nipkow@10984: apply (blast dest: monoD) nipkow@10984: apply (fastsimp intro!: lfp_lowerbound) nipkow@10984: apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset]) nipkow@10984: apply (clarsimp simp add: inv_image_def finite_psubset_def order_less_le) nipkow@10984: apply (blast intro!: finite_Diff dest: monoD) nipkow@10984: done nipkow@10984: nipkow@10984: nipkow@10984: text {* wenzelm@10997: An example of using the @{term while} combinator.\footnote{It is safe wenzelm@10997: to keep the example here, since there is no effect on the current wenzelm@10997: theory.} nipkow@10984: *} nipkow@10984: nipkow@10984: theorem "P (lfp (\N::int set. {#0} \ {(n + #2) mod #6 | n. n \ N})) = nipkow@10984: P {#0, #4, #2}" wenzelm@10997: proof - wenzelm@10997: have aux: "!!f A B. {f n | n. A n \ B n} = {f n | n. A n} \ {f n | n. B n}" nipkow@10984: apply blast wenzelm@10997: done wenzelm@10997: show ?thesis wenzelm@10997: apply (subst lfp_conv_while [where ?U = "{#0, #1, #2, #3, #4, #5}"]) wenzelm@10997: apply (rule monoI) wenzelm@10997: apply blast wenzelm@10997: apply simp wenzelm@10997: apply (simp add: aux set_eq_subset) wenzelm@10997: txt {* The fixpoint computation is performed purely by rewriting: *} wenzelm@10997: apply (simp add: while_unfold aux set_eq_subset del: subset_empty) wenzelm@10997: done wenzelm@10997: qed wenzelm@10251: wenzelm@10251: end