| author | wenzelm | 
| Thu, 28 Apr 2016 11:47:01 +0200 | |
| changeset 63067 | 0a8a75e400da | 
| parent 59807 | 22bc39064290 | 
| child 63146 | f1ecba0272f9 | 
| permissions | -rw-r--r-- | 
| 24147 | 1 | (* Title: HOL/UNITY/Project.thy | 
| 7630 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1999 University of Cambridge | |
| 4 | ||
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24147diff
changeset | 5 | Projections of state sets (also of actions and programs). | 
| 7630 | 6 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24147diff
changeset | 7 | Inheritance of GUARANTEES properties under extension. | 
| 7630 | 8 | *) | 
| 9 | ||
| 58889 | 10 | section{*Projections of State Sets*}
 | 
| 13798 | 11 | |
| 16417 | 12 | theory Project imports Extend begin | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 13 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 14 | definition projecting :: "['c program => 'c set, 'a*'b => 'c, | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 15 | 'a program, 'c program set, 'a program set] => bool" where | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8055diff
changeset | 16 | "projecting C h F X' X == | 
| 13819 | 17 | \<forall>G. extend h F\<squnion>G \<in> X' --> F\<squnion>project h (C G) G \<in> X" | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 18 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 19 | definition extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 20 | 'c program set, 'a program set] => bool" where | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8055diff
changeset | 21 | "extending C h F Y' Y == | 
| 13819 | 22 | \<forall>G. extend h F ok G --> F\<squnion>project h (C G) G \<in> Y | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24147diff
changeset | 23 | --> extend h F\<squnion>G \<in> Y'" | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8055diff
changeset | 24 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 25 | definition subset_closed :: "'a set set => bool" where | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 26 | "subset_closed U == \<forall>A \<in> U. Pow A \<subseteq> U" | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 27 | |
| 13790 | 28 | |
| 46912 | 29 | context Extend | 
| 30 | begin | |
| 31 | ||
| 32 | lemma project_extend_constrains_I: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 33 | "F \<in> A co B ==> project h C (extend h F) \<in> A co B" | 
| 13790 | 34 | apply (auto simp add: extend_act_def project_act_def constrains_def) | 
| 35 | done | |
| 36 | ||
| 37 | ||
| 13798 | 38 | subsection{*Safety*}
 | 
| 13790 | 39 | |
| 40 | (*used below to prove Join_project_ensures*) | |
| 46912 | 41 | lemma project_unless: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 42 | "[| G \<in> stable C; project h C G \<in> A unless B |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 43 | ==> G \<in> (C \<inter> extend_set h A) unless (extend_set h B)" | 
| 13790 | 44 | apply (simp add: unless_def project_constrains) | 
| 45 | apply (blast dest: stable_constrains_Int intro: constrains_weaken) | |
| 46 | done | |
| 47 | ||
| 13819 | 48 | (*Generalizes project_constrains to the program F\<squnion>project h C G | 
| 13790 | 49 | useful with guarantees reasoning*) | 
| 46912 | 50 | lemma Join_project_constrains: | 
| 13819 | 51 | "(F\<squnion>project h C G \<in> A co B) = | 
| 52 | (extend h F\<squnion>G \<in> (C \<inter> extend_set h A) co (extend_set h B) & | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 53 | F \<in> A co B)" | 
| 13790 | 54 | apply (simp (no_asm) add: project_constrains) | 
| 55 | apply (blast intro: extend_constrains [THEN iffD2, THEN constrains_weaken] | |
| 56 | dest: constrains_imp_subset) | |
| 57 | done | |
| 58 | ||
| 59 | (*The condition is required to prove the left-to-right direction | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 60 | could weaken it to G \<in> (C \<inter> extend_set h A) co C*) | 
| 46912 | 61 | lemma Join_project_stable: | 
| 13819 | 62 | "extend h F\<squnion>G \<in> stable C | 
| 63 | ==> (F\<squnion>project h C G \<in> stable A) = | |
| 64 | (extend h F\<squnion>G \<in> stable (C \<inter> extend_set h A) & | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 65 | F \<in> stable A)" | 
| 13790 | 66 | apply (unfold stable_def) | 
| 67 | apply (simp only: Join_project_constrains) | |
| 68 | apply (blast intro: constrains_weaken dest: constrains_Int) | |
| 69 | done | |
| 70 | ||
| 71 | (*For using project_guarantees in particular cases*) | |
| 46912 | 72 | lemma project_constrains_I: | 
| 13819 | 73 | "extend h F\<squnion>G \<in> extend_set h A co extend_set h B | 
| 74 | ==> F\<squnion>project h C G \<in> A co B" | |
| 13790 | 75 | apply (simp add: project_constrains extend_constrains) | 
| 76 | apply (blast intro: constrains_weaken dest: constrains_imp_subset) | |
| 77 | done | |
| 78 | ||
| 46912 | 79 | lemma project_increasing_I: | 
| 13819 | 80 | "extend h F\<squnion>G \<in> increasing (func o f) | 
| 81 | ==> F\<squnion>project h C G \<in> increasing func" | |
| 13790 | 82 | apply (unfold increasing_def stable_def) | 
| 83 | apply (simp del: Join_constrains | |
| 84 | add: project_constrains_I extend_set_eq_Collect) | |
| 85 | done | |
| 86 | ||
| 46912 | 87 | lemma Join_project_increasing: | 
| 13819 | 88 | "(F\<squnion>project h UNIV G \<in> increasing func) = | 
| 89 | (extend h F\<squnion>G \<in> increasing (func o f))" | |
| 13790 | 90 | apply (rule iffI) | 
| 91 | apply (erule_tac [2] project_increasing_I) | |
| 92 | apply (simp del: Join_stable | |
| 93 | add: increasing_def Join_project_stable) | |
| 94 | apply (auto simp add: extend_set_eq_Collect extend_stable [THEN iffD1]) | |
| 95 | done | |
| 96 | ||
| 97 | (*The UNIV argument is essential*) | |
| 46912 | 98 | lemma project_constrains_D: | 
| 13819 | 99 | "F\<squnion>project h UNIV G \<in> A co B | 
| 100 | ==> extend h F\<squnion>G \<in> extend_set h A co extend_set h B" | |
| 13790 | 101 | by (simp add: project_constrains extend_constrains) | 
| 102 | ||
| 46912 | 103 | end | 
| 104 | ||
| 13790 | 105 | |
| 13798 | 106 | subsection{*"projecting" and union/intersection (no converses)*}
 | 
