author | wenzelm |
Sat, 27 Oct 2001 00:00:05 +0200 | |
changeset 11954 | 3d1780208bf3 |
parent 10797 | 028d22926a41 |
permissions | -rw-r--r-- |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
1 |
(* Title: HOL/Induct/Exp |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
4 |
Copyright 1997 University of Cambridge |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
5 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
6 |
Example of Mutual Induction via Iteratived Inductive Definitions: Expressions |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
7 |
*) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
8 |
|
3145 | 9 |
AddSIs [eval.N, eval.X]; |
10 |
AddIs [eval.Op, eval.valOf]; |
|
3144 | 11 |
|
6141 | 12 |
val eval_elim_cases = |
13 |
map eval.mk_cases |
|
14 |
["(N(n),sigma) -|-> (n',s')", |
|
15 |
"(X(x),sigma) -|-> (n,s')", |
|
16 |
"(Op f a1 a2,sigma) -|-> (n,s')", |
|
17 |
"(VALOF c RESULTIS e, s) -|-> (n, s1)"]; |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
18 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
19 |
AddSEs eval_elim_cases; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
20 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
21 |
|
10759 | 22 |
Goal "(X x, s(x:=n)) -|-> (n, s(x:=n))"; |
23 |
by (rtac (fun_upd_same RS subst) 1 THEN resolve_tac eval.intrs 1); |
|
3145 | 24 |
qed "var_assign_eval"; |
25 |
||
26 |
AddSIs [var_assign_eval]; |
|
27 |
||
28 |
||
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
29 |
(** Make the induction rule look nicer -- though eta_contract makes the new |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
30 |
version look worse than it is...**) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
31 |
|
5069 | 32 |
Goal "{((e,s),(n,s')). P e s n s'} = \ |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
33 |
\ Collect (split (%v. split (split P v)))"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
34 |
by (rtac Collect_cong 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
35 |
by (split_all_tac 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
36 |
by (Simp_tac 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
37 |
val split_lemma = result(); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
38 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
39 |
(*New induction rule. Note the form of the VALOF induction hypothesis*) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
40 |
val major::prems = goal thy |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
41 |
"[| (e,s) -|-> (n,s'); \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
42 |
\ !!n s. P (N n) s n s; \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
43 |
\ !!s x. P (X x) s (s x) s; \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
44 |
\ !!e0 e1 f n0 n1 s s0 s1. \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
45 |
\ [| (e0,s) -|-> (n0,s0); P e0 s n0 s0; \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
46 |
\ (e1,s0) -|-> (n1,s1); P e1 s0 n1 s1 \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
47 |
\ |] ==> P (Op f e0 e1) s (f n0 n1) s1; \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
48 |
\ !!c e n s s0 s1. \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
49 |
\ [| (c,s) -[eval Int {((e,s),(n,s')). P e s n s'}]-> s0; \ |
3145 | 50 |
\ (c,s) -[eval]-> s0; \ |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
51 |
\ (e,s0) -|-> (n,s1); P e s0 n s1 |] \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
52 |
\ ==> P (VALOF c RESULTIS e) s n s1 \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
53 |
\ |] ==> P e s n s'"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
54 |
by (rtac (major RS eval.induct) 1); |
4089 | 55 |
by (blast_tac (claset() addIs prems) 1); |
56 |
by (blast_tac (claset() addIs prems) 1); |
|
57 |
by (blast_tac (claset() addIs prems) 1); |
|
3145 | 58 |
by (forward_tac [impOfSubs (Int_lower1 RS exec_mono)] 1); |
4264 | 59 |
by (auto_tac (claset() addIs prems,simpset() addsimps [split_lemma])); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
60 |
qed "eval_induct"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
61 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
62 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
63 |
(*Lemma for Function_eval. The major premise is that (c,s) executes to s1 |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
64 |
using eval restricted to its functional part. Note that the execution |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
65 |
(c,s) -[eval]-> s2 can use unrestricted eval! The reason is that |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
66 |
the execution (c,s) -[eval Int {...}]-> s1 assures us that execution is |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
67 |
functional on the argument (c,s). |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
68 |
*) |
10759 | 69 |
Goal "(c,s) -[eval Int {((e,s),(n,t)). ALL nt'. (e,s) -|-> nt' --> (n,t)=nt'}]-> s1 \ |
5148
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
70 |
\ ==> (ALL s2. (c,s) -[eval]-> s2 --> s2=s1)"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
71 |
by (etac exec.induct 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
72 |
by (ALLGOALS Full_simp_tac); |
10759 | 73 |
by (Blast_tac 1); |
74 |
by (Force_tac 1); |
|
75 |
by (Blast_tac 1); |
|
76 |
by (Blast_tac 1); |
|
77 |
by (Blast_tac 1); |
|
78 |
by (blast_tac (claset() addEs [exec_WHILE_case]) 1); |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
79 |
by (thin_tac "(?c,s2) -[?ev]-> s3" 1); |
3718 | 80 |
by (Clarify_tac 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
81 |
by (etac exec_WHILE_case 1); |
10759 | 82 |
by (ALLGOALS Blast_tac); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
83 |
qed "com_Unique"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
84 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
85 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
86 |
(*Expression evaluation is functional, or deterministic*) |
10797 | 87 |
Goalw [single_valued_def] "single_valued eval"; |
10759 | 88 |
by (EVERY1[rtac allI, rtac allI, rtac impI]); |
4264 | 89 |
by (split_all_tac 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
90 |
by (etac eval_induct 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
91 |
by (dtac com_Unique 4); |
10759 | 92 |
by (ALLGOALS Full_simp_tac); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
93 |
by (ALLGOALS Blast_tac); |
10797 | 94 |
qed "single_valued_eval"; |
3144 | 95 |
|
96 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
97 |
Goal "(e,s) -|-> (v,s') ==> (e = N n) --> (v=n & s'=s)"; |
3145 | 98 |
by (etac eval_induct 1); |
3144 | 99 |
by (ALLGOALS Asm_simp_tac); |
100 |
val lemma = result(); |
|
101 |
bind_thm ("eval_N_E", refl RSN (2, lemma RS mp)); |
|
102 |
||
103 |
AddSEs [eval_N_E]; |
|
104 |
||
105 |
||
106 |
(*This theorem says that "WHILE TRUE DO c" cannot terminate*) |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
107 |
Goal "(c', s) -[eval]-> t ==> (c' = WHILE (N 0) DO c) --> False"; |
3144 | 108 |
by (etac exec.induct 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4264
diff
changeset
|
109 |
by Auto_tac; |
3144 | 110 |
bind_thm ("while_true_E", refl RSN (2, result() RS mp)); |
111 |
||
112 |
||
113 |
(** Equivalence of IF e THEN c;;(WHILE e DO c) ELSE SKIP and WHILE e DO c **) |
|
114 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
115 |
Goal "(c',s) -[eval]-> t ==> \ |
3144 | 116 |
\ (c' = WHILE e DO c) --> \ |
117 |
\ (IF e THEN c;;c' ELSE SKIP, s) -[eval]-> t"; |
|
118 |
by (etac exec.induct 1); |
|
3145 | 119 |
by (ALLGOALS Asm_simp_tac); |
3144 | 120 |
by (ALLGOALS Blast_tac); |
121 |
bind_thm ("while_if1", refl RSN (2, result() RS mp)); |
|
122 |
||
123 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
124 |
Goal "(c',s) -[eval]-> t ==> \ |
3144 | 125 |
\ (c' = IF e THEN c;;(WHILE e DO c) ELSE SKIP) --> \ |
126 |
\ (WHILE e DO c, s) -[eval]-> t"; |
|
127 |
by (etac exec.induct 1); |
|
3145 | 128 |
by (ALLGOALS Asm_simp_tac); |
3144 | 129 |
by (ALLGOALS Blast_tac); |
130 |
bind_thm ("while_if2", refl RSN (2, result() RS mp)); |
|
131 |
||
132 |
||
5069 | 133 |
Goal "((IF e THEN c;;(WHILE e DO c) ELSE SKIP, s) -[eval]-> t) = \ |
3144 | 134 |
\ ((WHILE e DO c, s) -[eval]-> t)"; |
4089 | 135 |
by (blast_tac (claset() addIs [while_if1, while_if2]) 1); |
3144 | 136 |
qed "while_if"; |
137 |
||
138 |
||
139 |
||
3145 | 140 |
(** Equivalence of (IF e THEN c1 ELSE c2);;c |
141 |
and IF e THEN (c1;;c) ELSE (c2;;c) **) |
|
142 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
143 |
Goal "(c',s) -[eval]-> t ==> \ |
3145 | 144 |
\ (c' = (IF e THEN c1 ELSE c2);;c) --> \ |
145 |
\ (IF e THEN (c1;;c) ELSE (c2;;c), s) -[eval]-> t"; |
|
146 |
by (etac exec.induct 1); |
|
147 |
by (ALLGOALS Asm_simp_tac); |
|
148 |
by (Blast_tac 1); |
|
149 |
bind_thm ("if_semi1", refl RSN (2, result() RS mp)); |
|
150 |
||
151 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
152 |
Goal "(c',s) -[eval]-> t ==> \ |
3145 | 153 |
\ (c' = IF e THEN (c1;;c) ELSE (c2;;c)) --> \ |
154 |
\ ((IF e THEN c1 ELSE c2);;c, s) -[eval]-> t"; |
|
155 |
by (etac exec.induct 1); |
|
156 |
by (ALLGOALS Asm_simp_tac); |
|
157 |
by (ALLGOALS Blast_tac); |
|
158 |
bind_thm ("if_semi2", refl RSN (2, result() RS mp)); |
|
159 |
||
160 |
||
5069 | 161 |
Goal "(((IF e THEN c1 ELSE c2);;c, s) -[eval]-> t) = \ |
3145 | 162 |
\ ((IF e THEN (c1;;c) ELSE (c2;;c), s) -[eval]-> t)"; |
4089 | 163 |
by (blast_tac (claset() addIs [if_semi1, if_semi2]) 1); |
3145 | 164 |
qed "if_semi"; |
165 |
||
166 |
||
167 |
||
3144 | 168 |
(** Equivalence of VALOF c1 RESULTIS (VALOF c2 RESULTIS e) |
169 |
and VALOF c1;;c2 RESULTIS e |
|
170 |
**) |
|
171 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
172 |
Goal "(e',s) -|-> (v,s') ==> \ |
3144 | 173 |
\ (e' = VALOF c1 RESULTIS (VALOF c2 RESULTIS e)) --> \ |
174 |
\ (VALOF c1;;c2 RESULTIS e, s) -|-> (v,s')"; |
|
3145 | 175 |
by (etac eval_induct 1); |
176 |
by (ALLGOALS Asm_simp_tac); |
|
177 |
by (Blast_tac 1); |
|
3144 | 178 |
bind_thm ("valof_valof1", refl RSN (2, result() RS mp)); |
179 |
||
180 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
181 |
Goal "(e',s) -|-> (v,s') ==> \ |
3144 | 182 |
\ (e' = VALOF c1;;c2 RESULTIS e) --> \ |
183 |
\ (VALOF c1 RESULTIS (VALOF c2 RESULTIS e), s) -|-> (v,s')"; |
|
3145 | 184 |
by (etac eval_induct 1); |
185 |
by (ALLGOALS Asm_simp_tac); |
|
186 |
by (Blast_tac 1); |
|
3144 | 187 |
bind_thm ("valof_valof2", refl RSN (2, result() RS mp)); |
188 |
||
189 |
||
5069 | 190 |
Goal "((VALOF c1 RESULTIS (VALOF c2 RESULTIS e), s) -|-> (v,s')) = \ |
3144 | 191 |
\ ((VALOF c1;;c2 RESULTIS e, s) -|-> (v,s'))"; |
4089 | 192 |
by (blast_tac (claset() addIs [valof_valof1, valof_valof2]) 1); |
3144 | 193 |
qed "valof_valof"; |
194 |
||
3145 | 195 |
|
196 |
(** Equivalence of VALOF SKIP RESULTIS e and e **) |
|
197 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
198 |
Goal "(e',s) -|-> (v,s') ==> \ |
3145 | 199 |
\ (e' = VALOF SKIP RESULTIS e) --> \ |
200 |
\ (e, s) -|-> (v,s')"; |
|
201 |
by (etac eval_induct 1); |
|
202 |
by (ALLGOALS Asm_simp_tac); |
|
203 |
by (Blast_tac 1); |
|
204 |
bind_thm ("valof_skip1", refl RSN (2, result() RS mp)); |
|
205 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
206 |
Goal "(e,s) -|-> (v,s') ==> (VALOF SKIP RESULTIS e, s) -|-> (v,s')"; |
3145 | 207 |
by (Blast_tac 1); |
208 |
qed "valof_skip2"; |
|
209 |
||
5069 | 210 |
Goal "((VALOF SKIP RESULTIS e, s) -|-> (v,s')) = ((e, s) -|-> (v,s'))"; |
4089 | 211 |
by (blast_tac (claset() addIs [valof_skip1, valof_skip2]) 1); |
3145 | 212 |
qed "valof_skip"; |
213 |
||
214 |
||
215 |
(** Equivalence of VALOF x:=e RESULTIS x and e **) |
|
216 |
||
10759 | 217 |
Delsimps [fun_upd_apply]; |
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
218 |
Goal "(e',s) -|-> (v,s'') ==> \ |
3145 | 219 |
\ (e' = VALOF x:=e RESULTIS X x) --> \ |
10759 | 220 |
\ (EX s'. (e, s) -|-> (v,s') & (s'' = s'(x:=v)))"; |
3145 | 221 |
by (etac eval_induct 1); |
10759 | 222 |
by (ALLGOALS Asm_simp_tac); |
3145 | 223 |
by (thin_tac "?PP-->?QQ" 1); |
3718 | 224 |
by (Clarify_tac 1); |
3145 | 225 |
by (Simp_tac 1); |
226 |
by (Blast_tac 1); |
|
227 |
bind_thm ("valof_assign1", refl RSN (2, result() RS mp)); |
|
228 |
||
229 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
230 |
Goal "(e,s) -|-> (v,s') ==> \ |
10759 | 231 |
\ (VALOF x:=e RESULTIS X x, s) -|-> (v,s'(x:=v))"; |
3145 | 232 |
by (Blast_tac 1); |
233 |
qed "valof_assign2"; |