| author | wenzelm | 
| Sun, 04 Sep 2011 19:36:19 +0200 | |
| changeset 44706 | fe319b45315c | 
| parent 35416 | d8d7d1b785af | 
| child 44870 | 0d23a8ae14df | 
| permissions | -rw-r--r-- | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23767diff
changeset | 1 | (* Title: HOL/UNITY/Constrains.thy | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 3 | Copyright 1998 University of Cambridge | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 4 | |
| 13797 | 5 | Weak safety relations: restricted to the set of reachable states. | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 6 | *) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 7 | |
| 13798 | 8 | header{*Weak Safety*}
 | 
| 9 | ||
| 16417 | 10 | theory Constrains imports UNITY begin | 
| 6535 | 11 | |
| 12 | (*Initial states and program => (final state, reversed trace to it)... | |
| 13 | Arguments MUST be curried in an inductive definition*) | |
| 14 | ||
| 23767 | 15 | inductive_set | 
| 16 |   traces :: "['a set, ('a * 'a)set set] => ('a * 'a list) set"
 | |
| 17 |   for init :: "'a set" and acts :: "('a * 'a)set set"
 | |
| 18 | where | |
| 6535 | 19 | (*Initial trace is empty*) | 
| 13805 | 20 | Init: "s \<in> init ==> (s,[]) \<in> traces init acts" | 
| 6535 | 21 | |
| 23767 | 22 | | Acts: "[| act: acts; (s,evs) \<in> traces init acts; (s,s'): act |] | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23767diff
changeset | 23 | ==> (s', s#evs) \<in> traces init acts" | 
| 6535 | 24 | |
| 25 | ||
| 23767 | 26 | inductive_set | 
| 27 | reachable :: "'a program => 'a set" | |
| 28 | for F :: "'a program" | |
| 29 | where | |
| 13805 | 30 | Init: "s \<in> Init F ==> s \<in> reachable F" | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 31 | |
| 23767 | 32 | | Acts: "[| act: Acts F; s \<in> reachable F; (s,s'): act |] | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
23767diff
changeset | 33 | ==> s' \<in> reachable F" | 
| 6536 | 34 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 35 | definition Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) where | 
| 13805 | 36 |     "A Co B == {F. F \<in> (reachable F \<inter> A)  co  B}"
 | 
| 6536 | 37 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 38 | definition Unless :: "['a set, 'a set] => 'a program set" (infixl "Unless" 60) where | 
| 13805 | 39 | "A Unless B == (A-B) Co (A \<union> B)" | 
| 6536 | 40 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 41 | definition Stable :: "'a set => 'a program set" where | 
| 6536 | 42 | "Stable A == A Co A" | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 43 | |
| 6570 | 44 | (*Always is the weak form of "invariant"*) | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 45 | definition Always :: "'a set => 'a program set" where | 
| 13805 | 46 |     "Always A == {F. Init F \<subseteq> A} \<inter> Stable A"
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 47 | |
| 13805 | 48 | (*Polymorphic in both states and the meaning of \<le> *) | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 49 | definition Increasing :: "['a => 'b::{order}] => 'a program set" where
 | 
| 13805 | 50 |     "Increasing f == \<Inter>z. Stable {s. z \<le> f s}"
 | 
| 5784 | 51 | |
| 13797 | 52 | |
| 13798 | 53 | subsection{*traces and reachable*}
 | 
| 13797 | 54 | |
| 55 | lemma reachable_equiv_traces: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 56 |      "reachable F = {s. \<exists>evs. (s,evs) \<in> traces (Init F) (Acts F)}"
 | 
