| author | wenzelm | 
| Tue, 06 Sep 1994 13:10:38 +0200 | |
| changeset 583 | 550292083e66 | 
| parent 518 | 4530c45370b4 | 
| child 672 | 1922f98b8f7e | 
| permissions | -rw-r--r-- | 
| 482 | 1  | 
(* Title: ZF/IMP/Equiv.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Heiko Loetzbeyer & Robert Sandner, TUM  | 
|
4  | 
Copyright 1994 TUM  | 
|
5  | 
*)  | 
|
6  | 
||
| 511 | 7  | 
val prems = goal Equiv.thy "[| a: aexp; sigma: loc -> nat |] ==> \  | 
| 518 | 8  | 
\ <a,sigma> -a-> n <-> A(a,sigma) = n";  | 
| 511 | 9  | 
by (res_inst_tac [("x","n")] spec 1);                       (* quantify n *)
 | 
10  | 
by (res_inst_tac [("x","a")] aexp.induct 1);                (* struct. ind. *)
 | 
|
11  | 
by (resolve_tac prems 1); (* type prem. *)  | 
|
12  | 
by (rewrite_goals_tac A_rewrite_rules); (* rewr. Den. *)  | 
|
| 518 | 13  | 
by (TRYALL (fast_tac (ZF_cs addSIs (evala.intrs@prems)  | 
14  | 
addSEs aexp_elim_cases)));  | 
|
| 511 | 15  | 
val aexp_iff = result();  | 
| 482 | 16  | 
|
17  | 
||
| 518 | 18  | 
val aexp1 = prove_goal Equiv.thy  | 
19  | 
"[| <a,sigma> -a-> n; a: aexp; sigma: loc -> nat |] \  | 
|
20  | 
\ ==> A(a,sigma) = n" (* destruction rule *)  | 
|
21  | 
(fn prems => [(fast_tac (ZF_cs addSIs ((aexp_iff RS iffD1)::prems)) 1)]);  | 
|
| 482 | 22  | 
val aexp2 = aexp_iff RS iffD2;  | 
23  | 
||
24  | 
||
| 511 | 25  | 
val bexp_elim_cases =  | 
| 482 | 26  | 
[  | 
| 511 | 27  | 
evalb.mk_cases bexp.con_defs "<true,sigma> -b-> x",  | 
28  | 
evalb.mk_cases bexp.con_defs "<false,sigma> -b-> x",  | 
|
29  | 
evalb.mk_cases bexp.con_defs "<ROp(f,a0,a1),sigma> -b-> x",  | 
|
30  | 
evalb.mk_cases bexp.con_defs "<noti(b),sigma> -b-> x",  | 
|
31  | 
evalb.mk_cases bexp.con_defs "<b0 andi b1,sigma> -b-> x",  | 
|
32  | 
evalb.mk_cases bexp.con_defs "<b0 ori b1,sigma> -b-> x"  | 
|
| 482 | 33  | 
];  | 
34  | 
||
35  | 
||
36  | 
val prems = goal Equiv.thy "[| b: bexp; sigma: loc -> nat |] ==> \  | 
|
| 518 | 37  | 
\ <b,sigma> -b-> w <-> B(b,sigma) = w";  | 
| 
510
 
093665669f52
Simplified some proofs. Added some type assumptions to the introduction rules.
 
nipkow 
parents: 
500 
diff
changeset
 | 
38  | 
by (res_inst_tac [("x","w")] spec 1);			(* quantify w *)
 | 
| 511 | 39  | 
by (res_inst_tac [("x","b")] bexp.induct 1);		(* struct. ind. *)
 | 
| 
510
 
093665669f52
Simplified some proofs. Added some type assumptions to the introduction rules.
 
nipkow 
parents: 
500 
diff
changeset
 | 
40  | 
by (resolve_tac prems 1); (* type prem. *)  | 
| 
 
093665669f52
Simplified some proofs. Added some type assumptions to the introduction rules.
 
nipkow 
parents: 
500 
diff
changeset
 | 
