| author | wenzelm | 
| Wed, 06 Dec 2000 12:28:52 +0100 | |
| changeset 10605 | fe12dc60bc3c | 
| parent 10216 | e928bdf62014 | 
| child 11320 | 56aa53caf333 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/IMP/Equiv.ML | 
| 482 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Heiko Loetzbeyer & Robert Sandner, TUM | 
| 482 | 4 | Copyright 1994 TUM | 
| 5 | *) | |
| 6 | ||
| 5268 | 7 | Goal "[| a: aexp; sigma: loc -> nat |] ==> \ | 
| 4298 | 8 | \ <a,sigma> -a-> n <-> A(a,sigma) = n"; | 
| 511 | 9 | by (res_inst_tac [("x","n")] spec 1);                       (* quantify n *)
 | 
| 4298 | 10 | by (etac aexp.induct 1); | 
| 11 | by (ALLGOALS (fast_tac (claset() addSIs evala.intrs | |
| 12 | addSEs aexp_elim_cases | |
| 13 | addss (simpset())))); | |
| 760 | 14 | qed "aexp_iff"; | 
| 482 | 15 | |
| 16 | ||
| 4298 | 17 | val aexp1 = aexp_iff RS iffD1; | 
| 482 | 18 | val aexp2 = aexp_iff RS iffD2; | 
| 19 | ||
| 20 | ||
| 511 | 21 | val bexp_elim_cases = | 
| 482 | 22 | [ | 
| 6141 | 23 | evalb.mk_cases "<true,sigma> -b-> x", | 
| 24 | evalb.mk_cases "<false,sigma> -b-> x", | |
| 25 | evalb.mk_cases "<ROp(f,a0,a1),sigma> -b-> x", | |
| 26 | evalb.mk_cases "<noti(b),sigma> -b-> x", | |
| 27 | evalb.mk_cases "<b0 andi b1,sigma> -b-> x", | |
| 28 | evalb.mk_cases "<b0 ori b1,sigma> -b-> x" | |
| 482 | 29 | ]; | 
| 30 | ||
| 31 | ||
| 5268 | 32 | Goal "[| b: bexp; sigma: loc -> nat |] ==> \ | 
| 4298 | 33 | \ <b,sigma> -b-> w <-> B(b,sigma) = w"; | 
| 34 | by (res_inst_tac [("x","w")] spec 1);
 | |
| 35 | by (etac bexp.induct 1); | |
| 36 | by (ALLGOALS (fast_tac (claset() addSIs evalb.intrs | |
| 37 | addSEs bexp_elim_cases | |
| 38 | addss (simpset() addsimps [aexp_iff])))); | |
| 760 | 39 | qed "bexp_iff"; | 
| 482 | 40 | |
| 4298 | 41 | val bexp1 = bexp_iff RS iffD1; | 
| 518 | 42 | val bexp2 = bexp_iff RS iffD2; | 
| 482 | 43 | |
| 4298 | 44 | |
| 5137 | 45 | Goal "<c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)"; | 
| 1742 | 46 | by (etac evalc.induct 1); | 
| 4298 | 47 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [bexp1]))); | 
| 482 | 48 | (* skip *) | 
| 2469 | 49 | by (Fast_tac 1); | 
| 482 | 50 | (* assign *) | 
| 4298 | 51 | by (asm_full_simp_tac (simpset() addsimps | 
| 6047 
f2e0938ba38d
converted to use new primrec section and update operator
 paulson parents: 
5268diff
changeset | 52 | [aexp1, update_type] @ op_type_intrs) 1); | 
| 482 | 53 | (* comp *) | 
| 2469 | 54 | by (Fast_tac 1); | 
| 482 | 55 | (* while *) | 
| 10216 | 56 | by (etac (Gamma_bnd_mono RS lfp_unfold RS ssubst) 1); | 
| 4298 | 57 | by (asm_simp_tac (simpset() addsimps [Gamma_def, bexp1]) 1); | 
| 58 | by (blast_tac (claset() addSIs [bexp1]@evalb_type_intrs) 1); | |
| 59 | (* recursive case of while *) | |
| 10216 | 60 | by (etac (Gamma_bnd_mono RS lfp_unfold RS ssubst) 1); | 
| 4298 | 61 | by (asm_full_simp_tac (simpset() addsimps [Gamma_def, bexp1]) 1); | 
| 62 | by (blast_tac (claset() addSIs [bexp1]@evalb_type_intrs) 1); | |
| 500 | 63 | val com1 = result(); | 
| 482 | 64 | |
| 65 | ||
| 2469 | 66 | AddSIs [aexp2,bexp2,B_type,A_type]; | 
| 4298 | 67 | AddIs evalc.intrs; | 
| 68 | ||
| 482 | 69 | |
| 5137 | 70 | Goal "c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)"; | 
| 4298 | 71 | by (etac com.induct 1); | 
| 9178 | 72 | (* comp *) | 
| 73 | by (Force_tac 3); | |
| 4298 | 74 | (* assign *) | 
| 9178 | 75 | by (Force_tac 2); | 
| 76 | (* skip *) | |
| 77 | by (Force_tac 1); | |
| 4298 | 78 | (* while *) | 
| 4152 | 79 | by Safe_tac; | 
| 2469 | 80 | by (ALLGOALS Asm_full_simp_tac); | 
| 7499 | 81 | by (EVERY1 [ftac Gamma_bnd_mono , etac induct, atac]); | 
| 808 | 82 | by (rewtac Gamma_def); | 
| 4152 | 83 | by Safe_tac; | 
| 518 | 84 | by (EVERY1 [dtac bspec, atac]); | 
| 2469 | 85 | by (ALLGOALS Asm_full_simp_tac); | 
| 511 | 86 | (* while, if *) | 
| 4298 | 87 | by (ALLGOALS Blast_tac); | 
| 482 | 88 | val com2 = result(); | 
| 89 | ||
| 90 | ||
| 511 | 91 | (**** Proof of Equivalence ****) | 
| 482 | 92 | |
| 5268 | 93 | Goal "ALL c:com. C(c) = {io:(loc->nat)*(loc->nat). <c,fst(io)> -c-> snd(io)}";
 | 
| 4091 | 94 | by (fast_tac (claset() addIs [C_subset RS subsetD] | 
| 4298 | 95 | addEs [com2 RS bspec] | 
| 96 | addDs [com1] | |
| 97 | addss (simpset())) 1); | |
| 511 | 98 | val com_equivalence = result(); |