| 13797 | 57 | apply safe | 
| 58 | apply (erule_tac [2] traces.induct) | |
| 59 | apply (erule reachable.induct) | |
| 60 | apply (blast intro: reachable.intros traces.intros)+ | |
| 61 | done | |
| 62 | ||
| 13805 | 63 | lemma Init_subset_reachable: "Init F \<subseteq> reachable F" | 
| 13797 | 64 | by (blast intro: reachable.intros) | 
| 65 | ||
| 66 | lemma stable_reachable [intro!,simp]: | |
| 13805 | 67 | "Acts G \<subseteq> Acts F ==> G \<in> stable (reachable F)" | 
| 13797 | 68 | by (blast intro: stableI constrainsI reachable.intros) | 
| 69 | ||
| 70 | (*The set of all reachable states is an invariant...*) | |
| 13805 | 71 | lemma invariant_reachable: "F \<in> invariant (reachable F)" | 
| 13797 | 72 | apply (simp add: invariant_def) | 
| 73 | apply (blast intro: reachable.intros) | |
| 74 | done | |
| 75 | ||
| 76 | (*...in fact the strongest invariant!*) | |
| 13805 | 77 | lemma invariant_includes_reachable: "F \<in> invariant A ==> reachable F \<subseteq> A" | 
| 13797 | 78 | apply (simp add: stable_def constrains_def invariant_def) | 
| 79 | apply (rule subsetI) | |
| 80 | apply (erule reachable.induct) | |
| 81 | apply (blast intro: reachable.intros)+ | |
| 82 | done | |
| 83 | ||
| 84 | ||
| 13798 | 85 | subsection{*Co*}
 | 
| 13797 | 86 | |
| 13805 | 87 | (*F \<in> B co B' ==> F \<in> (reachable F \<inter> B) co (reachable F \<inter> B')*) | 
| 13797 | 88 | lemmas constrains_reachable_Int = | 
| 89 | subset_refl [THEN stable_reachable [unfolded stable_def], | |
| 90 | THEN constrains_Int, standard] | |
| 91 | ||
| 92 | (*Resembles the previous definition of Constrains*) | |
| 93 | lemma Constrains_eq_constrains: | |
| 13805 | 94 |      "A Co B = {F. F \<in> (reachable F  \<inter>  A) co (reachable F  \<inter>  B)}"
 | 
| 13797 | 95 | apply (unfold Constrains_def) | 
| 96 | apply (blast dest: constrains_reachable_Int intro: constrains_weaken) | |
| 97 | done | |
| 98 | ||
| 13805 | 99 | lemma constrains_imp_Constrains: "F \<in> A co A' ==> F \<in> A Co A'" | 
| 13797 | 100 | apply (unfold Constrains_def) | 
| 101 | apply (blast intro: constrains_weaken_L) | |
| 102 | done | |
| 103 | ||
| 13805 | 104 | lemma stable_imp_Stable: "F \<in> stable A ==> F \<in> Stable A" | 
| 13797 | 105 | apply (unfold stable_def Stable_def) | 
| 106 | apply (erule constrains_imp_Constrains) | |
| 107 | done | |
| 108 | ||
| 109 | lemma ConstrainsI: | |
| 13805 | 110 | "(!!act s s'. [| act: Acts F; (s,s') \<in> act; s \<in> A |] ==> s': A') | 
| 111 | ==> F \<in> A Co A'" | |
| 13797 | 112 | apply (rule constrains_imp_Constrains) | 
| 113 | apply (blast intro: constrainsI) | |
| 114 | done | |
| 115 | ||
| 13805 | 116 | lemma Constrains_empty [iff]: "F \<in> {} Co B"
 | 