41  | 
by (rewrite_goals_tac B_rewrite_rules); (* rewr. Den. *)  | 
| 518 | 42  | 
by (TRYALL (fast_tac (ZF_cs addSIs (evalb.intrs@prems@[aexp2])  | 
43  | 
addSDs [aexp1] addSEs bexp_elim_cases)));  | 
|
| 482 | 44  | 
val bexp_iff = result();  | 
45  | 
||
| 511 | 46  | 
val bexp1 = prove_goal Equiv.thy  | 
| 518 | 47  | 
"[| <b,sigma> -b-> w; b: bexp; sigma: loc -> nat |]\  | 
48  | 
\ ==> B(b,sigma) = w"  | 
|
49  | 
(fn prems => [(fast_tac (ZF_cs addSIs ((bexp_iff RS iffD1)::prems)) 1)]);  | 
|
50  | 
val bexp2 = bexp_iff RS iffD2;  | 
|
| 482 | 51  | 
|
| 518 | 52  | 
goal Equiv.thy "!!c. <c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";  | 
| 482 | 53  | 
|
| 500 | 54  | 
(* start with rule induction *)  | 
| 511 | 55  | 
be (evalc.mutual_induct RS spec RS spec RS spec RSN (2,rev_mp)) 1;  | 
| 482 | 56  | 
|
| 500 | 57  | 
by (rewrite_tac (Gamma_def::C_rewrite_rules));  | 
| 482 | 58  | 
(* skip *)  | 
| 500 | 59  | 
by (fast_tac comp_cs 1);  | 
60  | 
||
| 482 | 61  | 
(* assign *)  | 
62  | 
by (asm_full_simp_tac (ZF_ss addsimps [aexp1,assign_type] @  | 
|
63  | 
op_type_intrs) 1);  | 
|
64  | 
(* comp *)  | 
|
65  | 
by (fast_tac comp_cs 1);  | 
|
66  | 
||
67  | 
(* if *)  | 
|
| 500 | 68  | 
by (fast_tac (ZF_cs addSIs [bexp1] addIs [(fst_conv RS ssubst)]) 1);  | 
69  | 
by (fast_tac (ZF_cs addSIs [bexp1] addIs [(fst_conv RS ssubst)]) 1);  | 
|
| 482 | 70  | 
|
71  | 
(* while *)  | 
|
| 511 | 72  | 
by (etac (rewrite_rule [Gamma_def] (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1);  | 
73  | 
by (fast_tac (comp_cs addSIs [bexp1,idI]@evalb_type_intrs  | 
|
| 500 | 74  | 
addIs [(fst_conv RS ssubst)]) 1);  | 
| 482 | 75  | 
|
| 511 | 76  | 
by (etac (rewrite_rule [Gamma_def] (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1);  | 
77  | 
by (fast_tac (comp_cs addSIs [bexp1,compI]@evalb_type_intrs  | 
|
| 482 | 78  | 
addIs [(fst_conv RS ssubst)]) 1);  | 
79  | 
||
| 500 | 80  | 
val com1 = result();  | 
| 482 | 81  | 
|
82  | 
||
83  | 
val com_cs = ZF_cs addSIs [aexp2,bexp2,B_type,A_type]  | 
|
| 511 | 84  | 
addIs evalc.intrs  | 
| 482 | 85  | 
addSEs [idE,compE]  | 
86  | 
addEs [C_type,C_type_fst];  | 
|
87  | 
||
| 511 | 88  | 
val [prem] = goal Equiv.thy  | 
| 518 | 89  | 
"c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";  | 
| 511 | 90  | 
br (prem RS com.induct) 1;  | 
| 482 | 91  | 
by (rewrite_tac C_rewrite_rules);  | 
92  | 
by (safe_tac com_cs);  | 
|
| 500 | 93  | 
by (ALLGOALS (asm_full_simp_tac ZF_ss));  | 
| 482 | 94  | 
|
95  | 
(* skip *)  | 
|
96  | 
by (fast_tac com_cs 1);  | 
|
| 500 | 97  | 
|
| 482 | 98  | 
(* assign *)  | 
99  | 
by (fast_tac com_cs 1);  | 
|
| 500 | 100  | 
|
| 482 | 101  | 
(* comp *)  | 
| 518 | 102  | 
by (REPEAT (EVERY [(dtac bspec 1),(atac 1)]));  | 
| 482 | 103  | 
by (asm_full_simp_tac ZF_ss 1);  | 
104  | 
by (fast_tac com_cs 1);  | 
|
| 500 | 105  | 
|
| 482 | 106  | 
(* while *)  | 
| 518 | 107  | 
by (EVERY1 [forward_tac [Gamma_bnd_mono], etac induct, atac]);  | 
| 482 | 108  | 
by (rewrite_goals_tac [Gamma_def]);  | 
109  | 
by (safe_tac com_cs);  | 
|
| 518 | 110  | 
by (EVERY1 [dtac bspec, atac]);  | 
| 482 | 111  | 
by (ALLGOALS (asm_full_simp_tac ZF_ss));  | 
112  | 
||
| 511 | 113  | 
(* while, if *)  | 
| 482 | 114  | 
by (ALLGOALS (fast_tac com_cs));  | 
115  | 
val com2 = result();  | 
|
116  | 
||
117  | 
||
| 511 | 118  | 
(**** Proof of Equivalence ****)  | 
| 482 | 119  | 
|
120  | 
val com_iff_cs = ZF_cs addIs [C_subset RS subsetD]  | 
|
| 518 | 121  | 
addEs [com2 RS bspec]  | 
| 500 | 122  | 
addDs [com1];  | 
| 482 | 123  | 
|
| 511 | 124  | 
goal Equiv.thy  | 
125  | 
    "ALL c:com. C(c) = {io:(loc->nat)*(loc->nat). <c,fst(io)> -c-> snd(io)}";
 | 
|
| 482 | 126  | 
by (rtac ballI 1);  | 
127  | 
by (rtac equalityI 1);  | 
|
| 500 | 128  | 
(* => *)  | 
| 482 | 129  | 
by (fast_tac com_iff_cs 1);  | 
| 500 | 130  | 
(* <= *)  | 
131  | 
by (REPEAT (step_tac com_iff_cs 1));  | 
|
| 482 | 132  | 
by (asm_full_simp_tac ZF_ss 1);  | 
| 511 | 133  | 
val com_equivalence = result();  |