| 13790 | 107 | |
| 108 | lemma projecting_Int: | |
| 109 | "[| projecting C h F XA' XA; projecting C h F XB' XB |] | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 110 | ==> projecting C h F (XA' \<inter> XB') (XA \<inter> XB)" | 
| 13790 | 111 | by (unfold projecting_def, blast) | 
| 112 | ||
| 113 | lemma projecting_Un: | |
| 114 | "[| projecting C h F XA' XA; projecting C h F XB' XB |] | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 115 | ==> projecting C h F (XA' \<union> XB') (XA \<union> XB)" | 
| 13790 | 116 | by (unfold projecting_def, blast) | 
| 117 | ||
| 118 | lemma projecting_INT: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 119 | "[| !!i. i \<in> I ==> projecting C h F (X' i) (X i) |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 120 | ==> projecting C h F (\<Inter>i \<in> I. X' i) (\<Inter>i \<in> I. X i)" | 
| 13790 | 121 | by (unfold projecting_def, blast) | 
| 122 | ||
| 123 | lemma projecting_UN: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 124 | "[| !!i. i \<in> I ==> projecting C h F (X' i) (X i) |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 125 | ==> projecting C h F (\<Union>i \<in> I. X' i) (\<Union>i \<in> I. X i)" | 
| 13790 | 126 | by (unfold projecting_def, blast) | 
| 127 | ||
| 128 | lemma projecting_weaken: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 129 | "[| projecting C h F X' X; U'<=X'; X \<subseteq> U |] ==> projecting C h F U' U" | 
| 13790 | 130 | by (unfold projecting_def, auto) | 
| 131 | ||
| 132 | lemma projecting_weaken_L: | |
| 133 | "[| projecting C h F X' X; U'<=X' |] ==> projecting C h F U' X" | |
| 134 | by (unfold projecting_def, auto) | |
| 135 | ||
| 136 | lemma extending_Int: | |
| 137 | "[| extending C h F YA' YA; extending C h F YB' YB |] | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 138 | ==> extending C h F (YA' \<inter> YB') (YA \<inter> YB)" | 
| 13790 | 139 | by (unfold extending_def, blast) | 
| 140 | ||
| 141 | lemma extending_Un: | |
| 142 | "[| extending C h F YA' YA; extending C h F YB' YB |] | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 143 | ==> extending C h F (YA' \<union> YB') (YA \<union> YB)" | 
| 13790 | 144 | by (unfold extending_def, blast) | 
| 145 | ||
| 146 | lemma extending_INT: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 147 | "[| !!i. i \<in> I ==> extending C h F (Y' i) (Y i) |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 148 | ==> extending C h F (\<Inter>i \<in> I. Y' i) (\<Inter>i \<in> I. Y i)" | 
| 13790 | 149 | by (unfold extending_def, blast) | 
| 150 | ||
| 151 | lemma extending_UN: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 152 | "[| !!i. i \<in> I ==> extending C h F (Y' i) (Y i) |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 153 | ==> extending C h F (\<Union>i \<in> I. Y' i) (\<Union>i \<in> I. Y i)" | 
| 13790 | 154 | by (unfold extending_def, blast) | 
| 155 | ||
| 156 | lemma extending_weaken: | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 157 | "[| extending C h F Y' Y; Y'<=V'; V \<subseteq> Y |] ==> extending C h F V' V" | 
| 13790 | 158 | by (unfold extending_def, auto) | 
| 159 | ||
| 160 | lemma extending_weaken_L: | |
| 161 | "[| extending C h F Y' Y; Y'<=V' |] ==> extending C h F V' Y" | |
| 162 | by (unfold extending_def, auto) | |
| 163 | ||
| 164 | lemma projecting_UNIV: "projecting C h F X' UNIV" | |
| 165 | by (simp add: projecting_def) | |
| 166 | ||
| 46912 | 167 | context Extend | 
| 168 | begin | |
| 169 | ||
| 170 | lemma projecting_constrains: | |
| 13790 | 171 | "projecting C h F (extend_set h A co extend_set h B) (A co B)" | 
| 172 | apply (unfold projecting_def) | |
| 173 | apply (blast intro: project_constrains_I) | |
| 174 | done | |
| 175 | ||
| 46912 | 176 | lemma projecting_stable: | 
| 13790 | 177 | "projecting C h F (stable (extend_set h A)) (stable A)" | 
| 178 | apply (unfold stable_def) | |
| 179 | apply (rule projecting_constrains) | |
| 180 | done | |
| 181 | ||
| 46912 | 182 | lemma projecting_increasing: | 
| 13790 | 183 | "projecting C h F (increasing (func o f)) (increasing func)" | 
| 184 | apply (unfold projecting_def) | |
| 185 | apply (blast intro: project_increasing_I) | |
| 186 | done | |
| 187 | ||
| 46912 | 188 | lemma extending_UNIV: "extending C h F UNIV Y" | 
| 13790 | 189 | apply (simp (no_asm) add: extending_def) | 
| 190 | done | |
| 191 | ||
| 46912 | 192 | lemma extending_constrains: | 
| 13790 | 193 | "extending (%G. UNIV) h F (extend_set h A co extend_set h B) (A co B)" | 
| 194 | apply (unfold extending_def) | |
| 195 | apply (blast intro: project_constrains_D) | |
| 196 | done | |
| 197 | ||
| 46912 | 198 | lemma extending_stable: | 
| 13790 | 199 | "extending (%G. UNIV) h F (stable (extend_set h A)) (stable A)" | 
| 200 | apply (unfold stable_def) | |
| 201 | apply (rule extending_constrains) | |
| 202 | done | |
| 203 | ||
| 46912 | 204 | lemma extending_increasing: | 
| 13790 | 205 | "extending (%G. UNIV) h F (increasing (func o f)) (increasing func)" | 
| 206 | by (force simp only: extending_def Join_project_increasing) | |
| 207 | ||
| 208 | ||
| 13798 | 209 | subsection{*Reachability and project*}
 | 
