Theory FP
section‹Fixed Point of a Program›
theory FP imports UNITY begin
definition FP_Orig :: "'a program => 'a set" where
"FP_Orig F == ⋃{A. ∀B. F ∈ stable (A ∩ B)}"
definition FP :: "'a program => 'a set" where
"FP F == {s. F ∈ stable {s}}"
lemma stable_FP_Orig_Int: "F ∈ stable (FP_Orig F Int B)"
apply (simp only: FP_Orig_def stable_def Int_Union2)
apply (blast intro: constrains_UN)
done
lemma FP_Orig_weakest:
"(⋀B. F ∈ stable (A ∩ B)) ⟹ A <= FP_Orig F"
by (simp add: FP_Orig_def stable_def, blast)
lemma stable_FP_Int: "F ∈ stable (FP F ∩ B)"
proof -
have "F ∈ stable (⋃x∈B. FP F ∩ {x})"
apply (simp only: Int_insert_right FP_def stable_def)
apply (rule constrains_UN)
apply simp
done
also have "(⋃x∈B. FP F ∩ {x}) = FP F ∩ B"
by simp
finally show ?thesis .
qed
lemma FP_equivalence: "FP F = FP_Orig F"
apply (rule equalityI)
apply (rule stable_FP_Int [THEN FP_Orig_weakest])
apply (simp add: FP_Orig_def FP_def, clarify)
apply (drule_tac x = "{x}" in spec)
apply (simp add: Int_insert_right)
done
lemma FP_weakest:
"(⋀B. F ∈ stable (A Int B)) ⟹ A <= FP F"
by (simp add: FP_equivalence FP_Orig_weakest)
lemma Compl_FP:
"-(FP F) = (UN act: Acts F. -{s. act``{s} <= {s}})"
by (simp add: FP_def stable_def constrains_def, blast)
lemma Diff_FP: "A - (FP F) = (UN act: Acts F. A - {s. act``{s} <= {s}})"
by (simp add: Diff_eq Compl_FP)
lemma totalize_FP [simp]: "FP (totalize F) = FP F"
by (simp add: FP_def)
end