| author | wenzelm | 
| Wed, 06 Dec 2000 22:08:49 +0100 | |
| changeset 10621 | 3d15596ee644 | 
| parent 9870 | 2374ba026fc6 | 
| permissions | -rw-r--r-- | 
| 1700 | 1  | 
(* Title: HOL/IMP/Transition.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Tobias Nipkow & Robert Sandner, TUM  | 
|
4  | 
Copyright 1996 TUM  | 
|
5  | 
||
6  | 
Equivalence of Natural and Transition semantics  | 
|
7  | 
*)  | 
|
8  | 
||
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
9  | 
section "Winskel's Proof";  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
10  | 
|
| 1973 | 11  | 
AddSEs [rel_pow_0_E];  | 
| 1700 | 12  | 
|
| 6141 | 13  | 
val evalc1_SEs =  | 
14  | 
map evalc1.mk_cases  | 
|
15  | 
["(SKIP,s) -1-> t",  | 
|
| 9241 | 16  | 
"(x:==a,s) -1-> t",  | 
| 6141 | 17  | 
"(c1;c2, s) -1-> t",  | 
| 8016 | 18  | 
"(IF b THEN c1 ELSE c2, s) -1-> t",  | 
19  | 
"(WHILE b DO c, s) -1-> t"];  | 
|
| 6141 | 20  | 
|
21  | 
val evalc1_E = evalc1.mk_cases "(WHILE b DO c,s) -1-> t";  | 
|
| 1700 | 22  | 
|
| 1973 | 23  | 
AddSEs evalc1_SEs;  | 
24  | 
||
25  | 
AddIs evalc1.intrs;  | 
|
| 1700 | 26  | 
|
| 5278 | 27  | 
Goal "!s t u c d. (c,s) -n-> (SKIP,t) --> (d,t) -*-> (SKIP,u) --> \  | 
| 1700 | 28  | 
\ (c;d, s) -*-> (SKIP, u)";  | 
| 5183 | 29  | 
by (induct_tac "n" 1);  | 
| 4089 | 30  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2])1);  | 
| 
4772
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4746 
diff
changeset
 | 
