# HG changeset patch # User wenzelm # Date 971904589 -7200 # Node ID 5cc44cae9590be5f7e8e37545a99acb01f3ff3ff # Parent ca93fe25a84bac9969643c5cec7bbce52c7d3757 A general ``while'' combinator (from main HOL); diff -r ca93fe25a84b -r 5cc44cae9590 src/HOL/Library/While_Combinator.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/While_Combinator.thy Wed Oct 18 23:29:49 2000 +0200 @@ -0,0 +1,121 @@ +(* Title: HOL/Library/While.thy + ID: $Id$ + Author: Tobias Nipkow + Copyright 2000 TU Muenchen +*) + +header {* + \title{A general ``while'' combinator} + \author{Tobias Nipkow} +*} + +theory While_Combinator = Main: + +text {* + We define a while-combinator @{term while} and prove: (a) an + unrestricted unfolding law (even if while diverges!) (I got this + idea from Wolfgang Goerigk), and (b) the invariant rule for reasoning + about @{term while}. +*} + +consts while_aux :: "('a => bool) \ ('a => 'a) \ 'a => 'a" +recdef while_aux + "same_fst (\b. True) (\b. same_fst (\c. True) (\c. + {(t, s). b s \ c s = t \ + \ (\f. f 0 = s \ (\i. b (f i) \ c (f i) = f (i + 1)))}))" + "while_aux (b, c, s) = + (if (\f. f 0 = s \ (\i. b (f i) \ c (f i) = f (i + 1))) + then arbitrary + else if b s then while_aux (b, c, c s) + else s)" + +constdefs + while :: "('a => bool) => ('a => 'a) => 'a => 'a" + "while b c s == while_aux (b, c, s)" + +ML_setup {* + goalw_cterm [] (cterm_of (sign_of (the_context ())) + (HOLogic.mk_Trueprop (hd while_aux.tcs))); + br wf_same_fstI 1; + br wf_same_fstI 1; + by (asm_full_simp_tac (simpset() addsimps [wf_iff_no_infinite_down_chain]) 1); + by (Blast_tac 1); + qed "while_aux_tc"; +*} (* FIXME cannot prove recdef tcs in Isar yet! *) + +lemma while_aux_unfold: + "while_aux (b, c, s) = + (if \f. f 0 = s \ (\i. b (f i) \ c (f i) = f (i + 1)) + then arbitrary + else if b s then while_aux (b, c, c s) + else s)" + apply (rule while_aux_tc [THEN while_aux.simps [THEN trans]]) + apply (simp add: same_fst_def) + apply (rule refl) + done + +text {* + The recursion equation for @{term while}: directly executable! +*} + +theorem while_unfold: + "while b c s = (if b s then while b c (c s) else s)" + apply (unfold while_def) + apply (rule while_aux_unfold [THEN trans]) + apply auto + apply (subst while_aux_unfold) + apply simp + apply clarify + apply (erule_tac x = "\i. f (Suc i)" in allE) + apply blast + done + +text {* + The proof rule for @{term while}, where @{term P} is the invariant. +*} + +theorem while_rule [rule_format]: + "(!!s. P s ==> b s ==> P (c s)) ==> + (!!s. P s ==> \ b s ==> Q s) ==> + wf {(t, s). P s \ b s \ t = c s} ==> + P s --> Q (while b c s)" +proof - + case antecedent + assume wf: "wf {(t, s). P s \ b s \ t = c s}" + show ?thesis + apply (induct s rule: wf [THEN wf_induct]) + apply simp + apply clarify + apply (subst while_unfold) + apply (simp add: antecedent) + done +qed + +hide const while_aux + +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)" in while_rule) + apply (subst lfp_unfold) + apply assumption + apply clarsimp + apply (blast dest: monoD) + apply (fastsimp intro!: lfp_lowerbound) + apply (rule_tac r = "((Pow U <*> UNIV) <*> (Pow U <*> UNIV)) \ + inv_image finite_psubset (op - U o fst)" in wf_subset) + apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset]) + apply (clarsimp simp add: inv_image_def finite_psubset_def order_less_le) + apply (blast intro!: finite_Diff dest: monoD) + apply (subst lfp_unfold) + apply assumption + apply (simp add: monoD) + done + +end diff -r ca93fe25a84b -r 5cc44cae9590 src/HOL/Library/While_Combinator_Example.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/While_Combinator_Example.thy Wed Oct 18 23:29:49 2000 +0200 @@ -0,0 +1,25 @@ + +header {* \title{} *} + +theory While_Combinator_Example = While_Combinator: + +text {* + An example of using the @{term while} combinator. +*} + +lemma aux: "{f n| n. A n \ B n} = {f n| n. A n} \ {f n| n. B n}" + apply blast + done + +theorem "P (lfp (\N::int set. {#0} \ {(n + #2) mod #6| n. n \ N})) = + P {#0, #4, #2}" + 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 set_eq_subset del: subset_empty) + done + +end