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