| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 15 Dec 2023 10:48:05 +0100 | |
| changeset 79291 | e9a788a75775 | 
| parent 69313 | b021008c5397 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 1 | (* Title: HOL/UNITY/Guar.thy | 
| 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27682diff
changeset | 3 | Author: Sidi Ehmety | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 4 | |
| 11190 | 5 | From Chandy and Sanders, "Reasoning About Program Composition", | 
| 6 | Technical Report 2000-003, University of Florida, 2000. | |
| 7 | ||
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27682diff
changeset | 8 | Compatibility, weakest guarantees, etc. and Weakest existential | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27682diff
changeset | 9 | property, from Charpentier and Chandy "Theorems about Composition", | 
| 11190 | 10 | Fifth International Conference on Mathematics of Program, 2000. | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 11 | *) | 
| 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 12 | |
| 63146 | 13 | section\<open>Guarantees Specifications\<close> | 
| 13798 | 14 | |
| 27682 | 15 | theory Guar | 
| 16 | imports Comp | |
| 17 | begin | |
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 18 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11190diff
changeset | 19 | instance program :: (type) order | 
| 61169 | 20 | by standard (auto simp add: program_less_le dest: component_antisym intro: component_trans) | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 21 | |
| 63146 | 22 | text\<open>Existential and Universal properties. I formalize the two-program | 
| 23 | case, proving equivalence with Chandy and Sanders's n-ary definitions\<close> | |
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 24 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 25 | definition ex_prop :: "'a program set => bool" where | 
| 13819 | 26 | "ex_prop X == \<forall>F G. F ok G -->F \<in> X | G \<in> X --> (F\<squnion>G) \<in> X" | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 27 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 28 | definition strict_ex_prop :: "'a program set => bool" where | 
| 13819 | 29 | "strict_ex_prop X == \<forall>F G. F ok G --> (F \<in> X | G \<in> X) = (F\<squnion>G \<in> X)" | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 30 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 31 | definition uv_prop :: "'a program set => bool" where | 
| 13819 | 32 | "uv_prop X == SKIP \<in> X & (\<forall>F G. F ok G --> F \<in> X & G \<in> X --> (F\<squnion>G) \<in> X)" | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 33 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 34 | definition strict_uv_prop :: "'a program set => bool" where | 
| 13792 | 35 | "strict_uv_prop X == | 
| 13819 | 36 | SKIP \<in> X & (\<forall>F G. F ok G --> (F \<in> X & G \<in> X) = (F\<squnion>G \<in> X))" | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 37 | |
| 14112 | 38 | |
| 63146 | 39 | text\<open>Guarantees properties\<close> | 
| 14112 | 40 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 41 | definition guar :: "['a program set, 'a program set] => 'a program set" (infixl "guarantees" 55) where | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 42 | (*higher than membership, lower than Co*) | 
| 13819 | 43 |    "X guarantees Y == {F. \<forall>G. F ok G --> F\<squnion>G \<in> X --> F\<squnion>G \<in> Y}"
 | 
| 11190 | 44 | |
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 45 | |
| 11190 | 46 | (* Weakest guarantees *) | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 47 | definition wg :: "['a program, 'a program set] => 'a program set" where | 
| 61952 | 48 |   "wg F Y == \<Union>({X. F \<in> X guarantees Y})"
 | 
| 11190 | 49 | |
| 50 | (* Weakest existential property stronger than X *) | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 51 | definition wx :: "('a program) set => ('a program)set" where
 | 
| 61952 | 52 |    "wx X == \<Union>({Y. Y \<subseteq> X & ex_prop Y})"
 | 
| 11190 | 53 | |
| 54 | (*Ill-defined programs can arise through "Join"*) | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 55 | definition welldef :: "'a program set" where | 
| 13805 | 56 |   "welldef == {F. Init F \<noteq> {}}"
 | 
| 11190 | 57 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 58 | definition refines :: "['a program, 'a program, 'a program set] => bool" | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 59 |                         ("(3_ refines _ wrt _)" [10,10,10] 10) where
 | 
