| author | wenzelm | 
| Tue, 27 Apr 1999 10:50:50 +0200 | |
| changeset 6522 | 2f6cec5c046f | 
| parent 6141 | a6922171b396 | 
| child 10759 | 994877ee68cb | 
| 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 | |
| 5069 | 22 | Goal "(X x, s[n/x]) -|-> (n, s[n/x])"; | 
| 3145 | 23 | by (rtac (assign_same RS subst) 1 THEN resolve_tac eval.intrs 1); | 
| 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 | *) | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 69 | Goal "(c,s) -[eval Int {((e,s),(n,s')). Unique (e,s) (n,s') eval}]-> s1 \
 | 
| 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
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); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 73 | by (Blast_tac 3); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 74 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 75 | by (rewtac Unique_def); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 76 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 77 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 78 | by (Blast_tac 1); | 
| 4089 | 79 | by (blast_tac (claset() addEs [exec_WHILE_case]) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 80 | by (thin_tac "(?c,s2) -[?ev]-> s3" 1); | 
| 3718 | 81 | by (Clarify_tac 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 82 | by (etac exec_WHILE_case 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 83 | by (ALLGOALS Fast_tac); (*Blast_tac: proof fails*) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 84 | qed "com_Unique"; | 
| 
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 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 87 | (*Expression evaluation is functional, or deterministic*) | 
| 5069 | 88 | Goalw [Function_def] "Function eval"; | 
| 4264 | 89 | by (strip_tac 1); | 
| 90 | by (split_all_tac 1); | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 91 | by (etac eval_induct 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 92 | by (dtac com_Unique 4); | 
| 4089 | 93 | by (ALLGOALS (full_simp_tac (simpset() addsimps [Unique_def]))); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 94 | by (ALLGOALS Blast_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 95 | qed "Function_eval"; | 
| 3144 | 96 | |
| 97 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 98 | Goal "(e,s) -|-> (v,s') ==> (e = N n) --> (v=n & s'=s)"; | 
| 3145 | 99 | by (etac eval_induct 1); | 
| 3144 | 100 | by (ALLGOALS Asm_simp_tac); | 
| 101 | val lemma = result(); | |
| 102 | bind_thm ("eval_N_E", refl RSN (2, lemma RS mp));
 | |
| 103 | ||
| 104 | AddSEs [eval_N_E]; | |
| 105 | ||
| 106 | ||
| 107 | (*This theorem says that "WHILE TRUE DO c" cannot terminate*) | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 108 | Goal "(c', s) -[eval]-> t ==> (c' = WHILE (N 0) DO c) --> False"; | 
| 3144 | 109 | by (etac exec.induct 1); | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4264diff
changeset | 110 | by Auto_tac; | 
| 3144 | 111 | bind_thm ("while_true_E", refl RSN (2, result() RS mp));
 | 
| 112 | ||
| 113 | ||
| 114 | (** Equivalence of IF e THEN c;;(WHILE e DO c) ELSE SKIP and WHILE e DO c **) | |
| 115 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 116 | Goal "(c',s) -[eval]-> t ==> \ | 
| 3144 | 117 | \ (c' = WHILE e DO c) --> \ | 
| 118 | \ (IF e THEN c;;c' ELSE SKIP, s) -[eval]-> t"; | |
| 119 | by (etac exec.induct 1); | |
| 3145 | 120 | by (ALLGOALS Asm_simp_tac); | 
| 3144 | 121 | by (ALLGOALS Blast_tac); | 
| 122 | bind_thm ("while_if1", refl RSN (2, result() RS mp));
 | |
| 123 | ||
| 124 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 125 | Goal "(c',s) -[eval]-> t ==> \ | 
| 3144 | 126 | \ (c' = IF e THEN c;;(WHILE e DO c) ELSE SKIP) --> \ | 
| 127 | \ (WHILE e DO c, s) -[eval]-> t"; | |
| 128 | by (etac exec.induct 1); | |
| 3145 | 129 | by (ALLGOALS Asm_simp_tac); | 
| 3144 | 130 | by (ALLGOALS Blast_tac); | 
| 131 | bind_thm ("while_if2", refl RSN (2, result() RS mp));
 | |
| 132 | ||
| 133 | ||
| 5069 | 134 | Goal "((IF e THEN c;;(WHILE e DO c) ELSE SKIP, s) -[eval]-> t) = \ | 
| 3144 | 135 | \ ((WHILE e DO c, s) -[eval]-> t)"; | 
| 4089 | 136 | by (blast_tac (claset() addIs [while_if1, while_if2]) 1); | 
| 3144 | 137 | qed "while_if"; | 
| 138 | ||
| 139 | ||
| 140 | ||
| 3145 | 141 | (** Equivalence of (IF e THEN c1 ELSE c2);;c | 
| 142 | and IF e THEN (c1;;c) ELSE (c2;;c) **) | |
| 143 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 144 | Goal "(c',s) -[eval]-> t ==> \ | 
| 3145 | 145 | \ (c' = (IF e THEN c1 ELSE c2);;c) --> \ | 
| 146 | \ (IF e THEN (c1;;c) ELSE (c2;;c), s) -[eval]-> t"; | |
| 147 | by (etac exec.induct 1); | |
| 148 | by (ALLGOALS Asm_simp_tac); | |
| 149 | by (Blast_tac 1); | |
| 150 | bind_thm ("if_semi1", refl RSN (2, result() RS mp));
 | |
| 151 | ||
| 152 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 153 | Goal "(c',s) -[eval]-> t ==> \ | 
| 3145 | 154 | \ (c' = IF e THEN (c1;;c) ELSE (c2;;c)) --> \ | 
| 155 | \ ((IF e THEN c1 ELSE c2);;c, s) -[eval]-> t"; | |
| 156 | by (etac exec.induct 1); | |
| 157 | by (ALLGOALS Asm_simp_tac); | |
| 158 | by (ALLGOALS Blast_tac); | |
| 159 | bind_thm ("if_semi2", refl RSN (2, result() RS mp));
 | |
| 160 | ||
| 161 | ||
| 5069 | 162 | Goal "(((IF e THEN c1 ELSE c2);;c, s) -[eval]-> t) = \ | 
| 3145 | 163 | \ ((IF e THEN (c1;;c) ELSE (c2;;c), s) -[eval]-> t)"; | 
| 4089 | 164 | by (blast_tac (claset() addIs [if_semi1, if_semi2]) 1); | 
| 3145 | 165 | qed "if_semi"; | 
| 166 | ||
| 167 | ||
| 168 | ||
| 3144 | 169 | (** Equivalence of VALOF c1 RESULTIS (VALOF c2 RESULTIS e) | 
| 170 | and VALOF c1;;c2 RESULTIS e | |
| 171 | **) | |
| 172 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 173 | Goal "(e',s) -|-> (v,s') ==> \ | 
| 3144 | 174 | \ (e' = VALOF c1 RESULTIS (VALOF c2 RESULTIS e)) --> \ | 
| 175 | \ (VALOF c1;;c2 RESULTIS e, s) -|-> (v,s')"; | |
| 3145 | 176 | by (etac eval_induct 1); | 
| 177 | by (ALLGOALS Asm_simp_tac); | |
| 178 | by (Blast_tac 1); | |
| 3144 | 179 | bind_thm ("valof_valof1", refl RSN (2, result() RS mp));
 | 
| 180 | ||
| 181 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 182 | Goal "(e',s) -|-> (v,s') ==> \ | 
| 3144 | 183 | \ (e' = VALOF c1;;c2 RESULTIS e) --> \ | 
| 184 | \ (VALOF c1 RESULTIS (VALOF c2 RESULTIS e), s) -|-> (v,s')"; | |
| 3145 | 185 | by (etac eval_induct 1); | 
| 186 | by (ALLGOALS Asm_simp_tac); | |
| 187 | by (Blast_tac 1); | |
| 3144 | 188 | bind_thm ("valof_valof2", refl RSN (2, result() RS mp));
 | 
| 189 | ||
| 190 | ||
| 5069 | 191 | Goal "((VALOF c1 RESULTIS (VALOF c2 RESULTIS e), s) -|-> (v,s')) = \ | 
| 3144 | 192 | \ ((VALOF c1;;c2 RESULTIS e, s) -|-> (v,s'))"; | 
| 4089 | 193 | by (blast_tac (claset() addIs [valof_valof1, valof_valof2]) 1); | 
| 3144 | 194 | qed "valof_valof"; | 
| 195 | ||
| 3145 | 196 | |
| 197 | (** Equivalence of VALOF SKIP RESULTIS e and e **) | |
| 198 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 199 | Goal "(e',s) -|-> (v,s') ==> \ | 
| 3145 | 200 | \ (e' = VALOF SKIP RESULTIS e) --> \ | 
| 201 | \ (e, s) -|-> (v,s')"; | |
| 202 | by (etac eval_induct 1); | |
| 203 | by (ALLGOALS Asm_simp_tac); | |
| 204 | by (Blast_tac 1); | |
| 205 | bind_thm ("valof_skip1", refl RSN (2, result() RS mp));
 | |
| 206 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 207 | Goal "(e,s) -|-> (v,s') ==> (VALOF SKIP RESULTIS e, s) -|-> (v,s')"; | 
| 3145 | 208 | by (Blast_tac 1); | 
| 209 | qed "valof_skip2"; | |
| 210 | ||
| 5069 | 211 | Goal "((VALOF SKIP RESULTIS e, s) -|-> (v,s')) = ((e, s) -|-> (v,s'))"; | 
| 4089 | 212 | by (blast_tac (claset() addIs [valof_skip1, valof_skip2]) 1); | 
| 3145 | 213 | qed "valof_skip"; | 
| 214 | ||
| 215 | ||
| 216 | (** Equivalence of VALOF x:=e RESULTIS x and e **) | |
| 217 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 218 | Goal "(e',s) -|-> (v,s'') ==> \ | 
| 3145 | 219 | \ (e' = VALOF x:=e RESULTIS X x) --> \ | 
| 220 | \ (EX s'. (e, s) -|-> (v,s') & (s'' = s'[v/x]))"; | |
| 221 | by (etac eval_induct 1); | |
| 222 | by (ALLGOALS Asm_simp_tac); | |
| 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: 
5069diff
changeset | 230 | Goal "(e,s) -|-> (v,s') ==> \ | 
| 3145 | 231 | \ (VALOF x:=e RESULTIS X x, s) -|-> (v,s'[v/x])"; | 
| 232 | by (Blast_tac 1); | |
| 233 | qed "valof_assign2"; |