| 13797 | 117 | by (unfold Constrains_def constrains_def, blast) | 
| 118 | ||
| 13805 | 119 | lemma Constrains_UNIV [iff]: "F \<in> A Co UNIV" | 
| 13797 | 120 | by (blast intro: ConstrainsI) | 
| 121 | ||
| 122 | lemma Constrains_weaken_R: | |
| 13805 | 123 | "[| F \<in> A Co A'; A'<=B' |] ==> F \<in> A Co B'" | 
| 13797 | 124 | apply (unfold Constrains_def) | 
| 125 | apply (blast intro: constrains_weaken_R) | |
| 126 | done | |
| 127 | ||
| 128 | lemma Constrains_weaken_L: | |
| 13805 | 129 | "[| F \<in> A Co A'; B \<subseteq> A |] ==> F \<in> B Co A'" | 
| 13797 | 130 | apply (unfold Constrains_def) | 
| 131 | apply (blast intro: constrains_weaken_L) | |
| 132 | done | |
| 133 | ||
| 134 | lemma Constrains_weaken: | |
| 13805 | 135 | "[| F \<in> A Co A'; B \<subseteq> A; A'<=B' |] ==> F \<in> B Co B'" | 
| 13797 | 136 | apply (unfold Constrains_def) | 
| 137 | apply (blast intro: constrains_weaken) | |
| 138 | done | |
| 139 | ||
| 140 | (** Union **) | |
| 141 | ||
| 142 | lemma Constrains_Un: | |
| 13805 | 143 | "[| F \<in> A Co A'; F \<in> B Co B' |] ==> F \<in> (A \<union> B) Co (A' \<union> B')" | 
| 13797 | 144 | apply (unfold Constrains_def) | 
| 145 | apply (blast intro: constrains_Un [THEN constrains_weaken]) | |
| 146 | done | |
| 147 | ||
| 148 | lemma Constrains_UN: | |
| 13805 | 149 | assumes Co: "!!i. i \<in> I ==> F \<in> (A i) Co (A' i)" | 
| 150 | shows "F \<in> (\<Union>i \<in> I. A i) Co (\<Union>i \<in> I. A' i)" | |
| 13797 | 151 | apply (unfold Constrains_def) | 
| 152 | apply (rule CollectI) | |
| 153 | apply (rule Co [unfolded Constrains_def, THEN CollectD, THEN constrains_UN, | |
| 154 | THEN constrains_weaken], auto) | |
| 155 | done | |
| 156 | ||
| 157 | (** Intersection **) | |
| 158 | ||
| 159 | lemma Constrains_Int: | |
| 13805 | 160 | "[| F \<in> A Co A'; F \<in> B Co B' |] ==> F \<in> (A \<inter> B) Co (A' \<inter> B')" | 
| 13797 | 161 | apply (unfold Constrains_def) | 
| 162 | apply (blast intro: constrains_Int [THEN constrains_weaken]) | |
| 163 | done | |
| 164 | ||
| 165 | lemma Constrains_INT: | |
| 13805 | 166 | assumes Co: "!!i. i \<in> I ==> F \<in> (A i) Co (A' i)" | 
| 167 | shows "F \<in> (\<Inter>i \<in> I. A i) Co (\<Inter>i \<in> I. A' i)" | |
| 13797 | 168 | apply (unfold Constrains_def) | 
| 169 | apply (rule CollectI) | |
| 170 | apply (rule Co [unfolded Constrains_def, THEN CollectD, THEN constrains_INT, | |
| 171 | THEN constrains_weaken], auto) | |
| 172 | done | |
| 173 | ||
| 13805 | 174 | lemma Constrains_imp_subset: "F \<in> A Co A' ==> reachable F \<inter> A \<subseteq> A'" | 
| 13797 | 175 | by (simp add: constrains_imp_subset Constrains_def) | 
| 176 | ||
| 13805 | 177 | lemma Constrains_trans: "[| F \<in> A Co B; F \<in> B Co C |] ==> F \<in> A Co C" | 
| 13797 | 178 | apply (simp add: Constrains_eq_constrains) | 
| 179 | apply (blast intro: constrains_trans constrains_weaken) | |
| 180 | done | |
| 181 | ||
| 182 | lemma Constrains_cancel: | |
| 13805 | 183 | "[| F \<in> A Co (A' \<union> B); F \<in> B Co B' |] ==> F \<in> A Co (A' \<union> B')" | 
| 13797 | 184 | by (simp add: Constrains_eq_constrains constrains_def, blast) | 
| 185 | ||
| 186 | ||
| 13798 | 187 | subsection{*Stable*}
 | 