| 13790 | 210 | |
| 211 | (*In practice, C = reachable(...): the inclusion is equality*) | |
| 46912 | 212 | lemma reachable_imp_reachable_project: | 
| 13819 | 213 | "[| reachable (extend h F\<squnion>G) \<subseteq> C; | 
| 214 | z \<in> reachable (extend h F\<squnion>G) |] | |
| 215 | ==> f z \<in> reachable (F\<squnion>project h C G)" | |
| 13790 | 216 | apply (erule reachable.induct) | 
| 217 | apply (force intro!: reachable.Init simp add: split_extended_all, auto) | |
| 218 | apply (rule_tac act = x in reachable.Acts) | |
| 219 | apply auto | |
| 220 | apply (erule extend_act_D) | |
| 221 | apply (rule_tac act1 = "Restrict C act" | |
| 222 | in project_act_I [THEN [3] reachable.Acts], auto) | |
| 223 | done | |
| 224 | ||
| 46912 | 225 | lemma project_Constrains_D: | 
| 13819 | 226 | "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Co B | 
| 227 | ==> extend h F\<squnion>G \<in> (extend_set h A) Co (extend_set h B)" | |
| 13790 | 228 | apply (unfold Constrains_def) | 
| 229 | apply (simp del: Join_constrains | |
| 230 | add: Join_project_constrains, clarify) | |
| 231 | apply (erule constrains_weaken) | |
| 232 | apply (auto intro: reachable_imp_reachable_project) | |
| 233 | done | |
| 234 | ||
| 46912 | 235 | lemma project_Stable_D: | 
| 13819 | 236 | "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Stable A | 
| 237 | ==> extend h F\<squnion>G \<in> Stable (extend_set h A)" | |
| 13790 | 238 | apply (unfold Stable_def) | 
| 239 | apply (simp (no_asm_simp) add: project_Constrains_D) | |
| 240 | done | |
| 241 | ||
| 46912 | 242 | lemma project_Always_D: | 
| 13819 | 243 | "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Always A | 
| 244 | ==> extend h F\<squnion>G \<in> Always (extend_set h A)" | |
| 13790 | 245 | apply (unfold Always_def) | 
| 246 | apply (force intro: reachable.Init simp add: project_Stable_D split_extended_all) | |
| 247 | done | |
| 248 | ||
| 46912 | 249 | lemma project_Increasing_D: | 
| 13819 | 250 | "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Increasing func | 
| 251 | ==> extend h F\<squnion>G \<in> Increasing (func o f)" | |
| 13790 | 252 | apply (unfold Increasing_def, auto) | 
| 253 | apply (subst extend_set_eq_Collect [symmetric]) | |
| 254 | apply (simp (no_asm_simp) add: project_Stable_D) | |
| 255 | done | |
| 256 | ||
| 257 | ||
| 13798 | 258 | subsection{*Converse results for weak safety: benefits of the argument C *}
 | 
