| author | wenzelm | 
| Tue, 12 Feb 2002 20:35:35 +0100 | |
| changeset 12882 | e20f14f7ff35 | 
| parent 12537 | f2cda6fb1c9f | 
| child 13176 | 312bd350579b | 
| permissions | -rw-r--r-- | 
| 11479 | 1 | (* Title: ZF/UNITY/UNITY.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Sidi O Ehmety, Computer Laboratory | |
| 4 | Copyright 2001 University of Cambridge | |
| 5 | ||
| 6 | The basic UNITY theory (revised version, based upon the "co" operator) | |
| 7 | From Misra, "A Logic for Concurrent Programming", 1994 | |
| 8 | ||
| 9 | Proofs ported from HOL | |
| 10 | *) | |
| 11 | ||
| 12 | (** SKIP **) | |
| 12195 | 13 | Goalw [SKIP_def] "SKIP:program"; | 
| 14 | by (rewrite_goal_tac [program_def, mk_program_def] 1); | |
| 15 | by Auto_tac; | |
| 16 | qed "SKIP_in_program"; | |
| 17 | AddIffs [SKIP_in_program]; | |
| 18 | AddTCs [SKIP_in_program]; | |
| 11479 | 19 | |
| 20 | (** programify: coersion from anything to program **) | |
| 21 | ||
| 22 | Goalw [programify_def] | |
| 23 | "F:program ==> programify(F)=F"; | |
| 24 | by Auto_tac; | |
| 25 | qed "programify_program"; | |
| 26 | Addsimps [programify_program]; | |
| 27 | ||
| 28 | Goalw [programify_def] | |
| 29 | "programify(F):program"; | |
| 30 | by Auto_tac; | |
| 31 | qed "programify_in_program"; | |
| 32 | AddIffs [programify_in_program]; | |
| 33 | AddTCs [programify_in_program]; | |
| 34 | ||
| 35 | (** Collapsing rules: to remove programify from expressions **) | |
| 36 | Goalw [programify_def] | |
| 37 | "programify(programify(F))=programify(F)"; | |
| 38 | by Auto_tac; | |
| 39 | qed "programify_idem"; | |
| 12195 | 40 | AddIffs [programify_idem]; | 
| 11479 | 41 | |
| 42 | Goal | |
| 43 | "Init(programify(F)) = Init(F)"; | |
| 44 | by (simp_tac (simpset() addsimps [Init_def]) 1); | |
| 45 | qed "Init_programify"; | |
| 12195 | 46 | AddIffs [Init_programify]; | 
| 11479 | 47 | |
| 48 | Goal | |
| 49 | "Acts(programify(F)) = Acts(F)"; | |
| 50 | by (simp_tac (simpset() addsimps [Acts_def]) 1); | |
| 51 | qed "Acts_programify"; | |
| 12195 | 52 | AddIffs [Acts_programify]; | 
| 11479 | 53 | |
| 54 | Goal | |
| 55 | "AllowedActs(programify(F)) = AllowedActs(F)"; | |
| 56 | by (simp_tac (simpset() addsimps [AllowedActs_def]) 1); | |
| 57 | qed "AllowedActs_programify"; | |
| 12195 | 58 | AddIffs [AllowedActs_programify]; | 
| 11479 | 59 | |
| 12195 | 60 | (** program's inspectors **) | 
| 11479 | 61 | |
| 12195 | 62 | Goal "F:program ==>id(state):RawActs(F)"; | 
| 11479 | 63 | by (auto_tac (claset(), simpset() | 
| 64 | addsimps [program_def, RawActs_def])); | |
| 12195 | 65 | qed "id_in_RawActs"; | 
| 11479 | 66 | |
| 12195 | 67 | Goal "id(state):Acts(F)"; | 
| 11479 | 68 | by (simp_tac (simpset() | 
| 12195 | 69 | addsimps [id_in_RawActs, Acts_def]) 1); | 
| 70 | qed "id_in_Acts"; | |
| 11479 | 71 | |
| 12195 | 72 | Goal "F:program ==>id(state):RawAllowedActs(F)"; | 
| 11479 | 73 | by (auto_tac (claset(), simpset() | 
| 74 | addsimps [program_def, RawAllowedActs_def])); | |
| 12195 | 75 | qed "id_in_RawAllowedActs"; | 
| 11479 | 76 | |
| 12195 | 77 | Goal "id(state):AllowedActs(F)"; | 
| 11479 | 78 | by (simp_tac (simpset() | 
| 12195 | 79 | addsimps [id_in_RawAllowedActs, AllowedActs_def]) 1); | 
| 80 | qed "id_in_AllowedActs"; | |
| 11479 | 81 | |
| 12195 | 82 | AddIffs [id_in_Acts, id_in_AllowedActs]; | 
| 83 | AddTCs [id_in_Acts, id_in_AllowedActs]; | |
| 11479 | 84 | |
| 12195 | 85 | Goal "cons(id(state), Acts(F)) = Acts(F)"; | 
| 11479 | 86 | by (simp_tac (simpset() addsimps [cons_absorb]) 1); | 
| 12195 | 87 | qed "cons_id_Acts"; | 
| 11479 | 88 | |
| 12195 | 89 | Goal "cons(id(state), AllowedActs(F)) = AllowedActs(F)"; | 
| 11479 | 90 | by (simp_tac (simpset() addsimps [cons_absorb]) 1); | 
| 12195 | 91 | qed "cons_id_AllowedActs"; | 
| 11479 | 92 | |
| 12195 | 93 | AddIffs [cons_id_Acts, cons_id_AllowedActs]; | 
| 11479 | 94 | |
| 95 | (** inspectors's types **) | |
| 96 | Goal | |
| 12195 | 97 | "F:program ==> RawInit(F)<=state"; | 
| 11479 | 98 | by (auto_tac (claset(), simpset() | 
| 99 | addsimps [program_def, RawInit_def])); | |
| 100 | qed "RawInit_type"; | |
| 101 | ||
| 12195 | 102 | Goal | 
| 103 | "F:program ==> RawActs(F)<=Pow(state*state)"; | |
| 104 | by (auto_tac (claset(), simpset() | |
| 105 | addsimps [program_def, RawActs_def])); | |
| 106 | qed "RawActs_type"; | |
| 107 | ||
| 108 | Goal | |
| 109 | "F:program ==> RawAllowedActs(F)<=Pow(state*state)"; | |
| 110 | by (auto_tac (claset(), simpset() | |
| 111 | addsimps [program_def, RawAllowedActs_def])); | |
| 112 | qed "RawAllowedActs_type"; | |
| 113 | ||
| 114 | Goal "Init(F)<=state"; | |
| 11479 | 115 | by (simp_tac (simpset() | 
| 116 | addsimps [RawInit_type, Init_def]) 1); | |
| 117 | qed "Init_type"; | |
| 118 | ||
| 12195 | 119 | Goalw [st_set_def] "st_set(Init(F))"; | 
| 12484 | 120 | by (rtac Init_type 1); | 
| 12195 | 121 | qed "st_set_Init"; | 
| 122 | AddIffs [st_set_Init]; | |
| 11479 | 123 | |
| 124 | Goal | |
| 12195 | 125 | "Acts(F)<=Pow(state*state)"; | 
| 11479 | 126 | by (simp_tac (simpset() | 
| 127 | addsimps [RawActs_type, Acts_def]) 1); | |
| 128 | qed "Acts_type"; | |
| 129 | ||
| 130 | Goal | |
| 12195 | 131 | "AllowedActs(F)<=Pow(state*state)"; | 
| 11479 | 132 | by (simp_tac (simpset() | 
| 133 | addsimps [RawAllowedActs_type, AllowedActs_def]) 1); | |
| 134 | qed "AllowedActs_type"; | |
| 135 | ||
| 136 | (** More simplification rules involving state | |
| 137 | and Init, Acts, and AllowedActs **) | |
| 138 | ||
| 139 | Goal "state <= Init(F) <-> Init(F)=state"; | |
| 140 | by (cut_inst_tac [("F", "F")] Init_type 1);
 | |
| 12195 | 141 | by Auto_tac; | 
| 142 | qed "state_subset_is_Init_iff"; | |
| 143 | AddIffs [state_subset_is_Init_iff]; | |
| 11479 | 144 | |
| 12195 | 145 | Goal "Pow(state*state) <= Acts(F) <-> Acts(F)=Pow(state*state)"; | 
| 11479 | 146 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | 
| 12195 | 147 | by Auto_tac; | 
| 148 | qed "Pow_state_times_state_is_subset_Acts_iff"; | |
| 149 | AddIffs [Pow_state_times_state_is_subset_Acts_iff]; | |
| 11479 | 150 | |
| 12195 | 151 | Goal "Pow(state*state) <= AllowedActs(F) <-> AllowedActs(F)=Pow(state*state)"; | 
| 11479 | 152 | by (cut_inst_tac [("F", "F")] AllowedActs_type 1);
 | 
| 12195 | 153 | by Auto_tac; | 
| 154 | qed "Pow_state_times_state_is_subset_AllowedActs_iff"; | |
| 155 | AddIffs [Pow_state_times_state_is_subset_AllowedActs_iff]; | |
| 11479 | 156 | |
| 157 | (** Eliminating `Int state' from expressions **) | |
| 12195 | 158 | Goal "Init(F) Int state = Init(F)"; | 
| 11479 | 159 | by (cut_inst_tac [("F", "F")] Init_type 1);
 | 
| 12195 | 160 | by (Blast_tac 1); | 
| 11479 | 161 | qed "Init_Int_state"; | 
| 12195 | 162 | AddIffs [Init_Int_state]; | 
| 11479 | 163 | |
| 12195 | 164 | Goal "state Int Init(F) = Init(F)"; | 
| 165 | by (cut_inst_tac [("F", "F")] Init_type 1);
 | |
| 166 | by (Blast_tac 1); | |
| 167 | qed "state_Int_Init"; | |
| 168 | AddIffs [state_Int_Init]; | |
| 11479 | 169 | |
| 12195 | 170 | Goal "Acts(F) Int Pow(state*state) = Acts(F)"; | 
| 11479 | 171 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | 
| 12195 | 172 | by (Blast_tac 1); | 
| 173 | qed "Acts_Int_Pow_state_times_state"; | |
| 174 | AddIffs [Acts_Int_Pow_state_times_state]; | |
| 11479 | 175 | |
| 12195 | 176 | Goal "Pow(state*state) Int Acts(F) = Acts(F)"; | 
| 177 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | |
| 178 | by (Blast_tac 1); | |
| 179 | qed "state_times_state_Int_Acts"; | |
| 180 | AddIffs [state_times_state_Int_Acts]; | |
| 11479 | 181 | |
| 12195 | 182 | Goal "AllowedActs(F) Int Pow(state*state) = AllowedActs(F)"; | 
| 11479 | 183 | by (cut_inst_tac [("F", "F")] AllowedActs_type 1);
 | 
| 12195 | 184 | by (Blast_tac 1); | 
| 185 | qed "AllowedActs_Int_Pow_state_times_state"; | |
| 186 | AddIffs [AllowedActs_Int_Pow_state_times_state]; | |
| 11479 | 187 | |
| 12195 | 188 | Goal "Pow(state*state) Int AllowedActs(F) = AllowedActs(F)"; | 
| 189 | by (cut_inst_tac [("F", "F")] AllowedActs_type 1);
 | |
| 190 | by (Blast_tac 1); | |
| 191 | qed "state_times_state_Int_AllowedActs"; | |
| 192 | AddIffs [state_times_state_Int_AllowedActs]; | |
| 11479 | 193 | |
| 194 | (** mk_program **) | |
| 195 | ||
| 12195 | 196 | Goalw [mk_program_def, program_def] "mk_program(init, acts, allowed):program"; | 
| 197 | by Auto_tac; | |
| 198 | qed "mk_program_in_program"; | |
| 199 | AddIffs [mk_program_in_program]; | |
| 200 | AddTCs [mk_program_in_program]; | |
| 11479 | 201 | |
| 12195 | 202 | Goalw [RawInit_def, mk_program_def] | 
| 203 | "RawInit(mk_program(init, acts, allowed)) = init Int state"; | |
| 204 | by Auto_tac; | |
| 11479 | 205 | qed "RawInit_eq"; | 
| 12195 | 206 | AddIffs [RawInit_eq]; | 
| 11479 | 207 | |
| 12195 | 208 | Goalw [RawActs_def, mk_program_def] | 
| 209 | "RawActs(mk_program(init, acts, allowed)) = cons(id(state), acts Int Pow(state*state))"; | |
| 210 | by Auto_tac; | |
| 211 | qed "RawActs_eq"; | |
| 212 | AddIffs [RawActs_eq]; | |
| 11479 | 213 | |
| 12195 | 214 | Goalw [RawAllowedActs_def, mk_program_def] | 
| 215 | "RawAllowedActs(mk_program(init, acts, allowed)) = \ | |
| 216 | \ cons(id(state), allowed Int Pow(state*state))"; | |
| 217 | by Auto_tac; | |
| 218 | qed "RawAllowedActs_eq"; | |
| 219 | AddIffs [RawAllowedActs_eq]; | |
| 11479 | 220 | |
| 12195 | 221 | Goalw [Init_def] "Init(mk_program(init, acts, allowed)) = init Int state"; | 
| 222 | by (Simp_tac 1); | |
| 223 | qed "Init_eq"; | |
| 224 | AddIffs [Init_eq]; | |
| 225 | ||
| 226 | Goalw [Acts_def] | |
| 227 | "Acts(mk_program(init, acts, allowed)) = cons(id(state), acts Int Pow(state*state))"; | |
| 228 | by (Simp_tac 1); | |
| 11479 | 229 | qed "Acts_eq"; | 
| 12195 | 230 | AddIffs [Acts_eq]; | 
| 11479 | 231 | |
| 232 | Goalw [AllowedActs_def] | |
| 12195 | 233 | "AllowedActs(mk_program(init, acts, allowed))= \ | 
| 234 | \ cons(id(state), allowed Int Pow(state*state))"; | |
| 235 | by (Simp_tac 1); | |
| 11479 | 236 | qed "AllowedActs_eq"; | 
| 12195 | 237 | AddIffs [AllowedActs_eq]; | 
| 11479 | 238 | |
| 239 | (**Init, Acts, and AlowedActs of SKIP **) | |
| 240 | ||
| 12195 | 241 | Goalw [SKIP_def] "RawInit(SKIP) = state"; | 
| 242 | by Auto_tac; | |
| 11479 | 243 | qed "RawInit_SKIP"; | 
| 12195 | 244 | AddIffs [RawInit_SKIP]; | 
| 11479 | 245 | |
| 12195 | 246 | Goalw [SKIP_def] "RawAllowedActs(SKIP) = Pow(state*state)"; | 
| 247 | by Auto_tac; | |
| 248 | qed "RawAllowedActs_SKIP"; | |
| 249 | AddIffs [RawAllowedActs_SKIP]; | |
| 11479 | 250 | |
| 12195 | 251 | Goalw [SKIP_def] "RawActs(SKIP) = {id(state)}";
 | 
| 252 | by Auto_tac; | |
| 11479 | 253 | qed "RawActs_SKIP"; | 
| 12195 | 254 | AddIffs [RawActs_SKIP]; | 
| 11479 | 255 | |
| 12195 | 256 | Goalw [Init_def] "Init(SKIP) = state"; | 
| 257 | by Auto_tac; | |
| 258 | qed "Init_SKIP"; | |
| 259 | AddIffs [Init_SKIP]; | |
| 11479 | 260 | |
| 12195 | 261 | Goalw [Acts_def] "Acts(SKIP) = {id(state)}";
 | 
| 262 | by Auto_tac; | |
| 263 | qed "Acts_SKIP"; | |
| 264 | AddIffs [Acts_SKIP]; | |
| 11479 | 265 | |
| 12195 | 266 | Goalw [AllowedActs_def] "AllowedActs(SKIP) = Pow(state*state)"; | 
| 267 | by Auto_tac; | |
| 11479 | 268 | qed "AllowedActs_SKIP"; | 
| 12195 | 269 | AddIffs [AllowedActs_SKIP]; | 
| 11479 | 270 | |
| 12195 | 271 | (** Equality of UNITY programs **) | 
| 11479 | 272 | |
| 273 | Goal | |
| 274 | "F:program ==> mk_program(RawInit(F), RawActs(F), RawAllowedActs(F))=F"; | |
| 12195 | 275 | by (rewrite_goal_tac [program_def, mk_program_def,RawInit_def, | 
| 276 | RawActs_def, RawAllowedActs_def] 1); | |
| 277 | by Auto_tac; | |
| 11479 | 278 | by (REPEAT(Blast_tac 1)); | 
| 279 | qed "raw_surjective_mk_program"; | |
| 12195 | 280 | Addsimps [raw_surjective_mk_program]; | 
| 11479 | 281 | |
| 12195 | 282 | Goalw [Init_def, Acts_def, AllowedActs_def] | 
| 11479 | 283 | "mk_program(Init(F), Acts(F), AllowedActs(F)) = programify(F)"; | 
| 12195 | 284 | by Auto_tac; | 
| 11479 | 285 | qed "surjective_mk_program"; | 
| 12195 | 286 | AddIffs [surjective_mk_program]; | 
| 11479 | 287 | |
| 288 | Goal "[|Init(F) = Init(G); Acts(F) = Acts(G); \ | |
| 289 | \ AllowedActs(F) = AllowedActs(G); F:program; G:program |] ==> F = G"; | |
| 290 | by (stac (programify_program RS sym) 1); | |
| 291 | by (rtac sym 2); | |
| 292 | by (stac (programify_program RS sym) 2); | |
| 293 | by (stac (surjective_mk_program RS sym) 3); | |
| 294 | by (stac (surjective_mk_program RS sym) 3); | |
| 295 | by (ALLGOALS(Asm_simp_tac)); | |
| 296 | qed "program_equalityI"; | |
| 297 | ||
| 298 | val [major,minor] = | |
| 299 | Goal "[| F = G; \ | |
| 300 | \ [| Init(F) = Init(G); Acts(F) = Acts(G); AllowedActs(F) = AllowedActs(G) |]\ | |
| 301 | \ ==> P |] ==> P"; | |
| 302 | by (rtac minor 1); | |
| 303 | by (auto_tac (claset(), simpset() addsimps [major])); | |
| 304 | qed "program_equalityE"; | |
| 305 | ||
| 306 | ||
| 307 | Goal "[| F:program; G:program |] ==>(F=G) <-> \ | |
| 308 | \ (Init(F) = Init(G) & Acts(F) = Acts(G) & AllowedActs(F) = AllowedActs(G))"; | |
| 309 | by (blast_tac (claset() addIs [program_equalityI, program_equalityE]) 1); | |
| 310 | qed "program_equality_iff"; | |
| 311 | ||
| 312 | (*** These rules allow "lazy" definition expansion | |
| 313 | ||
| 314 | ...skipping 1 line | |
| 315 | ||
| 316 | ***) | |
| 317 | ||
| 318 | Goal "F == mk_program (init,acts,allowed) ==> Init(F) = init Int state"; | |
| 319 | by Auto_tac; | |
| 320 | qed "def_prg_Init"; | |
| 321 | ||
| 322 | ||
| 12195 | 323 | Goal "F == mk_program (init,acts,allowed) ==> \ | 
| 324 | \ Acts(F) = cons(id(state), acts Int Pow(state*state))"; | |
| 11479 | 325 | by Auto_tac; | 
| 326 | qed "def_prg_Acts"; | |
| 327 | ||
| 328 | ||
| 329 | Goal "F == mk_program (init,acts,allowed) ==> \ | |
| 12195 | 330 | \ AllowedActs(F) = cons(id(state), allowed Int Pow(state*state))"; | 
| 11479 | 331 | by Auto_tac; | 
| 332 | qed "def_prg_AllowedActs"; | |
| 333 | ||
| 334 | ||
| 335 | val [rew] = goal thy | |
| 336 | "[| F == mk_program (init,acts,allowed) |] \ | |
| 12195 | 337 | \ ==> Init(F) = init Int state & Acts(F) = cons(id(state), acts Int Pow(state*state)) & \ | 
| 338 | \ AllowedActs(F) = cons(id(state), allowed Int Pow(state*state)) "; | |
| 11479 | 339 | by (rewtac rew); | 
| 340 | by Auto_tac; | |
| 341 | qed "def_prg_simps"; | |
| 342 | ||
| 343 | ||
| 344 | (*An action is expanded only if a pair of states is being tested against it*) | |
| 345 | Goal "[| act == {<s,s'>:A*B. P(s, s')} |] ==> \
 | |
| 346 | \ (<s,s'>:act) <-> (<s,s'>:A*B & P(s, s'))"; | |
| 347 | by Auto_tac; | |
| 348 | qed "def_act_simp"; | |
| 349 | ||
| 350 | fun simp_of_act def = def RS def_act_simp; | |
| 351 | ||
| 352 | (*A set is expanded only if an element is being tested against it*) | |
| 353 | Goal "A == B ==> (x : A) <-> (x : B)"; | |
| 354 | by Auto_tac; | |
| 355 | qed "def_set_simp"; | |
| 356 | ||
| 357 | fun simp_of_set def = def RS def_set_simp; | |
| 358 | ||
| 359 | (*** co ***) | |
| 360 | ||
| 12195 | 361 | Goalw [constrains_def] | 
| 362 | "A co B <= program"; | |
| 363 | by Auto_tac; | |
| 364 | qed "constrains_type"; | |
| 365 | ||
| 366 | ||
| 11479 | 367 | val prems = Goalw [constrains_def] | 
| 368 | "[|(!!act s s'. [| act: Acts(F); <s,s'>:act; s:A|] ==> s':A'); \ | |
| 12195 | 369 | \ F:program; st_set(A) |] ==> F:A co A'"; | 
| 370 | by (auto_tac (claset() delrules [subsetI], simpset())); | |
| 371 | by (ALLGOALS(asm_full_simp_tac (simpset() addsimps prems))); | |
| 372 | by (Clarify_tac 1); | |
| 11479 | 373 | by (blast_tac(claset() addIs prems) 1); | 
| 374 | qed "constrainsI"; | |
| 375 | ||
| 376 | Goalw [constrains_def] | |
| 377 | "F:A co B ==> ALL act:Acts(F). act``A<=B"; | |
| 378 | by (Blast_tac 1); | |
| 379 | qed "constrainsD"; | |
| 380 | ||
| 381 | Goalw [constrains_def] | |
| 12195 | 382 | "F:A co B ==> F:program & st_set(A)"; | 
| 11479 | 383 | by (Blast_tac 1); | 
| 12195 | 384 | qed "constrainsD2"; | 
| 11479 | 385 | |
| 12195 | 386 | Goalw [constrains_def, st_set_def] "F : 0 co B <-> F:program"; | 
| 11479 | 387 | by (Blast_tac 1); | 
| 388 | qed "constrains_empty"; | |
| 389 | ||
| 12195 | 390 | Goalw [constrains_def, st_set_def] | 
| 391 | "(F : A co 0) <-> (A=0 & F:program)"; | |
| 392 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | |
| 393 | by Auto_tac; | |
| 394 | by (Blast_tac 1); | |
| 11479 | 395 | qed "constrains_empty2"; | 
| 396 | ||
| 12195 | 397 | Goalw [constrains_def, st_set_def] | 
| 398 | "(F: state co B) <-> (state<=B & F:program)"; | |
| 399 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | |
| 400 | by (Blast_tac 1); | |
| 11479 | 401 | qed "constrains_state"; | 
| 402 | ||
| 12195 | 403 | Goalw [constrains_def, st_set_def] "F:A co state <-> (F:program & st_set(A))"; | 
| 404 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | |
| 405 | by (Blast_tac 1); | |
| 11479 | 406 | qed "constrains_state2"; | 
| 407 | ||
| 12195 | 408 | AddIffs [constrains_empty, constrains_empty2, | 
| 11479 | 409 | constrains_state, constrains_state2]; | 
| 410 | ||
| 411 | (*monotonic in 2nd argument*) | |
| 412 | Goalw [constrains_def] | |
| 12195 | 413 | "[| F:A co A'; A'<=B' |] ==> F : A co B'"; | 
| 11479 | 414 | by (Blast_tac 1); | 
| 415 | qed "constrains_weaken_R"; | |
| 416 | ||
| 417 | (*anti-monotonic in 1st argument*) | |
| 12195 | 418 | Goalw [constrains_def, st_set_def] | 
| 11479 | 419 | "[| F : A co A'; B<=A |] ==> F : B co A'"; | 
| 420 | by (Blast_tac 1); | |
| 421 | qed "constrains_weaken_L"; | |
| 422 | ||
| 423 | Goal | |
| 12195 | 424 | "[| F : A co A'; B<=A; A'<=B' |] ==> F : B co B'"; | 
| 11479 | 425 | by (dtac constrains_weaken_R 1); | 
| 12195 | 426 | by (dtac constrains_weaken_L 2); | 
| 11479 | 427 | by (REPEAT(Blast_tac 1)); | 
| 428 | qed "constrains_weaken"; | |
| 429 | ||
| 430 | (** Union **) | |
| 431 | ||
| 12195 | 432 | Goalw [constrains_def, st_set_def] | 
| 11479 | 433 | "[| F : A co A'; F:B co B' |] ==> F:(A Un B) co (A' Un B')"; | 
| 434 | by Auto_tac; | |
| 12195 | 435 | by (Force_tac 1); | 
| 11479 | 436 | qed "constrains_Un"; | 
| 437 | ||
| 12195 | 438 | val major::minor::_ = Goalw [constrains_def, st_set_def] | 
| 439 | "[|(!!i. i:I ==> F:A(i) co A'(i)); F:program |]==> F:(UN i:I. A(i)) co (UN i:I. A'(i))"; | |
| 440 | by (cut_facts_tac [minor] 1); | |
| 441 | by Safe_tac; | |
| 12484 | 442 | by (ALLGOALS(ftac major )); | 
| 12195 | 443 | by (ALLGOALS(Asm_full_simp_tac)); | 
| 444 | by (REPEAT(Blast_tac 1)); | |
| 445 | qed "constrains_UN"; | |
| 11479 | 446 | |
| 12195 | 447 | Goalw [constrains_def, st_set_def] | 
| 11479 | 448 | "(A Un B) co C = (A co C) Int (B co C)"; | 
| 12195 | 449 | by Auto_tac; | 
| 450 | by (Force_tac 1); | |
| 11479 | 451 | qed "constrains_Un_distrib"; | 
| 452 | ||
| 12195 | 453 | Goalw [constrains_def, st_set_def] | 
| 454 | "i:I ==> (UN i:I. A(i)) co B = (INT i:I. A(i) co B)"; | |
| 455 | by (rtac equalityI 1); | |
| 456 | by (REPEAT(Force_tac 1)); | |
| 457 | qed "constrains_UN_distrib"; | |
| 11479 | 458 | |
| 12195 | 459 | (** Intersection **) | 
| 460 | Goalw [constrains_def, st_set_def] | |
| 461 | "C co (A Int B) = (C co A) Int (C co B)"; | |
| 462 | by (rtac equalityI 1); | |
| 463 | by (ALLGOALS(Clarify_tac)); (* to speed up the proof *) | |
| 464 | by (REPEAT(Blast_tac 1)); | |
| 465 | qed "constrains_Int_distrib"; | |
| 466 | ||
| 467 | Goalw [constrains_def, st_set_def] | |
| 468 | "x:I ==> A co (INT i:I. B(i)) = (INT i:I. A co B(i))"; | |
| 11479 | 469 | by (rtac equalityI 1); | 
| 470 | by Safe_tac; | |
| 12195 | 471 | by (REPEAT(Blast_tac 1)); | 
| 11479 | 472 | qed "constrains_INT_distrib"; | 
| 473 | ||
| 12195 | 474 | Goalw [constrains_def, st_set_def] | 
| 11479 | 475 | "[| F : A co A'; F : B co B' |] ==> F : (A Int B) co (A' Int B')"; | 
| 476 | by (Clarify_tac 1); | |
| 477 | by (Blast_tac 1); | |
| 478 | qed "constrains_Int"; | |
| 479 | ||
| 12195 | 480 | val major::minor::_ = Goalw [constrains_def, st_set_def] | 
| 481 | "[| (!!i. i:I==>F:A(i) co A'(i)); F:program|]==> F:(INT i:I. A(i)) co (INT i:I. A'(i))"; | |
| 482 | by (cut_facts_tac [minor] 1); | |
| 483 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | |
| 11479 | 484 | by (case_tac "I=0" 1); | 
| 485 | by (asm_full_simp_tac (simpset() addsimps [Inter_def]) 1); | |
| 486 | by (etac not_emptyE 1); | |
| 12195 | 487 | by Safe_tac; | 
| 488 | by (forw_inst_tac [("i", "xd")] major 1);
 | |
| 12484 | 489 | by (ftac major 2); | 
| 490 | by (ftac major 3); | |
| 12195 | 491 | by (REPEAT(Force_tac 1)); | 
| 492 | qed "constrains_INT"; | |
| 11479 | 493 | |
| 12195 | 494 | (* The rule below simulates the HOL's one for (INT z. A i) co (INT z. B i) *) | 
| 495 | Goalw [constrains_def] | |
| 496 | "[| ALL z. F:{s:state. P(s, z)} co {s:state. Q(s, z)}; F:program |]==>\
 | |
| 497 | \   F:{s:state. ALL z. P(s, z)} co {s:state. ALL z. Q(s, z)}";
 | |
| 11479 | 498 | by (Blast_tac 1); | 
| 12195 | 499 | qed "constrains_All"; | 
| 11479 | 500 | |
| 12195 | 501 | Goalw [constrains_def, st_set_def] | 
| 502 | "[| F:A co A' |] ==> A <= A'"; | |
| 503 | by Auto_tac; | |
| 504 | by (Blast_tac 1); | |
| 11479 | 505 | qed "constrains_imp_subset"; | 
| 506 | (*The reasoning is by subsets since "co" refers to single actions | |
| 507 | only. So this rule isn't that useful.*) | |
| 508 | ||
| 12195 | 509 | Goalw [constrains_def, st_set_def] | 
| 11479 | 510 | "[| F : A co B; F : B co C |] ==> F : A co C"; | 
| 511 | by Auto_tac; | |
| 12195 | 512 | by (Blast_tac 1); | 
| 11479 | 513 | qed "constrains_trans"; | 
| 514 | ||
| 12195 | 515 | Goal | 
| 516 | "[| F : A co (A' Un B); F : B co B' |] ==> F:A co (A' Un B')"; | |
| 517 | by (dres_inst_tac [("A", "B")] constrains_imp_subset 1);
 | |
| 518 | by (blast_tac (claset() addIs [constrains_weaken_R]) 1); | |
| 11479 | 519 | qed "constrains_cancel"; | 
| 520 | ||
| 521 | (*** unless ***) | |
| 522 | ||
| 12195 | 523 | Goalw [unless_def, constrains_def] | 
| 524 | "A unless B <= program"; | |
| 525 | by Auto_tac; | |
| 526 | qed "unless_type"; | |
| 527 | ||
| 528 | Goalw [unless_def] "[| F:(A-B) co (A Un B) |] ==> F : A unless B"; | |
| 529 | by (blast_tac (claset() addDs [constrainsD2]) 1); | |
| 11479 | 530 | qed "unlessI"; | 
| 531 | ||
| 532 | Goalw [unless_def] "F :A unless B ==> F : (A-B) co (A Un B)"; | |
| 12195 | 533 | by Auto_tac; | 
| 11479 | 534 | qed "unlessD"; | 
| 535 | ||
| 536 | (*** initially ***) | |
| 537 | ||
| 538 | Goalw [initially_def] | |
| 12195 | 539 | "initially(A) <= program"; | 
| 540 | by (Blast_tac 1); | |
| 541 | qed "initially_type"; | |
| 542 | ||
| 543 | Goalw [initially_def] | |
| 544 | "[| F:program; Init(F)<=A |] ==> F:initially(A)"; | |
| 11479 | 545 | by (Blast_tac 1); | 
| 546 | qed "initiallyI"; | |
| 547 | ||
| 548 | Goalw [initially_def] | |
| 549 | "F:initially(A) ==> Init(F)<=A"; | |
| 550 | by (Blast_tac 1); | |
| 551 | qed "initiallyD"; | |
| 552 | ||
| 12195 | 553 | (*** stable ***) | 
| 554 | ||
| 555 | Goalw [stable_def, constrains_def] | |
| 556 | "stable(A)<=program"; | |
| 11479 | 557 | by (Blast_tac 1); | 
| 12195 | 558 | qed "stable_type"; | 
| 11479 | 559 | |
| 560 | Goalw [stable_def] | |
| 561 | "F : A co A ==> F : stable(A)"; | |
| 562 | by (assume_tac 1); | |
| 563 | qed "stableI"; | |
| 564 | ||
| 12195 | 565 | Goalw [stable_def] "F:stable(A) ==> F : A co A"; | 
| 11479 | 566 | by (assume_tac 1); | 
| 567 | qed "stableD"; | |
| 568 | ||
| 12195 | 569 | Goalw [stable_def, constrains_def] "F:stable(A) ==> F:program & st_set(A)"; | 
| 570 | by Auto_tac; | |
| 11479 | 571 | qed "stableD2"; | 
| 572 | ||
| 12195 | 573 | Goalw [stable_def, constrains_def] "stable(state) = program"; | 
| 574 | by (auto_tac (claset() addDs [Acts_type RS subsetD], simpset())); | |
| 11479 | 575 | qed "stable_state"; | 
| 12195 | 576 | AddIffs [stable_state]; | 
| 11479 | 577 | |
| 578 | (** Union **) | |
| 579 | ||
| 580 | Goalw [stable_def] | |
| 12195 | 581 | "[| F : stable(A); F:stable(A') |] ==> F : stable(A Un A')"; | 
| 11479 | 582 | by (blast_tac (claset() addIs [constrains_Un]) 1); | 
| 583 | qed "stable_Un"; | |
| 584 | ||
| 585 | val [major, minor] = Goalw [stable_def] | |
| 12195 | 586 | "[|(!!i. i:I ==> F : stable(A(i))); F:program |] ==> F:stable (UN i:I. A(i))"; | 
| 11479 | 587 | by (cut_facts_tac [minor] 1); | 
| 588 | by (blast_tac (claset() addIs [constrains_UN, major]) 1); | |
| 589 | qed "stable_UN"; | |
| 590 | ||
| 591 | Goalw [stable_def] | |
| 592 | "[| F : stable(A); F : stable(A') |] ==> F : stable (A Int A')"; | |
| 593 | by (blast_tac (claset() addIs [constrains_Int]) 1); | |
| 594 | qed "stable_Int"; | |
| 595 | ||
| 596 | val [major, minor] = Goalw [stable_def] | |
| 12195 | 597 | "[| (!!i. i:I ==> F:stable(A(i))); F:program |] ==> F : stable (INT i:I. A(i))"; | 
| 11479 | 598 | by (cut_facts_tac [minor] 1); | 
| 599 | by (blast_tac (claset() addIs [constrains_INT, major]) 1); | |
| 600 | qed "stable_INT"; | |
| 601 | ||
| 602 | Goalw [stable_def] | |
| 12195 | 603 | "[|ALL z. F:stable({s:state. P(s, z)}); F:program|]==>F:stable({s:state. ALL z. P(s, z)})";
 | 
| 11479 | 604 | by (rtac constrains_All 1); | 
| 605 | by Auto_tac; | |
| 12195 | 606 | qed "stable_All"; | 
| 11479 | 607 | |
| 12195 | 608 | Goalw [stable_def, constrains_def, st_set_def] | 
| 609 | "[| F : stable(C); F : A co (C Un A') |] ==> F : (C Un A) co (C Un A')"; | |
| 610 | by (cut_inst_tac [("F", "F")] Acts_type 1);
 | |
| 611 | by Auto_tac; | |
| 612 | by (Force_tac 1); | |
| 11479 | 613 | qed "stable_constrains_Un"; | 
| 614 | ||
| 12195 | 615 | Goalw [stable_def, constrains_def, st_set_def] | 
| 11479 | 616 | "[| F : stable(C); F : (C Int A) co A' |] ==> F : (C Int A) co (C Int A')"; | 
| 617 | by (Clarify_tac 1); | |
| 618 | by (Blast_tac 1); | |
| 619 | qed "stable_constrains_Int"; | |
| 620 | ||
| 12195 | 621 | (* [| F:stable(C); F :(C Int A) co A |] ==> F:stable(C Int A) *) | 
| 11479 | 622 | bind_thm ("stable_constrains_stable", stable_constrains_Int RS stableI);
 | 
| 623 | ||
| 624 | (** invariant **) | |
| 625 | ||
| 12195 | 626 | Goalw [invariant_def] | 
| 627 | "invariant(A) <= program"; | |
| 628 | by (blast_tac (claset() addDs [stable_type RS subsetD]) 1); | |
| 629 | qed "invariant_type"; | |
| 11479 | 630 | |
| 12195 | 631 | Goalw [invariant_def, initially_def] | 
| 11479 | 632 | "[| Init(F)<=A; F:stable(A) |] ==> F : invariant(A)"; | 
| 12195 | 633 | by (forward_tac [stable_type RS subsetD] 1); | 
| 634 | by Auto_tac; | |
| 11479 | 635 | qed "invariantI"; | 
| 636 | ||
| 12195 | 637 | Goalw [invariant_def, initially_def] | 
| 11479 | 638 | "F:invariant(A) ==> Init(F)<=A & F:stable(A)"; | 
| 639 | by Auto_tac; | |
| 640 | qed "invariantD"; | |
| 641 | ||
| 12195 | 642 | Goalw [invariant_def] | 
| 643 | "F:invariant(A) ==> F:program & st_set(A)"; | |
| 11479 | 644 | by (blast_tac (claset() addDs [stableD2]) 1); | 
| 645 | qed "invariantD2"; | |
| 646 | ||
| 647 | (*Could also say "invariant A Int invariant B <= invariant (A Int B)"*) | |
| 12195 | 648 | Goalw [invariant_def, initially_def] | 
| 11479 | 649 | "[| F : invariant(A); F : invariant(B) |] ==> F : invariant(A Int B)"; | 
| 650 | by (asm_full_simp_tac (simpset() addsimps [stable_Int]) 1); | |
| 651 | by (Blast_tac 1); | |
| 652 | qed "invariant_Int"; | |
| 653 | ||
| 654 | (** The Elimination Theorem. The "free" m has become universally quantified! | |
| 12195 | 655 | Should the premise be !!m instead of ALL m ? Would make it harder | 
| 656 | to use in forward proof. **) | |
| 11479 | 657 | |
| 12195 | 658 | (* The general case easier to prove that le special case! *) | 
| 659 | Goalw [constrains_def, st_set_def] | |
| 660 |     "[| ALL m:M. F : {s:A. x(s) = m} co B(m); F:program  |] \
 | |
| 661 | \    ==> F:{s:A. x(s):M} co (UN m:M. B(m))";
 | |
| 11479 | 662 | by Safe_tac; | 
| 12195 | 663 | by Auto_tac; | 
| 11479 | 664 | by (Blast_tac 1); | 
| 665 | qed "elimination"; | |
| 666 | ||
| 12195 | 667 | (* As above, but for the special case of A=state *) | 
| 11479 | 668 | Goal "[| ALL m:M. F : {s:state. x(s) = m} co B(m); F:program  |] \
 | 
| 669 | \    ==> F:{s:state. x(s):M} co (UN m:M. B(m))";
 | |
| 670 | by (rtac elimination 1); | |
| 671 | by (ALLGOALS(Clarify_tac)); | |
| 672 | qed "eliminiation2"; | |
| 673 | ||
| 12195 | 674 | (** strongest_rhs **) | 
| 11479 | 675 | |
| 12195 | 676 | Goalw [constrains_def, strongest_rhs_def, st_set_def] | 
| 677 | "[| F:program; st_set(A) |] ==> F:A co (strongest_rhs(F,A))"; | |
| 678 | by (auto_tac (claset() addDs [Acts_type RS subsetD], simpset())); | |
| 11479 | 679 | qed "constrains_strongest_rhs"; | 
| 680 | ||
| 12195 | 681 | Goalw [constrains_def, strongest_rhs_def, st_set_def] | 
| 682 | "[| F:A co B; st_set(B) |] ==> strongest_rhs(F,A) <= B"; | |
| 11479 | 683 | by Safe_tac; | 
| 684 | by (dtac InterD 1); | |
| 12195 | 685 | by Auto_tac; | 
| 11479 | 686 | qed "strongest_rhs_is_strongest"; | 
| 687 | ||
| 688 | (*** increasing ***) | |
| 12195 | 689 | Goalw [increasing_def] "increasing(A, r, f) <= program"; | 
| 690 | by (case_tac "A=0" 1); | |
| 691 | by (etac not_emptyE 2); | |
| 692 | by (Clarify_tac 2); | |
| 693 | by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [Inter_iff, Inter_0]))); | |
| 694 | by (blast_tac (claset() addDs [stable_type RS subsetD]) 1); | |
| 695 | qed "increasing_type"; | |
| 11479 | 696 | |
| 12195 | 697 | Goalw [increasing_def] | 
| 698 |    "[| F:increasing(A, r, f); a:A |] ==> F:stable({s:state. <a, f`s>:r})";
 | |
| 699 | by (Blast_tac 1); | |
| 700 | qed "increasingD"; | |
| 11479 | 701 | |
| 12195 | 702 | Goalw [increasing_def] | 
| 703 | "F:increasing(A, r, f) ==> F:program & (EX a. a:A)"; | |
| 704 | by (auto_tac (claset() addDs [stable_type RS subsetD], | |
| 705 | simpset() addsimps [INT_iff])); | |
| 706 | qed "increasingD2"; | |
| 11479 | 707 | |
| 12195 | 708 | Goalw [increasing_def, stable_def] | 
| 709 | "F:increasing(A, r, lam s:state. c) <-> F:program & (EX a. a:A)"; | |
| 710 | by (auto_tac (claset() addDs [constrains_type RS subsetD], | |
| 711 | simpset() addsimps [INT_iff])); | |
| 712 | qed "increasing_constant"; | |
| 713 | AddIffs [increasing_constant]; | |
| 714 | ||
| 715 | Goalw [increasing_def, stable_def, constrains_def, st_set_def, part_order_def] | |
| 716 | "[| g:mono_map(A,r,A,r); part_order(A, r); f:state->A |] \ | |
| 717 | \ ==> increasing(A, r,f) <= increasing(A, r,g O f)"; | |
| 718 | by (case_tac "A=0" 1); | |
| 719 | by (Asm_full_simp_tac 1); | |
| 720 | by (etac not_emptyE 1); | |
| 721 | by (Clarify_tac 1); | |
| 722 | by (cut_inst_tac [("F", "xa")] Acts_type 1);
 | |
| 723 | by (asm_full_simp_tac (simpset() addsimps [Inter_iff, mono_map_def]) 1); | |
| 724 | by Auto_tac; | |
| 725 | by (dres_inst_tac [("psi", "ALL x:A. ALL xa:A. ?u(x,xa)")] asm_rl 1); 
 | |
| 726 | by (dres_inst_tac [("x", "f`xf")] bspec 1);
 | |
| 727 | by Auto_tac; | |
| 728 | by (dres_inst_tac [("psi", "ALL x:A. ALL xa:A. ?u(x,xa)")] asm_rl 1); 
 | |
| 12152 
46f128d8133c
Renamed some bound variables due to changes in simplifier.
 berghofe parents: 
11479diff
changeset | 729 | by (dres_inst_tac [("x", "act")] bspec 1);
 | 
| 12195 | 730 | by Auto_tac; | 
| 731 | by (dres_inst_tac [("psi", "Acts(?u) <= ?v")] asm_rl 1);
 | |
| 732 | by (dres_inst_tac [("psi", "?u <= state")] asm_rl 1);
 | |
| 733 | by (dres_inst_tac [("c", "xe")] subsetD 1);
 | |
| 11479 | 734 | by (rtac imageI 1); | 
| 735 | by Auto_tac; | |
| 12195 | 736 | by (asm_full_simp_tac (simpset() addsimps [refl_def, apply_type]) 1); | 
| 737 | by (dres_inst_tac [("x1", "f`xf"), ("x", "f`xe")] (bspec RS bspec) 1);
 | |
| 738 | by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [apply_type]))); | |
| 739 | by (res_inst_tac [("b", "g ` (f ` xf)")] trans_onD 1);
 | |