31  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2]addSDs [rel_pow_Suc_D2])1);  | 
| 1700 | 32  | 
qed_spec_mp "lemma1";  | 
33  | 
||
| 5117 | 34  | 
Goal "<c,s> -c-> s1 ==> (c,s) -*-> (SKIP,s1)";  | 
| 2031 | 35  | 
by (etac evalc.induct 1);  | 
| 1700 | 36  | 
|
37  | 
(* SKIP *)  | 
|
| 2031 | 38  | 
by (rtac rtrancl_refl 1);  | 
| 1700 | 39  | 
|
40  | 
(* ASSIGN *)  | 
|
| 4089 | 41  | 
by (fast_tac (claset() addSIs [r_into_rtrancl]) 1);  | 
| 1700 | 42  | 
|
43  | 
(* SEMI *)  | 
|
| 4089 | 44  | 
by (fast_tac (claset() addDs [rtrancl_imp_UN_rel_pow] addIs [lemma1]) 1);  | 
| 1700 | 45  | 
|
46  | 
(* IF *)  | 
|
| 4089 | 47  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2]) 1);  | 
48  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2]) 1);  | 
|
| 1700 | 49  | 
|
50  | 
(* WHILE *)  | 
|
| 4089 | 51  | 
by (fast_tac (claset() addSIs [r_into_rtrancl]) 1);  | 
52  | 
by (fast_tac (claset() addDs [rtrancl_imp_UN_rel_pow]  | 
|
| 1700 | 53  | 
addIs [rtrancl_into_rtrancl2,lemma1]) 1);  | 
54  | 
||
| 1730 | 55  | 
qed "evalc_impl_evalc1";  | 
| 1700 | 56  | 
|
57  | 
||
| 
9556
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
58  | 
Goal "(SKIP,s) -m-> (SKIP,t) ==> s = t & m = 0";  | 
| 
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
59  | 
by (etac rel_pow_E2 1);  | 
| 
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
60  | 
by (Asm_full_simp_tac 1);  | 
| 
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
61  | 
by (Fast_tac 1);  | 
| 
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
62  | 
val hlemma = result();  | 
| 
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
63  | 
|
| 5278 | 64  | 
Goal "!c d s u. (c;d,s) -n-> (SKIP,u) --> \  | 
| 1701 | 65  | 
\ (? t m. (c,s) -*-> (SKIP,t) & (d,t) -m-> (SKIP,u) & m <= n)";  | 
| 4897 | 66  | 
by (induct_tac "n" 1);  | 
| 1700 | 67  | 
(* case n = 0 *)  | 
| 4089 | 68  | 
by (fast_tac (claset() addss simpset()) 1);  | 
| 1700 | 69  | 
(* induction step *)  | 
| 4089 | 70  | 
by (fast_tac (claset() addSIs [le_SucI,le_refl]  | 
| 1700 | 71  | 
addSDs [rel_pow_Suc_D2]  | 
| 1973 | 72  | 
addSEs [rel_pow_imp_rtrancl,rtrancl_into_rtrancl2]) 1);  | 
| 1700 | 73  | 
qed_spec_mp "lemma2";  | 
74  | 
||
| 5069 | 75  | 
Goal "!s t. (c,s) -*-> (SKIP,t) --> <c,s> -c-> t";  | 
| 4897 | 76  | 
by (induct_tac "c" 1);  | 
| 4089 | 77  | 
by (safe_tac (claset() addSDs [rtrancl_imp_UN_rel_pow]));  | 
| 1700 | 78  | 
|
79  | 
(* SKIP *)  | 
|
| 4089 | 80  | 
by (fast_tac (claset() addSEs [rel_pow_E2]) 1);  | 
| 1700 | 81  | 
|
82  | 
(* ASSIGN *)  | 
|
| 4897 | 83  | 
by (fast_tac (claset() addSDs [hlemma] addSEs [rel_pow_E2]) 1);  | 
| 1700 | 84  | 
|
85  | 
(* SEMI *)  | 
|
| 4089 | 86  | 
by (fast_tac (claset() addSDs [lemma2,rel_pow_imp_rtrancl]) 1);  | 
| 1700 | 87  | 
|
88  | 
(* IF *)  | 
|
| 2031 | 89  | 
by (etac rel_pow_E2 1);  | 
| 1700 | 90  | 
by (Asm_full_simp_tac 1);  | 
| 4089 | 91  | 
by (fast_tac (claset() addSDs [rel_pow_imp_rtrancl]) 1);  | 
| 1700 | 92  | 
|
93  | 
(* WHILE, induction on the length of the computation *)  | 
|
| 4153 | 94  | 
by (eres_inst_tac [("P","?X -n-> ?Y")] rev_mp 1);
 | 
| 1700 | 95  | 
by (res_inst_tac [("x","s")] spec 1);
 | 