| 13790 | 259 | |
| 260 | (*In practice, C = reachable(...): the inclusion is equality*) | |
| 46912 | 261 | lemma reachable_project_imp_reachable: | 
| 13819 | 262 | "[| C \<subseteq> reachable(extend h F\<squnion>G); | 
| 263 | x \<in> reachable (F\<squnion>project h C G) |] | |
| 264 | ==> \<exists>y. h(x,y) \<in> reachable (extend h F\<squnion>G)" | |
| 13790 | 265 | apply (erule reachable.induct) | 
| 266 | apply (force intro: reachable.Init) | |
| 267 | apply (auto simp add: project_act_def) | |
| 268 | apply (force del: Id_in_Acts intro: reachable.Acts extend_act_D)+ | |
| 269 | done | |
| 270 | ||
| 46912 | 271 | lemma project_set_reachable_extend_eq: | 
| 13819 | 272 | "project_set h (reachable (extend h F\<squnion>G)) = | 
| 273 | reachable (F\<squnion>project h (reachable (extend h F\<squnion>G)) G)" | |
| 13790 | 274 | by (auto dest: subset_refl [THEN reachable_imp_reachable_project] | 
| 275 | subset_refl [THEN reachable_project_imp_reachable]) | |
| 276 | ||
| 277 | (*UNUSED*) | |
| 46912 | 278 | lemma reachable_extend_Join_subset: | 
| 13819 | 279 | "reachable (extend h F\<squnion>G) \<subseteq> C | 
| 280 | ==> reachable (extend h F\<squnion>G) \<subseteq> | |
| 281 | extend_set h (reachable (F\<squnion>project h C G))" | |
| 13790 | 282 | apply (auto dest: reachable_imp_reachable_project) | 
| 283 | done | |
| 284 | ||
| 46912 | 285 | lemma project_Constrains_I: | 
| 13819 | 286 | "extend h F\<squnion>G \<in> (extend_set h A) Co (extend_set h B) | 
| 287 | ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Co B" | |
| 13790 | 288 | apply (unfold Constrains_def) | 
| 289 | apply (simp del: Join_constrains | |
| 290 | add: Join_project_constrains extend_set_Int_distrib) | |
| 291 | apply (rule conjI) | |
| 292 | prefer 2 | |
| 293 | apply (force elim: constrains_weaken_L | |
| 294 | dest!: extend_constrains_project_set | |
| 295 | subset_refl [THEN reachable_project_imp_reachable]) | |
| 296 | apply (blast intro: constrains_weaken_L) | |
| 297 | done | |
| 298 | ||
| 46912 | 299 | lemma project_Stable_I: | 
| 13819 | 300 | "extend h F\<squnion>G \<in> Stable (extend_set h A) | 
| 301 | ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Stable A" | |
| 13790 | 302 | apply (unfold Stable_def) | 
| 303 | apply (simp (no_asm_simp) add: project_Constrains_I) | |
| 304 | done | |
| 305 | ||
| 46912 | 306 | lemma project_Always_I: | 
| 13819 | 307 | "extend h F\<squnion>G \<in> Always (extend_set h A) | 
| 308 | ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Always A" | |
| 13790 | 309 | apply (unfold Always_def) | 
| 310 | apply (auto simp add: project_Stable_I) | |
| 311 | apply (unfold extend_set_def, blast) | |
| 312 | done | |
| 313 | ||
| 46912 | 314 | lemma project_Increasing_I: | 
| 13819 | 315 | "extend h F\<squnion>G \<in> Increasing (func o f) | 
| 316 | ==> F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Increasing func" | |
| 13790 | 317 | apply (unfold Increasing_def, auto) | 
| 318 | apply (simp (no_asm_simp) add: extend_set_eq_Collect project_Stable_I) | |
| 319 | done | |
| 320 | ||
| 46912 | 321 | lemma project_Constrains: | 
| 13819 | 322 | "(F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Co B) = | 
| 323 | (extend h F\<squnion>G \<in> (extend_set h A) Co (extend_set h B))" | |
| 13790 | 324 | apply (blast intro: project_Constrains_I project_Constrains_D) | 
| 325 | done | |
| 326 | ||
| 46912 | 327 | lemma project_Stable: | 
| 13819 | 328 | "(F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Stable A) = | 
| 329 | (extend h F\<squnion>G \<in> Stable (extend_set h A))" | |
| 13790 | 330 | apply (unfold Stable_def) | 
| 331 | apply (rule project_Constrains) | |
| 332 | done | |
| 333 | ||
| 46912 | 334 | lemma project_Increasing: | 
| 13819 | 335 | "(F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> Increasing func) = | 
| 336 | (extend h F\<squnion>G \<in> Increasing (func o f))" | |
| 13790 | 337 | apply (simp (no_asm_simp) add: Increasing_def project_Stable extend_set_eq_Collect) | 
| 338 | done | |
| 339 | ||
| 13798 | 340 | subsection{*A lot of redundant theorems: all are proved to facilitate reasoning
 | 
| 341 | about guarantees.*} | |
| 13790 | 342 | |
| 46912 | 343 | lemma projecting_Constrains: | 
| 13819 | 344 | "projecting (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 345 | (extend_set h A Co extend_set h B) (A Co B)" | 
| 346 | ||
| 347 | apply (unfold projecting_def) | |
| 348 | apply (blast intro: project_Constrains_I) | |
| 349 | done | |
| 350 | ||
| 46912 | 351 | lemma projecting_Stable: | 
| 13819 | 352 | "projecting (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 353 | (Stable (extend_set h A)) (Stable A)" | 
| 354 | apply (unfold Stable_def) | |
| 355 | apply (rule projecting_Constrains) | |
| 356 | done | |
| 357 | ||
| 46912 | 358 | lemma projecting_Always: | 
| 13819 | 359 | "projecting (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 360 | (Always (extend_set h A)) (Always A)" | 
| 361 | apply (unfold projecting_def) | |
| 362 | apply (blast intro: project_Always_I) | |
| 363 | done | |
| 364 | ||
| 46912 | 365 | lemma projecting_Increasing: | 
| 13819 | 366 | "projecting (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 367 | (Increasing (func o f)) (Increasing func)" | 
| 368 | apply (unfold projecting_def) | |
| 369 | apply (blast intro: project_Increasing_I) | |
| 370 | done | |
| 371 | ||
| 46912 | 372 | lemma extending_Constrains: | 
| 13819 | 373 | "extending (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 374 | (extend_set h A Co extend_set h B) (A Co B)" | 
| 375 | apply (unfold extending_def) | |
| 376 | apply (blast intro: project_Constrains_D) | |
| 377 | done | |
| 378 | ||
| 46912 | 379 | lemma extending_Stable: | 
| 13819 | 380 | "extending (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 381 | (Stable (extend_set h A)) (Stable A)" | 
| 382 | apply (unfold extending_def) | |
| 383 | apply (blast intro: project_Stable_D) | |
| 384 | done | |
| 385 | ||
| 46912 | 386 | lemma extending_Always: | 
| 13819 | 387 | "extending (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 388 | (Always (extend_set h A)) (Always A)" | 
| 389 | apply (unfold extending_def) | |
| 390 | apply (blast intro: project_Always_D) | |
| 391 | done | |
| 392 | ||
| 46912 | 393 | lemma extending_Increasing: | 
| 13819 | 394 | "extending (%G. reachable (extend h F\<squnion>G)) h F | 
| 13790 | 395 | (Increasing (func o f)) (Increasing func)" | 
| 396 | apply (unfold extending_def) | |
| 397 | apply (blast intro: project_Increasing_D) | |
| 398 | done | |
| 399 | ||
| 400 | ||
| 13798 | 401 | subsection{*leadsETo in the precondition (??)*}
 | 
