paulson@7400: (* Title: HOL/UNITY/Guar.thy paulson@7400: Author: Lawrence C Paulson, Cambridge University Computer Laboratory wenzelm@32960: Author: Sidi Ehmety paulson@7400: ehmety@11190: From Chandy and Sanders, "Reasoning About Program Composition", ehmety@11190: Technical Report 2000-003, University of Florida, 2000. ehmety@11190: wenzelm@32960: Compatibility, weakest guarantees, etc. and Weakest existential wenzelm@32960: property, from Charpentier and Chandy "Theorems about Composition", ehmety@11190: Fifth International Conference on Mathematics of Program, 2000. paulson@7400: *) paulson@7400: paulson@13798: header{*Guarantees Specifications*} paulson@13798: haftmann@27682: theory Guar haftmann@27682: imports Comp haftmann@27682: begin paulson@7400: wenzelm@12338: instance program :: (type) order haftmann@27682: proof qed (auto simp add: program_less_le dest: component_antisym intro: component_refl component_trans) paulson@7400: paulson@14112: text{*Existential and Universal properties. I formalize the two-program paulson@14112: case, proving equivalence with Chandy and Sanders's n-ary definitions*} paulson@7400: haftmann@35416: definition ex_prop :: "'a program set => bool" where paulson@13819: "ex_prop X == \F G. F ok G -->F \ X | G \ X --> (F\G) \ X" paulson@7400: haftmann@35416: definition strict_ex_prop :: "'a program set => bool" where paulson@13819: "strict_ex_prop X == \F G. F ok G --> (F \ X | G \ X) = (F\G \ X)" paulson@7400: haftmann@35416: definition uv_prop :: "'a program set => bool" where paulson@13819: "uv_prop X == SKIP \ X & (\F G. F ok G --> F \ X & G \ X --> (F\G) \ X)" paulson@7400: haftmann@35416: definition strict_uv_prop :: "'a program set => bool" where paulson@13792: "strict_uv_prop X == paulson@13819: SKIP \ X & (\F G. F ok G --> (F \ X & G \ X) = (F\G \ X))" paulson@7400: paulson@14112: paulson@14112: text{*Guarantees properties*} paulson@14112: haftmann@35416: definition guar :: "['a program set, 'a program set] => 'a program set" (infixl "guarantees" 55) where haftmann@35416: (*higher than membership, lower than Co*) paulson@13819: "X guarantees Y == {F. \G. F ok G --> F\G \ X --> F\G \ Y}" ehmety@11190: paulson@7400: ehmety@11190: (* Weakest guarantees *) haftmann@35416: definition wg :: "['a program, 'a program set] => 'a program set" where paulson@13805: "wg F Y == Union({X. F \ X guarantees Y})" ehmety@11190: ehmety@11190: (* Weakest existential property stronger than X *) haftmann@35416: definition wx :: "('a program) set => ('a program)set" where paulson@13805: "wx X == Union({Y. Y \ X & ex_prop Y})" ehmety@11190: ehmety@11190: (*Ill-defined programs can arise through "Join"*) haftmann@35416: definition welldef :: "'a program set" where paulson@13805: "welldef == {F. Init F \ {}}" ehmety@11190: haftmann@35416: definition refines :: "['a program, 'a program, 'a program set] => bool" haftmann@35416: ("(3_ refines _ wrt _)" [10,10,10] 10) where ehmety@11190: "G refines F wrt X == paulson@14112: \H. (F ok H & G ok H & F\H \ welldef \ X) --> paulson@13819: (G\H \ welldef \ X)" paulson@7400: haftmann@35416: definition iso_refines :: "['a program, 'a program, 'a program set] => bool" haftmann@35416: ("(3_ iso'_refines _ wrt _)" [10,10,10] 10) where ehmety@11190: "G iso_refines F wrt X == paulson@13805: F \ welldef \ X --> G \ welldef \ X" paulson@7400: paulson@13792: paulson@13792: lemma OK_insert_iff: paulson@13792: "(OK (insert i I) F) = paulson@13805: (if i \ I then OK I F else OK I F & (F i ok JOIN I F))" paulson@13792: by (auto intro: ok_sym simp add: OK_iff_ok) paulson@13792: paulson@13792: paulson@14112: subsection{*Existential Properties*} paulson@14112: paulson@13798: lemma ex1 [rule_format]: paulson@13792: "[| ex_prop X; finite GG |] ==> paulson@13805: GG \ X \ {}--> OK GG (%G. G) --> (\G \ GG. G) \ X" paulson@13792: apply (unfold ex_prop_def) paulson@13792: apply (erule finite_induct) paulson@13792: apply (auto simp add: OK_insert_iff Int_insert_left) paulson@13792: done paulson@13792: paulson@13792: paulson@13792: lemma ex2: paulson@13805: "\GG. finite GG & GG \ X \ {} --> OK GG (%G. G) -->(\G \ GG. G):X paulson@13792: ==> ex_prop X" paulson@13792: apply (unfold ex_prop_def, clarify) paulson@13792: apply (drule_tac x = "{F,G}" in spec) paulson@13792: apply (auto dest: ok_sym simp add: OK_iff_ok) paulson@13792: done paulson@13792: paulson@13792: paulson@13792: (*Chandy & Sanders take this as a definition*) paulson@13792: lemma ex_prop_finite: paulson@13792: "ex_prop X = paulson@13805: (\GG. finite GG & GG \ X \ {} & OK GG (%G. G)--> (\G \ GG. G) \ X)" paulson@13792: by (blast intro: ex1 ex2) paulson@13792: paulson@13792: paulson@13792: (*Their "equivalent definition" given at the end of section 3*) paulson@13792: lemma ex_prop_equiv: paulson@13805: "ex_prop X = (\G. G \ X = (\H. (G component_of H) --> H \ X))" paulson@13792: apply auto paulson@14112: apply (unfold ex_prop_def component_of_def, safe, blast, blast) paulson@13792: apply (subst Join_commute) paulson@13792: apply (drule ok_sym, blast) paulson@13792: done paulson@13792: paulson@13792: paulson@14112: subsection{*Universal Properties*} paulson@14112: paulson@13792: lemma uv1 [rule_format]: paulson@13792: "[| uv_prop X; finite GG |] paulson@13805: ==> GG \ X & OK GG (%G. G) --> (\G \ GG. G) \ X" paulson@13792: apply (unfold uv_prop_def) paulson@13792: apply (erule finite_induct) paulson@13792: apply (auto simp add: Int_insert_left OK_insert_iff) paulson@13792: done paulson@13792: paulson@13792: lemma uv2: paulson@13805: "\GG. finite GG & GG \ X & OK GG (%G. G) --> (\G \ GG. G) \ X paulson@13792: ==> uv_prop X" paulson@13792: apply (unfold uv_prop_def) paulson@13792: apply (rule conjI) paulson@13792: apply (drule_tac x = "{}" in spec) paulson@13792: prefer 2 paulson@13792: apply clarify paulson@13792: apply (drule_tac x = "{F,G}" in spec) paulson@13792: apply (auto dest: ok_sym simp add: OK_iff_ok) paulson@13792: done paulson@13792: paulson@13792: (*Chandy & Sanders take this as a definition*) paulson@13792: lemma uv_prop_finite: paulson@13792: "uv_prop X = paulson@13805: (\GG. finite GG & GG \ X & OK GG (%G. G) --> (\G \ GG. G): X)" paulson@13792: by (blast intro: uv1 uv2) paulson@13792: paulson@14112: subsection{*Guarantees*} paulson@13792: paulson@13792: lemma guaranteesI: paulson@14112: "(!!G. [| F ok G; F\G \ X |] ==> F\G \ Y) ==> F \ X guarantees Y" paulson@13792: by (simp add: guar_def component_def) paulson@13792: paulson@13792: lemma guaranteesD: paulson@14112: "[| F \ X guarantees Y; F ok G; F\G \ X |] ==> F\G \ Y" paulson@13792: by (unfold guar_def component_def, blast) paulson@13792: paulson@13792: (*This version of guaranteesD matches more easily in the conclusion paulson@13805: The major premise can no longer be F \ H since we need to reason about G*) paulson@13792: lemma component_guaranteesD: paulson@14112: "[| F \ X guarantees Y; F\G = H; H \ X; F ok G |] ==> H \ Y" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: lemma guarantees_weaken: paulson@13805: "[| F \ X guarantees X'; Y \ X; X' \ Y' |] ==> F \ Y guarantees Y'" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13805: lemma subset_imp_guarantees_UNIV: "X \ Y ==> X guarantees Y = UNIV" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: (*Equivalent to subset_imp_guarantees_UNIV but more intuitive*) paulson@13805: lemma subset_imp_guarantees: "X \ Y ==> F \ X guarantees Y" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: (*Remark at end of section 4.1 *) paulson@13792: paulson@13792: lemma ex_prop_imp: "ex_prop Y ==> (Y = UNIV guarantees Y)" paulson@13792: apply (simp (no_asm_use) add: guar_def ex_prop_equiv) paulson@13792: apply safe paulson@13792: apply (drule_tac x = x in spec) paulson@13792: apply (drule_tac [2] x = x in spec) paulson@13792: apply (drule_tac [2] sym) paulson@13792: apply (auto simp add: component_of_def) paulson@13792: done paulson@13792: paulson@13792: lemma guarantees_imp: "(Y = UNIV guarantees Y) ==> ex_prop(Y)" paulson@14112: by (auto simp add: guar_def ex_prop_equiv component_of_def dest: sym) paulson@13792: paulson@13792: lemma ex_prop_equiv2: "(ex_prop Y) = (Y = UNIV guarantees Y)" paulson@13792: apply (rule iffI) paulson@13792: apply (rule ex_prop_imp) paulson@13792: apply (auto simp add: guarantees_imp) paulson@13792: done paulson@13792: paulson@13792: paulson@14112: subsection{*Distributive Laws. Re-Orient to Perform Miniscoping*} paulson@13792: paulson@13792: lemma guarantees_UN_left: paulson@13805: "(\i \ I. X i) guarantees Y = (\i \ I. X i guarantees Y)" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: lemma guarantees_Un_left: paulson@13805: "(X \ Y) guarantees Z = (X guarantees Z) \ (Y guarantees Z)" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: lemma guarantees_INT_right: paulson@13805: "X guarantees (\i \ I. Y i) = (\i \ I. X guarantees Y i)" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: lemma guarantees_Int_right: paulson@13805: "Z guarantees (X \ Y) = (Z guarantees X) \ (Z guarantees Y)" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: lemma guarantees_Int_right_I: paulson@13805: "[| F \ Z guarantees X; F \ Z guarantees Y |] paulson@13805: ==> F \ Z guarantees (X \ Y)" paulson@13792: by (simp add: guarantees_Int_right) paulson@13792: paulson@13792: lemma guarantees_INT_right_iff: paulson@13805: "(F \ X guarantees (INTER I Y)) = (\i\I. F \ X guarantees (Y i))" paulson@13792: by (simp add: guarantees_INT_right) paulson@13792: paulson@13805: lemma shunting: "(X guarantees Y) = (UNIV guarantees (-X \ Y))" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: lemma contrapositive: "(X guarantees Y) = -Y guarantees -X" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: (** The following two can be expressed using intersection and subset, which paulson@13792: is more faithful to the text but looks cryptic. paulson@13792: **) paulson@13792: paulson@13792: lemma combining1: paulson@13805: "[| F \ V guarantees X; F \ (X \ Y) guarantees Z |] paulson@13805: ==> F \ (V \ Y) guarantees Z" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: lemma combining2: paulson@13805: "[| F \ V guarantees (X \ Y); F \ Y guarantees Z |] paulson@13805: ==> F \ V guarantees (X \ Z)" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: (** The following two follow Chandy-Sanders, but the use of object-quantifiers paulson@13792: does not suit Isabelle... **) paulson@13792: paulson@13805: (*Premise should be (!!i. i \ I ==> F \ X guarantees Y i) *) paulson@13792: lemma all_guarantees: paulson@13805: "\i\I. F \ X guarantees (Y i) ==> F \ X guarantees (\i \ I. Y i)" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13805: (*Premises should be [| F \ X guarantees Y i; i \ I |] *) paulson@13792: lemma ex_guarantees: paulson@13805: "\i\I. F \ X guarantees (Y i) ==> F \ X guarantees (\i \ I. Y i)" paulson@13792: by (unfold guar_def, blast) paulson@13792: paulson@13792: paulson@14112: subsection{*Guarantees: Additional Laws (by lcp)*} paulson@13792: paulson@13792: lemma guarantees_Join_Int: paulson@13805: "[| F \ U guarantees V; G \ X guarantees Y; F ok G |] paulson@13819: ==> F\G \ (U \ X) guarantees (V \ Y)" paulson@14112: apply (simp add: guar_def, safe) paulson@14112: apply (simp add: Join_assoc) paulson@13819: apply (subgoal_tac "F\G\Ga = G\(F\Ga) ") paulson@14112: apply (simp add: ok_commute) paulson@14112: apply (simp add: Join_ac) paulson@13792: done paulson@13792: paulson@13792: lemma guarantees_Join_Un: paulson@13805: "[| F \ U guarantees V; G \ X guarantees Y; F ok G |] paulson@13819: ==> F\G \ (U \ X) guarantees (V \ Y)" paulson@14112: apply (simp add: guar_def, safe) paulson@14112: apply (simp add: Join_assoc) paulson@13819: apply (subgoal_tac "F\G\Ga = G\(F\Ga) ") paulson@14112: apply (simp add: ok_commute) paulson@14112: apply (simp add: Join_ac) paulson@13792: done paulson@13792: paulson@13792: lemma guarantees_JN_INT: paulson@13805: "[| \i\I. F i \ X i guarantees Y i; OK I F |] paulson@13805: ==> (JOIN I F) \ (INTER I X) guarantees (INTER I Y)" paulson@13792: apply (unfold guar_def, auto) paulson@13792: apply (drule bspec, assumption) paulson@13792: apply (rename_tac "i") paulson@13819: apply (drule_tac x = "JOIN (I-{i}) F\G" in spec) paulson@13792: apply (auto intro: OK_imp_ok paulson@13792: simp add: Join_assoc [symmetric] JN_Join_diff JN_absorb) paulson@13792: done paulson@13792: paulson@13792: lemma guarantees_JN_UN: paulson@13805: "[| \i\I. F i \ X i guarantees Y i; OK I F |] paulson@13805: ==> (JOIN I F) \ (UNION I X) guarantees (UNION I Y)" paulson@13792: apply (unfold guar_def, auto) paulson@13792: apply (drule bspec, assumption) paulson@13792: apply (rename_tac "i") paulson@13819: apply (drule_tac x = "JOIN (I-{i}) F\G" in spec) paulson@13792: apply (auto intro: OK_imp_ok paulson@13792: simp add: Join_assoc [symmetric] JN_Join_diff JN_absorb) paulson@13792: done paulson@13792: paulson@13792: paulson@14112: subsection{*Guarantees Laws for Breaking Down the Program (by lcp)*} paulson@13792: paulson@13792: lemma guarantees_Join_I1: paulson@13819: "[| F \ X guarantees Y; F ok G |] ==> F\G \ X guarantees Y" paulson@14112: by (simp add: guar_def Join_assoc) paulson@13792: paulson@14112: lemma guarantees_Join_I2: paulson@13819: "[| G \ X guarantees Y; F ok G |] ==> F\G \ X guarantees Y" paulson@13792: apply (simp add: Join_commute [of _ G] ok_commute [of _ G]) paulson@13792: apply (blast intro: guarantees_Join_I1) paulson@13792: done paulson@13792: paulson@13792: lemma guarantees_JN_I: paulson@13805: "[| i \ I; F i \ X guarantees Y; OK I F |] paulson@13805: ==> (\i \ I. (F i)) \ X guarantees Y" paulson@13792: apply (unfold guar_def, clarify) paulson@13819: apply (drule_tac x = "JOIN (I-{i}) F\G" in spec) paulson@14112: apply (auto intro: OK_imp_ok paulson@14112: simp add: JN_Join_diff JN_Join_diff Join_assoc [symmetric]) paulson@13792: done paulson@13792: paulson@13792: paulson@13792: (*** well-definedness ***) paulson@13792: paulson@13819: lemma Join_welldef_D1: "F\G \ welldef ==> F \ welldef" paulson@13792: by (unfold welldef_def, auto) paulson@13792: paulson@13819: lemma Join_welldef_D2: "F\G \ welldef ==> G \ welldef" paulson@13792: by (unfold welldef_def, auto) paulson@13792: paulson@13792: (*** refinement ***) paulson@13792: paulson@13792: lemma refines_refl: "F refines F wrt X" paulson@13792: by (unfold refines_def, blast) paulson@13792: paulson@14112: (*We'd like transitivity, but how do we get it?*) paulson@14112: lemma refines_trans: paulson@13792: "[| H refines G wrt X; G refines F wrt X |] ==> H refines F wrt X" paulson@14112: apply (simp add: refines_def) paulson@14112: oops paulson@13792: paulson@13792: paulson@13792: lemma strict_ex_refine_lemma: paulson@13792: "strict_ex_prop X paulson@13819: ==> (\H. F ok H & G ok H & F\H \ X --> G\H \ X) paulson@13805: = (F \ X --> G \ X)" paulson@13792: by (unfold strict_ex_prop_def, auto) paulson@13792: paulson@13792: lemma strict_ex_refine_lemma_v: paulson@13792: "strict_ex_prop X paulson@13819: ==> (\H. F ok H & G ok H & F\H \ welldef & F\H \ X --> G\H \ X) = paulson@13805: (F \ welldef \ X --> G \ X)" paulson@13792: apply (unfold strict_ex_prop_def, safe) paulson@13792: apply (erule_tac x = SKIP and P = "%H. ?PP H --> ?RR H" in allE) paulson@13792: apply (auto dest: Join_welldef_D1 Join_welldef_D2) paulson@13792: done paulson@13792: paulson@13792: lemma ex_refinement_thm: paulson@13792: "[| strict_ex_prop X; paulson@13819: \H. F ok H & G ok H & F\H \ welldef \ X --> G\H \ welldef |] paulson@13792: ==> (G refines F wrt X) = (G iso_refines F wrt X)" paulson@13792: apply (rule_tac x = SKIP in allE, assumption) paulson@13792: apply (simp add: refines_def iso_refines_def strict_ex_refine_lemma_v) paulson@13792: done paulson@13792: paulson@13792: paulson@13792: lemma strict_uv_refine_lemma: paulson@13792: "strict_uv_prop X ==> paulson@13819: (\H. F ok H & G ok H & F\H \ X --> G\H \ X) = (F \ X --> G \ X)" paulson@13792: by (unfold strict_uv_prop_def, blast) paulson@13792: paulson@13792: lemma strict_uv_refine_lemma_v: paulson@13792: "strict_uv_prop X paulson@13819: ==> (\H. F ok H & G ok H & F\H \ welldef & F\H \ X --> G\H \ X) = paulson@13805: (F \ welldef \ X --> G \ X)" paulson@13792: apply (unfold strict_uv_prop_def, safe) paulson@13792: apply (erule_tac x = SKIP and P = "%H. ?PP H --> ?RR H" in allE) paulson@13792: apply (auto dest: Join_welldef_D1 Join_welldef_D2) paulson@13792: done paulson@13792: paulson@13792: lemma uv_refinement_thm: paulson@13792: "[| strict_uv_prop X; paulson@13819: \H. F ok H & G ok H & F\H \ welldef \ X --> paulson@13819: G\H \ welldef |] paulson@13792: ==> (G refines F wrt X) = (G iso_refines F wrt X)" paulson@13792: apply (rule_tac x = SKIP in allE, assumption) paulson@13792: apply (simp add: refines_def iso_refines_def strict_uv_refine_lemma_v) paulson@13792: done paulson@13792: paulson@13792: (* Added by Sidi Ehmety from Chandy & Sander, section 6 *) paulson@13792: lemma guarantees_equiv: paulson@13805: "(F \ X guarantees Y) = (\H. H \ X \ (F component_of H \ H \ Y))" paulson@13792: by (unfold guar_def component_of_def, auto) paulson@13792: paulson@14112: lemma wg_weakest: "!!X. F\ (X guarantees Y) ==> X \ (wg F Y)" paulson@13792: by (unfold wg_def, auto) paulson@13792: paulson@14112: lemma wg_guarantees: "F\ ((wg F Y) guarantees Y)" paulson@13792: by (unfold wg_def guar_def, blast) paulson@13792: paulson@14112: lemma wg_equiv: "(H \ wg F X) = (F component_of H --> H \ X)" paulson@14112: by (simp add: guarantees_equiv wg_def, blast) paulson@13792: paulson@13805: lemma component_of_wg: "F component_of H ==> (H \ wg F X) = (H \ X)" paulson@13792: by (simp add: wg_equiv) paulson@13792: paulson@13792: lemma wg_finite: paulson@13805: "\FF. finite FF & FF \ X \ {} --> OK FF (%F. F) paulson@13805: --> (\F\FF. ((\F \ FF. F): wg F X) = ((\F \ FF. F):X))" paulson@13792: apply clarify paulson@13805: apply (subgoal_tac "F component_of (\F \ FF. F) ") paulson@13792: apply (drule_tac X = X in component_of_wg, simp) paulson@13792: apply (simp add: component_of_def) paulson@13805: apply (rule_tac x = "\F \ (FF-{F}) . F" in exI) paulson@13792: apply (auto intro: JN_Join_diff dest: ok_sym simp add: OK_iff_ok) paulson@13792: done paulson@13792: paulson@13805: lemma wg_ex_prop: "ex_prop X ==> (F \ X) = (\H. H \ wg F X)" paulson@13792: apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv) paulson@13792: apply blast paulson@13792: done paulson@13792: paulson@13792: (** From Charpentier and Chandy "Theorems About Composition" **) paulson@13792: (* Proposition 2 *) paulson@13792: lemma wx_subset: "(wx X)<=X" paulson@13792: by (unfold wx_def, auto) paulson@13792: paulson@13792: lemma wx_ex_prop: "ex_prop (wx X)" berghofe@16647: apply (simp add: wx_def ex_prop_equiv cong: bex_cong, safe, blast) paulson@14112: apply force paulson@13792: done paulson@13792: paulson@13805: lemma wx_weakest: "\Z. Z<= X --> ex_prop Z --> Z \ wx X" paulson@14112: by (auto simp add: wx_def) paulson@13792: paulson@13792: (* Proposition 6 *) paulson@13819: lemma wx'_ex_prop: "ex_prop({F. \G. F ok G --> F\G \ X})" paulson@13792: apply (unfold ex_prop_def, safe) paulson@14112: apply (drule_tac x = "G\Ga" in spec) paulson@14112: apply (force simp add: ok_Join_iff1 Join_assoc) paulson@13819: apply (drule_tac x = "F\Ga" in spec) paulson@14112: apply (simp add: ok_Join_iff1 ok_commute Join_ac) paulson@13792: done paulson@13792: paulson@14112: text{* Equivalence with the other definition of wx *} paulson@13792: paulson@14112: lemma wx_equiv: "wx X = {F. \G. F ok G --> (F\G) \ X}" paulson@13792: apply (unfold wx_def, safe) paulson@14112: apply (simp add: ex_prop_def, blast) paulson@13792: apply (simp (no_asm)) paulson@13819: apply (rule_tac x = "{F. \G. F ok G --> F\G \ X}" in exI, safe) paulson@13792: apply (rule_tac [2] wx'_ex_prop) paulson@14112: apply (drule_tac x = SKIP in spec)+ paulson@14112: apply auto paulson@13792: done paulson@13792: paulson@13792: paulson@14112: text{* Propositions 7 to 11 are about this second definition of wx. paulson@14112: They are the same as the ones proved for the first definition of wx, paulson@14112: by equivalence *} paulson@13792: paulson@13792: (* Proposition 12 *) paulson@13792: (* Main result of the paper *) paulson@14112: lemma guarantees_wx_eq: "(X guarantees Y) = wx(-X \ Y)" paulson@14112: by (simp add: guar_def wx_equiv) paulson@13792: paulson@13792: paulson@13792: (* Rules given in section 7 of Chandy and Sander's paulson@13792: Reasoning About Program composition paper *) paulson@13792: lemma stable_guarantees_Always: paulson@14112: "Init F \ A ==> F \ (stable A) guarantees (Always A)" paulson@13792: apply (rule guaranteesI) paulson@14112: apply (simp add: Join_commute) paulson@13792: apply (rule stable_Join_Always1) paulson@14112: apply (simp_all add: invariant_def Join_stable) paulson@13792: done paulson@13792: paulson@13792: lemma constrains_guarantees_leadsTo: paulson@13805: "F \ transient A ==> F \ (A co A \ B) guarantees (A leadsTo (B-A))" paulson@13792: apply (rule guaranteesI) paulson@13792: apply (rule leadsTo_Basis') paulson@14112: apply (drule constrains_weaken_R) paulson@14112: prefer 2 apply assumption paulson@14112: apply blast paulson@13792: apply (blast intro: Join_transient_I1) paulson@13792: done paulson@13792: paulson@7400: end