| 11190 | 60 | "G refines F wrt X == | 
| 14112 | 61 | \<forall>H. (F ok H & G ok H & F\<squnion>H \<in> welldef \<inter> X) --> | 
| 13819 | 62 | (G\<squnion>H \<in> welldef \<inter> X)" | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 63 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 64 | definition iso_refines :: "['a program, 'a program, 'a program set] => bool" | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 65 |                               ("(3_ iso'_refines _ wrt _)" [10,10,10] 10) where
 | 
| 11190 | 66 | "G iso_refines F wrt X == | 
| 13805 | 67 | F \<in> welldef \<inter> X --> G \<in> welldef \<inter> X" | 
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 68 | |
| 13792 | 69 | |
| 70 | lemma OK_insert_iff: | |
| 71 | "(OK (insert i I) F) = | |
| 13805 | 72 | (if i \<in> I then OK I F else OK I F & (F i ok JOIN I F))" | 
| 13792 | 73 | by (auto intro: ok_sym simp add: OK_iff_ok) | 
| 74 | ||
| 75 | ||
| 63146 | 76 | subsection\<open>Existential Properties\<close> | 
| 14112 | 77 | |
| 45477 | 78 | lemma ex1: | 
| 79 | assumes "ex_prop X" and "finite GG" | |
| 80 |   shows "GG \<inter> X \<noteq> {} \<Longrightarrow> OK GG (%G. G) \<Longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X"
 | |
| 81 | apply (atomize (full)) | |
| 82 | using assms(2) apply induct | |
| 83 | using assms(1) apply (unfold ex_prop_def) | |
| 84 | apply (auto simp add: OK_insert_iff Int_insert_left) | |
| 85 | done | |
| 13792 | 86 | |
| 87 | lemma ex2: | |
| 67613 | 88 |      "\<forall>GG. finite GG & GG \<inter> X \<noteq> {} \<longrightarrow> OK GG (\<lambda>G. G) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X 
 | 
| 89 | \<Longrightarrow> ex_prop X" | |
| 13792 | 90 | apply (unfold ex_prop_def, clarify) | 
| 91 | apply (drule_tac x = "{F,G}" in spec)
 | |
| 92 | apply (auto dest: ok_sym simp add: OK_iff_ok) | |
| 93 | done | |
| 94 | ||
| 95 | ||
| 96 | (*Chandy & Sanders take this as a definition*) | |
| 97 | lemma ex_prop_finite: | |
| 98 | "ex_prop X = | |
| 13805 | 99 |       (\<forall>GG. finite GG & GG \<inter> X \<noteq> {} & OK GG (%G. G)--> (\<Squnion>G \<in> GG. G) \<in> X)"
 | 
| 13792 | 100 | by (blast intro: ex1 ex2) | 
| 101 | ||
| 102 | ||
| 103 | (*Their "equivalent definition" given at the end of section 3*) | |
| 104 | lemma ex_prop_equiv: | |
| 13805 | 105 | "ex_prop X = (\<forall>G. G \<in> X = (\<forall>H. (G component_of H) --> H \<in> X))" | 
| 13792 | 106 | apply auto | 
| 14112 | 107 | apply (unfold ex_prop_def component_of_def, safe, blast, blast) | 
| 13792 | 108 | apply (subst Join_commute) | 
| 109 | apply (drule ok_sym, blast) | |
| 110 | done | |
| 111 | ||
| 112 | ||
| 63146 | 113 | subsection\<open>Universal Properties\<close> | 
| 14112 | 114 | |
| 45477 | 115 | lemma uv1: | 
| 116 | assumes "uv_prop X" | |
| 117 | and "finite GG" | |
| 118 | and "GG \<subseteq> X" | |
| 119 | and "OK GG (%G. G)" | |
| 120 | shows "(\<Squnion>G \<in> GG. G) \<in> X" | |
| 121 | using assms(2-) | |
| 122 | apply induct | |
| 123 | using assms(1) | |
| 124 | apply (unfold uv_prop_def) | |
| 125 | apply (auto simp add: Int_insert_left OK_insert_iff) | |
| 126 | done | |
| 13792 | 127 | |
| 128 | lemma uv2: | |
| 13805 | 129 | "\<forall>GG. finite GG & GG \<subseteq> X & OK GG (%G. G) --> (\<Squnion>G \<in> GG. G) \<in> X | 
| 13792 | 130 | ==> uv_prop X" | 
| 131 | apply (unfold uv_prop_def) | |
| 132 | apply (rule conjI) | |
| 133 |  apply (drule_tac x = "{}" in spec)
 | |