| 13790 | 402 | |
| 13798 | 403 | subsubsection{*transient*}
 | 
| 13790 | 404 | |
| 46912 | 405 | lemma transient_extend_set_imp_project_transient: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 406 | "[| G \<in> transient (C \<inter> extend_set h A); G \<in> stable C |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 407 | ==> project h C G \<in> transient (project_set h C \<inter> A)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 408 | apply (auto simp add: transient_def Domain_project_act) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 409 | apply (subgoal_tac "act `` (C \<inter> extend_set h A) \<subseteq> - extend_set h A") | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 410 | prefer 2 | 
| 13790 | 411 | apply (simp add: stable_def constrains_def, blast) | 
| 412 | (*back to main goal*) | |
| 59807 | 413 | apply (erule_tac V = "AA \<subseteq> -C \<union> BB" for AA BB in thin_rl) | 
| 13790 | 414 | apply (drule bspec, assumption) | 
| 415 | apply (simp add: extend_set_def project_act_def, blast) | |
| 416 | done | |
| 417 | ||
| 418 | (*converse might hold too?*) | |
| 46912 | 419 | lemma project_extend_transient_D: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 420 | "project h C (extend h F) \<in> transient (project_set h C \<inter> D) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 421 | ==> F \<in> transient (project_set h C \<inter> D)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 422 | apply (simp add: transient_def Domain_project_act, safe) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 423 | apply blast+ | 
| 13790 | 424 | done | 
| 425 | ||
| 426 | ||
| 13798 | 427 | subsubsection{*ensures -- a primitive combining progress with safety*}
 | 