| 9870 | 96  | 
by (induct_thm_tac nat_less_induct "n" 1);  | 
| 1700 | 97  | 
by (strip_tac 1);  | 
| 2031 | 98  | 
by (etac rel_pow_E2 1);  | 
| 1700 | 99  | 
by (Asm_full_simp_tac 1);  | 
| 6141 | 100  | 
by (etac evalc1_E 1);  | 
| 1700 | 101  | 
|
102  | 
(* WhileFalse *)  | 
|
| 4089 | 103  | 
by (fast_tac (claset() addSDs [hlemma]) 1);  | 
| 1700 | 104  | 
|
105  | 
(* WhileTrue *)  | 
|
| 4089 | 106  | 
by (fast_tac(claset() addSDs[lemma2,le_imp_less_or_eq,less_Suc_eq RS iffD2])1);  | 
| 1700 | 107  | 
|
108  | 
qed_spec_mp "evalc1_impl_evalc";  | 
|
109  | 
||
110  | 
||
111  | 
(**** proof of the equivalence of evalc and evalc1 ****)  | 
|
112  | 
||
| 5069 | 113  | 
Goal "((c, s) -*-> (SKIP, t)) = (<c,s> -c-> t)";  | 
| 1700 | 114  | 
by (fast_tac (HOL_cs addSEs [evalc1_impl_evalc,evalc_impl_evalc1]) 1);  | 
115  | 
qed "evalc1_eq_evalc";  | 
|
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
116  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
117  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
118  | 
section "A Proof Without -n->";  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
119  | 
|
| 5278 | 120  | 
Goal "(c1,s1) -*-> (SKIP,s2) ==> \  | 
| 1939 | 121  | 
\ (c2,s2) -*-> cs3 --> (c1;c2,s1) -*-> cs3";  | 
| 4746 | 122  | 
by (etac converse_rtrancl_induct2 1);  | 
| 4089 | 123  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2]) 1);  | 
124  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2]) 1);  | 
|
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
125  | 
qed_spec_mp "my_lemma1";  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
126  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
127  | 
|
| 5117 | 128  | 
Goal "<c,s> -c-> s1 ==> (c,s) -*-> (SKIP,s1)";  | 
| 2031 | 129  | 
by (etac evalc.induct 1);  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
130  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
131  | 
(* SKIP *)  | 
| 2031 | 132  | 
by (rtac rtrancl_refl 1);  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
133  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
134  | 
(* ASSIGN *)  | 
| 4089 | 135  | 
by (fast_tac (claset() addSIs [r_into_rtrancl]) 1);  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
136  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
137  | 
(* SEMI *)  | 
| 4089 | 138  | 
by (fast_tac (claset() addIs [my_lemma1]) 1);  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
139  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
140  | 
(* IF *)  | 
| 4089 | 141  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2]) 1);  | 
142  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2]) 1);  | 
|
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
143  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
144  | 
(* WHILE *)  | 
| 4089 | 145  | 
by (fast_tac (claset() addSIs [r_into_rtrancl]) 1);  | 
146  | 
by (fast_tac (claset() addIs [rtrancl_into_rtrancl2,my_lemma1]) 1);  | 
|
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
147  | 
|
| 1730 | 148  | 
qed "evalc_impl_evalc1";  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
149  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
150  | 
(* The opposite direction is based on a Coq proof done by Ranan Fraer and  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
151  | 
Yves Bertot. The following sketch is from an email by Ranan Fraer.  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
152  | 
*)  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
153  | 
(*  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
154  | 
First we've broke it into 2 lemmas:  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
155  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
156  | 
Lemma 1  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
157  | 
((c,s) --> (SKIP,t)) => (<c,s> -c-> t)  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
158  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
159  | 
This is a quick one, dealing with the cases skip, assignment  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
160  | 
and while_false.  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
161  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
162  | 
Lemma 2  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
163  | 
((c,s) -*-> (c',s')) /\ <c',s'> -c'-> t  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
164  | 
=>  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
165  | 
<c,s> -c-> t  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
166  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
167  | 
This is proved by rule induction on the -*-> relation  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
168  | 
and the induction step makes use of a third lemma:  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
169  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
170  | 
Lemma 3  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
171  | 
((c,s) --> (c',s')) /\ <c',s'> -c'-> t  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
172  | 
=>  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
173  | 
<c,s> -c-> t  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
174  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
175  | 
This captures the essence of the proof, as it shows that <c',s'>  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
176  | 
behaves as the continuation of <c,s> with respect to the natural  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
177  | 
semantics.  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
178  | 
The proof of Lemma 3 goes by rule induction on the --> relation,  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
179  | 
dealing with the cases sequence1, sequence2, if_true, if_false and  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
180  | 
while_true. In particular in the case (sequence1) we make use again  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
181  | 
of Lemma 1.  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
182  | 
*)  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
183  | 
|
| 4897 | 184  | 
(*Delsimps [update_apply];*)  | 
| 5278 | 185  | 
Goal "((c,s) -1-> (c',s')) ==> (!t. <c',s'> -c-> t --> <c,s> -c-> t)";  | 
| 2031 | 186  | 
by (etac evalc1.induct 1);  | 
| 
4477
 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 
