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