| 13797 | 188 | |
| 189 | (*Useful because there's no Stable_weaken. [Tanja Vos]*) | |
| 13805 | 190 | lemma Stable_eq: "[| F \<in> Stable A; A = B |] ==> F \<in> Stable B" | 
| 13797 | 191 | by blast | 
| 192 | ||
| 13805 | 193 | lemma Stable_eq_stable: "(F \<in> Stable A) = (F \<in> stable (reachable F \<inter> A))" | 
| 13797 | 194 | by (simp add: Stable_def Constrains_eq_constrains stable_def) | 
| 195 | ||
| 13805 | 196 | lemma StableI: "F \<in> A Co A ==> F \<in> Stable A" | 
| 13797 | 197 | by (unfold Stable_def, assumption) | 
| 198 | ||
| 13805 | 199 | lemma StableD: "F \<in> Stable A ==> F \<in> A Co A" | 
| 13797 | 200 | by (unfold Stable_def, assumption) | 
| 201 | ||
| 202 | lemma Stable_Un: | |
| 13805 | 203 | "[| F \<in> Stable A; F \<in> Stable A' |] ==> F \<in> Stable (A \<union> A')" | 
| 13797 | 204 | apply (unfold Stable_def) | 
| 205 | apply (blast intro: Constrains_Un) | |
| 206 | done | |
| 207 | ||
| 208 | lemma Stable_Int: | |
| 13805 | 209 | "[| F \<in> Stable A; F \<in> Stable A' |] ==> F \<in> Stable (A \<inter> A')" | 
| 13797 | 210 | apply (unfold Stable_def) | 
| 211 | apply (blast intro: Constrains_Int) | |
| 212 | done | |
| 213 | ||
| 214 | lemma Stable_Constrains_Un: | |
| 13805 | 215 | "[| F \<in> Stable C; F \<in> A Co (C \<union> A') |] | 
| 216 | ==> F \<in> (C \<union> A) Co (C \<union> A')" | |
| 13797 | 217 | apply (unfold Stable_def) | 
| 218 | apply (blast intro: Constrains_Un [THEN Constrains_weaken]) | |
| 219 | done | |
| 220 | ||
| 221 | lemma Stable_Constrains_Int: | |
| 13805 | 222 | "[| F \<in> Stable C; F \<in> (C \<inter> A) Co A' |] | 
| 223 | ==> F \<in> (C \<inter> A) Co (C \<inter> A')" | |
| 13797 | 224 | apply (unfold Stable_def) | 
| 225 | apply (blast intro: Constrains_Int [THEN Constrains_weaken]) | |
| 226 | done | |
| 227 | ||
| 228 | lemma Stable_UN: | |
| 13805 | 229 | "(!!i. i \<in> I ==> F \<in> Stable (A i)) ==> F \<in> Stable (\<Union>i \<in> I. A i)" | 
| 13797 | 230 | by (simp add: Stable_def Constrains_UN) | 
| 231 | ||
| 232 | lemma Stable_INT: | |
| 13805 | 233 | "(!!i. i \<in> I ==> F \<in> Stable (A i)) ==> F \<in> Stable (\<Inter>i \<in> I. A i)" | 
| 13797 | 234 | by (simp add: Stable_def Constrains_INT) | 
| 235 | ||
| 13805 | 236 | lemma Stable_reachable: "F \<in> Stable (reachable F)" | 
| 13797 | 237 | by (simp add: Stable_eq_stable) | 
| 238 | ||
| 239 | ||
| 240 | ||
| 13798 | 241 | subsection{*Increasing*}
 | 
| 13797 | 242 | |
| 243 | lemma IncreasingD: | |
| 13805 | 244 |      "F \<in> Increasing f ==> F \<in> Stable {s. x \<le> f s}"
 | 
| 13797 | 245 | by (unfold Increasing_def, blast) | 
| 246 | ||
| 247 | lemma mono_Increasing_o: | |
| 13805 | 248 | "mono g ==> Increasing f \<subseteq> Increasing (g o f)" | 
| 13797 | 249 | apply (simp add: Increasing_def Stable_def Constrains_def stable_def | 
| 250 | constrains_def) | |
| 251 | apply (blast intro: monoD order_trans) | |
| 252 | done | |
| 253 | ||
| 254 | lemma strict_IncreasingD: | |
| 13805 | 255 |      "!!z::nat. F \<in> Increasing f ==> F \<in> Stable {s. z < f s}"
 | 
| 13797 | 256 | by (simp add: Increasing_def Suc_le_eq [symmetric]) | 
| 257 | ||
| 258 | lemma increasing_imp_Increasing: | |
| 13805 | 259 | "F \<in> increasing f ==> F \<in> Increasing f" | 
| 13797 | 260 | apply (unfold increasing_def Increasing_def) | 
| 261 | apply (blast intro: stable_imp_Stable) | |
| 262 | done | |
| 263 | ||
| 264 | lemmas Increasing_constant = | |
| 265 | increasing_constant [THEN increasing_imp_Increasing, standard, iff] | |
| 266 | ||
| 267 | ||
| 13798 | 268 | subsection{*The Elimination Theorem*}
 | 
| 269 | ||
| 270 | (*The "free" m has become universally quantified! Should the premise be !!m | |
| 13805 | 271 | instead of \<forall>m ? Would make it harder to use in forward proof.*) | 
| 13797 | 272 | |
| 273 | lemma Elimination: | |
| 13805 | 274 |     "[| \<forall>m. F \<in> {s. s x = m} Co (B m) |]  
 | 
| 275 |      ==> F \<in> {s. s x \<in> M} Co (\<Union>m \<in> M. B m)"
 | |
| 13797 | 276 | by (unfold Constrains_def constrains_def, blast) | 
| 277 | ||
| 278 | (*As above, but for the trivial case of a one-variable state, in which the | |
| 279 | state is identified with its one variable.*) | |
| 280 | lemma Elimination_sing: | |
| 13805 | 281 |     "(\<forall>m. F \<in> {m} Co (B m)) ==> F \<in> M Co (\<Union>m \<in> M. B m)"
 | 
| 13797 | 282 | by (unfold Constrains_def constrains_def, blast) | 
| 283 | ||
| 284 | ||
| 13798 | 285 | subsection{*Specialized laws for handling Always*}
 | 
| 13797 | 286 | |
| 287 | (** Natural deduction rules for "Always A" **) | |
| 288 | ||
| 13805 | 289 | lemma AlwaysI: "[| Init F \<subseteq> A; F \<in> Stable A |] ==> F \<in> Always A" | 
| 13797 | 290 | by (simp add: Always_def) | 
| 291 | ||
| 13805 | 292 | lemma AlwaysD: "F \<in> Always A ==> Init F \<subseteq> A & F \<in> Stable A" | 
| 13797 | 293 | by (simp add: Always_def) | 
| 294 | ||
| 295 | lemmas AlwaysE = AlwaysD [THEN conjE, standard] | |
| 296 | lemmas Always_imp_Stable = AlwaysD [THEN conjunct2, standard] | |
| 297 | ||
| 298 | ||
| 299 | (*The set of all reachable states is Always*) | |
| 13805 | 300 | lemma Always_includes_reachable: "F \<in> Always A ==> reachable F \<subseteq> A" | 
| 13797 | 301 | apply (simp add: Stable_def Constrains_def constrains_def Always_def) | 
| 302 | apply (rule subsetI) | |
| 303 | apply (erule reachable.induct) | |
| 304 | apply (blast intro: reachable.intros)+ | |
| 305 | done | |
| 306 | ||
| 307 | lemma invariant_imp_Always: | |
| 13805 | 308 | "F \<in> invariant A ==> F \<in> Always A" | 
| 13797 | 309 | apply (unfold Always_def invariant_def Stable_def stable_def) | 
| 310 | apply (blast intro: constrains_imp_Constrains) | |
| 311 | done | |
| 312 | ||
| 313 | lemmas Always_reachable = | |
| 314 | invariant_reachable [THEN invariant_imp_Always, standard] | |
| 315 | ||
| 316 | lemma Always_eq_invariant_reachable: | |
| 13805 | 317 |      "Always A = {F. F \<in> invariant (reachable F \<inter> A)}"
 | 
| 13797 | 318 | apply (simp add: Always_def invariant_def Stable_def Constrains_eq_constrains | 
| 319 | stable_def) | |
| 320 | apply (blast intro: reachable.intros) | |
| 321 | done | |
| 322 | ||
| 323 | (*the RHS is the traditional definition of the "always" operator*) | |
| 13805 | 324 | lemma Always_eq_includes_reachable: "Always A = {F. reachable F \<subseteq> A}"
 | 
| 13797 | 325 | by (auto dest: invariant_includes_reachable simp add: Int_absorb2 invariant_reachable Always_eq_invariant_reachable) | 
| 326 | ||
| 327 | lemma Always_UNIV_eq [simp]: "Always UNIV = UNIV" | |
| 328 | by (auto simp add: Always_eq_includes_reachable) | |
| 329 | ||
| 13805 | 330 | lemma UNIV_AlwaysI: "UNIV \<subseteq> A ==> F \<in> Always A" | 
| 13797 | 331 | by (auto simp add: Always_eq_includes_reachable) | 
| 332 | ||
| 13805 | 333 | lemma Always_eq_UN_invariant: "Always A = (\<Union>I \<in> Pow A. invariant I)" | 
| 13797 | 334 | apply (simp add: Always_eq_includes_reachable) | 
| 335 | apply (blast intro: invariantI Init_subset_reachable [THEN subsetD] | |
| 336 | invariant_includes_reachable [THEN subsetD]) | |
| 337 | done | |
| 338 | ||
| 13805 | 339 | lemma Always_weaken: "[| F \<in> Always A; A \<subseteq> B |] ==> F \<in> Always B" | 
| 13797 | 340 | by (auto simp add: Always_eq_includes_reachable) | 
| 341 | ||
| 342 | ||
| 13798 | 343 | subsection{*"Co" rules involving Always*}
 | 
| 13797 | 344 | |
| 345 | lemma Always_Constrains_pre: | |
| 13805 | 346 | "F \<in> Always INV ==> (F \<in> (INV \<inter> A) Co A') = (F \<in> A Co A')" | 
| 13797 | 347 | by (simp add: Always_includes_reachable [THEN Int_absorb2] Constrains_def | 
| 348 | Int_assoc [symmetric]) | |
| 349 | ||
| 350 | lemma Always_Constrains_post: | |
| 13805 | 351 | "F \<in> Always INV ==> (F \<in> A Co (INV \<inter> A')) = (F \<in> A Co A')" | 
| 13797 | 352 | by (simp add: Always_includes_reachable [THEN Int_absorb2] | 
| 353 | Constrains_eq_constrains Int_assoc [symmetric]) | |
| 354 | ||
| 13805 | 355 | (* [| F \<in> Always INV; F \<in> (INV \<inter> A) Co A' |] ==> F \<in> A Co A' *) | 
| 13797 | 356 | lemmas Always_ConstrainsI = Always_Constrains_pre [THEN iffD1, standard] | 
| 357 | ||
| 13805 | 358 | (* [| F \<in> Always INV; F \<in> A Co A' |] ==> F \<in> A Co (INV \<inter> A') *) | 
| 13797 | 359 | lemmas Always_ConstrainsD = Always_Constrains_post [THEN iffD2, standard] | 
| 360 | ||
| 361 | (*The analogous proof of Always_LeadsTo_weaken doesn't terminate*) | |
| 362 | lemma Always_Constrains_weaken: | |
| 13805 | 363 | "[| F \<in> Always C; F \<in> A Co A'; | 
| 364 | C \<inter> B \<subseteq> A; C \<inter> A' \<subseteq> B' |] | |
| 365 | ==> F \<in> B Co B'" | |
| 13797 | 366 | apply (rule Always_ConstrainsI, assumption) | 
| 367 | apply (drule Always_ConstrainsD, assumption) | |
| 368 | apply (blast intro: Constrains_weaken) | |
| 369 | done | |
| 370 | ||
| 371 | ||
| 372 | (** Conjoining Always properties **) | |
| 373 | ||
| 13805 | 374 | lemma Always_Int_distrib: "Always (A \<inter> B) = Always A \<inter> Always B" | 
| 13797 | 375 | by (auto simp add: Always_eq_includes_reachable) | 
| 376 | ||
| 13805 | 377 | lemma Always_INT_distrib: "Always (INTER I A) = (\<Inter>i \<in> I. Always (A i))" | 
| 13797 | 378 | by (auto simp add: Always_eq_includes_reachable) | 
| 379 | ||
| 380 | lemma Always_Int_I: | |
| 13805 | 381 | "[| F \<in> Always A; F \<in> Always B |] ==> F \<in> Always (A \<inter> B)" | 
| 13797 | 382 | by (simp add: Always_Int_distrib) | 
| 383 | ||
| 384 | (*Allows a kind of "implication introduction"*) | |
| 385 | lemma Always_Compl_Un_eq: | |
| 13805 | 386 | "F \<in> Always A ==> (F \<in> Always (-A \<union> B)) = (F \<in> Always B)" | 
| 13797 | 387 | by (auto simp add: Always_eq_includes_reachable) | 
| 388 | ||
| 389 | (*Delete the nearest invariance assumption (which will be the second one | |
| 390 | used by Always_Int_I) *) | |
| 13805 | 391 | lemmas Always_thin = thin_rl [of "F \<in> Always A", standard] | 
| 13797 | 392 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 393 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 394 | subsection{*Totalize*}
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 395 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 396 | lemma reachable_imp_reachable_tot: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 397 | "s \<in> reachable F ==> s \<in> reachable (totalize F)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 398 | apply (erule reachable.induct) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 399 | apply (rule reachable.Init) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 400 | apply simp | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 401 | apply (rule_tac act = "totalize_act act" in reachable.Acts) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 402 | apply (auto simp add: totalize_act_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 403 | done | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 404 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 405 | lemma reachable_tot_imp_reachable: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 406 | "s \<in> reachable (totalize F) ==> s \<in> reachable F" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 407 | apply (erule reachable.induct) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 408 | apply (rule reachable.Init, simp) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 409 | apply (force simp add: totalize_act_def intro: reachable.Acts) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 410 | done | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 411 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 412 | lemma reachable_tot_eq [simp]: "reachable (totalize F) = reachable F" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 413 | by (blast intro: reachable_imp_reachable_tot reachable_tot_imp_reachable) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 414 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 415 | lemma totalize_Constrains_iff [simp]: "(totalize F \<in> A Co B) = (F \<in> A Co B)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 416 | by (simp add: Constrains_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 417 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 418 | lemma totalize_Stable_iff [simp]: "(totalize F \<in> Stable A) = (F \<in> Stable A)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 419 | by (simp add: Stable_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 420 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 421 | lemma totalize_Always_iff [simp]: "(totalize F \<in> Always A) = (F \<in> Always A)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 422 | by (simp add: Always_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 423 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 424 | end |