10251
|
1 |
header {* \title{} *}
|
|
2 |
|
|
3 |
theory While_Combinator_Example = While_Combinator:
|
|
4 |
|
|
5 |
text {*
|
10653
|
6 |
\medskip An application: computation of the @{term lfp} on finite
|
|
7 |
sets via iteration.
|
|
8 |
*}
|
|
9 |
|
|
10 |
theorem lfp_conv_while:
|
|
11 |
"mono f ==> finite U ==> f U = U ==>
|
|
12 |
lfp f = fst (while (\<lambda>(A, fA). A \<noteq> fA) (\<lambda>(A, fA). (fA, f fA)) ({}, f {}))"
|
|
13 |
apply (rule_tac P = "\<lambda>(A, B). (A \<subseteq> U \<and> B = f A \<and> A \<subseteq> B \<and> B \<subseteq> lfp f)" and
|
|
14 |
r = "((Pow U <*> UNIV) <*> (Pow U <*> UNIV)) \<inter>
|
|
15 |
inv_image finite_psubset (op - U o fst)" in while_rule)
|
|
16 |
apply (subst lfp_unfold)
|
|
17 |
apply assumption
|
|
18 |
apply (simp add: monoD)
|
|
19 |
apply (subst lfp_unfold)
|
|
20 |
apply assumption
|
|
21 |
apply clarsimp
|
|
22 |
apply (blast dest: monoD)
|
|
23 |
apply (fastsimp intro!: lfp_lowerbound)
|
|
24 |
apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset])
|
|
25 |
apply (clarsimp simp add: inv_image_def finite_psubset_def order_less_le)
|
|
26 |
apply (blast intro!: finite_Diff dest: monoD)
|
|
27 |
done
|
|
28 |
|
|
29 |
text {*
|
10251
|
30 |
An example of using the @{term while} combinator.
|
|
31 |
*}
|
|
32 |
|
10392
|
33 |
lemma aux: "{f n | n. A n \<or> B n} = {f n | n. A n} \<union> {f n | n. B n}"
|
10251
|
34 |
apply blast
|
|
35 |
done
|
|
36 |
|
10392
|
37 |
theorem "P (lfp (\<lambda>N::int set. {#0} \<union> {(n + #2) mod #6 | n. n \<in> N})) =
|
10251
|
38 |
P {#0, #4, #2}"
|
|
39 |
apply (subst lfp_conv_while [where ?U = "{#0, #1, #2, #3, #4, #5}"])
|
|
40 |
apply (rule monoI)
|
|
41 |
apply blast
|
|
42 |
apply simp
|
|
43 |
apply (simp add: aux set_eq_subset)
|
|
44 |
txt {* The fixpoint computation is performed purely by rewriting: *}
|
|
45 |
apply (simp add: while_unfold aux set_eq_subset del: subset_empty)
|
|
46 |
done
|
|
47 |
|
|
48 |
end
|