wenzelm@32960: (* Title: HOL/UNITY/Constrains.thy paulson@5313: Author: Lawrence C Paulson, Cambridge University Computer Laboratory paulson@5313: Copyright 1998 University of Cambridge paulson@5313: paulson@13797: Weak safety relations: restricted to the set of reachable states. paulson@5313: *) paulson@5313: paulson@13798: header{*Weak Safety*} paulson@13798: haftmann@16417: theory Constrains imports UNITY begin paulson@6535: paulson@6535: (*Initial states and program => (final state, reversed trace to it)... paulson@6535: Arguments MUST be curried in an inductive definition*) paulson@6535: berghofe@23767: inductive_set berghofe@23767: traces :: "['a set, ('a * 'a)set set] => ('a * 'a list) set" berghofe@23767: for init :: "'a set" and acts :: "('a * 'a)set set" berghofe@23767: where paulson@6535: (*Initial trace is empty*) paulson@13805: Init: "s \ init ==> (s,[]) \ traces init acts" paulson@6535: berghofe@23767: | Acts: "[| act: acts; (s,evs) \ traces init acts; (s,s'): act |] wenzelm@32960: ==> (s', s#evs) \ traces init acts" paulson@6535: paulson@6535: berghofe@23767: inductive_set berghofe@23767: reachable :: "'a program => 'a set" berghofe@23767: for F :: "'a program" berghofe@23767: where paulson@13805: Init: "s \ Init F ==> s \ reachable F" paulson@5313: berghofe@23767: | Acts: "[| act: Acts F; s \ reachable F; (s,s'): act |] wenzelm@32960: ==> s' \ reachable F" paulson@6536: haftmann@35416: definition Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) where paulson@13805: "A Co B == {F. F \ (reachable F \ A) co B}" paulson@6536: haftmann@35416: definition Unless :: "['a set, 'a set] => 'a program set" (infixl "Unless" 60) where paulson@13805: "A Unless B == (A-B) Co (A \ B)" paulson@6536: haftmann@35416: definition Stable :: "'a set => 'a program set" where paulson@6536: "Stable A == A Co A" paulson@5313: paulson@6570: (*Always is the weak form of "invariant"*) haftmann@35416: definition Always :: "'a set => 'a program set" where paulson@13805: "Always A == {F. Init F \ A} \ Stable A" paulson@5313: paulson@13805: (*Polymorphic in both states and the meaning of \ *) haftmann@35416: definition Increasing :: "['a => 'b::{order}] => 'a program set" where paulson@13805: "Increasing f == \z. Stable {s. z \ f s}" paulson@5784: paulson@13797: paulson@13798: subsection{*traces and reachable*} paulson@13797: paulson@13797: lemma reachable_equiv_traces: paulson@13812: "reachable F = {s. \evs. (s,evs) \ traces (Init F) (Acts F)}" paulson@13797: apply safe paulson@13797: apply (erule_tac [2] traces.induct) paulson@13797: apply (erule reachable.induct) paulson@13797: apply (blast intro: reachable.intros traces.intros)+ paulson@13797: done paulson@13797: paulson@13805: lemma Init_subset_reachable: "Init F \ reachable F" paulson@13797: by (blast intro: reachable.intros) paulson@13797: paulson@13797: lemma stable_reachable [intro!,simp]: paulson@13805: "Acts G \ Acts F ==> G \ stable (reachable F)" paulson@13797: by (blast intro: stableI constrainsI reachable.intros) paulson@13797: paulson@13797: (*The set of all reachable states is an invariant...*) paulson@13805: lemma invariant_reachable: "F \ invariant (reachable F)" paulson@13797: apply (simp add: invariant_def) paulson@13797: apply (blast intro: reachable.intros) paulson@13797: done paulson@13797: paulson@13797: (*...in fact the strongest invariant!*) paulson@13805: lemma invariant_includes_reachable: "F \ invariant A ==> reachable F \ A" paulson@13797: apply (simp add: stable_def constrains_def invariant_def) paulson@13797: apply (rule subsetI) paulson@13797: apply (erule reachable.induct) paulson@13797: apply (blast intro: reachable.intros)+ paulson@13797: done paulson@13797: paulson@13797: paulson@13798: subsection{*Co*} paulson@13797: paulson@13805: (*F \ B co B' ==> F \ (reachable F \ B) co (reachable F \ B')*) paulson@13797: lemmas constrains_reachable_Int = paulson@13797: subset_refl [THEN stable_reachable [unfolded stable_def], paulson@13797: THEN constrains_Int, standard] paulson@13797: paulson@13797: (*Resembles the previous definition of Constrains*) paulson@13797: lemma Constrains_eq_constrains: paulson@13805: "A Co B = {F. F \ (reachable F \ A) co (reachable F \ B)}" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (blast dest: constrains_reachable_Int intro: constrains_weaken) paulson@13797: done paulson@13797: paulson@13805: lemma constrains_imp_Constrains: "F \ A co A' ==> F \ A Co A'" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (blast intro: constrains_weaken_L) paulson@13797: done paulson@13797: paulson@13805: lemma stable_imp_Stable: "F \ stable A ==> F \ Stable A" paulson@13797: apply (unfold stable_def Stable_def) paulson@13797: apply (erule constrains_imp_Constrains) paulson@13797: done paulson@13797: paulson@13797: lemma ConstrainsI: paulson@13805: "(!!act s s'. [| act: Acts F; (s,s') \ act; s \ A |] ==> s': A') paulson@13805: ==> F \ A Co A'" paulson@13797: apply (rule constrains_imp_Constrains) paulson@13797: apply (blast intro: constrainsI) paulson@13797: done paulson@13797: paulson@13805: lemma Constrains_empty [iff]: "F \ {} Co B" paulson@13797: by (unfold Constrains_def constrains_def, blast) paulson@13797: paulson@13805: lemma Constrains_UNIV [iff]: "F \ A Co UNIV" paulson@13797: by (blast intro: ConstrainsI) paulson@13797: paulson@13797: lemma Constrains_weaken_R: paulson@13805: "[| F \ A Co A'; A'<=B' |] ==> F \ A Co B'" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (blast intro: constrains_weaken_R) paulson@13797: done paulson@13797: paulson@13797: lemma Constrains_weaken_L: paulson@13805: "[| F \ A Co A'; B \ A |] ==> F \ B Co A'" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (blast intro: constrains_weaken_L) paulson@13797: done paulson@13797: paulson@13797: lemma Constrains_weaken: paulson@13805: "[| F \ A Co A'; B \ A; A'<=B' |] ==> F \ B Co B'" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (blast intro: constrains_weaken) paulson@13797: done paulson@13797: paulson@13797: (** Union **) paulson@13797: paulson@13797: lemma Constrains_Un: paulson@13805: "[| F \ A Co A'; F \ B Co B' |] ==> F \ (A \ B) Co (A' \ B')" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (blast intro: constrains_Un [THEN constrains_weaken]) paulson@13797: done paulson@13797: paulson@13797: lemma Constrains_UN: paulson@13805: assumes Co: "!!i. i \ I ==> F \ (A i) Co (A' i)" paulson@13805: shows "F \ (\i \ I. A i) Co (\i \ I. A' i)" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (rule CollectI) paulson@13797: apply (rule Co [unfolded Constrains_def, THEN CollectD, THEN constrains_UN, paulson@13797: THEN constrains_weaken], auto) paulson@13797: done paulson@13797: paulson@13797: (** Intersection **) paulson@13797: paulson@13797: lemma Constrains_Int: paulson@13805: "[| F \ A Co A'; F \ B Co B' |] ==> F \ (A \ B) Co (A' \ B')" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (blast intro: constrains_Int [THEN constrains_weaken]) paulson@13797: done paulson@13797: paulson@13797: lemma Constrains_INT: paulson@13805: assumes Co: "!!i. i \ I ==> F \ (A i) Co (A' i)" paulson@13805: shows "F \ (\i \ I. A i) Co (\i \ I. A' i)" paulson@13797: apply (unfold Constrains_def) paulson@13797: apply (rule CollectI) paulson@13797: apply (rule Co [unfolded Constrains_def, THEN CollectD, THEN constrains_INT, paulson@13797: THEN constrains_weaken], auto) paulson@13797: done paulson@13797: paulson@13805: lemma Constrains_imp_subset: "F \ A Co A' ==> reachable F \ A \ A'" paulson@13797: by (simp add: constrains_imp_subset Constrains_def) paulson@13797: paulson@13805: lemma Constrains_trans: "[| F \ A Co B; F \ B Co C |] ==> F \ A Co C" paulson@13797: apply (simp add: Constrains_eq_constrains) paulson@13797: apply (blast intro: constrains_trans constrains_weaken) paulson@13797: done paulson@13797: paulson@13797: lemma Constrains_cancel: paulson@13805: "[| F \ A Co (A' \ B); F \ B Co B' |] ==> F \ A Co (A' \ B')" paulson@13797: by (simp add: Constrains_eq_constrains constrains_def, blast) paulson@13797: paulson@13797: paulson@13798: subsection{*Stable*} paulson@13797: paulson@13797: (*Useful because there's no Stable_weaken. [Tanja Vos]*) paulson@13805: lemma Stable_eq: "[| F \ Stable A; A = B |] ==> F \ Stable B" paulson@13797: by blast paulson@13797: paulson@13805: lemma Stable_eq_stable: "(F \ Stable A) = (F \ stable (reachable F \ A))" paulson@13797: by (simp add: Stable_def Constrains_eq_constrains stable_def) paulson@13797: paulson@13805: lemma StableI: "F \ A Co A ==> F \ Stable A" paulson@13797: by (unfold Stable_def, assumption) paulson@13797: paulson@13805: lemma StableD: "F \ Stable A ==> F \ A Co A" paulson@13797: by (unfold Stable_def, assumption) paulson@13797: paulson@13797: lemma Stable_Un: paulson@13805: "[| F \ Stable A; F \ Stable A' |] ==> F \ Stable (A \ A')" paulson@13797: apply (unfold Stable_def) paulson@13797: apply (blast intro: Constrains_Un) paulson@13797: done paulson@13797: paulson@13797: lemma Stable_Int: paulson@13805: "[| F \ Stable A; F \ Stable A' |] ==> F \ Stable (A \ A')" paulson@13797: apply (unfold Stable_def) paulson@13797: apply (blast intro: Constrains_Int) paulson@13797: done paulson@13797: paulson@13797: lemma Stable_Constrains_Un: paulson@13805: "[| F \ Stable C; F \ A Co (C \ A') |] paulson@13805: ==> F \ (C \ A) Co (C \ A')" paulson@13797: apply (unfold Stable_def) paulson@13797: apply (blast intro: Constrains_Un [THEN Constrains_weaken]) paulson@13797: done paulson@13797: paulson@13797: lemma Stable_Constrains_Int: paulson@13805: "[| F \ Stable C; F \ (C \ A) Co A' |] paulson@13805: ==> F \ (C \ A) Co (C \ A')" paulson@13797: apply (unfold Stable_def) paulson@13797: apply (blast intro: Constrains_Int [THEN Constrains_weaken]) paulson@13797: done paulson@13797: paulson@13797: lemma Stable_UN: paulson@13805: "(!!i. i \ I ==> F \ Stable (A i)) ==> F \ Stable (\i \ I. A i)" paulson@13797: by (simp add: Stable_def Constrains_UN) paulson@13797: paulson@13797: lemma Stable_INT: paulson@13805: "(!!i. i \ I ==> F \ Stable (A i)) ==> F \ Stable (\i \ I. A i)" paulson@13797: by (simp add: Stable_def Constrains_INT) paulson@13797: paulson@13805: lemma Stable_reachable: "F \ Stable (reachable F)" paulson@13797: by (simp add: Stable_eq_stable) paulson@13797: paulson@13797: paulson@13797: paulson@13798: subsection{*Increasing*} paulson@13797: paulson@13797: lemma IncreasingD: paulson@13805: "F \ Increasing f ==> F \ Stable {s. x \ f s}" paulson@13797: by (unfold Increasing_def, blast) paulson@13797: paulson@13797: lemma mono_Increasing_o: paulson@13805: "mono g ==> Increasing f \ Increasing (g o f)" paulson@13797: apply (simp add: Increasing_def Stable_def Constrains_def stable_def paulson@13797: constrains_def) paulson@13797: apply (blast intro: monoD order_trans) paulson@13797: done paulson@13797: paulson@13797: lemma strict_IncreasingD: paulson@13805: "!!z::nat. F \ Increasing f ==> F \ Stable {s. z < f s}" paulson@13797: by (simp add: Increasing_def Suc_le_eq [symmetric]) paulson@13797: paulson@13797: lemma increasing_imp_Increasing: paulson@13805: "F \ increasing f ==> F \ Increasing f" paulson@13797: apply (unfold increasing_def Increasing_def) paulson@13797: apply (blast intro: stable_imp_Stable) paulson@13797: done paulson@13797: paulson@13797: lemmas Increasing_constant = paulson@13797: increasing_constant [THEN increasing_imp_Increasing, standard, iff] paulson@13797: paulson@13797: paulson@13798: subsection{*The Elimination Theorem*} paulson@13798: paulson@13798: (*The "free" m has become universally quantified! Should the premise be !!m paulson@13805: instead of \m ? Would make it harder to use in forward proof.*) paulson@13797: paulson@13797: lemma Elimination: paulson@13805: "[| \m. F \ {s. s x = m} Co (B m) |] paulson@13805: ==> F \ {s. s x \ M} Co (\m \ M. B m)" paulson@13797: by (unfold Constrains_def constrains_def, blast) paulson@13797: paulson@13797: (*As above, but for the trivial case of a one-variable state, in which the paulson@13797: state is identified with its one variable.*) paulson@13797: lemma Elimination_sing: paulson@13805: "(\m. F \ {m} Co (B m)) ==> F \ M Co (\m \ M. B m)" paulson@13797: by (unfold Constrains_def constrains_def, blast) paulson@13797: paulson@13797: paulson@13798: subsection{*Specialized laws for handling Always*} paulson@13797: paulson@13797: (** Natural deduction rules for "Always A" **) paulson@13797: paulson@13805: lemma AlwaysI: "[| Init F \ A; F \ Stable A |] ==> F \ Always A" paulson@13797: by (simp add: Always_def) paulson@13797: paulson@13805: lemma AlwaysD: "F \ Always A ==> Init F \ A & F \ Stable A" paulson@13797: by (simp add: Always_def) paulson@13797: paulson@13797: lemmas AlwaysE = AlwaysD [THEN conjE, standard] paulson@13797: lemmas Always_imp_Stable = AlwaysD [THEN conjunct2, standard] paulson@13797: paulson@13797: paulson@13797: (*The set of all reachable states is Always*) paulson@13805: lemma Always_includes_reachable: "F \ Always A ==> reachable F \ A" paulson@13797: apply (simp add: Stable_def Constrains_def constrains_def Always_def) paulson@13797: apply (rule subsetI) paulson@13797: apply (erule reachable.induct) paulson@13797: apply (blast intro: reachable.intros)+ paulson@13797: done paulson@13797: paulson@13797: lemma invariant_imp_Always: paulson@13805: "F \ invariant A ==> F \ Always A" paulson@13797: apply (unfold Always_def invariant_def Stable_def stable_def) paulson@13797: apply (blast intro: constrains_imp_Constrains) paulson@13797: done paulson@13797: paulson@13797: lemmas Always_reachable = paulson@13797: invariant_reachable [THEN invariant_imp_Always, standard] paulson@13797: paulson@13797: lemma Always_eq_invariant_reachable: paulson@13805: "Always A = {F. F \ invariant (reachable F \ A)}" paulson@13797: apply (simp add: Always_def invariant_def Stable_def Constrains_eq_constrains paulson@13797: stable_def) paulson@13797: apply (blast intro: reachable.intros) paulson@13797: done paulson@13797: paulson@13797: (*the RHS is the traditional definition of the "always" operator*) paulson@13805: lemma Always_eq_includes_reachable: "Always A = {F. reachable F \ A}" paulson@13797: by (auto dest: invariant_includes_reachable simp add: Int_absorb2 invariant_reachable Always_eq_invariant_reachable) paulson@13797: paulson@13797: lemma Always_UNIV_eq [simp]: "Always UNIV = UNIV" paulson@13797: by (auto simp add: Always_eq_includes_reachable) paulson@13797: paulson@13805: lemma UNIV_AlwaysI: "UNIV \ A ==> F \ Always A" paulson@13797: by (auto simp add: Always_eq_includes_reachable) paulson@13797: paulson@13805: lemma Always_eq_UN_invariant: "Always A = (\I \ Pow A. invariant I)" paulson@13797: apply (simp add: Always_eq_includes_reachable) paulson@13797: apply (blast intro: invariantI Init_subset_reachable [THEN subsetD] paulson@13797: invariant_includes_reachable [THEN subsetD]) paulson@13797: done paulson@13797: paulson@13805: lemma Always_weaken: "[| F \ Always A; A \ B |] ==> F \ Always B" paulson@13797: by (auto simp add: Always_eq_includes_reachable) paulson@13797: paulson@13797: paulson@13798: subsection{*"Co" rules involving Always*} paulson@13797: paulson@13797: lemma Always_Constrains_pre: paulson@13805: "F \ Always INV ==> (F \ (INV \ A) Co A') = (F \ A Co A')" paulson@13797: by (simp add: Always_includes_reachable [THEN Int_absorb2] Constrains_def paulson@13797: Int_assoc [symmetric]) paulson@13797: paulson@13797: lemma Always_Constrains_post: paulson@13805: "F \ Always INV ==> (F \ A Co (INV \ A')) = (F \ A Co A')" paulson@13797: by (simp add: Always_includes_reachable [THEN Int_absorb2] paulson@13797: Constrains_eq_constrains Int_assoc [symmetric]) paulson@13797: paulson@13805: (* [| F \ Always INV; F \ (INV \ A) Co A' |] ==> F \ A Co A' *) paulson@13797: lemmas Always_ConstrainsI = Always_Constrains_pre [THEN iffD1, standard] paulson@13797: paulson@13805: (* [| F \ Always INV; F \ A Co A' |] ==> F \ A Co (INV \ A') *) paulson@13797: lemmas Always_ConstrainsD = Always_Constrains_post [THEN iffD2, standard] paulson@13797: paulson@13797: (*The analogous proof of Always_LeadsTo_weaken doesn't terminate*) paulson@13797: lemma Always_Constrains_weaken: paulson@13805: "[| F \ Always C; F \ A Co A'; paulson@13805: C \ B \ A; C \ A' \ B' |] paulson@13805: ==> F \ B Co B'" paulson@13797: apply (rule Always_ConstrainsI, assumption) paulson@13797: apply (drule Always_ConstrainsD, assumption) paulson@13797: apply (blast intro: Constrains_weaken) paulson@13797: done paulson@13797: paulson@13797: paulson@13797: (** Conjoining Always properties **) paulson@13797: paulson@13805: lemma Always_Int_distrib: "Always (A \ B) = Always A \ Always B" paulson@13797: by (auto simp add: Always_eq_includes_reachable) paulson@13797: paulson@13805: lemma Always_INT_distrib: "Always (INTER I A) = (\i \ I. Always (A i))" paulson@13797: by (auto simp add: Always_eq_includes_reachable) paulson@13797: paulson@13797: lemma Always_Int_I: paulson@13805: "[| F \ Always A; F \ Always B |] ==> F \ Always (A \ B)" paulson@13797: by (simp add: Always_Int_distrib) paulson@13797: paulson@13797: (*Allows a kind of "implication introduction"*) paulson@13797: lemma Always_Compl_Un_eq: paulson@13805: "F \ Always A ==> (F \ Always (-A \ B)) = (F \ Always B)" paulson@13797: by (auto simp add: Always_eq_includes_reachable) paulson@13797: paulson@13797: (*Delete the nearest invariance assumption (which will be the second one paulson@13797: used by Always_Int_I) *) paulson@13805: lemmas Always_thin = thin_rl [of "F \ Always A", standard] paulson@13797: paulson@13812: paulson@13812: subsection{*Totalize*} paulson@13812: paulson@13812: lemma reachable_imp_reachable_tot: paulson@13812: "s \ reachable F ==> s \ reachable (totalize F)" paulson@13812: apply (erule reachable.induct) paulson@13812: apply (rule reachable.Init) paulson@13812: apply simp paulson@13812: apply (rule_tac act = "totalize_act act" in reachable.Acts) paulson@13812: apply (auto simp add: totalize_act_def) paulson@13812: done paulson@13812: paulson@13812: lemma reachable_tot_imp_reachable: paulson@13812: "s \ reachable (totalize F) ==> s \ reachable F" paulson@13812: apply (erule reachable.induct) paulson@13812: apply (rule reachable.Init, simp) paulson@13812: apply (force simp add: totalize_act_def intro: reachable.Acts) paulson@13812: done paulson@13812: paulson@13812: lemma reachable_tot_eq [simp]: "reachable (totalize F) = reachable F" paulson@13812: by (blast intro: reachable_imp_reachable_tot reachable_tot_imp_reachable) paulson@13812: paulson@13812: lemma totalize_Constrains_iff [simp]: "(totalize F \ A Co B) = (F \ A Co B)" paulson@13812: by (simp add: Constrains_def) paulson@13812: paulson@13812: lemma totalize_Stable_iff [simp]: "(totalize F \ Stable A) = (F \ Stable A)" paulson@13812: by (simp add: Stable_def) paulson@13812: paulson@13812: lemma totalize_Always_iff [simp]: "(totalize F \ Always A) = (F \ Always A)" paulson@13812: by (simp add: Always_def) paulson@13812: paulson@5313: end