1(*  Title:      HOL/UNITY/FP.thy
2    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3    Copyright   1998  University of Cambridge
4
5From Misra, "A Logic for Concurrent Programming", 1994
6*)
7
8section\<open>Fixed Point of a Program\<close>
9
10theory FP imports UNITY begin
11
12definition FP_Orig :: "'a program => 'a set" where
13    "FP_Orig F == \<Union>{A. \<forall>B. F \<in> stable (A \<inter> B)}"
14
15definition FP :: "'a program => 'a set" where
16    "FP F == {s. F \<in> stable {s}}"
17
18lemma stable_FP_Orig_Int: "F \<in> stable (FP_Orig F Int B)"
19apply (simp only: FP_Orig_def stable_def Int_Union2)
20apply (blast intro: constrains_UN)
21done
22
23lemma FP_Orig_weakest:
24    "(\<And>B. F \<in> stable (A \<inter> B)) \<Longrightarrow> A <= FP_Orig F"
25by (simp add: FP_Orig_def stable_def, blast)
26
27lemma stable_FP_Int: "F \<in> stable (FP F \<inter> B)"
28proof -
29  have "F \<in> stable (\<Union>x\<in>B. FP F \<inter> {x})"
30    apply (simp only: Int_insert_right FP_def stable_def)
31    apply (rule constrains_UN)
32    apply simp
33    done
34  also have "(\<Union>x\<in>B. FP F \<inter> {x}) = FP F \<inter> B"
35    by simp
36  finally show ?thesis .
37qed
38
39lemma FP_equivalence: "FP F = FP_Orig F"
40apply (rule equalityI) 
41 apply (rule stable_FP_Int [THEN FP_Orig_weakest])
42apply (simp add: FP_Orig_def FP_def, clarify)
43apply (drule_tac x = "{x}" in spec)
44apply (simp add: Int_insert_right)
45done
46
47lemma FP_weakest:
48    "(\<And>B. F \<in> stable (A Int B)) \<Longrightarrow> A <= FP F"
49by (simp add: FP_equivalence FP_Orig_weakest)
50
51lemma Compl_FP: 
52    "-(FP F) = (UN act: Acts F. -{s. act``{s} <= {s}})"
53by (simp add: FP_def stable_def constrains_def, blast)
54
55lemma Diff_FP: "A - (FP F) = (UN act: Acts F. A - {s. act``{s} <= {s}})"
56by (simp add: Diff_eq Compl_FP)
57
58lemma totalize_FP [simp]: "FP (totalize F) = FP F"
59by (simp add: FP_def)
60
61end
62