| author | paulson |
| Wed, 29 Jan 2003 16:34:51 +0100 | |
| changeset 13792 | d1811693899c |
| parent 13601 | fd3e3d6b37b2 |
| permissions | -rw-r--r-- |
| 4776 | 1 |
(* Title: HOL/UNITY/SubstAx |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
6 |
LeadsTo relation, restricted to the set of reachable states. |
| 4776 | 7 |
*) |
8 |
||
| 6575 | 9 |
(*Resembles the previous definition of LeadsTo*) |
10 |
Goalw [LeadsTo_def] |
|
11 |
"A LeadsTo B = {F. F : (reachable F Int A) leadsTo (reachable F Int B)}";
|
|
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8041
diff
changeset
|
12 |
by (blast_tac (claset() addDs [psp_stable2] addIs [leadsTo_weaken]) 1); |
| 6575 | 13 |
qed "LeadsTo_eq_leadsTo"; |
14 |
||
15 |
||
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
16 |
(*** Specialized laws for handling invariants ***) |
|
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
17 |
|
| 6570 | 18 |
(** Conjoining an Always property **) |
| 5544 | 19 |
|
|
6811
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
20 |
Goal "F : Always INV ==> (F : (INV Int A) LeadsTo A') = (F : A LeadsTo A')"; |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
21 |
by (asm_full_simp_tac |
| 6570 | 22 |
(simpset() addsimps [LeadsTo_def, Always_eq_includes_reachable, |
23 |
Int_absorb2, Int_assoc RS sym]) 1); |
|
|
6811
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
24 |
qed "Always_LeadsTo_pre"; |
| 5544 | 25 |
|
|
6811
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
26 |
Goal "F : Always INV ==> (F : A LeadsTo (INV Int A')) = (F : A LeadsTo A')"; |
| 5544 | 27 |
by (asm_full_simp_tac |
| 6575 | 28 |
(simpset() addsimps [LeadsTo_eq_leadsTo, Always_eq_includes_reachable, |
| 6570 | 29 |
Int_absorb2, Int_assoc RS sym]) 1); |
|
6811
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
30 |
qed "Always_LeadsTo_post"; |
|
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
31 |
|
|
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
32 |
(* [| F : Always C; F : (C Int A) LeadsTo A' |] ==> F : A LeadsTo A' *) |
|
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
33 |
bind_thm ("Always_LeadsToI", Always_LeadsTo_pre RS iffD1);
|
|
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
34 |
|
|
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
35 |
(* [| F : Always INV; F : A LeadsTo A' |] ==> F : A LeadsTo (INV Int A') *) |
|
4700ca722bbd
Always_LeadsTo_pre. Always_LeadsTo_post: new equivalences suggested by Misra
paulson
parents:
6710
diff
changeset
|
36 |
bind_thm ("Always_LeadsToD", Always_LeadsTo_post RS iffD2);
|
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
37 |
|
|
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
38 |
|
| 4776 | 39 |
(*** Introduction rules: Basis, Trans, Union ***) |
40 |
||
| 6536 | 41 |
Goal "F : A leadsTo B ==> F : A LeadsTo B"; |
| 5111 | 42 |
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1); |
| 6575 | 43 |
by (blast_tac (claset() addIs [leadsTo_weaken_L]) 1); |
| 4776 | 44 |
qed "leadsTo_imp_LeadsTo"; |
45 |
||
| 6536 | 46 |
Goal "[| F : A LeadsTo B; F : B LeadsTo C |] ==> F : A LeadsTo C"; |
| 6575 | 47 |
by (full_simp_tac (simpset() addsimps [LeadsTo_eq_leadsTo]) 1); |
| 4776 | 48 |
by (blast_tac (claset() addIs [leadsTo_Trans]) 1); |
49 |
qed "LeadsTo_Trans"; |
|
50 |
||
| 5648 | 51 |
val prems = Goalw [LeadsTo_def] |
| 6536 | 52 |
"(!!A. A : S ==> F : A LeadsTo B) ==> F : (Union S) LeadsTo B"; |
| 5111 | 53 |
by (Simp_tac 1); |
| 4776 | 54 |
by (stac Int_Union 1); |
| 5648 | 55 |
by (blast_tac (claset() addIs [leadsTo_UN] addDs prems) 1); |
| 4776 | 56 |
qed "LeadsTo_Union"; |
57 |
||
58 |
||
59 |
(*** Derived rules ***) |
|
60 |
||
| 6536 | 61 |
Goal "F : A LeadsTo UNIV"; |
| 6575 | 62 |
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1); |
| 4776 | 63 |
qed "LeadsTo_UNIV"; |
64 |
Addsimps [LeadsTo_UNIV]; |
|
65 |
||
66 |
(*Useful with cancellation, disjunction*) |
|
| 6536 | 67 |
Goal "F : A LeadsTo (A' Un A') ==> F : A LeadsTo A'"; |
| 4776 | 68 |
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); |
69 |
qed "LeadsTo_Un_duplicate"; |
|
70 |
||
| 6536 | 71 |
Goal "F : A LeadsTo (A' Un C Un C) ==> F : A LeadsTo (A' Un C)"; |
| 4776 | 72 |
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); |
73 |
qed "LeadsTo_Un_duplicate2"; |
|
74 |
||
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
75 |
val prems = |
| 6536 | 76 |
Goal "(!!i. i : I ==> F : (A i) LeadsTo B) ==> F : (UN i:I. A i) LeadsTo B"; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
5804
diff
changeset
|
77 |
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1); |
| 4776 | 78 |
by (blast_tac (claset() addIs (LeadsTo_Union::prems)) 1); |
79 |
qed "LeadsTo_UN"; |
|
80 |
||
81 |
(*Binary union introduction rule*) |
|
| 6536 | 82 |
Goal "[| F : A LeadsTo C; F : B LeadsTo C |] ==> F : (A Un B) LeadsTo C"; |
| 4776 | 83 |
by (stac Un_eq_Union 1); |
84 |
by (blast_tac (claset() addIs [LeadsTo_Union]) 1); |
|
85 |
qed "LeadsTo_Un"; |
|
86 |
||
|
6710
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
87 |
(*Lets us look at the starting state*) |
|
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
88 |
val prems = |
|
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
89 |
Goal "(!!s. s : A ==> F : {s} LeadsTo B) ==> F : A LeadsTo B";
|
|
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
90 |
by (stac (UN_singleton RS sym) 1 THEN rtac LeadsTo_UN 1); |
|
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
91 |
by (blast_tac (claset() addIs prems) 1); |
|
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
92 |
qed "single_LeadsTo_I"; |
|
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
93 |
|
| 6536 | 94 |
Goal "A <= B ==> F : A LeadsTo B"; |
| 5111 | 95 |
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1); |
| 4776 | 96 |
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1); |
97 |
qed "subset_imp_LeadsTo"; |
|
98 |
||
99 |
bind_thm ("empty_LeadsTo", empty_subsetI RS subset_imp_LeadsTo);
|
|
100 |
Addsimps [empty_LeadsTo]; |
|
101 |
||
| 6536 | 102 |
Goal "[| F : A LeadsTo A'; A' <= B' |] ==> F : A LeadsTo B'"; |
| 5111 | 103 |
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1); |
| 4776 | 104 |
by (blast_tac (claset() addIs [leadsTo_weaken_R]) 1); |
105 |
qed_spec_mp "LeadsTo_weaken_R"; |
|
106 |
||
| 6536 | 107 |
Goal "[| F : A LeadsTo A'; B <= A |] \ |
108 |
\ ==> F : B LeadsTo A'"; |
|
| 5111 | 109 |
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1); |
| 4776 | 110 |
by (blast_tac (claset() addIs [leadsTo_weaken_L]) 1); |
111 |
qed_spec_mp "LeadsTo_weaken_L"; |
|
112 |
||
| 6536 | 113 |
Goal "[| F : A LeadsTo A'; \ |
| 5340 | 114 |
\ B <= A; A' <= B' |] \ |
| 6536 | 115 |
\ ==> F : B LeadsTo B'"; |
| 5340 | 116 |
by (blast_tac (claset() addIs [LeadsTo_weaken_R, LeadsTo_weaken_L, |
117 |
LeadsTo_Trans]) 1); |
|
118 |
qed "LeadsTo_weaken"; |
|
| 4776 | 119 |
|
| 6570 | 120 |
Goal "[| F : Always C; F : A LeadsTo A'; \ |
| 5544 | 121 |
\ C Int B <= A; C Int A' <= B' |] \ |
| 6536 | 122 |
\ ==> F : B LeadsTo B'"; |
| 6570 | 123 |
by (blast_tac (claset() addDs [Always_LeadsToI] addIs[LeadsTo_weaken] |
124 |
addIs [Always_LeadsToD]) 1); |
|
125 |
qed "Always_LeadsTo_weaken"; |
|
| 5340 | 126 |
|
127 |
(** Two theorems for "proof lattices" **) |
|
128 |
||
| 8216 | 129 |
Goal "F : A LeadsTo B ==> F : (A Un B) LeadsTo B"; |
| 5340 | 130 |
by (blast_tac (claset() addIs [LeadsTo_Un, subset_imp_LeadsTo]) 1); |
131 |
qed "LeadsTo_Un_post"; |
|
132 |
||
| 6536 | 133 |
Goal "[| F : A LeadsTo B; F : B LeadsTo C |] \ |
134 |
\ ==> F : (A Un B) LeadsTo C"; |
|
| 5340 | 135 |
by (blast_tac (claset() addIs [LeadsTo_Un, subset_imp_LeadsTo, |
136 |
LeadsTo_weaken_L, LeadsTo_Trans]) 1); |
|
137 |
qed "LeadsTo_Trans_Un"; |
|
138 |
||
139 |
||
140 |
(** Distributive laws **) |
|
141 |
||
| 6536 | 142 |
Goal "(F : (A Un B) LeadsTo C) = (F : A LeadsTo C & F : B LeadsTo C)"; |
| 4776 | 143 |
by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken_L]) 1); |
144 |
qed "LeadsTo_Un_distrib"; |
|
145 |
||
| 6536 | 146 |
Goal "(F : (UN i:I. A i) LeadsTo B) = (ALL i : I. F : (A i) LeadsTo B)"; |
| 4776 | 147 |
by (blast_tac (claset() addIs [LeadsTo_UN, LeadsTo_weaken_L]) 1); |
148 |
qed "LeadsTo_UN_distrib"; |
|
149 |
||
| 6536 | 150 |
Goal "(F : (Union S) LeadsTo B) = (ALL A : S. F : A LeadsTo B)"; |
| 4776 | 151 |
by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_L]) 1); |
152 |
qed "LeadsTo_Union_distrib"; |
|
153 |
||
154 |
||
| 6570 | 155 |
(** More rules using the premise "Always INV" **) |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
156 |
|
|
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
157 |
Goal "F : A Ensures B ==> F : A LeadsTo B"; |
| 6575 | 158 |
by (asm_full_simp_tac |
|
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
159 |
(simpset() addsimps [Ensures_def, LeadsTo_def, leadsTo_Basis]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
160 |
qed "LeadsTo_Basis"; |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
161 |
|
|
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
162 |
Goal "[| F : (A-B) Co (A Un B); F : transient (A-B) |] \ |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
163 |
\ ==> F : A Ensures B"; |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
164 |
by (asm_full_simp_tac |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
165 |
(simpset() addsimps [Ensures_def, Constrains_eq_constrains]) 1); |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
166 |
by (blast_tac (claset() addIs [ensuresI, constrains_weaken, |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
167 |
transient_strengthen]) 1); |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
168 |
qed "EnsuresI"; |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
169 |
|
| 6570 | 170 |
Goal "[| F : Always INV; \ |
| 6536 | 171 |
\ F : (INV Int (A-A')) Co (A Un A'); \ |
| 5648 | 172 |
\ F : transient (INV Int (A-A')) |] \ |
| 6536 | 173 |
\ ==> F : A LeadsTo A'"; |
| 6570 | 174 |
by (rtac Always_LeadsToI 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
175 |
by (assume_tac 1); |
|
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
176 |
by (blast_tac (claset() addIs [EnsuresI, LeadsTo_Basis, |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
177 |
Always_ConstrainsD RS Constrains_weaken, |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
178 |
transient_strengthen]) 1); |
| 6570 | 179 |
qed "Always_LeadsTo_Basis"; |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
180 |
|
| 5253 | 181 |
(*Set difference: maybe combine with leadsTo_weaken_L?? |
182 |
This is the most useful form of the "disjunction" rule*) |
|
| 6536 | 183 |
Goal "[| F : (A-B) LeadsTo C; F : (A Int B) LeadsTo C |] \ |
184 |
\ ==> F : A LeadsTo C"; |
|
| 5479 | 185 |
by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken]) 1); |
| 4776 | 186 |
qed "LeadsTo_Diff"; |
187 |
||
188 |
||
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
189 |
val prems = |
| 6536 | 190 |
Goal "(!! i. i:I ==> F : (A i) LeadsTo (A' i)) \ |
191 |
\ ==> F : (UN i:I. A i) LeadsTo (UN i:I. A' i)"; |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
5804
diff
changeset
|
192 |
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1); |
| 4776 | 193 |
by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_R] |
194 |
addIs prems) 1); |
|
195 |
qed "LeadsTo_UN_UN"; |
|
196 |
||
197 |
||
198 |
(*Version with no index set*) |
|
| 5257 | 199 |
val prems = |
| 6536 | 200 |
Goal "(!! i. F : (A i) LeadsTo (A' i)) \ |
201 |
\ ==> F : (UN i. A i) LeadsTo (UN i. A' i)"; |
|
| 4776 | 202 |
by (blast_tac (claset() addIs [LeadsTo_UN_UN] |
203 |
addIs prems) 1); |
|
204 |
qed "LeadsTo_UN_UN_noindex"; |
|
205 |
||
206 |
(*Version with no index set*) |
|
| 6536 | 207 |
Goal "ALL i. F : (A i) LeadsTo (A' i) \ |
208 |
\ ==> F : (UN i. A i) LeadsTo (UN i. A' i)"; |
|
| 4776 | 209 |
by (blast_tac (claset() addIs [LeadsTo_UN_UN]) 1); |
210 |
qed "all_LeadsTo_UN_UN"; |
|
211 |
||
212 |
||
213 |
(*Binary union version*) |
|
| 6536 | 214 |
Goal "[| F : A LeadsTo A'; F : B LeadsTo B' |] \ |
215 |
\ ==> F : (A Un B) LeadsTo (A' Un B')"; |
|
| 4776 | 216 |
by (blast_tac (claset() addIs [LeadsTo_Un, |
217 |
LeadsTo_weaken_R]) 1); |
|
218 |
qed "LeadsTo_Un_Un"; |
|
219 |
||
220 |
||
221 |
(** The cancellation law **) |
|
222 |
||
| 6536 | 223 |
Goal "[| F : A LeadsTo (A' Un B); F : B LeadsTo B' |] \ |
224 |
\ ==> F : A LeadsTo (A' Un B')"; |
|
| 4776 | 225 |
by (blast_tac (claset() addIs [LeadsTo_Un_Un, |
226 |
subset_imp_LeadsTo, LeadsTo_Trans]) 1); |
|
227 |
qed "LeadsTo_cancel2"; |
|
228 |
||
| 6536 | 229 |
Goal "[| F : A LeadsTo (A' Un B); F : (B-A') LeadsTo B' |] \ |
230 |
\ ==> F : A LeadsTo (A' Un B')"; |
|
| 4776 | 231 |
by (rtac LeadsTo_cancel2 1); |
232 |
by (assume_tac 2); |
|
233 |
by (ALLGOALS Asm_simp_tac); |
|
234 |
qed "LeadsTo_cancel_Diff2"; |
|
235 |
||
| 6536 | 236 |
Goal "[| F : A LeadsTo (B Un A'); F : B LeadsTo B' |] \ |
237 |
\ ==> F : A LeadsTo (B' Un A')"; |
|
| 4776 | 238 |
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1); |
239 |
by (blast_tac (claset() addSIs [LeadsTo_cancel2]) 1); |
|
240 |
qed "LeadsTo_cancel1"; |
|
241 |
||
| 6536 | 242 |
Goal "[| F : A LeadsTo (B Un A'); F : (B-A') LeadsTo B' |] \ |
243 |
\ ==> F : A LeadsTo (B' Un A')"; |
|
| 4776 | 244 |
by (rtac LeadsTo_cancel1 1); |
245 |
by (assume_tac 2); |
|
246 |
by (ALLGOALS Asm_simp_tac); |
|
247 |
qed "LeadsTo_cancel_Diff1"; |
|
248 |
||
249 |
||
250 |
(** The impossibility law **) |
|
251 |
||
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
252 |
(*The set "A" may be non-empty, but it contains no reachable states*) |
| 6570 | 253 |
Goal "F : A LeadsTo {} ==> F : Always (-A)";
|
254 |
by (full_simp_tac (simpset() addsimps [LeadsTo_def, |
|
255 |
Always_eq_includes_reachable]) 1); |
|
256 |
by (dtac leadsTo_empty 1); |
|
257 |
by Auto_tac; |
|
| 4776 | 258 |
qed "LeadsTo_empty"; |
259 |
||
260 |
||
261 |
(** PSP: Progress-Safety-Progress **) |
|
262 |
||
| 5639 | 263 |
(*Special case of PSP: Misra's "stable conjunction"*) |
| 6536 | 264 |
Goal "[| F : A LeadsTo A'; F : Stable B |] \ |
265 |
\ ==> F : (A Int B) LeadsTo (A' Int B)"; |
|
| 6575 | 266 |
by (full_simp_tac |
267 |
(simpset() addsimps [LeadsTo_eq_leadsTo, Stable_eq_stable]) 1); |
|
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
268 |
by (dtac psp_stable 1); |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
269 |
by (assume_tac 1); |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
270 |
by (asm_full_simp_tac (simpset() addsimps Int_ac) 1); |
|
6710
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
271 |
qed "PSP_Stable"; |
| 4776 | 272 |
|
| 6536 | 273 |
Goal "[| F : A LeadsTo A'; F : Stable B |] \ |
274 |
\ ==> F : (B Int A) LeadsTo (B Int A')"; |
|
|
6710
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
275 |
by (asm_simp_tac (simpset() addsimps PSP_Stable::Int_ac) 1); |
|
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
276 |
qed "PSP_Stable2"; |
| 4776 | 277 |
|
| 6575 | 278 |
Goal "[| F : A LeadsTo A'; F : B Co B' |] \ |
|
6710
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
279 |
\ ==> F : (A Int B') LeadsTo ((A' Int B) Un (B' - B))"; |
| 6575 | 280 |
by (full_simp_tac |
281 |
(simpset() addsimps [LeadsTo_def, Constrains_eq_constrains]) 1); |
|
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
282 |
by (blast_tac (claset() addDs [psp] addIs [leadsTo_weaken]) 1); |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
283 |
qed "PSP"; |
| 4776 | 284 |
|
| 6536 | 285 |
Goal "[| F : A LeadsTo A'; F : B Co B' |] \ |
|
6710
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
286 |
\ ==> F : (B' Int A) LeadsTo ((B Int A') Un (B' - B))"; |
| 5536 | 287 |
by (asm_simp_tac (simpset() addsimps PSP::Int_ac) 1); |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
288 |
qed "PSP2"; |
| 4776 | 289 |
|
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
290 |
Goalw [Unless_def] |
| 6536 | 291 |
"[| F : A LeadsTo A'; F : B Unless B' |] \ |
292 |
\ ==> F : (A Int B) LeadsTo ((A' Int B) Un B')"; |
|
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
293 |
by (dtac PSP 1); |
| 4776 | 294 |
by (assume_tac 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
295 |
by (blast_tac (claset() addIs [LeadsTo_Diff, LeadsTo_weaken, |
| 5584 | 296 |
subset_imp_LeadsTo]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
297 |
qed "PSP_Unless"; |
| 4776 | 298 |
|
299 |
||
|
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5648
diff
changeset
|
300 |
Goal "[| F : Stable A; F : transient C; \ |
| 6570 | 301 |
\ F : Always (-A Un B Un C) |] ==> F : A LeadsTo B"; |
302 |
by (etac Always_LeadsTo_weaken 1); |
|
|
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5648
diff
changeset
|
303 |
by (rtac LeadsTo_Diff 1); |
|
6710
4d438b714571
new rule single_LeadsTo_I; stronger PSP rule; PSP_stable2->PSP_Stable2
paulson
parents:
6575
diff
changeset
|
304 |
by (etac (transient_imp_leadsTo RS leadsTo_imp_LeadsTo RS PSP_Stable2) 2); |
|
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5648
diff
changeset
|
305 |
by (ALLGOALS (blast_tac (claset() addIs [subset_imp_LeadsTo]))); |
| 6570 | 306 |
qed "Stable_transient_Always_LeadsTo"; |
|
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5648
diff
changeset
|
307 |
|
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5648
diff
changeset
|
308 |
|
| 4776 | 309 |
(*** Induction rules ***) |
310 |
||
311 |
(** Meta or object quantifier ????? **) |
|
| 5232 | 312 |
Goal "[| wf r; \ |
| 10834 | 313 |
\ ALL m. F : (A Int f-`{m}) LeadsTo \
|
314 |
\ ((A Int f-`(r^-1 `` {m})) Un B) |] \
|
|
| 6536 | 315 |
\ ==> F : A LeadsTo B"; |
| 6575 | 316 |
by (full_simp_tac (simpset() addsimps [LeadsTo_eq_leadsTo]) 1); |
| 4776 | 317 |
by (etac leadsTo_wf_induct 1); |
318 |
by (blast_tac (claset() addIs [leadsTo_weaken]) 1); |
|
319 |
qed "LeadsTo_wf_induct"; |
|
320 |
||
321 |
||
| 5232 | 322 |
Goal "[| wf r; \ |
| 10834 | 323 |
\ ALL m:I. F : (A Int f-`{m}) LeadsTo \
|
324 |
\ ((A Int f-`(r^-1 `` {m})) Un B) |] \
|
|
325 |
\ ==> F : A LeadsTo ((A - (f-`I)) Un B)"; |
|
| 4776 | 326 |
by (etac LeadsTo_wf_induct 1); |
327 |
by Safe_tac; |
|
328 |
by (case_tac "m:I" 1); |
|
329 |
by (blast_tac (claset() addIs [LeadsTo_weaken]) 1); |
|
330 |
by (blast_tac (claset() addIs [subset_imp_LeadsTo]) 1); |
|
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
331 |
qed "Bounded_induct"; |
| 4776 | 332 |
|
333 |
||
| 8216 | 334 |
val prems = |
| 10834 | 335 |
Goal "(!!m::nat. F : (A Int f-`{m}) LeadsTo ((A Int f-`(lessThan m)) Un B)) \
|
| 6536 | 336 |
\ ==> F : A LeadsTo B"; |
| 4776 | 337 |
by (rtac (wf_less_than RS LeadsTo_wf_induct) 1); |
| 8216 | 338 |
by (auto_tac (claset() addIs prems, simpset())); |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
339 |
qed "LessThan_induct"; |
| 4776 | 340 |
|
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
341 |
(*Integer version. Could generalize from 0 to any lower bound*) |
| 5584 | 342 |
val [reach, prem] = |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
343 |
Goal "[| F : Always {s. (0::int) <= f s}; \
|
| 6536 | 344 |
\ !! z. F : (A Int {s. f s = z}) LeadsTo \
|
| 5584 | 345 |
\ ((A Int {s. f s < z}) Un B) |] \
|
| 6536 | 346 |
\ ==> F : A LeadsTo B"; |
| 8216 | 347 |
by (res_inst_tac [("f", "nat o f")] LessThan_induct 1);
|
| 5544 | 348 |
by (simp_tac (simpset() addsimps [vimage_def]) 1); |
| 6570 | 349 |
by (rtac ([reach, prem] MRS Always_LeadsTo_weaken) 1); |
| 5584 | 350 |
by (auto_tac (claset(), simpset() addsimps [nat_eq_iff, nat_less_iff])); |
| 5544 | 351 |
qed "integ_0_le_induct"; |
352 |
||
| 10834 | 353 |
Goal "!!l::nat. [| ALL m:(greaterThan l). F : (A Int f-`{m}) LeadsTo \
|
354 |
\ ((A Int f-`(lessThan m)) Un B) |] \ |
|
355 |
\ ==> F : A LeadsTo ((A Int (f-`(atMost l))) Un B)"; |
|
| 4776 | 356 |
by (simp_tac (HOL_ss addsimps [Diff_eq RS sym, vimage_Compl, Compl_greaterThan RS sym]) 1); |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
357 |
by (rtac (wf_less_than RS Bounded_induct) 1); |
| 4776 | 358 |
by (Asm_simp_tac 1); |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
359 |
qed "LessThan_bounded_induct"; |
| 4776 | 360 |
|
| 10834 | 361 |
Goal "!!l::nat. [| ALL m:(lessThan l). F : (A Int f-`{m}) LeadsTo \
|
362 |
\ ((A Int f-`(greaterThan m)) Un B) |] \ |
|
363 |
\ ==> F : A LeadsTo ((A Int (f-`(atLeast l))) Un B)"; |
|
| 4776 | 364 |
by (res_inst_tac [("f","f"),("f1", "%k. l - k")]
|
365 |
(wf_less_than RS wf_inv_image RS LeadsTo_wf_induct) 1); |
|
366 |
by (simp_tac (simpset() addsimps [inv_image_def, Image_singleton]) 1); |
|
367 |
by (Clarify_tac 1); |
|
368 |
by (case_tac "m<l" 1); |
|
369 |
by (blast_tac (claset() addIs [not_leE, subset_imp_LeadsTo]) 2); |
|
370 |
by (blast_tac (claset() addIs [LeadsTo_weaken_R, diff_less_mono2]) 1); |
|
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
371 |
qed "GreaterThan_bounded_induct"; |
| 4776 | 372 |
|
373 |
||
374 |
(*** Completion: Binary and General Finite versions ***) |
|
375 |
||
| 6536 | 376 |
Goal "[| F : A LeadsTo (A' Un C); F : A' Co (A' Un C); \ |
377 |
\ F : B LeadsTo (B' Un C); F : B' Co (B' Un C) |] \ |
|
378 |
\ ==> F : (A Int B) LeadsTo ((A' Int B') Un C)"; |
|
| 6575 | 379 |
by (full_simp_tac |
380 |
(simpset() addsimps [LeadsTo_eq_leadsTo, Constrains_eq_constrains, |
|
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
381 |
Int_Un_distrib]) 1); |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
382 |
by (blast_tac (claset() addIs [completion, leadsTo_weaken]) 1); |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
383 |
qed "Completion"; |
| 4776 | 384 |
|
| 6564 | 385 |
Goal "finite I \ |
| 6536 | 386 |
\ ==> (ALL i:I. F : (A i) LeadsTo (A' i Un C)) --> \ |
387 |
\ (ALL i:I. F : (A' i) Co (A' i Un C)) --> \ |
|
388 |
\ F : (INT i:I. A i) LeadsTo ((INT i:I. A' i) Un C)"; |
|
| 4776 | 389 |
by (etac finite_induct 1); |
|
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
390 |
by Auto_tac; |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
391 |
by (rtac Completion 1); |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
392 |
by (simp_tac (HOL_ss addsimps (INT_simps RL [sym])) 4); |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
393 |
by (rtac Constrains_INT 4); |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
394 |
by Auto_tac; |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
395 |
val lemma = result(); |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
396 |
|
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
397 |
val prems = Goal |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
398 |
"[| finite I; \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
399 |
\ !!i. i:I ==> F : (A i) LeadsTo (A' i Un C); \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
400 |
\ !!i. i:I ==> F : (A' i) Co (A' i Un C) |] \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
401 |
\ ==> F : (INT i:I. A i) LeadsTo ((INT i:I. A' i) Un C)"; |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
402 |
by (blast_tac (claset() addIs (lemma RS mp RS mp)::prems) 1); |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
403 |
qed "Finite_completion"; |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
404 |
|
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
405 |
Goalw [Stable_def] |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
406 |
"[| F : A LeadsTo A'; F : Stable A'; \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
407 |
\ F : B LeadsTo B'; F : Stable B' |] \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
408 |
\ ==> F : (A Int B) LeadsTo (A' Int B')"; |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
409 |
by (res_inst_tac [("C1", "{}")] (Completion RS LeadsTo_weaken_R) 1);
|
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
410 |
by (REPEAT (Force_tac 1)); |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
411 |
qed "Stable_completion"; |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
412 |
|
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
413 |
val prems = Goalw [Stable_def] |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
414 |
"[| finite I; \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
415 |
\ !!i. i:I ==> F : (A i) LeadsTo (A' i); \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
416 |
\ !!i. i:I ==> F : Stable (A' i) |] \ |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
417 |
\ ==> F : (INT i:I. A i) LeadsTo (INT i:I. A' i)"; |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
418 |
by (res_inst_tac [("C1", "{}")] (Finite_completion RS LeadsTo_weaken_R) 1);
|
| 4776 | 419 |
by (ALLGOALS Asm_simp_tac); |
|
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
420 |
by (ALLGOALS (blast_tac (claset() addIs prems))); |
|
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
8216
diff
changeset
|
421 |
qed "Finite_stable_completion"; |
| 5232 | 422 |
|
423 |
||
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
424 |
(*proves "ensures/leadsTo" properties when the program is specified*) |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5422
diff
changeset
|
425 |
fun ensures_tac sact = |
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
426 |
SELECT_GOAL |
| 6570 | 427 |
(EVERY [REPEAT (Always_Int_tac 1), |
428 |
etac Always_LeadsTo_Basis 1 |
|
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
429 |
ORELSE (*subgoal may involve LeadsTo, leadsTo or ensures*) |
|
7522
d93b52bda2dd
ensures_tac now handles leadsTo as well as LeadsTo
paulson
parents:
6909
diff
changeset
|
430 |
REPEAT (ares_tac [LeadsTo_Basis, leadsTo_Basis, |
|
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
431 |
EnsuresI, ensuresI] 1), |
| 6536 | 432 |
(*now there are two subgoals: co & transient*) |
| 5648 | 433 |
simp_tac (simpset() addsimps !program_defs_ref) 2, |
| 8041 | 434 |
res_inst_tac [("act", sact)] transientI 2,
|
| 5340 | 435 |
(*simplify the command's domain*) |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5422
diff
changeset
|
436 |
simp_tac (simpset() addsimps [Domain_def]) 3, |
|
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5422
diff
changeset
|
437 |
constrains_tac 1, |
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
438 |
ALLGOALS Clarify_tac, |
| 13601 | 439 |
ALLGOALS Asm_lr_simp_tac]); |