| author | regensbu | 
| Thu, 29 Jun 1995 16:28:40 +0200 | |
| changeset 1168 | 74be52691d62 | 
| parent 1043 | ffa68eb2730b | 
| child 1267 | bca91b4e1710 | 
| permissions | -rw-r--r-- | 
| 244 | 1  | 
(* Title: HOLCF/ex/loop.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Franz Regensburger  | 
|
4  | 
Copyright 1993 Technische Universitaet Muenchen  | 
|
5  | 
||
6  | 
Lemmas for theory loop.thy  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
open Loop;  | 
|
10  | 
||
11  | 
(* --------------------------------------------------------------------------- *)  | 
|
12  | 
(* access to definitions *)  | 
|
13  | 
(* --------------------------------------------------------------------------- *)  | 
|
14  | 
||
| 1043 | 15  | 
val step_def2 = prove_goalw Loop.thy [step_def]  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
16  | 
"step`b`g`x = If b`x then g`x else x fi"  | 
| 244 | 17  | 
(fn prems =>  | 
18  | 
[  | 
|
19  | 
(simp_tac Cfun_ss 1)  | 
|
20  | 
]);  | 
|
21  | 
||
| 1043 | 22  | 
val while_def2 = prove_goalw Loop.thy [while_def]  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
23  | 
"while`b`g = fix`(LAM f x. If b`x then f`(g`x) else x fi)"  | 
| 244 | 24  | 
(fn prems =>  | 
25  | 
[  | 
|
26  | 
(simp_tac Cfun_ss 1)  | 
|
27  | 
]);  | 
|
28  | 
||
29  | 
||
30  | 
(* ------------------------------------------------------------------------- *)  | 
|
31  | 
(* rekursive properties of while *)  | 
|
32  | 
(* ------------------------------------------------------------------------- *)  | 
|
33  | 
||
| 1043 | 34  | 
val while_unfold = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
35  | 
"while`b`g`x = If b`x then while`b`g`(g`x) else x fi"  | 
| 244 | 36  | 
(fn prems =>  | 
37  | 
[  | 
|
38  | 
(fix_tac5 while_def2 1),  | 
|
39  | 
(simp_tac Cfun_ss 1)  | 
|
40  | 
]);  | 
|
41  | 
||
| 1043 | 42  | 
val while_unfold2 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
43  | 
"!x.while`b`g`x = while`b`g`(iterate k (step`b`g) x)"  | 
| 244 | 44  | 
(fn prems =>  | 
45  | 
[  | 
|
46  | 
(nat_ind_tac "k" 1),  | 
|
47  | 
(simp_tac (HOLCF_ss addsimps [iterate_0,iterate_Suc]) 1),  | 
|
48  | 
(rtac allI 1),  | 
|
49  | 
(rtac trans 1),  | 
|
50  | 
(rtac (while_unfold RS ssubst) 1),  | 
|
51  | 
(rtac refl 2),  | 
|
52  | 
(rtac (iterate_Suc2 RS ssubst) 1),  | 
|
53  | 
(rtac trans 1),  | 
|
54  | 
(etac spec 2),  | 
|
55  | 
(rtac (step_def2 RS ssubst) 1),  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
56  | 
	(res_inst_tac [("p","b`x")] trE 1),
 | 
| 244 | 57  | 
(asm_simp_tac HOLCF_ss 1),  | 
58  | 
(rtac (while_unfold RS ssubst) 1),  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
59  | 
	(res_inst_tac [("s","UU"),("t","b`UU")]ssubst 1),
 | 
| 244 | 60  | 
(etac (flat_tr RS flat_codom RS disjE) 1),  | 
61  | 
(atac 1),  | 
|
62  | 
(etac spec 1),  | 
|
63  | 
(simp_tac HOLCF_ss 1),  | 
|
64  | 
(asm_simp_tac HOLCF_ss 1),  | 
|
65  | 
(asm_simp_tac HOLCF_ss 1),  | 
|
66  | 
(rtac (while_unfold RS ssubst) 1),  | 
|
67  | 
(asm_simp_tac HOLCF_ss 1)  | 
|
68  | 
]);  | 
|
69  | 
||
| 1043 | 70  | 
val while_unfold3 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
71  | 
"while`b`g`x = while`b`g`(step`b`g`x)"  | 
| 244 | 72  | 
(fn prems =>  | 
73  | 
[  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
74  | 
	(res_inst_tac [("s",
 | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
75  | 
"while`b`g`(iterate (Suc 0) (step`b`g) x)")] trans 1),  | 
| 244 | 76  | 
(rtac (while_unfold2 RS spec) 1),  | 
77  | 
(simp_tac iterate_ss 1)  | 
|
78  | 
]);  | 
|
79  | 
||
80  | 
||
81  | 
(* --------------------------------------------------------------------------- *)  | 
|
82  | 
(* properties of while and iterations *)  | 
|
83  | 
(* --------------------------------------------------------------------------- *)  | 
|
84  | 
||
| 1043 | 85  | 
val loop_lemma1 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
86  | 
"[|? y.b`y=FF; iterate k (step`b`g) x = UU|]==>iterate(Suc k) (step`b`g) x=UU"  | 
| 244 | 87  | 
(fn prems =>  | 
88  | 
[  | 
|
89  | 
(cut_facts_tac prems 1),  | 
|
90  | 
(simp_tac iterate_ss 1),  | 
|
91  | 
(rtac trans 1),  | 
|
92  | 
(rtac step_def2 1),  | 
|
93  | 
(asm_simp_tac HOLCF_ss 1),  | 
|
94  | 
(etac exE 1),  | 
|
95  | 
(etac (flat_tr RS flat_codom RS disjE) 1),  | 
|
96  | 
(asm_simp_tac HOLCF_ss 1),  | 
|
97  | 
(asm_simp_tac HOLCF_ss 1)  | 
|
98  | 
]);  | 
|
99  | 
||
| 1043 | 100  | 
val loop_lemma2 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
101  | 
"[|? y.b`y=FF;iterate (Suc k) (step`b`g) x ~=UU |]==>\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
102  | 
\iterate k (step`b`g) x ~=UU"  | 
| 244 | 103  | 
(fn prems =>  | 
104  | 
[  | 
|
105  | 
(cut_facts_tac prems 1),  | 
|
106  | 
(rtac contrapos 1),  | 
|
107  | 
(etac loop_lemma1 2),  | 
|
108  | 
(atac 1),  | 
|
109  | 
(atac 1)  | 
|
110  | 
]);  | 
|
111  | 
||
| 1043 | 112  | 
val loop_lemma3 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
113  | 
"[|!x. INV x & b`x=TT & g`x~=UU --> INV (g`x);\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
114  | 
\? y.b`y=FF; INV x|] ==>\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
115  | 
\iterate k (step`b`g) x ~=UU --> INV (iterate k (step`b`g) x)"  | 
| 244 | 116  | 
(fn prems =>  | 
117  | 
[  | 
|
118  | 
(cut_facts_tac prems 1),  | 
|
119  | 
(nat_ind_tac "k" 1),  | 
|
120  | 
(asm_simp_tac iterate_ss 1),  | 
|
121  | 
(strip_tac 1),  | 
|
122  | 
(simp_tac (iterate_ss addsimps [step_def2]) 1),  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
123  | 
	(res_inst_tac [("p","b`(iterate k1 (step`b`g) x)")] trE 1),
 | 
| 244 | 124  | 
(etac notE 1),  | 
125  | 
(asm_simp_tac (HOLCF_ss addsimps [step_def2,iterate_Suc] ) 1),  | 
|
126  | 
(asm_simp_tac HOLCF_ss 1),  | 
|
127  | 
(rtac mp 1),  | 
|
128  | 
(etac spec 1),  | 
|
129  | 
(asm_simp_tac (HOLCF_ss addsimps [loop_lemma2] ) 1),  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
130  | 
	(res_inst_tac [("s","iterate (Suc k1) (step`b`g) x"),
 | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
131  | 
		("t","g`(iterate k1 (step`b`g) x)")] ssubst 1),
 | 
| 244 | 132  | 
(atac 2),  | 
133  | 
(asm_simp_tac (HOLCF_ss addsimps [iterate_Suc,step_def2] ) 1),  | 
|
134  | 
(asm_simp_tac (HOLCF_ss addsimps [loop_lemma2] ) 1)  | 
|
135  | 
]);  | 
|
136  | 
||
137  | 
||
| 1043 | 138  | 
val loop_lemma4 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
139  | 
"!x. b`(iterate k (step`b`g) x)=FF --> while`b`g`x= iterate k (step`b`g) x"  | 
| 244 | 140  | 
(fn prems =>  | 
141  | 
[  | 
|
142  | 
(nat_ind_tac "k" 1),  | 
|
143  | 
(simp_tac iterate_ss 1),  | 
|
144  | 
(strip_tac 1),  | 
|
145  | 
(rtac (while_unfold RS ssubst) 1),  | 
|
146  | 
(asm_simp_tac HOLCF_ss 1),  | 
|
147  | 
(rtac allI 1),  | 
|
148  | 
(rtac (iterate_Suc2 RS ssubst) 1),  | 
|
149  | 
(strip_tac 1),  | 
|
150  | 
(rtac trans 1),  | 
|
151  | 
(rtac while_unfold3 1),  | 
|
152  | 
(asm_simp_tac HOLCF_ss 1)  | 
|
153  | 
]);  | 
|
154  | 
||
| 1043 | 155  | 
val loop_lemma5 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
156  | 
"!k. b`(iterate k (step`b`g) x) ~= FF ==>\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
157  | 
\ !m. while`b`g`(iterate m (step`b`g) x)=UU"  | 
| 244 | 158  | 
(fn prems =>  | 
159  | 
[  | 
|
160  | 
(cut_facts_tac prems 1),  | 
|
161  | 
(rtac (while_def2 RS ssubst) 1),  | 
|
162  | 
(rtac fix_ind 1),  | 
|
163  | 
(rtac (allI RS adm_all) 1),  | 
|
164  | 
(rtac adm_eq 1),  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
165  | 
(cont_tacR 1),  | 
| 244 | 166  | 
(simp_tac HOLCF_ss 1),  | 
167  | 
(rtac allI 1),  | 
|
168  | 
(simp_tac HOLCF_ss 1),  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
169  | 
	(res_inst_tac [("p","b`(iterate m (step`b`g) x)")] trE 1),
 | 
