| author | wenzelm | 
| Thu, 19 Jan 2006 21:22:15 +0100 | |
| changeset 18710 | 527aa560a9e0 | 
| parent 17778 | 93d7e524417a | 
| child 20943 | cf19faf11bbd | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/ex/MT.ML | 
| 969 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Jacob Frost, Cambridge University Computer Laboratory | 
| 969 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | Based upon the article | |
| 7 | Robin Milner and Mads Tofte, | |
| 8 | Co-induction in Relational Semantics, | |
| 9 | Theoretical Computer Science 87 (1991), pages 209-220. | |
| 10 | ||
| 11 | Written up as | |
| 12 | Jacob Frost, A Case Study of Co-induction in Isabelle/HOL | |
| 13 | Report 308, Computer Lab, University of Cambridge (1993). | |
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 14 | |
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 15 | NEEDS TO USE INDUCTIVE DEFS PACKAGE | 
| 969 | 16 | *) | 
| 17 | ||
| 18 | (* ############################################################ *) | |
| 19 | (* Inference systems *) | |
| 20 | (* ############################################################ *) | |
| 21 | ||
| 15386 | 22 | val lfp_lemma2 = thm "lfp_lemma2"; | 
| 23 | val lfp_lemma3 = thm "lfp_lemma3"; | |
| 24 | val gfp_lemma2 = thm "gfp_lemma2"; | |
| 25 | val gfp_lemma3 = thm "gfp_lemma3"; | |
| 26 | ||
| 2935 | 27 | val infsys_mono_tac = (REPEAT (ares_tac (basic_monos@[allI,impI]) 1)); | 
| 969 | 28 | |
| 17289 | 29 | val prems = goal (the_context ()) "P a b ==> P (fst (a,b)) (snd (a,b))"; | 
| 4089 | 30 | by (simp_tac (simpset() addsimps prems) 1); | 
| 969 | 31 | qed "infsys_p1"; | 
| 32 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 33 | Goal "P (fst (a,b)) (snd (a,b)) ==> P a b"; | 
| 1266 | 34 | by (Asm_full_simp_tac 1); | 
| 969 | 35 | qed "infsys_p2"; | 
| 36 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 37 | Goal "P a b c ==> P (fst(fst((a,b),c))) (snd(fst ((a,b),c))) (snd ((a,b),c))"; | 
| 1266 | 38 | by (Asm_full_simp_tac 1); | 
| 969 | 39 | qed "infsys_pp1"; | 
| 40 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 41 | Goal "P (fst(fst((a,b),c))) (snd(fst((a,b),c))) (snd((a,b),c)) ==> P a b c"; | 
| 1266 | 42 | by (Asm_full_simp_tac 1); | 
| 969 | 43 | qed "infsys_pp2"; | 
| 44 | ||
| 45 | (* ############################################################ *) | |
| 46 | (* Fixpoints *) | |
| 47 | (* ############################################################ *) | |
| 48 | ||
| 49 | (* Least fixpoints *) | |
| 50 | ||
| 17289 | 51 | val prems = goal (the_context ()) "[| mono(f); x:f(lfp(f)) |] ==> x:lfp(f)"; | 
| 969 | 52 | by (rtac subsetD 1); | 
| 53 | by (rtac lfp_lemma2 1); | |
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 54 | by (resolve_tac prems 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 55 | by (resolve_tac prems 1); | 
| 969 | 56 | qed "lfp_intro2"; | 
| 57 | ||
| 17289 | 58 | val prems = goal (the_context ()) | 
| 969 | 59 | " [| x:lfp(f); mono(f); !!y. y:f(lfp(f)) ==> P(y) |] ==> \ | 
| 60 | \ P(x)"; | |
| 61 | by (cut_facts_tac prems 1); | |
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 62 | by (resolve_tac prems 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 63 | by (rtac subsetD 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 64 | by (rtac lfp_lemma3 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 65 | by (assume_tac 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 66 | by (assume_tac 1); | 
| 969 | 67 | qed "lfp_elim2"; | 
| 68 | ||
| 17289 | 69 | val prems = goal (the_context ()) | 
| 3842 | 70 |   " [| x:lfp(f); mono(f); !!y. y:f(lfp(f) Int {x. P(x)}) ==> P(y) |] ==> \
 | 
| 969 | 71 | \ P(x)"; | 
| 72 | by (cut_facts_tac prems 1); | |
| 10202 | 73 | by (etac lfp_induct 1); | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 74 | by (assume_tac 1); | 
| 969 | 75 | by (eresolve_tac prems 1); | 
| 76 | qed "lfp_ind2"; | |
| 77 | ||
| 78 | (* Greatest fixpoints *) | |
| 79 | ||
| 80 | (* Note : "[| x:S; S <= f(S Un gfp(f)); mono(f) |] ==> x:gfp(f)" *) | |
| 81 | ||
| 17289 | 82 | val [cih,monoh] = goal (the_context ()) "[| x:f({x} Un gfp(f)); mono(f) |] ==> x:gfp(f)";
 | 
| 969 | 83 | by (rtac (cih RSN (2,gfp_upperbound RS subsetD)) 1); | 
| 84 | by (rtac (monoh RS monoD) 1); | |
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 85 | by (rtac (UnE RS subsetI) 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 86 | by (assume_tac 1); | 
| 4089 | 87 | by (blast_tac (claset() addSIs [cih]) 1); | 
| 969 | 88 | by (rtac (monoh RS monoD RS subsetD) 1); | 
| 89 | by (rtac Un_upper2 1); | |
| 90 | by (etac (monoh RS gfp_lemma2 RS subsetD) 1); | |
| 91 | qed "gfp_coind2"; | |
| 92 | ||
| 17289 | 93 | val [gfph,monoh,caseh] = goal (the_context ()) | 
| 969 | 94 | "[| x:gfp(f); mono(f); !! y. y:f(gfp(f)) ==> P(y) |] ==> P(x)"; | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 95 | by (rtac caseh 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 96 | by (rtac subsetD 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 97 | by (rtac gfp_lemma2 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 98 | by (rtac monoh 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 99 | by (rtac gfph 1); | 
| 969 | 100 | qed "gfp_elim2"; | 
| 101 | ||
| 102 | (* ############################################################ *) | |
| 103 | (* Expressions *) | |
| 104 | (* ############################################################ *) | |
| 105 | ||
| 106 | val e_injs = [e_const_inj, e_var_inj, e_fn_inj, e_fix_inj, e_app_inj]; | |
| 107 | ||
| 17289 | 108 | val e_disjs = | 
| 109 | [ e_disj_const_var, | |
| 110 | e_disj_const_fn, | |
| 111 | e_disj_const_fix, | |
| 969 | 112 | e_disj_const_app, | 
| 17289 | 113 | e_disj_var_fn, | 
| 114 | e_disj_var_fix, | |
| 115 | e_disj_var_app, | |
| 116 | e_disj_fn_fix, | |
| 117 | e_disj_fn_app, | |
| 969 | 118 | e_disj_fix_app | 
| 119 | ]; | |
| 120 | ||
| 121 | val e_disj_si = e_disjs @ (e_disjs RL [not_sym]); | |
| 122 | val e_disj_se = (e_disj_si RL [notE]); | |
| 123 | ||
| 124 | fun e_ext_cs cs = cs addSIs e_disj_si addSEs e_disj_se addSDs e_injs; | |
| 125 | ||
| 126 | (* ############################################################ *) | |
| 127 | (* Values *) | |
| 128 | (* ############################################################ *) | |
| 129 | ||
| 130 | val v_disjs = [v_disj_const_clos]; | |
| 131 | val v_disj_si = v_disjs @ (v_disjs RL [not_sym]); | |
| 132 | val v_disj_se = (v_disj_si RL [notE]); | |
| 133 | ||
| 134 | val v_injs = [v_const_inj, v_clos_inj]; | |
| 135 | ||
| 136 | fun v_ext_cs cs = cs addSIs v_disj_si addSEs v_disj_se addSDs v_injs; | |
| 137 | ||
| 138 | (* ############################################################ *) | |
| 139 | (* Evaluations *) | |
| 140 | (* ############################################################ *) | |
| 141 | ||
| 142 | (* Monotonicity of eval_fun *) | |
| 143 | ||
| 5069 | 144 | Goalw [mono_def, eval_fun_def] "mono(eval_fun)"; | 
| 969 | 145 | by infsys_mono_tac; | 
| 146 | qed "eval_fun_mono"; | |
| 147 | ||
| 148 | (* Introduction rules *) | |
| 149 | ||
| 5069 | 150 | Goalw [eval_def, eval_rel_def] "ve |- e_const(c) ---> v_const(c)"; | 
| 969 | 151 | by (rtac lfp_intro2 1); | 
| 152 | by (rtac eval_fun_mono 1); | |
| 153 | by (rewtac eval_fun_def); | |
| 17289 | 154 | (*Naughty! But the quantifiers are nested VERY deeply...*) | 
| 4089 | 155 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 156 | qed "eval_const"; | 
| 157 | ||
| 17289 | 158 | Goalw [eval_def, eval_rel_def] | 
| 969 | 159 | "ev:ve_dom(ve) ==> ve |- e_var(ev) ---> ve_app ve ev"; | 
| 160 | by (rtac lfp_intro2 1); | |
| 161 | by (rtac eval_fun_mono 1); | |
| 162 | by (rewtac eval_fun_def); | |
| 4089 | 163 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 1266 | 164 | qed "eval_var2"; | 
| 969 | 165 | |
| 17289 | 166 | Goalw [eval_def, eval_rel_def] | 
| 969 | 167 | "ve |- fn ev => e ---> v_clos(<|ev,e,ve|>)"; | 
| 168 | by (rtac lfp_intro2 1); | |
| 169 | by (rtac eval_fun_mono 1); | |
| 170 | by (rewtac eval_fun_def); | |
| 4089 | 171 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 172 | qed "eval_fn"; | 
| 173 | ||
| 17289 | 174 | Goalw [eval_def, eval_rel_def] | 
| 969 | 175 |   " cl = <| ev1, e, ve + {ev2 |-> v_clos(cl)} |> ==> \
 | 
| 176 | \ ve |- fix ev2(ev1) = e ---> v_clos(cl)"; | |
| 177 | by (rtac lfp_intro2 1); | |
| 178 | by (rtac eval_fun_mono 1); | |
| 179 | by (rewtac eval_fun_def); | |
| 4089 | 180 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 181 | qed "eval_fix"; | 
| 182 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 183 | Goalw [eval_def, eval_rel_def] | 
| 969 | 184 | " [| ve |- e1 ---> v_const(c1); ve |- e2 ---> v_const(c2) |] ==> \ | 
| 17289 | 185 | \ ve |- e1 @@ e2 ---> v_const(c_app c1 c2)"; | 
| 969 | 186 | by (rtac lfp_intro2 1); | 
| 187 | by (rtac eval_fun_mono 1); | |
| 188 | by (rewtac eval_fun_def); | |
| 4089 | 189 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 190 | qed "eval_app1"; | 
| 191 | ||
| 17289 | 192 | Goalw [eval_def, eval_rel_def] | 
| 969 | 193 | " [| ve |- e1 ---> v_clos(<|xm,em,vem|>); \ | 
| 194 | \ ve |- e2 ---> v2; \ | |
| 195 | \       vem + {xm |-> v2} |- em ---> v \
 | |
| 196 | \ |] ==> \ | |
| 17289 | 197 | \ ve |- e1 @@ e2 ---> v"; | 
| 969 | 198 | by (rtac lfp_intro2 1); | 
| 199 | by (rtac eval_fun_mono 1); | |
| 200 | by (rewtac eval_fun_def); | |
| 4089 | 201 | by (blast_tac (claset() addSIs [disjI2]) 1); | 
| 969 | 202 | qed "eval_app2"; | 
| 203 | ||
| 204 | (* Strong elimination, induction on evaluations *) | |
| 205 | ||
| 17289 | 206 | val prems = goalw (the_context ()) [eval_def, eval_rel_def] | 
| 969 | 207 | " [| ve |- e ---> v; \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 208 | \ !!ve c. P(((ve,e_const(c)),v_const(c))); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 209 | \ !!ev ve. ev:ve_dom(ve) ==> P(((ve,e_var(ev)),ve_app ve ev)); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 210 | \ !!ev ve e. P(((ve,fn ev => e),v_clos(<|ev,e,ve|>))); \ | 
| 969 | 211 | \ !!ev1 ev2 ve cl e. \ | 
| 212 | \        cl = <| ev1, e, ve + {ev2 |-> v_clos(cl)} |> ==> \
 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 213 | \ P(((ve,fix ev2(ev1) = e),v_clos(cl))); \ | 
| 969 | 214 | \ !!ve c1 c2 e1 e2. \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 215 | \ [| P(((ve,e1),v_const(c1))); P(((ve,e2),v_const(c2))) |] ==> \ | 
| 17289 | 216 | \ P(((ve,e1 @@ e2),v_const(c_app c1 c2))); \ | 
| 969 | 217 | \ !!ve vem xm e1 e2 em v v2. \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 218 | \ [| P(((ve,e1),v_clos(<|xm,em,vem|>))); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 219 | \ P(((ve,e2),v2)); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 220 | \            P(((vem + {xm |-> v2},em),v)) \
 | 
| 969 | 221 | \ |] ==> \ | 
| 17289 | 222 | \ P(((ve,e1 @@ e2),v)) \ | 
| 969 | 223 | \ |] ==> \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 224 | \ P(((ve,e),v))"; | 
| 969 | 225 | by (resolve_tac (prems RL [lfp_ind2]) 1); | 
| 226 | by (rtac eval_fun_mono 1); | |
| 227 | by (rewtac eval_fun_def); | |
| 228 | by (dtac CollectD 1); | |
| 4153 | 229 | by Safe_tac; | 
| 969 | 230 | by (ALLGOALS (resolve_tac prems)); | 
| 2935 | 231 | by (ALLGOALS (Blast_tac)); | 
| 969 | 232 | qed "eval_ind0"; | 
| 233 | ||
| 17289 | 234 | val prems = goal (the_context ()) | 
| 969 | 235 | " [| ve |- e ---> v; \ | 
| 236 | \ !!ve c. P ve (e_const c) (v_const c); \ | |
| 237 | \ !!ev ve. ev:ve_dom(ve) ==> P ve (e_var ev) (ve_app ve ev); \ | |
| 238 | \ !!ev ve e. P ve (fn ev => e) (v_clos <|ev,e,ve|>); \ | |
| 239 | \ !!ev1 ev2 ve cl e. \ | |
| 240 | \        cl = <| ev1, e, ve + {ev2 |-> v_clos(cl)} |> ==> \
 | |
| 241 | \ P ve (fix ev2(ev1) = e) (v_clos cl); \ | |
| 242 | \ !!ve c1 c2 e1 e2. \ | |
| 243 | \ [| P ve e1 (v_const c1); P ve e2 (v_const c2) |] ==> \ | |
| 17289 | 244 | \ P ve (e1 @@ e2) (v_const(c_app c1 c2)); \ | 
| 969 | 245 | \ !!ve vem evm e1 e2 em v v2. \ | 
| 246 | \ [| P ve e1 (v_clos <|evm,em,vem|>); \ | |
| 247 | \ P ve e2 v2; \ | |
| 248 | \            P (vem + {evm |-> v2}) em v \
 | |
| 17289 | 249 | \ |] ==> P ve (e1 @@ e2) v \ | 
| 969 | 250 | \ |] ==> P ve e v"; | 
| 251 | by (res_inst_tac [("P","P")] infsys_pp2 1);
 | |
| 252 | by (rtac eval_ind0 1); | |
| 253 | by (ALLGOALS (rtac infsys_pp1)); | |
| 254 | by (ALLGOALS (resolve_tac prems)); | |
| 255 | by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_pp2 1))); | |
| 256 | qed "eval_ind"; | |
| 257 | ||
| 258 | (* ############################################################ *) | |
| 259 | (* Elaborations *) | |
| 260 | (* ############################################################ *) | |
| 261 | ||
| 5069 | 262 | Goalw [mono_def, elab_fun_def] "mono(elab_fun)"; | 
| 969 | 263 | by infsys_mono_tac; | 
| 264 | qed "elab_fun_mono"; | |
| 265 | ||
| 266 | (* Introduction rules *) | |
| 267 | ||
| 17289 | 268 | Goalw [elab_def, elab_rel_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 269 | "c isof ty ==> te |- e_const(c) ===> ty"; | 
| 969 | 270 | by (rtac lfp_intro2 1); | 
| 271 | by (rtac elab_fun_mono 1); | |
| 272 | by (rewtac elab_fun_def); | |
| 4089 | 273 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 274 | qed "elab_const"; | 
| 275 | ||
| 17289 | 276 | Goalw [elab_def, elab_rel_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 277 | "x:te_dom(te) ==> te |- e_var(x) ===> te_app te x"; | 
| 969 | 278 | by (rtac lfp_intro2 1); | 
| 279 | by (rtac elab_fun_mono 1); | |
| 280 | by (rewtac elab_fun_def); | |
| 4089 | 281 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 282 | qed "elab_var"; | 
| 283 | ||
| 17289 | 284 | Goalw [elab_def, elab_rel_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 285 |   "te + {x |=> ty1} |- e ===> ty2 ==> te |- fn x => e ===> ty1->ty2";
 | 
| 969 | 286 | by (rtac lfp_intro2 1); | 
| 287 | by (rtac elab_fun_mono 1); | |
| 288 | by (rewtac elab_fun_def); | |
| 4089 | 289 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 290 | qed "elab_fn"; | 
| 291 | ||
| 5069 | 292 | Goalw [elab_def, elab_rel_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 293 |   "te + {f |=> ty1->ty2} + {x |=> ty1} |- e ===> ty2 ==> \
 | 
| 2935 | 294 | \ te |- fix f(x) = e ===> ty1->ty2"; | 
| 969 | 295 | by (rtac lfp_intro2 1); | 
| 296 | by (rtac elab_fun_mono 1); | |
| 297 | by (rewtac elab_fun_def); | |
| 4089 | 298 | by (blast_tac (claset() addSIs [exI]) 1); | 
| 969 | 299 | qed "elab_fix"; | 
| 300 | ||
| 17289 | 301 | Goalw [elab_def, elab_rel_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 302 | "[| te |- e1 ===> ty1->ty2; te |- e2 ===> ty1 |] ==> \ | 
| 17289 | 303 | \ te |- e1 @@ e2 ===> ty2"; | 
| 969 | 304 | by (rtac lfp_intro2 1); | 
| 305 | by (rtac elab_fun_mono 1); | |
| 306 | by (rewtac elab_fun_def); | |
| 4089 | 307 | by (blast_tac (claset() addSIs [disjI2]) 1); | 
| 969 | 308 | qed "elab_app"; | 
| 309 | ||
| 310 | (* Strong elimination, induction on elaborations *) | |
| 311 | ||
| 17289 | 312 | val prems = goalw (the_context ()) [elab_def, elab_rel_def] | 
| 969 | 313 | " [| te |- e ===> t; \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 314 | \ !!te c t. c isof t ==> P(((te,e_const(c)),t)); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 315 | \ !!te x. x:te_dom(te) ==> P(((te,e_var(x)),te_app te x)); \ | 
| 969 | 316 | \ !!te x e t1 t2. \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 317 | \        [| te + {x |=> t1} |- e ===> t2; P(((te + {x |=> t1},e),t2)) |] ==> \
 | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 318 | \ P(((te,fn x => e),t1->t2)); \ | 
| 969 | 319 | \ !!te f x e t1 t2. \ | 
| 320 | \        [| te + {f |=> t1->t2} + {x |=> t1} |- e ===> t2; \
 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 321 | \           P(((te + {f |=> t1->t2} + {x |=> t1},e),t2)) \
 | 
| 969 | 322 | \ |] ==> \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 323 | \ P(((te,fix f(x) = e),t1->t2)); \ | 
| 969 | 324 | \ !!te e1 e2 t1 t2. \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 325 | \ [| te |- e1 ===> t1->t2; P(((te,e1),t1->t2)); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 326 | \ te |- e2 ===> t1; P(((te,e2),t1)) \ | 
| 969 | 327 | \ |] ==> \ | 
| 17289 | 328 | \ P(((te,e1 @@ e2),t2)) \ | 
| 969 | 329 | \ |] ==> \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 330 | \ P(((te,e),t))"; | 
| 969 | 331 | by (resolve_tac (prems RL [lfp_ind2]) 1); | 
| 332 | by (rtac elab_fun_mono 1); | |
| 333 | by (rewtac elab_fun_def); | |
| 334 | by (dtac CollectD 1); | |
| 4153 | 335 | by Safe_tac; | 
| 969 | 336 | by (ALLGOALS (resolve_tac prems)); | 
| 2935 | 337 | by (ALLGOALS (Blast_tac)); | 
| 969 | 338 | qed "elab_ind0"; | 
| 339 | ||
| 17289 | 340 | val prems = goal (the_context ()) | 
| 969 | 341 | " [| te |- e ===> t; \ | 
| 342 | \ !!te c t. c isof t ==> P te (e_const c) t; \ | |
| 343 | \ !!te x. x:te_dom(te) ==> P te (e_var x) (te_app te x); \ | |
| 344 | \ !!te x e t1 t2. \ | |
| 345 | \        [| te + {x |=> t1} |- e ===> t2; P (te + {x |=> t1}) e t2 |] ==> \
 | |
| 346 | \ P te (fn x => e) (t1->t2); \ | |
| 347 | \ !!te f x e t1 t2. \ | |
| 348 | \        [| te + {f |=> t1->t2} + {x |=> t1} |- e ===> t2; \
 | |
| 349 | \           P (te + {f |=> t1->t2} + {x |=> t1}) e t2 \
 | |
| 350 | \ |] ==> \ | |
| 351 | \ P te (fix f(x) = e) (t1->t2); \ | |
| 352 | \ !!te e1 e2 t1 t2. \ | |
| 353 | \ [| te |- e1 ===> t1->t2; P te e1 (t1->t2); \ | |
| 354 | \ te |- e2 ===> t1; P te e2 t1 \ | |
| 355 | \ |] ==> \ | |
| 17289 | 356 | \ P te (e1 @@ e2) t2 \ | 
| 969 | 357 | \ |] ==> \ | 
| 358 | \ P te e t"; | |
| 359 | by (res_inst_tac [("P","P")] infsys_pp2 1);
 | |
| 360 | by (rtac elab_ind0 1); | |
| 361 | by (ALLGOALS (rtac infsys_pp1)); | |
| 362 | by (ALLGOALS (resolve_tac prems)); | |
| 363 | by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_pp2 1))); | |
| 364 | qed "elab_ind"; | |
| 365 | ||
| 366 | (* Weak elimination, case analysis on elaborations *) | |
| 367 | ||
| 17289 | 368 | val prems = goalw (the_context ()) [elab_def, elab_rel_def] | 
| 969 | 369 | " [| te |- e ===> t; \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 370 | \ !!te c t. c isof t ==> P(((te,e_const(c)),t)); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 371 | \ !!te x. x:te_dom(te) ==> P(((te,e_var(x)),te_app te x)); \ | 
| 969 | 372 | \ !!te x e t1 t2. \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 373 | \        te + {x |=> t1} |- e ===> t2 ==> P(((te,fn x => e),t1->t2)); \
 | 
| 969 | 374 | \ !!te f x e t1 t2. \ | 
| 375 | \        te + {f |=> t1->t2} + {x |=> t1} |- e ===> t2 ==> \
 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 376 | \ P(((te,fix f(x) = e),t1->t2)); \ | 
| 969 | 377 | \ !!te e1 e2 t1 t2. \ | 
| 378 | \ [| te |- e1 ===> t1->t2; te |- e2 ===> t1 |] ==> \ | |
| 17289 | 379 | \ P(((te,e1 @@ e2),t2)) \ | 
| 969 | 380 | \ |] ==> \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 381 | \ P(((te,e),t))"; | 
| 969 | 382 | by (resolve_tac (prems RL [lfp_elim2]) 1); | 
| 383 | by (rtac elab_fun_mono 1); | |
| 384 | by (rewtac elab_fun_def); | |
| 385 | by (dtac CollectD 1); | |
| 4153 | 386 | by Safe_tac; | 
| 969 | 387 | by (ALLGOALS (resolve_tac prems)); | 
| 2935 | 388 | by (ALLGOALS (Blast_tac)); | 
| 969 | 389 | qed "elab_elim0"; | 
| 390 | ||
| 17289 | 391 | val prems = goal (the_context ()) | 
| 969 | 392 | " [| te |- e ===> t; \ | 
| 393 | \ !!te c t. c isof t ==> P te (e_const c) t; \ | |
| 394 | \ !!te x. x:te_dom(te) ==> P te (e_var x) (te_app te x); \ | |
| 395 | \ !!te x e t1 t2. \ | |
| 396 | \        te + {x |=> t1} |- e ===> t2 ==> P te (fn x => e) (t1->t2); \
 | |
| 397 | \ !!te f x e t1 t2. \ | |
| 398 | \        te + {f |=> t1->t2} + {x |=> t1} |- e ===> t2 ==> \
 | |
| 399 | \ P te (fix f(x) = e) (t1->t2); \ | |
| 400 | \ !!te e1 e2 t1 t2. \ | |
| 401 | \ [| te |- e1 ===> t1->t2; te |- e2 ===> t1 |] ==> \ | |
| 17289 | 402 | \ P te (e1 @@ e2) t2 \ | 
| 969 | 403 | \ |] ==> \ | 
| 404 | \ P te e t"; | |
| 405 | by (res_inst_tac [("P","P")] infsys_pp2 1);
 | |
| 406 | by (rtac elab_elim0 1); | |
| 407 | by (ALLGOALS (rtac infsys_pp1)); | |
| 408 | by (ALLGOALS (resolve_tac prems)); | |
| 409 | by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_pp2 1))); | |
| 410 | qed "elab_elim"; | |
| 411 | ||
| 412 | (* Elimination rules for each expression *) | |
| 413 | ||
| 17289 | 414 | fun elab_e_elim_tac p = | 
| 415 | ( (rtac elab_elim 1) THEN | |
| 416 | (resolve_tac p 1) THEN | |
| 4353 
d565d2197a5f
updated for latest Blast_tac, which treats equality differently
 paulson parents: 
4153diff
changeset | 417 | (REPEAT (fast_tac (e_ext_cs HOL_cs) 1)) | 
| 969 | 418 | ); | 
| 419 | ||
| 17289 | 420 | val prems = goal (the_context ()) "te |- e ===> t ==> (e = e_const(c) --> c isof t)"; | 
| 969 | 421 | by (elab_e_elim_tac prems); | 
| 422 | qed "elab_const_elim_lem"; | |
| 423 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 424 | Goal "te |- e_const(c) ===> t ==> c isof t"; | 
| 969 | 425 | by (dtac elab_const_elim_lem 1); | 
| 2935 | 426 | by (Blast_tac 1); | 
| 969 | 427 | qed "elab_const_elim"; | 
| 428 | ||
| 17289 | 429 | val prems = goal (the_context ()) | 
| 969 | 430 | "te |- e ===> t ==> (e = e_var(x) --> t=te_app te x & x:te_dom(te))"; | 
| 431 | by (elab_e_elim_tac prems); | |
| 432 | qed "elab_var_elim_lem"; | |
| 433 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 434 | Goal "te |- e_var(ev) ===> t ==> t=te_app te ev & ev : te_dom(te)"; | 
| 969 | 435 | by (dtac elab_var_elim_lem 1); | 
| 2935 | 436 | by (Blast_tac 1); | 
| 969 | 437 | qed "elab_var_elim"; | 
| 438 | ||
| 17289 | 439 | val prems = goal (the_context ()) | 
| 969 | 440 | " te |- e ===> t ==> \ | 
| 441 | \ ( e = fn x1 => e1 --> \ | |
| 3842 | 442 | \     (? t1 t2. t=t_fun t1 t2 & te + {x1 |=> t1} |- e1 ===> t2) \
 | 
| 969 | 443 | \ )"; | 
| 444 | by (elab_e_elim_tac prems); | |
| 445 | qed "elab_fn_elim_lem"; | |
| 446 | ||
| 5278 | 447 | Goal " te |- fn x1 => e1 ===> t ==> \ | 
| 969 | 448 | \   (? t1 t2. t=t1->t2 & te + {x1 |=> t1} |- e1 ===> t2)";
 | 
| 449 | by (dtac elab_fn_elim_lem 1); | |
| 2935 | 450 | by (Blast_tac 1); | 
| 969 | 451 | qed "elab_fn_elim"; | 
| 452 | ||
| 17289 | 453 | val prems = goal (the_context ()) | 
| 969 | 454 | " te |- e ===> t ==> \ | 
| 455 | \ (e = fix f(x) = e1 --> \ | |
| 17289 | 456 | \   (? t1 t2. t=t1->t2 & te + {f |=> t1->t2} + {x |=> t1} |- e1 ===> t2))";
 | 
| 969 | 457 | by (elab_e_elim_tac prems); | 
| 458 | qed "elab_fix_elim_lem"; | |
| 459 | ||
| 5278 | 460 | Goal " te |- fix ev1(ev2) = e1 ===> t ==> \ | 
| 969 | 461 | \   (? t1 t2. t=t1->t2 & te + {ev1 |=> t1->t2} + {ev2 |=> t1} |- e1 ===> t2)";
 | 
| 462 | by (dtac elab_fix_elim_lem 1); | |
| 2935 | 463 | by (Blast_tac 1); | 
| 969 | 464 | qed "elab_fix_elim"; | 
| 465 | ||
| 17289 | 466 | val prems = goal (the_context ()) | 
| 969 | 467 | " te |- e ===> t2 ==> \ | 
| 17289 | 468 | \ (e = e1 @@ e2 --> (? t1 . te |- e1 ===> t1->t2 & te |- e2 ===> t1))"; | 
| 969 | 469 | by (elab_e_elim_tac prems); | 
| 470 | qed "elab_app_elim_lem"; | |
| 471 | ||
| 17289 | 472 | Goal "te |- e1 @@ e2 ===> t2 ==> (? t1 . te |- e1 ===> t1->t2 & te |- e2 ===> t1)"; | 
| 969 | 473 | by (dtac elab_app_elim_lem 1); | 
| 2935 | 474 | by (Blast_tac 1); | 
| 969 | 475 | qed "elab_app_elim"; | 
| 476 | ||
| 477 | (* ############################################################ *) | |
| 478 | (* The extended correspondence relation *) | |
| 479 | (* ############################################################ *) | |
| 480 | ||
| 481 | (* Monotonicity of hasty_fun *) | |
| 482 | ||
| 17289 | 483 | Goalw [mono_def, hasty_fun_def] "mono(hasty_fun)"; | 
| 969 | 484 | by infsys_mono_tac; | 
| 2935 | 485 | by (Blast_tac 1); | 
| 486 | qed "mono_hasty_fun"; | |
| 969 | 487 | |
| 17289 | 488 | (* | 
| 489 | Because hasty_rel has been defined as the greatest fixpoint of hasty_fun it | |
| 969 | 490 | enjoys two strong indtroduction (co-induction) rules and an elimination rule. | 
| 491 | *) | |
| 492 | ||
| 493 | (* First strong indtroduction (co-induction) rule for hasty_rel *) | |
| 494 | ||
| 5069 | 495 | Goalw [hasty_rel_def] "c isof t ==> (v_const(c),t) : hasty_rel"; | 
| 969 | 496 | by (rtac gfp_coind2 1); | 
| 17289 | 497 | by (rewtac hasty_fun_def); | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 498 | by (rtac CollectI 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 499 | by (rtac disjI1 1); | 
| 2935 | 500 | by (Blast_tac 1); | 
| 969 | 501 | by (rtac mono_hasty_fun 1); | 
| 502 | qed "hasty_rel_const_coind"; | |
| 503 | ||
| 504 | (* Second strong introduction (co-induction) rule for hasty_rel *) | |
| 505 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 506 | Goalw [hasty_rel_def] | 
| 969 | 507 | " [| te |- fn ev => e ===> t; \ | 
| 508 | \ ve_dom(ve) = te_dom(te); \ | |
| 509 | \ ! ev1. \ | |
| 510 | \ ev1:ve_dom(ve) --> \ | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 511 | \         (ve_app ve ev1,te_app te ev1) : {(v_clos(<|ev,e,ve|>),t)} Un hasty_rel \
 | 
| 969 | 512 | \ |] ==> \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 513 | \ (v_clos(<|ev,e,ve|>),t) : hasty_rel"; | 
| 969 | 514 | by (rtac gfp_coind2 1); | 
| 515 | by (rewtac hasty_fun_def); | |
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 516 | by (rtac CollectI 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 517 | by (rtac disjI2 1); | 
| 2935 | 518 | by (blast_tac HOL_cs 1); | 
| 969 | 519 | by (rtac mono_hasty_fun 1); | 
| 520 | qed "hasty_rel_clos_coind"; | |
| 521 | ||
| 522 | (* Elimination rule for hasty_rel *) | |
| 523 | ||
| 17289 | 524 | val prems = goalw (the_context ()) [hasty_rel_def] | 
| 3842 | 525 | " [| !! c t. c isof t ==> P((v_const(c),t)); \ | 
| 969 | 526 | \ !! te ev e t ve. \ | 
| 527 | \ [| te |- fn ev => e ===> t; \ | |
| 528 | \ ve_dom(ve) = te_dom(te); \ | |
| 3842 | 529 | \ !ev1. ev1:ve_dom(ve) --> (ve_app ve ev1,te_app te ev1) : hasty_rel \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 530 | \ |] ==> P((v_clos(<|ev,e,ve|>),t)); \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 531 | \ (v,t) : hasty_rel \ | 
| 8114 | 532 | \ |] ==> P(v,t)"; | 
| 969 | 533 | by (cut_facts_tac prems 1); | 
| 534 | by (etac gfp_elim2 1); | |
| 535 | by (rtac mono_hasty_fun 1); | |
| 536 | by (rewtac hasty_fun_def); | |
| 537 | by (dtac CollectD 1); | |
| 538 | by (fold_goals_tac [hasty_fun_def]); | |
| 4153 | 539 | by Safe_tac; | 
| 2935 | 540 | by (REPEAT (ares_tac prems 1)); | 
| 969 | 541 | qed "hasty_rel_elim0"; | 
| 542 | ||
| 17289 | 543 | val prems = goal (the_context ()) | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 544 | " [| (v,t) : hasty_rel; \ | 
| 3842 | 545 | \ !! c t. c isof t ==> P (v_const c) t; \ | 
| 969 | 546 | \ !! te ev e t ve. \ | 
| 547 | \ [| te |- fn ev => e ===> t; \ | |
| 548 | \ ve_dom(ve) = te_dom(te); \ | |
| 3842 | 549 | \ !ev1. ev1:ve_dom(ve) --> (ve_app ve ev1,te_app te ev1) : hasty_rel \ | 
| 969 | 550 | \ |] ==> P (v_clos <|ev,e,ve|>) t \ | 
| 551 | \ |] ==> P v t"; | |
| 552 | by (res_inst_tac [("P","P")] infsys_p2 1);
 | |
| 553 | by (rtac hasty_rel_elim0 1); | |
| 554 | by (ALLGOALS (rtac infsys_p1)); | |
| 555 | by (ALLGOALS (resolve_tac prems)); | |
| 556 | by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_p2 1))); | |
| 557 | qed "hasty_rel_elim"; | |
| 558 | ||
| 559 | (* Introduction rules for hasty *) | |
| 560 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 561 | Goalw [hasty_def] "c isof t ==> v_const(c) hasty t"; | 
| 2935 | 562 | by (etac hasty_rel_const_coind 1); | 
| 969 | 563 | qed "hasty_const"; | 
| 564 | ||
| 17289 | 565 | Goalw [hasty_def,hasty_env_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 566 | "te |- fn ev => e ===> t & ve hastyenv te ==> v_clos(<|ev,e,ve|>) hasty t"; | 
| 969 | 567 | by (rtac hasty_rel_clos_coind 1); | 
| 4089 | 568 | by (ALLGOALS (blast_tac (claset() delrules [equalityI]))); | 
| 969 | 569 | qed "hasty_clos"; | 
| 570 | ||
| 571 | (* Elimination on constants for hasty *) | |
| 572 | ||
| 17289 | 573 | Goalw [hasty_def] | 
| 574 | "v hasty t ==> (!c.(v = v_const(c) --> c isof t))"; | |
| 969 | 575 | by (rtac hasty_rel_elim 1); | 
| 2935 | 576 | by (ALLGOALS (blast_tac (v_ext_cs HOL_cs))); | 
| 969 | 577 | qed "hasty_elim_const_lem"; | 
| 578 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 579 | Goal "v_const(c) hasty t ==> c isof t"; | 
| 2935 | 580 | by (dtac hasty_elim_const_lem 1); | 
| 581 | by (Blast_tac 1); | |
| 969 | 582 | qed "hasty_elim_const"; | 
| 583 | ||
| 584 | (* Elimination on closures for hasty *) | |
| 585 | ||
| 17289 | 586 | Goalw [hasty_env_def,hasty_def] | 
| 969 | 587 | " v hasty t ==> \ | 
| 588 | \ ! x e ve. \ | |
| 3842 | 589 | \ v=v_clos(<|x,e,ve|>) --> (? te. te |- fn x => e ===> t & ve hastyenv te)"; | 
| 969 | 590 | by (rtac hasty_rel_elim 1); | 
| 2935 | 591 | by (ALLGOALS (blast_tac (v_ext_cs HOL_cs))); | 
| 969 | 592 | qed "hasty_elim_clos_lem"; | 
| 593 | ||
| 5278 | 594 | Goal "v_clos(<|ev,e,ve|>) hasty t ==> \ | 
| 3842 | 595 | \ ? te. te |- fn ev => e ===> t & ve hastyenv te "; | 
| 2935 | 596 | by (dtac hasty_elim_clos_lem 1); | 
| 597 | by (Blast_tac 1); | |
| 969 | 598 | qed "hasty_elim_clos"; | 
| 599 | ||
| 600 | (* ############################################################ *) | |
| 601 | (* The pointwise extension of hasty to environments *) | |
| 602 | (* ############################################################ *) | |
| 603 | ||
| 5278 | 604 | Goal "[| ve hastyenv te; v hasty t |] ==> \ | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 605 | \        ve + {ev |-> v} hastyenv te + {ev |=> t}";
 | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 606 | by (rewtac hasty_env_def); | 
| 4089 | 607 | by (asm_full_simp_tac (simpset() delsimps mem_simps | 
| 1266 | 608 | addsimps [ve_dom_owr, te_dom_owr]) 1); | 
| 2935 | 609 | by (safe_tac HOL_cs); | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 610 | by (excluded_middle_tac "ev=x" 1); | 
| 4089 | 611 | by (asm_full_simp_tac (simpset() addsimps [ve_app_owr2, te_app_owr2]) 1); | 
| 612 | by (asm_simp_tac (simpset() addsimps [ve_app_owr1, te_app_owr1]) 1); | |
| 969 | 613 | qed "hasty_env1"; | 
| 614 | ||
| 615 | (* ############################################################ *) | |
| 616 | (* The Consistency theorem *) | |
| 617 | (* ############################################################ *) | |
| 618 | ||
| 5278 | 619 | Goal "[| ve hastyenv te ; te |- e_const(c) ===> t |] ==> v_const(c) hasty t"; | 
| 969 | 620 | by (dtac elab_const_elim 1); | 
| 621 | by (etac hasty_const 1); | |
| 622 | qed "consistency_const"; | |
| 623 | ||
| 5069 | 624 | Goalw [hasty_env_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 625 | "[| ev : ve_dom(ve); ve hastyenv te ; te |- e_var(ev) ===> t |] ==> \ | 
| 2935 | 626 | \ ve_app ve ev hasty t"; | 
| 969 | 627 | by (dtac elab_var_elim 1); | 
| 2935 | 628 | by (Blast_tac 1); | 
| 969 | 629 | qed "consistency_var"; | 
| 630 | ||
| 5278 | 631 | Goal "[| ve hastyenv te ; te |- fn ev => e ===> t |] ==> \ | 
| 2935 | 632 | \ v_clos(<| ev, e, ve |>) hasty t"; | 
| 969 | 633 | by (rtac hasty_clos 1); | 
| 2935 | 634 | by (Blast_tac 1); | 
| 969 | 635 | qed "consistency_fn"; | 
| 636 | ||
| 5069 | 637 | Goalw [hasty_env_def,hasty_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 638 |   "[| cl = <| ev1, e, ve + { ev2 |-> v_clos(cl) } |>; \
 | 
| 969 | 639 | \ ve hastyenv te ; \ | 
| 640 | \ te |- fix ev2 ev1 = e ===> t \ | |
| 641 | \ |] ==> \ | |
| 642 | \ v_clos(cl) hasty t"; | |
| 643 | by (dtac elab_fix_elim 1); | |
| 2935 | 644 | by (safe_tac HOL_cs); | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 645 | (*Do a single unfolding of cl*) | 
| 7499 | 646 | by ((ftac ssubst 1) THEN (assume_tac 2)); | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 647 | by (rtac hasty_rel_clos_coind 1); | 
| 969 | 648 | by (etac elab_fn 1); | 
| 4089 | 649 | by (asm_simp_tac (simpset() addsimps [ve_dom_owr, te_dom_owr]) 1); | 
| 969 | 650 | |
| 4089 | 651 | by (asm_simp_tac (simpset() delsimps mem_simps addsimps [ve_dom_owr]) 1); | 
| 2935 | 652 | by (safe_tac HOL_cs); | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 653 | by (excluded_middle_tac "ev2=ev1a" 1); | 
| 4089 | 654 | by (asm_full_simp_tac (simpset() addsimps [ve_app_owr2, te_app_owr2]) 1); | 
| 969 | 655 | |
| 4089 | 656 | by (asm_simp_tac (simpset() delsimps mem_simps | 
| 1266 | 657 | addsimps [ve_app_owr1, te_app_owr1]) 1); | 
| 2935 | 658 | by (Blast_tac 1); | 
| 969 | 659 | qed "consistency_fix"; | 
| 660 | ||
| 5278 | 661 | Goal "[| ! t te. ve hastyenv te --> te |- e1 ===> t --> v_const(c1) hasty t;\ | 
| 969 | 662 | \ ! t te. ve hastyenv te --> te |- e2 ===> t --> v_const(c2) hasty t; \ | 
| 17289 | 663 | \ ve hastyenv te ; te |- e1 @@ e2 ===> t \ | 
| 969 | 664 | \ |] ==> \ | 
| 665 | \ v_const(c_app c1 c2) hasty t"; | |
| 666 | by (dtac elab_app_elim 1); | |
| 4153 | 667 | by Safe_tac; | 
| 969 | 668 | by (rtac hasty_const 1); | 
| 669 | by (rtac isof_app 1); | |
| 670 | by (rtac hasty_elim_const 1); | |
| 2935 | 671 | by (Blast_tac 1); | 
| 969 | 672 | by (rtac hasty_elim_const 1); | 
| 2935 | 673 | by (Blast_tac 1); | 
| 969 | 674 | qed "consistency_app1"; | 
| 675 | ||
| 5278 | 676 | Goal "[| ! t te. \ | 
| 969 | 677 | \ ve hastyenv te --> \ | 
| 678 | \ te |- e1 ===> t --> v_clos(<|evm, em, vem|>) hasty t; \ | |
| 679 | \ ! t te. ve hastyenv te --> te |- e2 ===> t --> v2 hasty t; \ | |
| 680 | \ ! t te. \ | |
| 681 | \        vem + { evm |-> v2 } hastyenv te  --> te |- em ===> t --> v hasty t; \
 | |
| 682 | \ ve hastyenv te ; \ | |
| 17289 | 683 | \ te |- e1 @@ e2 ===> t \ | 
| 969 | 684 | \ |] ==> \ | 
| 685 | \ v hasty t"; | |
| 686 | by (dtac elab_app_elim 1); | |
| 4153 | 687 | by Safe_tac; | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 688 | by ((etac allE 1) THEN (etac allE 1) THEN (etac impE 1)); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 689 | by (assume_tac 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 690 | by (etac impE 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 691 | by (assume_tac 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 692 | by ((etac allE 1) THEN (etac allE 1) THEN (etac impE 1)); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 693 | by (assume_tac 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 694 | by (etac impE 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 695 | by (assume_tac 1); | 
| 969 | 696 | by (dtac hasty_elim_clos 1); | 
| 4153 | 697 | by Safe_tac; | 
| 969 | 698 | by (dtac elab_fn_elim 1); | 
| 4089 | 699 | by (blast_tac (claset() addIs [hasty_env1] addSDs [t_fun_inj]) 1); | 
| 969 | 700 | qed "consistency_app2"; | 
| 701 | ||
| 5278 | 702 | Goal "ve |- e ---> v ==> \ | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 703 | \ (! t te. ve hastyenv te --> te |- e ===> t --> v hasty t)"; | 
| 969 | 704 | |
| 705 | (* Proof by induction on the structure of evaluations *) | |
| 706 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 707 | by (etac eval_ind 1); | 
| 4153 | 708 | by Safe_tac; | 
| 17289 | 709 | by (DEPTH_SOLVE | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 710 | (ares_tac [consistency_const, consistency_var, consistency_fn, | 
| 1465 | 711 | consistency_fix, consistency_app1, consistency_app2] 1)); | 
| 969 | 712 | qed "consistency"; | 
| 713 | ||
| 714 | (* ############################################################ *) | |
| 715 | (* The Basic Consistency theorem *) | |
| 716 | (* ############################################################ *) | |
| 717 | ||
| 17289 | 718 | Goalw [isof_env_def,hasty_env_def] | 
| 969 | 719 | "ve isofenv te ==> ve hastyenv te"; | 
| 4153 | 720 | by Safe_tac; | 
| 1047 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 721 | by (etac allE 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 722 | by (etac impE 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 723 | by (assume_tac 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 724 | by (etac exE 1); | 
| 
5133479a37cf
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
972diff
changeset | 725 | by (etac conjE 1); | 
| 969 | 726 | by (dtac hasty_const 1); | 
| 1266 | 727 | by (Asm_simp_tac 1); | 
| 969 | 728 | qed "basic_consistency_lem"; | 
| 729 | ||
| 5278 | 730 | Goal "[| ve isofenv te; ve |- e ---> v_const(c); te |- e ===> t |] ==> c isof t"; | 
| 969 | 731 | by (rtac hasty_elim_const 1); | 
| 732 | by (dtac consistency 1); | |
| 4089 | 733 | by (blast_tac (claset() addSIs [basic_consistency_lem]) 1); | 
| 969 | 734 | qed "basic_consistency"; |