| 134 | prefer 2 | |
| 135 | apply clarify | |
| 136 |  apply (drule_tac x = "{F,G}" in spec)
 | |
| 137 | apply (auto dest: ok_sym simp add: OK_iff_ok) | |
| 138 | done | |
| 139 | ||
| 140 | (*Chandy & Sanders take this as a definition*) | |
| 141 | lemma uv_prop_finite: | |
| 142 | "uv_prop X = | |
| 67613 | 143 | (\<forall>GG. finite GG \<and> GG \<subseteq> X \<and> OK GG (\<lambda>G. G) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)" | 
| 13792 | 144 | by (blast intro: uv1 uv2) | 
| 145 | ||
| 63146 | 146 | subsection\<open>Guarantees\<close> | 
| 13792 | 147 | |
| 148 | lemma guaranteesI: | |
| 14112 | 149 | "(!!G. [| F ok G; F\<squnion>G \<in> X |] ==> F\<squnion>G \<in> Y) ==> F \<in> X guarantees Y" | 
| 13792 | 150 | by (simp add: guar_def component_def) | 
| 151 | ||
| 152 | lemma guaranteesD: | |
| 14112 | 153 | "[| F \<in> X guarantees Y; F ok G; F\<squnion>G \<in> X |] ==> F\<squnion>G \<in> Y" | 
| 13792 | 154 | by (unfold guar_def component_def, blast) | 
| 155 | ||
| 156 | (*This version of guaranteesD matches more easily in the conclusion | |
| 13805 | 157 | The major premise can no longer be F \<subseteq> H since we need to reason about G*) | 
| 13792 | 158 | lemma component_guaranteesD: | 
| 14112 | 159 | "[| F \<in> X guarantees Y; F\<squnion>G = H; H \<in> X; F ok G |] ==> H \<in> Y" | 
| 13792 | 160 | by (unfold guar_def, blast) | 
| 161 | ||
| 162 | lemma guarantees_weaken: | |
| 13805 | 163 | "[| F \<in> X guarantees X'; Y \<subseteq> X; X' \<subseteq> Y' |] ==> F \<in> Y guarantees Y'" | 
| 13792 | 164 | by (unfold guar_def, blast) | 
| 165 | ||
| 13805 | 166 | lemma subset_imp_guarantees_UNIV: "X \<subseteq> Y ==> X guarantees Y = UNIV" | 
| 13792 | 167 | by (unfold guar_def, blast) | 
| 168 | ||
| 169 | (*Equivalent to subset_imp_guarantees_UNIV but more intuitive*) | |
| 13805 | 170 | lemma subset_imp_guarantees: "X \<subseteq> Y ==> F \<in> X guarantees Y" | 
| 13792 | 171 | by (unfold guar_def, blast) | 
| 172 | ||
| 173 | (*Remark at end of section 4.1 *) | |
| 174 | ||
| 175 | lemma ex_prop_imp: "ex_prop Y ==> (Y = UNIV guarantees Y)" | |
| 176 | apply (simp (no_asm_use) add: guar_def ex_prop_equiv) | |
| 177 | apply safe | |
| 178 | apply (drule_tac x = x in spec) | |
| 179 | apply (drule_tac [2] x = x in spec) | |
| 180 | apply (drule_tac [2] sym) | |
| 181 | apply (auto simp add: component_of_def) | |
| 182 | done | |
| 183 | ||
| 184 | lemma guarantees_imp: "(Y = UNIV guarantees Y) ==> ex_prop(Y)" | |
| 14112 | 185 | by (auto simp add: guar_def ex_prop_equiv component_of_def dest: sym) | 
| 13792 | 186 | |
| 187 | lemma ex_prop_equiv2: "(ex_prop Y) = (Y = UNIV guarantees Y)" | |
| 188 | apply (rule iffI) | |
| 189 | apply (rule ex_prop_imp) | |
| 190 | apply (auto simp add: guarantees_imp) | |
| 191 | done | |
| 192 | ||
| 193 | ||
| 63146 | 194 | subsection\<open>Distributive Laws. Re-Orient to Perform Miniscoping\<close> | 
| 13792 | 195 | |
| 196 | lemma guarantees_UN_left: | |
| 13805 | 197 | "(\<Union>i \<in> I. X i) guarantees Y = (\<Inter>i \<in> I. X i guarantees Y)" | 
| 13792 | 198 | by (unfold guar_def, blast) | 
| 199 | ||
| 200 | lemma guarantees_Un_left: | |
| 13805 | 201 | "(X \<union> Y) guarantees Z = (X guarantees Z) \<inter> (Y guarantees Z)" | 
| 13792 | 202 | by (unfold guar_def, blast) | 
| 203 | ||
| 204 | lemma guarantees_INT_right: | |
| 13805 | 205 | "X guarantees (\<Inter>i \<in> I. Y i) = (\<Inter>i \<in> I. X guarantees Y i)" | 
| 13792 | 206 | by (unfold guar_def, blast) | 
| 207 | ||
| 208 | lemma guarantees_Int_right: | |
| 13805 | 209 | "Z guarantees (X \<inter> Y) = (Z guarantees X) \<inter> (Z guarantees Y)" | 
| 13792 | 210 | by (unfold guar_def, blast) | 
| 211 | ||
| 212 | lemma guarantees_Int_right_I: | |
| 13805 | 213 | "[| F \<in> Z guarantees X; F \<in> Z guarantees Y |] | 
| 214 | ==> F \<in> Z guarantees (X \<inter> Y)" | |
| 13792 | 215 | by (simp add: guarantees_Int_right) | 
| 216 | ||
| 217 | lemma guarantees_INT_right_iff: | |
| 69313 | 218 | "(F \<in> X guarantees (\<Inter>(Y ` I))) = (\<forall>i\<in>I. F \<in> X guarantees (Y i))" | 
| 13792 | 219 | by (simp add: guarantees_INT_right) | 
| 220 | ||
| 13805 | 221 | lemma shunting: "(X guarantees Y) = (UNIV guarantees (-X \<union> Y))" | 
| 13792 | 222 | by (unfold guar_def, blast) | 
| 223 | ||
| 224 | lemma contrapositive: "(X guarantees Y) = -Y guarantees -X" | |
| 225 | by (unfold guar_def, blast) | |
| 226 | ||
| 227 | (** The following two can be expressed using intersection and subset, which | |
| 228 | is more faithful to the text but looks cryptic. | |
| 229 | **) | |
| 230 | ||
| 231 | lemma combining1: | |
| 13805 | 232 | "[| F \<in> V guarantees X; F \<in> (X \<inter> Y) guarantees Z |] | 
| 233 | ==> F \<in> (V \<inter> Y) guarantees Z" | |
| 13792 | 234 | by (unfold guar_def, blast) | 
| 235 | ||
| 236 | lemma combining2: | |
| 13805 | 237 | "[| F \<in> V guarantees (X \<union> Y); F \<in> Y guarantees Z |] | 
| 238 | ==> F \<in> V guarantees (X \<union> Z)" | |
| 13792 | 239 | by (unfold guar_def, blast) | 
| 240 | ||
| 241 | (** The following two follow Chandy-Sanders, but the use of object-quantifiers | |
| 242 | does not suit Isabelle... **) | |
| 243 | ||
| 13805 | 244 | (*Premise should be (!!i. i \<in> I ==> F \<in> X guarantees Y i) *) | 
| 13792 | 245 | lemma all_guarantees: | 
| 13805 | 246 | "\<forall>i\<in>I. F \<in> X guarantees (Y i) ==> F \<in> X guarantees (\<Inter>i \<in> I. Y i)" | 
| 13792 | 247 | by (unfold guar_def, blast) | 
| 248 | ||
| 13805 | 249 | (*Premises should be [| F \<in> X guarantees Y i; i \<in> I |] *) | 
| 13792 | 250 | lemma ex_guarantees: | 
| 13805 | 251 | "\<exists>i\<in>I. F \<in> X guarantees (Y i) ==> F \<in> X guarantees (\<Union>i \<in> I. Y i)" | 
| 13792 | 252 | by (unfold guar_def, blast) | 
| 253 | ||
| 254 | ||
| 63146 | 255 | subsection\<open>Guarantees: Additional Laws (by lcp)\<close> | 
| 13792 | 256 | |
| 257 | lemma guarantees_Join_Int: | |
| 13805 | 258 | "[| F \<in> U guarantees V; G \<in> X guarantees Y; F ok G |] | 
| 13819 | 259 | ==> F\<squnion>G \<in> (U \<inter> X) guarantees (V \<inter> Y)" | 
| 14112 | 260 | apply (simp add: guar_def, safe) | 
| 261 | apply (simp add: Join_assoc) | |
| 13819 | 262 | apply (subgoal_tac "F\<squnion>G\<squnion>Ga = G\<squnion>(F\<squnion>Ga) ") | 
| 14112 | 263 | apply (simp add: ok_commute) | 
| 264 | apply (simp add: Join_ac) | |
| 13792 | 265 | done | 
| 266 | ||
| 267 | lemma guarantees_Join_Un: | |
| 13805 | 268 | "[| F \<in> U guarantees V; G \<in> X guarantees Y; F ok G |] | 
| 13819 | 269 | ==> F\<squnion>G \<in> (U \<union> X) guarantees (V \<union> Y)" | 
| 14112 | 270 | apply (simp add: guar_def, safe) | 
| 271 | apply (simp add: Join_assoc) | |
| 13819 | 272 | apply (subgoal_tac "F\<squnion>G\<squnion>Ga = G\<squnion>(F\<squnion>Ga) ") | 
| 14112 | 273 | apply (simp add: ok_commute) | 
| 274 | apply (simp add: Join_ac) | |
| 13792 | 275 | done | 
| 276 | ||
| 277 | lemma guarantees_JN_INT: | |
| 13805 | 278 | "[| \<forall>i\<in>I. F i \<in> X i guarantees Y i; OK I F |] | 
| 69313 | 279 | ==> (JOIN I F) \<in> (\<Inter>(X ` I)) guarantees (\<Inter>(Y ` I))" | 
| 13792 | 280 | apply (unfold guar_def, auto) | 
| 281 | apply (drule bspec, assumption) | |
| 282 | apply (rename_tac "i") | |
| 13819 | 283 | apply (drule_tac x = "JOIN (I-{i}) F\<squnion>G" in spec)
 | 