| 740 | by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [apply_type]))); | |
| 741 | qed "mono_increasing_comp"; | |
| 11479 | 742 | |
| 743 | (*Holds by the theorem (succ(m) le n) = (m < n) *) | |
| 12195 | 744 | Goalw [increasing_def] | 
| 745 |      "[| F:increasing(nat, {<m,n>:nat*nat. m le n}, f); f:state->nat; k:nat |] \
 | |
| 746 | \  ==> F: stable({s:state. k < f`s})";
 | |
| 11479 | 747 | by (Clarify_tac 1); | 
| 748 | by (asm_full_simp_tac (simpset() addsimps [INT_iff]) 1); | |
| 749 | by Safe_tac; | |
| 12195 | 750 | by (dres_inst_tac [("x", "succ(k)")] bspec 1);
 | 
| 11479 | 751 | by (auto_tac (claset(), simpset() addsimps [apply_type, Collect_conj_eq])); | 
| 12195 | 752 | by (subgoal_tac "{x: state . f`x : nat} = state" 1);
 | 
| 11479 | 753 | by Auto_tac; | 
| 12195 | 754 | qed "strict_increasingD"; | 
| 755 | ||
| 756 | ||
| 757 | (* Used in WFair.thy *) | |
| 758 | Goal "A:Pow(Pow(B)) ==> Union(A):Pow(B)"; | |
| 759 | by Auto_tac; | |
| 760 | qed "Union_PowI"; |