| 13790 | 428 | |
| 429 | (*Used to prove project_leadsETo_I*) | |
| 46912 | 430 | lemma ensures_extend_set_imp_project_ensures: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 431 | "[| extend h F \<in> stable C; G \<in> stable C; | 
| 13819 | 432 | extend h F\<squnion>G \<in> A ensures B; A-B = C \<inter> extend_set h D |] | 
| 433 | ==> F\<squnion>project h C G | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 434 | \<in> (project_set h C \<inter> project_set h A) ensures (project_set h B)" | 
| 46912 | 435 | apply (simp add: ensures_def project_constrains extend_transient, | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 436 | clarify) | 
| 13790 | 437 | apply (intro conjI) | 
| 438 | (*first subgoal*) | |
| 439 | apply (blast intro: extend_stable_project_set | |
| 440 | [THEN stableD, THEN constrains_Int, THEN constrains_weaken] | |
| 441 | dest!: extend_constrains_project_set equalityD1) | |
| 442 | (*2nd subgoal*) | |
| 443 | apply (erule stableD [THEN constrains_Int, THEN constrains_weaken]) | |
| 444 | apply assumption | |
| 445 | apply (simp (no_asm_use) add: extend_set_def) | |
| 446 | apply blast | |
| 447 | apply (simp add: extend_set_Int_distrib extend_set_Un_distrib) | |
| 448 | apply (blast intro!: extend_set_project_set [THEN subsetD], blast) | |
| 449 | (*The transient part*) | |
| 450 | apply auto | |
| 451 | prefer 2 | |
| 452 | apply (force dest!: equalityD1 | |
| 453 | intro: transient_extend_set_imp_project_transient | |
| 454 | [THEN transient_strengthen]) | |
| 455 | apply (simp (no_asm_use) add: Int_Diff) | |
| 456 | apply (force dest!: equalityD1 | |
| 457 | intro: transient_extend_set_imp_project_transient | |
| 458 | [THEN project_extend_transient_D, THEN transient_strengthen]) | |
| 459 | done | |
| 460 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 461 | text{*Transferring a transient property upwards*}
 | 
| 46912 | 462 | lemma project_transient_extend_set: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 463 | "project h C G \<in> transient (project_set h C \<inter> A - B) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 464 | ==> G \<in> transient (C \<inter> extend_set h A - extend_set h B)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 465 | apply (simp add: transient_def project_set_def extend_set_def project_act_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 466 | apply (elim disjE bexE) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 467 | apply (rule_tac x=Id in bexI) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 468 | apply (blast intro!: rev_bexI )+ | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 469 | done | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 470 | |
| 46912 | 471 | lemma project_unless2: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 472 | "[| G \<in> stable C; project h C G \<in> (project_set h C \<inter> A) unless B |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 473 | ==> G \<in> (C \<inter> extend_set h A) unless (extend_set h B)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 474 | by (auto dest: stable_constrains_Int intro: constrains_weaken | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 475 | simp add: unless_def project_constrains Diff_eq Int_assoc | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 476 | Int_extend_set_lemma) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 477 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 478 | |
| 46912 | 479 | lemma extend_unless: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 480 | "[|extend h F \<in> stable C; F \<in> A unless B|] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 481 | ==> extend h F \<in> C \<inter> extend_set h A unless extend_set h B" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 482 | apply (simp add: unless_def stable_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 483 | apply (drule constrains_Int) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 484 | apply (erule extend_constrains [THEN iffD2]) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 485 | apply (erule constrains_weaken, blast) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 486 | apply blast | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 487 | done | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 488 | |
| 13790 | 489 | (*Used to prove project_leadsETo_D*) | 
| 46912 | 490 | lemma Join_project_ensures: | 
| 13819 | 491 | "[| extend h F\<squnion>G \<in> stable C; | 
| 492 | F\<squnion>project h C G \<in> A ensures B |] | |
| 493 | ==> extend h F\<squnion>G \<in> (C \<inter> extend_set h A) ensures (extend_set h B)" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 494 | apply (auto simp add: ensures_eq extend_unless project_unless) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 495 | apply (blast intro: extend_transient [THEN iffD2] transient_strengthen) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 496 | apply (blast intro: project_transient_extend_set transient_strengthen) | 
| 13790 | 497 | done | 
| 498 | ||
| 13798 | 499 | text{*Lemma useful for both STRONG and WEAK progress, but the transient
 | 
| 500 | condition's very strong*} | |
| 13790 | 501 | |
| 502 | (*The strange induction formula allows induction over the leadsTo | |
| 503 | assumption's non-atomic precondition*) | |
| 46912 | 504 | lemma PLD_lemma: | 
| 13819 | 505 | "[| extend h F\<squnion>G \<in> stable C; | 
| 506 | F\<squnion>project h C G \<in> (project_set h C \<inter> A) leadsTo B |] | |
| 507 | ==> extend h F\<squnion>G \<in> | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 508 | C \<inter> extend_set h (project_set h C \<inter> A) leadsTo (extend_set h B)" | 
| 13790 | 509 | apply (erule leadsTo_induct) | 
| 46912 | 510 | apply (blast intro: Join_project_ensures) | 
| 13790 | 511 | apply (blast intro: psp_stable2 [THEN leadsTo_weaken_L] leadsTo_Trans) | 
| 512 | apply (simp del: UN_simps add: Int_UN_distrib leadsTo_UN extend_set_Union) | |
| 513 | done | |
| 514 | ||
| 46912 | 515 | lemma project_leadsTo_D_lemma: | 
| 13819 | 516 | "[| extend h F\<squnion>G \<in> stable C; | 
| 517 | F\<squnion>project h C G \<in> (project_set h C \<inter> A) leadsTo B |] | |
| 518 | ==> extend h F\<squnion>G \<in> (C \<inter> extend_set h A) leadsTo (extend_set h B)" | |
| 13790 | 519 | apply (rule PLD_lemma [THEN leadsTo_weaken]) | 
| 520 | apply (auto simp add: split_extended_all) | |
| 521 | done | |
| 522 | ||
| 46912 | 523 | lemma Join_project_LeadsTo: | 
| 13819 | 524 | "[| C = (reachable (extend h F\<squnion>G)); | 
| 525 | F\<squnion>project h C G \<in> A LeadsTo B |] | |
| 526 | ==> extend h F\<squnion>G \<in> (extend_set h A) LeadsTo (extend_set h B)" | |
| 13790 | 527 | by (simp del: Join_stable add: LeadsTo_def project_leadsTo_D_lemma | 
| 528 | project_set_reachable_extend_eq) | |
| 529 | ||
| 530 | ||
| 13798 | 531 | subsection{*Towards the theorem @{text project_Ensures_D}*}
 | 
