--- a/src/HOL/Library/Library.thy Fri Jan 26 15:50:28 2001 +0100
+++ b/src/HOL/Library/Library.thy Fri Jan 26 15:50:52 2001 +0100
@@ -7,6 +7,6 @@
Nested_Environment +
Accessible_Part +
Multiset +
- While_Combinator + While_Combinator_Example:
+ While_Combinator:
end
(*>*)
--- a/src/HOL/Library/ROOT.ML Fri Jan 26 15:50:28 2001 +0100
+++ b/src/HOL/Library/ROOT.ML Fri Jan 26 15:50:52 2001 +0100
@@ -1,2 +1,1 @@
-
use_thy "Library";
--- a/src/HOL/Library/While_Combinator.thy Fri Jan 26 15:50:28 2001 +0100
+++ b/src/HOL/Library/While_Combinator.thy Fri Jan 26 15:50:52 2001 +0100
@@ -67,14 +67,16 @@
apply blast
done
+hide const while_aux
+
text {*
The proof rule for @{term while}, where @{term P} is the invariant.
*}
theorem while_rule_lemma[rule_format]:
- "(!!s. P s ==> b s ==> P (c s)) ==>
- (!!s. P s ==> \<not> b s ==> Q s) ==>
- wf {(t, s). P s \<and> b s \<and> t = c s} ==>
+ "[| !!s. P s ==> b s ==> P (c s);
+ !!s. P s ==> \<not> b s ==> Q s;
+ wf {(t, s). P s \<and> b s \<and> t = c s} |] ==>
P s --> Q (while b c s)"
proof -
case antecedent
@@ -89,10 +91,12 @@
qed
theorem while_rule:
- "[| P s; !!s. [| P s; b s |] ==> P (c s);
- !!s. [| P s; \<not> b s |] ==> Q s;
- wf r; !!s. [| P s; b s |] ==> (c s, s) \<in> r |] ==>
- Q (while b c s)"
+ "[| P s;
+ !!s. [| P s; b s |] ==> P (c s);
+ !!s. [| P s; \<not> b s |] ==> Q s;
+ wf r;
+ !!s. [| P s; b s |] ==> (c s, s) \<in> r |] ==>
+ Q (while b c s)"
apply (rule while_rule_lemma)
prefer 4 apply assumption
apply blast
@@ -101,6 +105,50 @@
apply blast
done
-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 (\<lambda>(A, fA). A \<noteq> fA) (\<lambda>(A, fA). (fA, f fA)) ({}, f {}))"
+apply (rule_tac P = "\<lambda>(A, B). (A \<subseteq> U \<and> B = f A \<and> A \<subseteq> B \<and> B \<subseteq> lfp f)" and
+ r = "((Pow U <*> UNIV) <*> (Pow U <*> UNIV)) \<inter>
+ 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: inv_image_def finite_psubset_def order_less_le)
+apply (blast intro!: finite_Diff dest: monoD)
+done
+
+
+(*
+text {*
+ An example of using the @{term while} combinator.
+*}
+
+lemma aux: "{f n | n. A n \<or> B n} = {f n | n. A n} \<union> {f n | n. B n}"
+ apply blast
+ done
+
+theorem "P (lfp (\<lambda>N::int set. {#0} \<union> {(n + #2) mod #6 | n. n \<in> 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
--- a/src/HOL/Library/While_Combinator_Example.thy Fri Jan 26 15:50:28 2001 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,49 +0,0 @@
-
-header {* \title{} *}
-
-theory While_Combinator_Example = While_Combinator:
-
-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 (\<lambda>(A, fA). A \<noteq> fA) (\<lambda>(A, fA). (fA, f fA)) ({}, f {}))"
-apply (rule_tac P = "\<lambda>(A, B). (A \<subseteq> U \<and> B = f A \<and> A \<subseteq> B \<and> B \<subseteq> lfp f)" and
- r = "((Pow U <*> UNIV) <*> (Pow U <*> UNIV)) \<inter>
- 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: inv_image_def finite_psubset_def order_less_le)
-apply (blast intro!: finite_Diff dest: monoD)
-done
-
-text {*
- An example of using the @{term while} combinator.
-*}
-
-lemma aux: "{f n | n. A n \<or> B n} = {f n | n. A n} \<union> {f n | n. B n}"
- apply blast
- done
-
-theorem "P (lfp (\<lambda>N::int set. {#0} \<union> {(n + #2) mod #6 | n. n \<in> 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