src/HOL/UNITY/FP.thy
changeset 67613 ce654b0e6d69
parent 63146 f1ecba0272f9
child 69661 a03a63b81f44
equal deleted inserted replaced
67610:4939494ed791 67613:ce654b0e6d69
     8 section\<open>Fixed Point of a Program\<close>
     8 section\<open>Fixed Point of a Program\<close>
     9 
     9 
    10 theory FP imports UNITY begin
    10 theory FP imports UNITY begin
    11 
    11 
    12 definition FP_Orig :: "'a program => 'a set" where
    12 definition FP_Orig :: "'a program => 'a set" where
    13     "FP_Orig F == \<Union>{A. ALL B. F : stable (A Int B)}"
    13     "FP_Orig F == \<Union>{A. \<forall>B. F \<in> stable (A \<inter> B)}"
    14 
    14 
    15 definition FP :: "'a program => 'a set" where
    15 definition FP :: "'a program => 'a set" where
    16     "FP F == {s. F : stable {s}}"
    16     "FP F == {s. F \<in> stable {s}}"
    17 
    17 
    18 lemma stable_FP_Orig_Int: "F : stable (FP_Orig F Int B)"
    18 lemma stable_FP_Orig_Int: "F \<in> stable (FP_Orig F Int B)"
    19 apply (simp only: FP_Orig_def stable_def Int_Union2)
    19 apply (simp only: FP_Orig_def stable_def Int_Union2)
    20 apply (blast intro: constrains_UN)
    20 apply (blast intro: constrains_UN)
    21 done
    21 done
    22 
    22 
    23 lemma FP_Orig_weakest:
    23 lemma FP_Orig_weakest:
    24     "(!!B. F : stable (A Int B)) ==> A <= FP_Orig F"
    24     "(\<And>B. F \<in> stable (A \<inter> B)) \<Longrightarrow> A <= FP_Orig F"
    25 by (simp add: FP_Orig_def stable_def, blast)
    25 by (simp add: FP_Orig_def stable_def, blast)
    26 
    26 
    27 lemma stable_FP_Int: "F : stable (FP F Int B)"
    27 lemma stable_FP_Int: "F \<in> stable (FP F \<inter> B)"
    28 apply (subgoal_tac "FP F Int B = (UN x:B. FP F Int {x}) ")
    28 apply (subgoal_tac "FP F Int B = (UN x:B. FP F Int {x}) ")
    29 prefer 2 apply blast
    29 prefer 2 apply blast
    30 apply (simp (no_asm_simp) add: Int_insert_right)
    30 apply (simp (no_asm_simp) add: Int_insert_right)
    31 apply (simp add: FP_def stable_def)
    31 apply (simp add: FP_def stable_def)
    32 apply (rule constrains_UN)
    32 apply (rule constrains_UN)
    40 apply (drule_tac x = "{x}" in spec)
    40 apply (drule_tac x = "{x}" in spec)
    41 apply (simp add: Int_insert_right)
    41 apply (simp add: Int_insert_right)
    42 done
    42 done
    43 
    43 
    44 lemma FP_weakest:
    44 lemma FP_weakest:
    45     "(!!B. F : stable (A Int B)) ==> A <= FP F"
    45     "(\<And>B. F \<in> stable (A Int B)) \<Longrightarrow> A <= FP F"
    46 by (simp add: FP_equivalence FP_Orig_weakest)
    46 by (simp add: FP_equivalence FP_Orig_weakest)
    47 
    47 
    48 lemma Compl_FP: 
    48 lemma Compl_FP: 
    49     "-(FP F) = (UN act: Acts F. -{s. act``{s} <= {s}})"
    49     "-(FP F) = (UN act: Acts F. -{s. act``{s} <= {s}})"
    50 by (simp add: FP_def stable_def constrains_def, blast)
    50 by (simp add: FP_def stable_def constrains_def, blast)