| 13792 | 284 | apply (auto intro: OK_imp_ok | 
| 285 | simp add: Join_assoc [symmetric] JN_Join_diff JN_absorb) | |
| 286 | done | |
| 287 | ||
| 288 | lemma guarantees_JN_UN: | |
| 13805 | 289 | "[| \<forall>i\<in>I. F i \<in> X i guarantees Y i; OK I F |] | 
| 69313 | 290 | ==> (JOIN I F) \<in> (\<Union>(X ` I)) guarantees (\<Union>(Y ` I))" | 
| 13792 | 291 | apply (unfold guar_def, auto) | 
| 292 | apply (drule bspec, assumption) | |
| 293 | apply (rename_tac "i") | |
| 13819 | 294 | apply (drule_tac x = "JOIN (I-{i}) F\<squnion>G" in spec)
 | 
| 13792 | 295 | apply (auto intro: OK_imp_ok | 
| 296 | simp add: Join_assoc [symmetric] JN_Join_diff JN_absorb) | |
| 297 | done | |
| 298 | ||
| 299 | ||
| 63146 | 300 | subsection\<open>Guarantees Laws for Breaking Down the Program (by lcp)\<close> | 
| 13792 | 301 | |
| 302 | lemma guarantees_Join_I1: | |
| 13819 | 303 | "[| F \<in> X guarantees Y; F ok G |] ==> F\<squnion>G \<in> X guarantees Y" | 
| 14112 | 304 | by (simp add: guar_def Join_assoc) | 
| 13792 | 305 | |
| 14112 | 306 | lemma guarantees_Join_I2: | 
| 13819 | 307 | "[| G \<in> X guarantees Y; F ok G |] ==> F\<squnion>G \<in> X guarantees Y" | 
| 13792 | 308 | apply (simp add: Join_commute [of _ G] ok_commute [of _ G]) | 
| 309 | apply (blast intro: guarantees_Join_I1) | |
| 310 | done | |
| 311 | ||
| 312 | lemma guarantees_JN_I: | |
| 13805 | 313 | "[| i \<in> I; F i \<in> X guarantees Y; OK I F |] | 
| 314 | ==> (\<Squnion>i \<in> I. (F i)) \<in> X guarantees Y" | |
| 13792 | 315 | apply (unfold guar_def, clarify) | 
| 13819 | 316 | apply (drule_tac x = "JOIN (I-{i}) F\<squnion>G" in spec)
 | 