| 244 | 170  | 
(asm_simp_tac HOLCF_ss 1),  | 
171  | 
(asm_simp_tac HOLCF_ss 1),  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
172  | 
	(res_inst_tac [("s","xa`(iterate (Suc m) (step`b`g) x)")] trans 1),
 | 
| 244 | 173  | 
(etac spec 2),  | 
174  | 
(rtac cfun_arg_cong 1),  | 
|
175  | 
(rtac trans 1),  | 
|
176  | 
(rtac (iterate_Suc RS sym) 2),  | 
|
177  | 
(asm_simp_tac (HOLCF_ss addsimps [step_def2]) 1),  | 
|
178  | 
(dtac spec 1),  | 
|
179  | 
(contr_tac 1)  | 
|
180  | 
]);  | 
|
181  | 
||
182  | 
||
| 1043 | 183  | 
val loop_lemma6 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
184  | 
"!k. b`(iterate k (step`b`g) x) ~= FF ==> while`b`g`x=UU"  | 
| 244 | 185  | 
(fn prems =>  | 
186  | 
[  | 
|
187  | 
	(res_inst_tac [("t","x")] (iterate_0 RS subst) 1),
 | 
|
188  | 
(rtac (loop_lemma5 RS spec) 1),  | 
|
189  | 
(resolve_tac prems 1)  | 
|
190  | 
]);  | 
|
191  | 
||
| 1043 | 192  | 
val loop_lemma7 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
193  | 
"while`b`g`x ~= UU ==> ? k. b`(iterate k (step`b`g) x) = FF"  | 
| 244 | 194  | 
(fn prems =>  | 
195  | 
[  | 
|
196  | 
(cut_facts_tac prems 1),  | 
|
197  | 
(etac swap 1),  | 
|
198  | 
(rtac loop_lemma6 1),  | 
|
199  | 
(fast_tac HOL_cs 1)  | 
|
200  | 
]);  | 
|
201  | 
||
| 1043 | 202  | 
val loop_lemma8 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
203  | 
"while`b`g`x ~= UU ==> ? y. b`y=FF"  | 
| 244 | 204  | 
(fn prems =>  | 
205  | 
[  | 
|
206  | 
(cut_facts_tac prems 1),  | 
|
207  | 
(dtac loop_lemma7 1),  | 
|
208  | 
(fast_tac HOL_cs 1)  | 
|
209  | 
]);  | 
|
210  | 
||
211  | 
||
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
212  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
213  | 
(* an invariant rule for loops *)  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
214  | 
(* ------------------------------------------------------------------------- *)  | 
| 244 | 215  | 
|
| 1043 | 216  | 
val loop_inv2 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
217  | 
"[| (!y. INV y & b`y=TT & g`y ~= UU --> INV (g`y));\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
218  | 
\ (!y. INV y & b`y=FF --> Q y);\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
219  | 
\ INV x; while`b`g`x~=UU |] ==> Q (while`b`g`x)"  | 
| 244 | 220  | 
(fn prems =>  | 
221  | 
[  | 
|
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
222  | 
	(res_inst_tac [("P","%k. b`(iterate k (step`b`g) x)=FF")] exE 1),
 | 
| 244 | 223  | 
(rtac loop_lemma7 1),  | 
224  | 
(resolve_tac prems 1),  | 
|
225  | 
(rtac ((loop_lemma4 RS spec RS mp) RS ssubst) 1),  | 
|
226  | 
(atac 1),  | 
|
227  | 
(rtac (nth_elem (1,prems) RS spec RS mp) 1),  | 
|
228  | 
(rtac conjI 1),  | 
|
229  | 
(atac 2),  | 
|
230  | 
(rtac (loop_lemma3 RS mp) 1),  | 
|
231  | 
(resolve_tac prems 1),  | 
|
232  | 
(rtac loop_lemma8 1),  | 
|
233  | 
(resolve_tac prems 1),  | 
|
234  | 
(resolve_tac prems 1),  | 
|
235  | 
(rtac classical3 1),  | 
|
236  | 
(resolve_tac prems 1),  | 
|
237  | 
(etac box_equals 1),  | 
|
238  | 
(rtac (loop_lemma4 RS spec RS mp RS sym) 1),  | 
|
239  | 
(atac 1),  | 
|
240  | 
(rtac refl 1)  | 
|
241  | 
]);  | 
|
242  | 
||
| 1043 | 243  | 
val loop_inv3 = prove_goal Loop.thy  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
244  | 
"[| !!y.[| INV y; b`y=TT; g`y~=UU|] ==> INV (g`y);\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
245  | 
\ !!y.[| INV y; b`y=FF|]==> Q y;\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
246  | 
\ INV x; while`b`g`x~=UU |] ==> Q (while`b`g`x)"  | 
| 244 | 247  | 
(fn prems =>  | 
248  | 
[  | 
|
249  | 
(rtac loop_inv2 1),  | 
|
250  | 
(rtac allI 1),  | 
|
251  | 
(rtac impI 1),  | 
|
252  | 
(resolve_tac prems 1),  | 
|
253  | 
(fast_tac HOL_cs 1),  | 
|
254  | 
(fast_tac HOL_cs 1),  | 
|
255  | 
(fast_tac HOL_cs 1),  | 
|
256  | 
(rtac allI 1),  | 
|
257  | 
(rtac impI 1),  | 
|
258  | 
(resolve_tac prems 1),  | 
|
259  | 
(fast_tac HOL_cs 1),  | 
|
260  | 
(fast_tac HOL_cs 1),  | 
|
261  | 
(resolve_tac prems 1),  | 
|
262  | 
(resolve_tac prems 1)  | 
|
263  | 
]);  | 
|
264  | 
||
| 1043 | 265  | 
val loop_inv = prove_goal Loop.thy  | 
| 244 | 266  | 
"[| P(x);\  | 
| 
1168
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
267  | 
\ !!y.P y ==> INV y;\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
268  | 
\ !!y.[| INV y; b`y=TT; g`y~=UU|] ==> INV (g`y);\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
269  | 
\ !!y.[| INV y; b`y=FF|]==> Q y;\  | 
| 
 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 
regensbu 
parents: 
1043 
diff
changeset
 | 
270  | 
\ while`b`g`x ~= UU |] ==> Q (while`b`g`x)"  | 
| 244 | 271  | 
(fn prems =>  | 
272  | 
[  | 
|
273  | 
(rtac loop_inv3 1),  | 
|
274  | 
(eresolve_tac prems 1),  | 
|
275  | 
(atac 1),  | 
|
276  | 
(atac 1),  | 
|
277  | 
(resolve_tac prems 1),  | 
|
278  | 
(atac 1),  | 
|
279  | 
(atac 1),  | 
|
280  | 
(resolve_tac prems 1),  | 
|
281  | 
(resolve_tac prems 1),  | 
|
282  | 
(resolve_tac prems 1)  | 
|
283  | 
]);  |