| author | krauss | 
| Mon, 05 Jun 2006 14:22:58 +0200 | |
| changeset 19769 | c40ce2de2020 | 
| parent 16417 | 9bc16273c2d4 | 
| child 24893 | b8ef7afe3a6b | 
| permissions | -rw-r--r-- | 
| 11479 | 1 | (* Title: ZF/UNITY/Guar.thy | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 2 | ID: $Id \<in> Guar.thy,v 1.3 2001/11/15 14:51:43 ehmety Exp $ | 
| 11479 | 3 | Author: Sidi O Ehmety, Computer Laboratory | 
| 4 | Copyright 2001 University of Cambridge | |
| 5 | ||
| 6 | Guarantees, etc. | |
| 7 | ||
| 8 | From Chandy and Sanders, "Reasoning About Program Composition", | |
| 9 | Technical Report 2000-003, University of Florida, 2000. | |
| 10 | ||
| 11 | Revised by Sidi Ehmety on January 2001 | |
| 12 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 13 | Added \<in> Compatibility, weakest guarantees, etc. | 
| 11479 | 14 | |
| 15 | and Weakest existential property, | |
| 16 | from Charpentier and Chandy "Theorems about Composition", | |
| 17 | Fifth International Conference on Mathematics of Program, 2000. | |
| 18 | ||
| 19 | Theory ported from HOL. | |
| 20 | *) | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 21 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 22 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 23 | header{*The Chandy-Sanders Guarantees Operator*}
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 24 | |
| 16417 | 25 | theory Guar imports Comp begin | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 26 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 27 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 28 | (* To be moved to theory WFair???? *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 29 | lemma leadsTo_Basis': "[| F \<in> A co A Un B; F \<in> transient(A); st_set(B) |] ==> F \<in> A leadsTo B" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 30 | apply (frule constrainsD2) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 31 | apply (drule_tac B = "A-B" in constrains_weaken_L, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 32 | apply (drule_tac B = "A-B" in transient_strengthen, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 33 | apply (blast intro: ensuresI [THEN leadsTo_Basis]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 34 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 35 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 36 | |
| 11479 | 37 | constdefs | 
| 38 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 39 | (*Existential and Universal properties. We formalize the two-program | 
| 11479 | 40 | case, proving equivalence with Chandy and Sanders's n-ary definitions*) | 
| 41 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 42 | ex_prop :: "i => o" | 
| 12195 | 43 | "ex_prop(X) == X<=program & | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 44 | (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> F \<in> X | G \<in> X --> (F Join G) \<in> X)" | 
| 11479 | 45 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 46 | strict_ex_prop :: "i => o" | 
| 12195 | 47 | "strict_ex_prop(X) == X<=program & | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 48 | (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> (F \<in> X | G \<in> X) <-> (F Join G \<in> X))" | 
| 11479 | 49 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 50 | uv_prop :: "i => o" | 
| 12195 | 51 | "uv_prop(X) == X<=program & | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 52 | (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> F \<in> X & G \<in> X --> (F Join G) \<in> X))" | 
| 11479 | 53 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 54 | strict_uv_prop :: "i => o" | 
| 12195 | 55 | "strict_uv_prop(X) == X<=program & | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 56 | (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G -->(F \<in> X & G \<in> X) <-> (F Join G \<in> X)))" | 
| 11479 | 57 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 58 | guar :: "[i, i] => i" (infixl "guarantees" 55) | 
| 11479 | 59 | (*higher than membership, lower than Co*) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 60 |   "X guarantees Y == {F \<in> program. \<forall>G \<in> program. F ok G --> F Join G \<in> X --> F Join G \<in> Y}"
 | 
| 11479 | 61 | |
| 62 | (* Weakest guarantees *) | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 63 | wg :: "[i,i] => i" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 64 |   "wg(F,Y) == Union({X \<in> Pow(program). F:(X guarantees Y)})"
 | 
| 11479 | 65 | |
| 66 | (* Weakest existential property stronger than X *) | |
| 67 | wx :: "i =>i" | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 68 |    "wx(X) == Union({Y \<in> Pow(program). Y<=X & ex_prop(Y)})"
 | 
| 11479 | 69 | |
| 70 | (*Ill-defined programs can arise through "Join"*) | |
| 71 | welldef :: i | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 72 |   "welldef == {F \<in> program. Init(F) \<noteq> 0}"
 | 
| 11479 | 73 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 74 |   refines :: "[i, i, i] => o" ("(3_ refines _ wrt _)" [10,10,10] 10)
 | 
| 11479 | 75 | "G refines F wrt X == | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 76 | \<forall>H \<in> program. (F ok H & G ok H & F Join H \<in> welldef Int X) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 77 | --> (G Join H \<in> welldef Int X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 78 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 79 |   iso_refines :: "[i,i, i] => o"  ("(3_ iso'_refines _ wrt _)" [10,10,10] 10)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 80 | "G iso_refines F wrt X == F \<in> welldef Int X --> G \<in> welldef Int X" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 81 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 82 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 83 | (*** existential properties ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 84 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 85 | lemma ex_imp_subset_program: "ex_prop(X) ==> X\<subseteq>program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 86 | by (simp add: ex_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 87 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 88 | lemma ex1 [rule_format]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 89 | "GG \<in> Fin(program) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 90 | ==> ex_prop(X) --> GG Int X\<noteq>0 --> OK(GG, (%G. G)) -->(\<Squnion>G \<in> GG. G) \<in> X" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 91 | apply (unfold ex_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 92 | apply (erule Fin_induct) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 93 | apply (simp_all add: OK_cons_iff) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 94 | apply (safe elim!: not_emptyE, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 95 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 96 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 97 | lemma ex2 [rule_format]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 98 | "X \<subseteq> program ==> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 99 | (\<forall>GG \<in> Fin(program). GG Int X \<noteq> 0 --> OK(GG,(%G. G))-->(\<Squnion>G \<in> GG. G) \<in> X) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 100 | --> ex_prop(X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 101 | apply (unfold ex_prop_def, clarify) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 102 | apply (drule_tac x = "{F,G}" in bspec)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 103 | apply (simp_all add: OK_iff_ok) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 104 | apply (auto intro: ok_sym) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 105 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 106 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 107 | (*Chandy & Sanders take this as a definition*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 108 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 109 | lemma ex_prop_finite: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 110 | "ex_prop(X) <-> (X\<subseteq>program & | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 111 | (\<forall>GG \<in> Fin(program). GG Int X \<noteq> 0 & OK(GG,(%G. G))-->(\<Squnion>G \<in> GG. G) \<in> X))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 112 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 113 | apply (blast intro: ex1 ex2 dest: ex_imp_subset_program)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 114 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 115 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 116 | (* Equivalent definition of ex_prop given at the end of section 3*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 117 | lemma ex_prop_equiv: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 118 | "ex_prop(X) <-> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 119 | X\<subseteq>program & (\<forall>G \<in> program. (G \<in> X <-> (\<forall>H \<in> program. (G component_of H) --> H \<in> X)))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 120 | apply (unfold ex_prop_def component_of_def, safe, force, force, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 121 | apply (subst Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 122 | apply (blast intro: ok_sym) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 123 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 124 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 125 | (*** universal properties ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 126 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 127 | lemma uv_imp_subset_program: "uv_prop(X)==> X\<subseteq>program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 128 | apply (unfold uv_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 129 | apply (simp (no_asm_simp)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 130 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 131 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 132 | lemma uv1 [rule_format]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 133 | "GG \<in> Fin(program) ==> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 134 | (uv_prop(X)--> GG \<subseteq> X & OK(GG, (%G. G)) --> (\<Squnion>G \<in> GG. G) \<in> X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 135 | apply (unfold uv_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 136 | apply (erule Fin_induct) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 137 | apply (auto simp add: OK_cons_iff) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 138 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 139 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 140 | lemma uv2 [rule_format]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 141 | "X\<subseteq>program ==> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 142 | (\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG,(%G. G)) --> (\<Squnion>G \<in> GG. G) \<in> X) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 143 | --> uv_prop(X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 144 | apply (unfold uv_prop_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 145 | apply (drule_tac x = 0 in bspec, simp+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 146 | apply (drule_tac x = "{F,G}" in bspec, simp) 
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 147 | apply (force dest: ok_sym simp add: OK_iff_ok) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 148 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 149 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 150 | (*Chandy & Sanders take this as a definition*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 151 | lemma uv_prop_finite: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 152 | "uv_prop(X) <-> X\<subseteq>program & | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 153 | (\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG, %G. G) --> (\<Squnion>G \<in> GG. G) \<in> X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 154 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 155 | apply (blast dest: uv_imp_subset_program) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 156 | apply (blast intro: uv1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 157 | apply (blast intro!: uv2 dest:) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 158 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 159 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 160 | (*** guarantees ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 161 | lemma guaranteesI: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 162 | "[| (!!G. [| F ok G; F Join G \<in> X; G \<in> program |] ==> F Join G \<in> Y); | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 163 | F \<in> program |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 164 | ==> F \<in> X guarantees Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 165 | by (simp add: guar_def component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 166 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 167 | lemma guaranteesD: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 168 | "[| F \<in> X guarantees Y; F ok G; F Join G \<in> X; G \<in> program |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 169 | ==> F Join G \<in> Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 170 | by (simp add: guar_def component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 171 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 172 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 173 | (*This version of guaranteesD matches more easily in the conclusion | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 174 | The major premise can no longer be F\<subseteq>H since we need to reason about G*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 175 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 176 | lemma component_guaranteesD: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 177 | "[| F \<in> X guarantees Y; F Join G = H; H \<in> X; F ok G; G \<in> program |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 178 | ==> H \<in> Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 179 | by (simp add: guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 180 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 181 | lemma guarantees_weaken: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 182 | "[| F \<in> X guarantees X'; Y \<subseteq> X; X' \<subseteq> Y' |] ==> F \<in> Y guarantees Y'" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 183 | by (simp add: guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 184 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 185 | lemma subset_imp_guarantees_program: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 186 | "X \<subseteq> Y ==> X guarantees Y = program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 187 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 188 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 189 | (*Equivalent to subset_imp_guarantees_UNIV but more intuitive*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 190 | lemma subset_imp_guarantees: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 191 | "[| X \<subseteq> Y; F \<in> program |] ==> F \<in> X guarantees Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 192 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 193 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 194 | lemma component_of_Join1: "F ok G ==> F component_of (F Join G)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 195 | by (unfold component_of_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 196 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 197 | lemma component_of_Join2: "F ok G ==> G component_of (F Join G)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 198 | apply (subst Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 199 | apply (blast intro: ok_sym component_of_Join1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 200 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 201 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 202 | (*Remark at end of section 4.1 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 203 | lemma ex_prop_imp: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 204 | "ex_prop(Y) ==> (Y = (program guarantees Y))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 205 | apply (simp (no_asm_use) add: ex_prop_equiv guar_def component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 206 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 207 | apply (rule equalityI, blast, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 208 | apply (drule_tac x = x in bspec, assumption, force) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 209 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 210 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 211 | lemma guarantees_imp: "(Y = program guarantees Y) ==> ex_prop(Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 212 | apply (unfold guar_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 213 | apply (simp (no_asm_simp) add: ex_prop_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 214 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 215 | apply (blast intro: elim: equalityE) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 216 | apply (simp_all (no_asm_use) add: component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 217 | apply (force elim: equalityE)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 218 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 219 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 220 | lemma ex_prop_equiv2: "(ex_prop(Y)) <-> (Y = program guarantees Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 221 | by (blast intro: ex_prop_imp guarantees_imp) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 222 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 223 | (** Distributive laws. Re-orient to perform miniscoping **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 224 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 225 | lemma guarantees_UN_left: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 226 | "i \<in> I ==>(\<Union>i \<in> I. X(i)) guarantees Y = (\<Inter>i \<in> I. X(i) guarantees Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 227 | apply (unfold guar_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 228 | apply (rule equalityI, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 229 | prefer 2 apply force | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 230 | apply blast+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 231 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 232 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 233 | lemma guarantees_Un_left: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 234 | "(X Un Y) guarantees Z = (X guarantees Z) Int (Y guarantees Z)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 235 | apply (unfold guar_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 236 | apply (rule equalityI, safe, blast+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 237 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 238 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 239 | lemma guarantees_INT_right: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 240 | "i \<in> I ==> X guarantees (\<Inter>i \<in> I. Y(i)) = (\<Inter>i \<in> I. X guarantees Y(i))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 241 | apply (unfold guar_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 242 | apply (rule equalityI, safe, blast+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 243 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 244 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 245 | lemma guarantees_Int_right: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 246 | "Z guarantees (X Int Y) = (Z guarantees X) Int (Z guarantees Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 247 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 248 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 249 | lemma guarantees_Int_right_I: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 250 | "[| F \<in> Z guarantees X; F \<in> Z guarantees Y |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 251 | ==> F \<in> Z guarantees (X Int Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 252 | by (simp (no_asm_simp) add: guarantees_Int_right) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 253 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 254 | lemma guarantees_INT_right_iff: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 255 | "i \<in> I==> (F \<in> X guarantees (\<Inter>i \<in> I. Y(i))) <-> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 256 | (\<forall>i \<in> I. F \<in> X guarantees Y(i))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 257 | by (simp add: guarantees_INT_right INT_iff, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 258 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 259 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 260 | lemma shunting: "(X guarantees Y) = (program guarantees ((program-X) Un Y))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 261 | by (unfold guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 262 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 263 | lemma contrapositive: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 264 | "(X guarantees Y) = (program - Y) guarantees (program -X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 265 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 266 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 267 | (** The following two can be expressed using intersection and subset, which | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 268 | is more faithful to the text but looks cryptic. | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 269 | **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 270 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 271 | lemma combining1: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 272 | "[| F \<in> V guarantees X; F \<in> (X Int Y) guarantees Z |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 273 | ==> F \<in> (V Int Y) guarantees Z" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 274 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 275 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 276 | lemma combining2: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 277 | "[| F \<in> V guarantees (X Un Y); F \<in> Y guarantees Z |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 278 | ==> F \<in> V guarantees (X Un Z)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 279 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 280 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 281 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 282 | (** The following two follow Chandy-Sanders, but the use of object-quantifiers | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 283 | does not suit Isabelle... **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 284 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 285 | (*Premise should be (!!i. i \<in> I ==> F \<in> X guarantees Y i) *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 286 | lemma all_guarantees: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 287 | "[| \<forall>i \<in> I. F \<in> X guarantees Y(i); i \<in> I |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 288 | ==> F \<in> X guarantees (\<Inter>i \<in> I. Y(i))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 289 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 290 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 291 | (*Premises should be [| F \<in> X guarantees Y i; i \<in> I |] *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 292 | lemma ex_guarantees: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 293 | "\<exists>i \<in> I. F \<in> X guarantees Y(i) ==> F \<in> X guarantees (\<Union>i \<in> I. Y(i))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 294 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 295 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 296 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 297 | (*** Additional guarantees laws, by lcp ***) | 
| 11479 | 298 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 299 | lemma guarantees_Join_Int: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 300 | "[| F \<in> U guarantees V; G \<in> X guarantees Y; F ok G |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 301 | ==> F Join G: (U Int X) guarantees (V Int Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 302 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 303 | apply (unfold guar_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 304 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 305 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 306 | apply (simp add: Join_assoc) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 307 | apply (subgoal_tac "F Join G Join Ga = G Join (F Join Ga) ") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 308 | apply (simp add: ok_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 309 | apply (simp (no_asm_simp) add: Join_ac) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 310 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 311 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 312 | lemma guarantees_Join_Un: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 313 | "[| F \<in> U guarantees V; G \<in> X guarantees Y; F ok G |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 314 | ==> F Join G: (U Un X) guarantees (V Un Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 315 | apply (unfold guar_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 316 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 317 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 318 | apply (simp add: Join_assoc) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 319 | apply (subgoal_tac "F Join G Join Ga = G Join (F Join Ga) ") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 320 | apply (rotate_tac 4) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 321 | apply (drule_tac x = "F Join Ga" in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 322 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 323 | apply (force simp add: ok_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 324 | apply (simp (no_asm_simp) add: Join_ac) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 325 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 326 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 327 | lemma guarantees_JN_INT: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 328 | "[| \<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F); i \<in> I |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 329 | ==> (\<Squnion>i \<in> I. F(i)) \<in> (\<Inter>i \<in> I. X(i)) guarantees (\<Inter>i \<in> I. Y(i))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 330 | apply (unfold guar_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 331 | prefer 2 apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 332 | apply (drule_tac x = xa in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 333 | apply (simp_all add: INT_iff, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 334 | apply (drule_tac x = "(\<Squnion>x \<in> (I-{xa}) . F (x)) Join G" and A=program in bspec)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 335 | apply (auto intro: OK_imp_ok simp add: Join_assoc [symmetric] JN_Join_diff JN_absorb) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 336 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 337 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 338 | lemma guarantees_JN_UN: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 339 | "[| \<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F) |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 340 | ==> JOIN(I,F) \<in> (\<Union>i \<in> I. X(i)) guarantees (\<Union>i \<in> I. Y(i))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 341 | apply (unfold guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 342 | apply (drule_tac x = y in bspec, simp_all, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 343 | apply (rename_tac G y) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 344 | apply (drule_tac x = "JOIN (I-{y}, F) Join G" and A=program in bspec)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 345 | apply (auto intro: OK_imp_ok simp add: Join_assoc [symmetric] JN_Join_diff JN_absorb) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 346 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 347 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 348 | (*** guarantees laws for breaking down the program, by lcp ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 349 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 350 | lemma guarantees_Join_I1: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 351 | "[| F \<in> X guarantees Y; F ok G |] ==> F Join G \<in> X guarantees Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 352 | apply (simp add: guar_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 353 | apply (simp add: Join_assoc) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 354 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 355 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 356 | lemma guarantees_Join_I2: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 357 | "[| G \<in> X guarantees Y; F ok G |] ==> F Join G \<in> X guarantees Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 358 | apply (simp add: Join_commute [of _ G] ok_commute [of _ G]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 359 | apply (blast intro: guarantees_Join_I1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 360 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 361 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 362 | lemma guarantees_JN_I: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 363 | "[| i \<in> I; F(i) \<in> X guarantees Y; OK(I,F) |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 364 | ==> (\<Squnion>i \<in> I. F(i)) \<in> X guarantees Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 365 | apply (unfold guar_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 366 | apply (drule_tac x = "JOIN (I-{i},F) Join G" in bspec)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 367 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 368 | apply (auto intro: OK_imp_ok simp add: JN_Join_diff Join_assoc [symmetric]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 369 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 370 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 371 | (*** well-definedness ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 372 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 373 | lemma Join_welldef_D1: "F Join G \<in> welldef ==> programify(F) \<in> welldef" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 374 | by (unfold welldef_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 375 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 376 | lemma Join_welldef_D2: "F Join G \<in> welldef ==> programify(G) \<in> welldef" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 377 | by (unfold welldef_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 378 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 379 | (*** refinement ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 380 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 381 | lemma refines_refl: "F refines F wrt X" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 382 | by (unfold refines_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 383 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 384 | (* More results on guarantees, added by Sidi Ehmety from Chandy & Sander, section 6 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 385 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 386 | lemma wg_type: "wg(F, X) \<subseteq> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 387 | by (unfold wg_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 388 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 389 | lemma guarantees_type: "X guarantees Y \<subseteq> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 390 | by (unfold guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 391 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 392 | lemma wgD2: "G \<in> wg(F, X) ==> G \<in> program & F \<in> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 393 | apply (unfold wg_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 394 | apply (blast dest: guarantees_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 395 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 396 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 397 | lemma guarantees_equiv: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 398 | "(F \<in> X guarantees Y) <-> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 399 | F \<in> program & (\<forall>H \<in> program. H \<in> X --> (F component_of H --> H \<in> Y))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 400 | by (unfold guar_def component_of_def, force) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 401 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 402 | lemma wg_weakest: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 403 | "!!X. [| F \<in> (X guarantees Y); X \<subseteq> program |] ==> X \<subseteq> wg(F,Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 404 | by (unfold wg_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 405 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 406 | lemma wg_guarantees: "F \<in> program ==> F \<in> wg(F,Y) guarantees Y" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 407 | by (unfold wg_def guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 408 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 409 | lemma wg_equiv: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 410 | "H \<in> wg(F,X) <-> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 411 | ((F component_of H --> H \<in> X) & F \<in> program & H \<in> program)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 412 | apply (simp add: wg_def guarantees_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 413 | apply (rule iffI, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 414 | apply (rule_tac [4] x = "{H}" in bexI)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 415 | apply (rule_tac [3] x = "{H}" in bexI, blast+)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 416 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 417 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 418 | lemma component_of_wg: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 419 | "F component_of H ==> H \<in> wg(F,X) <-> (H \<in> X & F \<in> program & H \<in> program)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 420 | by (simp (no_asm_simp) add: wg_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 421 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 422 | lemma wg_finite [rule_format]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 423 | "\<forall>FF \<in> Fin(program). FF Int X \<noteq> 0 --> OK(FF, %F. F) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 424 | --> (\<forall>F \<in> FF. ((\<Squnion>F \<in> FF. F) \<in> wg(F,X)) <-> ((\<Squnion>F \<in> FF. F) \<in> X))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 425 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 426 | apply (subgoal_tac "F component_of (\<Squnion>F \<in> FF. F) ") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 427 | apply (drule_tac X = X in component_of_wg) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 428 | apply (force dest!: Fin.dom_subset [THEN subsetD, THEN PowD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 429 | apply (simp_all add: component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 430 | apply (rule_tac x = "\<Squnion>F \<in> (FF-{F}) . F" in exI)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 431 | apply (auto intro: JN_Join_diff dest: ok_sym simp add: OK_iff_ok) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 432 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 433 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 434 | lemma wg_ex_prop: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 435 | "ex_prop(X) ==> (F \<in> X) <-> (\<forall>H \<in> program. H \<in> wg(F,X) & F \<in> program)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 436 | apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 437 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 438 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 439 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 440 | (** From Charpentier and Chandy "Theorems About Composition" **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 441 | (* Proposition 2 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 442 | lemma wx_subset: "wx(X)\<subseteq>X" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 443 | by (unfold wx_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 444 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 445 | lemma wx_ex_prop: "ex_prop(wx(X))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 446 | apply (simp (no_asm_use) add: ex_prop_def wx_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 447 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 448 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 449 | apply (rule_tac x=x in bexI, force, simp)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 450 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 451 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 452 | lemma wx_weakest: "\<forall>Z. Z\<subseteq>program --> Z\<subseteq> X --> ex_prop(Z) --> Z \<subseteq> wx(X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 453 | by (unfold wx_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 454 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 455 | (* Proposition 6 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 456 | lemma wx'_ex_prop: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 457 |  "ex_prop({F \<in> program. \<forall>G \<in> program. F ok G --> F Join G \<in> X})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 458 | apply (unfold ex_prop_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 459 | apply (drule_tac x = "G Join Ga" in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 460 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 461 | apply (force simp add: Join_assoc) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 462 | apply (drule_tac x = "F Join Ga" in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 463 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 464 | apply (simp (no_asm_use)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 465 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 466 | apply (simp (no_asm_simp) add: ok_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 467 | apply (subgoal_tac "F Join G = G Join F") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 468 | apply (simp (no_asm_simp) add: Join_assoc) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 469 | apply (simp (no_asm) add: Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 470 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 471 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 472 | (* Equivalence with the other definition of wx *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 473 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 474 | lemma wx_equiv: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 475 |      "wx(X) = {F \<in> program. \<forall>G \<in> program. F ok G --> (F Join G) \<in> X}"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 476 | apply (unfold wx_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 477 | apply (rule equalityI, safe, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 478 | apply (simp (no_asm_use) add: ex_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 479 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 480 | apply (rule_tac B = "{F \<in> program. \<forall>G \<in> program. F ok G --> F Join G \<in> X}" 
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 481 | in UnionI, | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 482 | safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 483 | apply (rule_tac [2] wx'_ex_prop) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 484 | apply (drule_tac x=SKIP in bspec, simp)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 485 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 486 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 487 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 488 | (* Propositions 7 to 11 are all about this second definition of wx. And | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 489 | by equivalence between the two definition, they are the same as the ones proved *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 490 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 491 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 492 | (* Proposition 12 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 493 | (* Main result of the paper *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 494 | lemma guarantees_wx_eq: "(X guarantees Y) = wx((program-X) Un Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 495 | by (auto simp add: guar_def wx_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 496 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 497 | (* | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 498 | {* Corollary, but this result has already been proved elsewhere *}
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 499 | "ex_prop(X guarantees Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 500 | *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 501 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 502 | (* Rules given in section 7 of Chandy and Sander's | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 503 | Reasoning About Program composition paper *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 504 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 505 | lemma stable_guarantees_Always: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 506 | "[| Init(F) \<subseteq> A; F \<in> program |] ==> F \<in> stable(A) guarantees Always(A)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 507 | apply (rule guaranteesI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 508 | prefer 2 apply assumption | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 509 | apply (simp (no_asm) add: Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 510 | apply (rule stable_Join_Always1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 511 | apply (simp_all add: invariant_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 512 | apply (auto simp add: programify_def initially_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 513 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 514 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 515 | lemma constrains_guarantees_leadsTo: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 516 | "[| F \<in> transient(A); st_set(B) |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 517 | ==> F: (A co A Un B) guarantees (A leadsTo (B-A))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 518 | apply (rule guaranteesI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 519 | prefer 2 apply (blast dest: transient_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 520 | apply (rule leadsTo_Basis') | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 521 | apply (blast intro: constrains_weaken_R) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 522 | apply (blast intro!: Join_transient_I1, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 523 | done | 
| 11479 | 524 | |
| 525 | end |