| 44871 | 317 | apply (auto intro: OK_imp_ok simp add: JN_Join_diff Join_assoc [symmetric]) | 
| 13792 | 318 | done | 
| 319 | ||
| 320 | ||
| 321 | (*** well-definedness ***) | |
| 322 | ||
| 13819 | 323 | lemma Join_welldef_D1: "F\<squnion>G \<in> welldef ==> F \<in> welldef" | 
| 13792 | 324 | by (unfold welldef_def, auto) | 
| 325 | ||
| 13819 | 326 | lemma Join_welldef_D2: "F\<squnion>G \<in> welldef ==> G \<in> welldef" | 
| 13792 | 327 | by (unfold welldef_def, auto) | 
| 328 | ||
| 329 | (*** refinement ***) | |
| 330 | ||
| 331 | lemma refines_refl: "F refines F wrt X" | |
| 332 | by (unfold refines_def, blast) | |
| 333 | ||
| 14112 | 334 | (*We'd like transitivity, but how do we get it?*) | 
| 335 | lemma refines_trans: | |
| 13792 | 336 | "[| H refines G wrt X; G refines F wrt X |] ==> H refines F wrt X" | 
| 14112 | 337 | apply (simp add: refines_def) | 
| 338 | oops | |
| 13792 | 339 | |
| 340 | ||
| 341 | lemma strict_ex_refine_lemma: | |
| 342 | "strict_ex_prop X | |
| 13819 | 343 | ==> (\<forall>H. F ok H & G ok H & F\<squnion>H \<in> X --> G\<squnion>H \<in> X) | 
| 13805 | 344 | = (F \<in> X --> G \<in> X)" | 
| 13792 | 345 | by (unfold strict_ex_prop_def, auto) | 
| 346 | ||
| 347 | lemma strict_ex_refine_lemma_v: | |
| 348 | "strict_ex_prop X | |
| 13819 | 349 | ==> (\<forall>H. F ok H & G ok H & F\<squnion>H \<in> welldef & F\<squnion>H \<in> X --> G\<squnion>H \<in> X) = | 
| 13805 | 350 | (F \<in> welldef \<inter> X --> G \<in> X)" | 
| 13792 | 351 | apply (unfold strict_ex_prop_def, safe) | 
| 59807 | 352 | apply (erule_tac x = SKIP and P = "%H. PP H --> RR H" for PP RR in allE) | 
| 13792 | 353 | apply (auto dest: Join_welldef_D1 Join_welldef_D2) | 
| 354 | done | |
| 355 | ||
| 356 | lemma ex_refinement_thm: | |
| 357 | "[| strict_ex_prop X; | |
| 13819 | 358 | \<forall>H. F ok H & G ok H & F\<squnion>H \<in> welldef \<inter> X --> G\<squnion>H \<in> welldef |] | 
| 13792 | 359 | ==> (G refines F wrt X) = (G iso_refines F wrt X)" | 
| 360 | apply (rule_tac x = SKIP in allE, assumption) | |
| 361 | apply (simp add: refines_def iso_refines_def strict_ex_refine_lemma_v) | |
| 362 | done | |
| 363 | ||
| 364 | ||
| 365 | lemma strict_uv_refine_lemma: | |
| 366 | "strict_uv_prop X ==> | |
| 13819 | 367 | (\<forall>H. F ok H & G ok H & F\<squnion>H \<in> X --> G\<squnion>H \<in> X) = (F \<in> X --> G \<in> X)" | 
| 13792 | 368 | by (unfold strict_uv_prop_def, blast) | 
| 369 | ||
| 370 | lemma strict_uv_refine_lemma_v: | |
| 371 | "strict_uv_prop X | |
| 13819 | 372 | ==> (\<forall>H. F ok H & G ok H & F\<squnion>H \<in> welldef & F\<squnion>H \<in> X --> G\<squnion>H \<in> X) = | 
| 13805 | 373 | (F \<in> welldef \<inter> X --> G \<in> X)" | 
| 13792 | 374 | apply (unfold strict_uv_prop_def, safe) | 
| 59807 | 375 | apply (erule_tac x = SKIP and P = "%H. PP H --> RR H" for PP RR in allE) | 
| 13792 | 376 | apply (auto dest: Join_welldef_D1 Join_welldef_D2) | 
| 377 | done | |
| 378 | ||
| 379 | lemma uv_refinement_thm: | |
| 380 | "[| strict_uv_prop X; | |
| 13819 | 381 | \<forall>H. F ok H & G ok H & F\<squnion>H \<in> welldef \<inter> X --> | 
| 382 | G\<squnion>H \<in> welldef |] | |
| 13792 | 383 | ==> (G refines F wrt X) = (G iso_refines F wrt X)" | 
| 384 | apply (rule_tac x = SKIP in allE, assumption) | |
| 385 | apply (simp add: refines_def iso_refines_def strict_uv_refine_lemma_v) | |
| 386 | done | |
| 387 | ||
| 388 | (* Added by Sidi Ehmety from Chandy & Sander, section 6 *) | |
| 389 | lemma guarantees_equiv: | |
| 13805 | 390 | "(F \<in> X guarantees Y) = (\<forall>H. H \<in> X \<longrightarrow> (F component_of H \<longrightarrow> H \<in> Y))" | 
| 13792 | 391 | by (unfold guar_def component_of_def, auto) | 
| 392 | ||
| 14112 | 393 | lemma wg_weakest: "!!X. F\<in> (X guarantees Y) ==> X \<subseteq> (wg F Y)" | 
| 13792 | 394 | by (unfold wg_def, auto) | 
| 395 | ||
| 14112 | 396 | lemma wg_guarantees: "F\<in> ((wg F Y) guarantees Y)" | 
| 13792 | 397 | by (unfold wg_def guar_def, blast) | 
| 398 | ||
| 14112 | 399 | lemma wg_equiv: "(H \<in> wg F X) = (F component_of H --> H \<in> X)" | 
| 400 | by (simp add: guarantees_equiv wg_def, blast) | |
| 13792 | 401 | |
| 13805 | 402 | lemma component_of_wg: "F component_of H ==> (H \<in> wg F X) = (H \<in> X)" | 
| 13792 | 403 | by (simp add: wg_equiv) | 
| 404 | ||
| 405 | lemma wg_finite: | |
| 67613 | 406 |     "\<forall>FF. finite FF \<and> FF \<inter> X \<noteq> {} \<longrightarrow> OK FF (\<lambda>F. F)
 | 