paulson 
parents: 
4153 
diff
changeset
 | 
187  | 
by Auto_tac;  | 
| 1730 | 188  | 
qed_spec_mp "FB_lemma3";  | 
| 4897 | 189  | 
(*Addsimps [update_apply];*)  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
190  | 
|
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
191  | 
val [major] = goal Transition.thy  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
192  | 
"(c,s) -*-> (c',s') ==> <c',s'> -c-> t --> <c,s> -c-> t";  | 
| 2031 | 193  | 
by (rtac (major RS rtrancl_induct2) 1);  | 
| 4897 | 194  | 
by (Fast_tac 1);  | 
| 4651 | 195  | 
by (fast_tac (claset() addIs [FB_lemma3]) 1);  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
196  | 
qed_spec_mp "FB_lemma2";  | 
| 
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
197  | 
|
| 5117 | 198  | 
Goal "(c,s) -*-> (SKIP,t) ==> <c,s> -c-> t";  | 
| 4089 | 199  | 
by (fast_tac (claset() addEs [FB_lemma2]) 1);  | 
| 
1707
 
e1a64a6c454d
Added an equivalence proof which avoids the use of -n->
 
nipkow 
parents: 
1701 
diff
changeset
 | 
200  | 
qed "evalc1_impl_evalc";  | 
| 8016 | 201  | 
|
202  | 
||
203  | 
section "The proof in Nielson and Nielson";  | 
|
204  | 
||
205  | 
(* The more precise n=i1+i2+1 is proved by the same script but complicates  | 
|
206  | 
life further down, where i1,i2 < n is needed.  | 
|
207  | 
*)  | 
|
208  | 
Goal "!c1 s. (c1;c2,s) -n-> (SKIP,t) --> \  | 
|
209  | 
\ (? i1 i2 u. (c1,s) -i1-> (SKIP,u) & (c2,u) -i2-> (SKIP,t) & i1<n & i2<n)";  | 
|
| 8064 | 210  | 
by (induct_tac "n" 1);  | 
| 
9556
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
211  | 
by (Fast_tac 1);  | 
| 8064 | 212  | 
by (fast_tac (claset() addSIs [rel_pow_0_I,rel_pow_Suc_I2]  | 
| 8016 | 213  | 
addSDs [rel_pow_Suc_D2] addss simpset()) 1);  | 
214  | 
qed_spec_mp "comp_decomp_lemma";  | 
|
215  | 
||
216  | 
Goal "!c s t. (c,s) -n-> (SKIP,t) --> <c,s> -c-> t";  | 
|
| 9870 | 217  | 
by (induct_thm_tac nat_less_induct "n" 1);  | 
| 8064 | 218  | 
by (Clarify_tac 1);  | 
219  | 
by (etac rel_pow_E2 1);  | 
|
| 8016 | 220  | 
by (asm_full_simp_tac (simpset() addsimps evalc.intrs) 1);  | 
| 
8442
 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 
wenzelm 
parents: 
8423 
diff
changeset
 | 
221  | 
by (case_tac "c" 1);  | 
| 
9556
 
dcdcfb0545e0
moved Hoare_example to Examples; other minor improvements
 
oheimb 
parents: 
9241 
diff
changeset
 | 
222  | 
by (Fast_tac 1);  | 
| 8064 | 223  | 
by (Blast_tac 1);  | 
224  | 
by (blast_tac (claset() addSDs [rel_pow_Suc_I2 RS comp_decomp_lemma]) 1);  | 
|
225  | 
by (Blast_tac 1);  | 
|
226  | 
by (Blast_tac 1);  | 
|
| 8016 | 227  | 
qed_spec_mp "evalc1_impl_evalc";  |