| 13790 | 532 | |
| 46912 | 533 | lemma project_ensures_D_lemma: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 534 | "[| G \<in> stable ((C \<inter> extend_set h A) - (extend_set h B)); | 
| 13819 | 535 | F\<squnion>project h C G \<in> (project_set h C \<inter> A) ensures B; | 
| 536 | extend h F\<squnion>G \<in> stable C |] | |
| 537 | ==> extend h F\<squnion>G \<in> (C \<inter> extend_set h A) ensures (extend_set h B)" | |
| 13790 | 538 | (*unless*) | 
| 539 | apply (auto intro!: project_unless2 [unfolded unless_def] | |
| 540 | intro: project_extend_constrains_I | |
| 541 | simp add: ensures_def) | |
| 542 | (*transient*) | |
| 543 | (*A G-action cannot occur*) | |
| 544 | prefer 2 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 545 | apply (blast intro: project_transient_extend_set) | 
| 13790 | 546 | (*An F-action*) | 
| 547 | apply (force elim!: extend_transient [THEN iffD2, THEN transient_strengthen] | |
| 548 | simp add: split_extended_all) | |
| 549 | done | |
| 550 | ||
| 46912 | 551 | lemma project_ensures_D: | 
| 13819 | 552 | "[| F\<squnion>project h UNIV G \<in> A ensures B; | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 553 | G \<in> stable (extend_set h A - extend_set h B) |] | 
| 13819 | 554 | ==> extend h F\<squnion>G \<in> (extend_set h A) ensures (extend_set h B)" | 
| 46471 | 555 | apply (rule project_ensures_D_lemma [of _ UNIV, elim_format], auto) | 
| 13790 | 556 | done | 
| 557 | ||
| 46912 | 558 | lemma project_Ensures_D: | 
| 13819 | 559 | "[| F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A Ensures B; | 
| 560 | G \<in> stable (reachable (extend h F\<squnion>G) \<inter> extend_set h A - | |
| 13790 | 561 | extend_set h B) |] | 
| 13819 | 562 | ==> extend h F\<squnion>G \<in> (extend_set h A) Ensures (extend_set h B)" | 
| 13790 | 563 | apply (unfold Ensures_def) | 
| 46471 | 564 | apply (rule project_ensures_D_lemma [elim_format]) | 
| 13790 | 565 | apply (auto simp add: project_set_reachable_extend_eq [symmetric]) | 
| 566 | done | |
| 567 | ||
| 568 | ||
| 13798 | 569 | subsection{*Guarantees*}
 | 
| 13790 | 570 | |
| 46912 | 571 | lemma project_act_Restrict_subset_project_act: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 572 | "project_act h (Restrict C act) \<subseteq> project_act h act" | 
| 13790 | 573 | apply (auto simp add: project_act_def) | 
| 574 | done | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24147diff
changeset | 575 | |
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24147diff
changeset | 576 | |
| 46912 | 577 | lemma subset_closed_ok_extend_imp_ok_project: | 
| 13790 | 578 | "[| extend h F ok G; subset_closed (AllowedActs F) |] | 
| 579 | ==> F ok project h C G" | |
| 580 | apply (auto simp add: ok_def) | |
| 581 | apply (rename_tac act) | |
| 582 | apply (drule subsetD, blast) | |
| 583 | apply (rule_tac x = "Restrict C (extend_act h act)" in rev_image_eqI) | |
| 584 | apply simp + | |
| 585 | apply (cut_tac project_act_Restrict_subset_project_act) | |
| 586 | apply (auto simp add: subset_closed_def) | |
| 587 | done | |
| 588 | ||
| 589 | ||
| 590 | (*Weak precondition and postcondition | |
| 591 | Not clear that it has a converse [or that we want one!]*) | |
| 592 | ||
| 593 | (*The raw version; 3rd premise could be weakened by adding the | |
| 13819 | 594 | precondition extend h F\<squnion>G \<in> X' *) | 
| 46912 | 595 | lemma project_guarantees_raw: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 596 | assumes xguary: "F \<in> X guarantees Y" | 
| 13790 | 597 | and closed: "subset_closed (AllowedActs F)" | 
| 13819 | 598 | and project: "!!G. extend h F\<squnion>G \<in> X' | 
| 599 | ==> F\<squnion>project h (C G) G \<in> X" | |
| 600 | and extend: "!!G. [| F\<squnion>project h (C G) G \<in> Y |] | |
| 601 | ==> extend h F\<squnion>G \<in> Y'" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 602 | shows "extend h F \<in> X' guarantees Y'" | 
| 13790 | 603 | apply (rule xguary [THEN guaranteesD, THEN extend, THEN guaranteesI]) | 
| 604 | apply (blast intro: closed subset_closed_ok_extend_imp_ok_project) | |
| 605 | apply (erule project) | |
| 606 | done | |
| 607 | ||
| 46912 | 608 | lemma project_guarantees: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 609 | "[| F \<in> X guarantees Y; subset_closed (AllowedActs F); | 
| 13790 | 610 | projecting C h F X' X; extending C h F Y' Y |] | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 611 | ==> extend h F \<in> X' guarantees Y'" | 
| 13790 | 612 | apply (rule guaranteesI) | 
| 613 | apply (auto simp add: guaranteesD projecting_def extending_def | |
| 614 | subset_closed_ok_extend_imp_ok_project) | |
| 615 | done | |
| 616 | ||
| 617 | ||
| 618 | (*It seems that neither "guarantees" law can be proved from the other.*) | |
| 619 | ||
| 620 | ||
| 13798 | 621 | subsection{*guarantees corollaries*}
 | 