| 407 | \<longrightarrow> (\<forall>F\<in>FF. ((\<Squnion>F \<in> FF. F) \<in> wg F X) = ((\<Squnion>F \<in> FF. F) \<in> X))" | |
| 13792 | 408 | apply clarify | 
| 13805 | 409 | apply (subgoal_tac "F component_of (\<Squnion>F \<in> FF. F) ") | 
| 13792 | 410 | apply (drule_tac X = X in component_of_wg, simp) | 
| 411 | apply (simp add: component_of_def) | |
| 13805 | 412 | apply (rule_tac x = "\<Squnion>F \<in> (FF-{F}) . F" in exI)
 | 
| 13792 | 413 | apply (auto intro: JN_Join_diff dest: ok_sym simp add: OK_iff_ok) | 
| 414 | done | |
| 415 | ||
| 13805 | 416 | lemma wg_ex_prop: "ex_prop X ==> (F \<in> X) = (\<forall>H. H \<in> wg F X)" | 
| 13792 | 417 | apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv) | 
| 418 | apply blast | |
| 419 | done | |
| 420 | ||
| 421 | (** From Charpentier and Chandy "Theorems About Composition" **) | |
| 422 | (* Proposition 2 *) | |
| 423 | lemma wx_subset: "(wx X)<=X" | |
| 424 | by (unfold wx_def, auto) | |
| 425 | ||
| 426 | lemma wx_ex_prop: "ex_prop (wx X)" | |
| 16647 
c6d81ddebb0e
Proof of wx_ex_prop must now use old bex_cong to prevent simplifier from looping.
 berghofe parents: 
16417diff
changeset | 427 | apply (simp add: wx_def ex_prop_equiv cong: bex_cong, safe, blast) | 
| 14112 | 428 | apply force | 
| 13792 | 429 | done | 
| 430 | ||
| 13805 | 431 | lemma wx_weakest: "\<forall>Z. Z<= X --> ex_prop Z --> Z \<subseteq> wx X" | 
| 14112 | 432 | by (auto simp add: wx_def) | 
| 13792 | 433 | |
| 434 | (* Proposition 6 *) | |
| 13819 | 435 | lemma wx'_ex_prop: "ex_prop({F. \<forall>G. F ok G --> F\<squnion>G \<in> X})"
 | 
| 13792 | 436 | apply (unfold ex_prop_def, safe) | 
| 14112 | 437 | apply (drule_tac x = "G\<squnion>Ga" in spec) | 
| 44871 | 438 | apply (force simp add: Join_assoc) | 
| 13819 | 439 | apply (drule_tac x = "F\<squnion>Ga" in spec) | 
| 44871 | 440 | apply (simp add: ok_commute Join_ac) | 
| 13792 | 441 | done | 
| 442 | ||
| 63146 | 443 | text\<open>Equivalence with the other definition of wx\<close> | 
| 13792 | 444 | |
| 14112 | 445 | lemma wx_equiv: "wx X = {F. \<forall>G. F ok G --> (F\<squnion>G) \<in> X}"
 | 