| 13790 | 622 | |
| 13798 | 623 | subsubsection{*Some could be deleted: the required versions are easy to prove*}
 | 
| 13790 | 624 | |
| 46912 | 625 | lemma extend_guar_increasing: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 626 | "[| F \<in> UNIV guarantees increasing func; | 
| 13790 | 627 | subset_closed (AllowedActs F) |] | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 628 | ==> extend h F \<in> X' guarantees increasing (func o f)" | 
| 13790 | 629 | apply (erule project_guarantees) | 
| 630 | apply (rule_tac [3] extending_increasing) | |
| 631 | apply (rule_tac [2] projecting_UNIV, auto) | |
| 632 | done | |
| 633 | ||
| 46912 | 634 | lemma extend_guar_Increasing: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 635 | "[| F \<in> UNIV guarantees Increasing func; | 
| 13790 | 636 | subset_closed (AllowedActs F) |] | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 637 | ==> extend h F \<in> X' guarantees Increasing (func o f)" | 
| 13790 | 638 | apply (erule project_guarantees) | 
| 639 | apply (rule_tac [3] extending_Increasing) | |
| 640 | apply (rule_tac [2] projecting_UNIV, auto) | |
| 641 | done | |
| 642 | ||
| 46912 | 643 | lemma extend_guar_Always: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 644 | "[| F \<in> Always A guarantees Always B; | 
| 13790 | 645 | subset_closed (AllowedActs F) |] | 
| 646 | ==> extend h F | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 647 | \<in> Always(extend_set h A) guarantees Always(extend_set h B)" | 
| 13790 | 648 | apply (erule project_guarantees) | 
| 649 | apply (rule_tac [3] extending_Always) | |
| 650 | apply (rule_tac [2] projecting_Always, auto) | |
| 651 | done | |
| 652 | ||
| 653 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 654 | subsubsection{*Guarantees with a leadsTo postcondition*}
 | 
| 13790 | 655 | |
| 46912 | 656 | lemma project_leadsTo_D: | 
| 13819 | 657 | "F\<squnion>project h UNIV G \<in> A leadsTo B | 
| 658 | ==> extend h F\<squnion>G \<in> (extend_set h A) leadsTo (extend_set h B)" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 659 | apply (rule_tac C1 = UNIV in project_leadsTo_D_lemma [THEN leadsTo_weaken], auto) | 
| 13790 | 660 | done | 
| 661 | ||
| 46912 | 662 | lemma project_LeadsTo_D: | 
| 13819 | 663 | "F\<squnion>project h (reachable (extend h F\<squnion>G)) G \<in> A LeadsTo B | 
| 664 | ==> extend h F\<squnion>G \<in> (extend_set h A) LeadsTo (extend_set h B)" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 665 | apply (rule refl [THEN Join_project_LeadsTo], auto) | 
| 13790 | 666 | done | 
| 667 | ||
| 46912 | 668 | lemma extending_leadsTo: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 669 | "extending (%G. UNIV) h F | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 670 | (extend_set h A leadsTo extend_set h B) (A leadsTo B)" | 
| 13790 | 671 | apply (unfold extending_def) | 
| 672 | apply (blast intro: project_leadsTo_D) | |
| 673 | done | |
| 674 | ||
| 46912 | 675 | lemma extending_LeadsTo: | 
| 13819 | 676 | "extending (%G. reachable (extend h F\<squnion>G)) h F | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13798diff
changeset | 677 | (extend_set h A LeadsTo extend_set h B) (A LeadsTo B)" | 
| 13790 | 678 | apply (unfold extending_def) | 
| 679 | apply (blast intro: project_LeadsTo_D) | |
| 680 | done | |
| 681 | ||
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 682 | end | 
| 46912 | 683 | |
| 684 | end |