| 13792 | 446 | apply (unfold wx_def, safe) | 
| 14112 | 447 | apply (simp add: ex_prop_def, blast) | 
| 13792 | 448 | apply (simp (no_asm)) | 
| 13819 | 449 | apply (rule_tac x = "{F. \<forall>G. F ok G --> F\<squnion>G \<in> X}" in exI, safe)
 | 
| 13792 | 450 | apply (rule_tac [2] wx'_ex_prop) | 
| 14112 | 451 | apply (drule_tac x = SKIP in spec)+ | 
| 452 | apply auto | |
| 13792 | 453 | done | 
| 454 | ||
| 455 | ||
| 63146 | 456 | text\<open>Propositions 7 to 11 are about this second definition of wx. | 
| 14112 | 457 | They are the same as the ones proved for the first definition of wx, | 
| 63146 | 458 | by equivalence\<close> | 
| 13792 | 459 | |
| 460 | (* Proposition 12 *) | |
| 461 | (* Main result of the paper *) | |
| 14112 | 462 | lemma guarantees_wx_eq: "(X guarantees Y) = wx(-X \<union> Y)" | 
| 463 | by (simp add: guar_def wx_equiv) | |
| 13792 | 464 | |
| 465 | ||
| 466 | (* Rules given in section 7 of Chandy and Sander's | |
| 467 | Reasoning About Program composition paper *) | |
| 468 | lemma stable_guarantees_Always: | |
| 14112 | 469 | "Init F \<subseteq> A ==> F \<in> (stable A) guarantees (Always A)" | 
| 13792 | 470 | apply (rule guaranteesI) | 
| 14112 | 471 | apply (simp add: Join_commute) | 
| 13792 | 472 | apply (rule stable_Join_Always1) | 
| 44871 | 473 | apply (simp_all add: invariant_def) | 
| 13792 | 474 | done | 
| 475 | ||
| 476 | lemma constrains_guarantees_leadsTo: | |
| 13805 | 477 | "F \<in> transient A ==> F \<in> (A co A \<union> B) guarantees (A leadsTo (B-A))" | 
| 13792 | 478 | apply (rule guaranteesI) | 
| 479 | apply (rule leadsTo_Basis') | |
| 14112 | 480 | apply (drule constrains_weaken_R) | 
| 481 | prefer 2 apply assumption | |
| 482 | apply blast | |
| 13792 | 483 | apply (blast intro: Join_transient_I1) | 
| 484 | done | |
| 485 | ||
| 7400 
fbd5582761e6
new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
 paulson parents: diff
changeset | 486 | end |