1461
|
1 |
(* Title: ZF/Coind/ECR.ML
|
915
|
2 |
ID: $Id$
|
1461
|
3 |
Author: Jacob Frost, Cambridge University Computer Laboratory
|
915
|
4 |
Copyright 1995 University of Cambridge
|
|
5 |
*)
|
|
6 |
|
|
7 |
open Dynamic ECR;
|
|
8 |
|
|
9 |
(* Specialised co-induction rule *)
|
|
10 |
|
|
11 |
goal ECR.thy
|
|
12 |
"!!x.[| x:ExVar; e:Exp; t:Ty; ve:ValEnv; te:TyEnv; \
|
|
13 |
\ <te, e_fn(x, e), t>:ElabRel; ve_dom(ve) = te_dom(te); \
|
|
14 |
\ {<ve_app(ve,y),te_app(te,y)>.y:ve_dom(ve)}: \
|
|
15 |
\ Pow({<v_clos(x,e,ve),t>} Un HasTyRel) |] ==> \
|
|
16 |
\ <v_clos(x, e, ve),t>:HasTyRel";
|
|
17 |
by (rtac HasTyRel.coinduct 1);
|
|
18 |
by (rtac singletonI 1);
|
4091
|
19 |
by (fast_tac (claset() addIs Val_ValEnv.intrs) 1);
|
915
|
20 |
by (rtac disjI2 1);
|
|
21 |
by (etac singletonE 1);
|
|
22 |
by (REPEAT_FIRST (resolve_tac [conjI,exI]));
|
|
23 |
by (REPEAT (atac 1));
|
|
24 |
qed "htr_closCI";
|
|
25 |
|
|
26 |
(* Specialised elimination rules *)
|
|
27 |
|
|
28 |
val htr_constE =
|
|
29 |
(HasTyRel.mk_cases Val_ValEnv.con_defs "<v_const(c),t>:HasTyRel");
|
|
30 |
|
|
31 |
val htr_closE =
|
|
32 |
(HasTyRel.mk_cases Val_ValEnv.con_defs "<v_clos(x,e,ve),t>:HasTyRel");
|
|
33 |
|
|
34 |
(* Classical reasoning sets *)
|
|
35 |
|
|
36 |
fun mk_htr_cs cs =
|
|
37 |
let open HasTyRel in
|
|
38 |
( cs
|
|
39 |
addSIs [htr_constI]
|
|
40 |
addSEs [htr_constE]
|
|
41 |
addIs [htr_closI,htr_closCI]
|
|
42 |
addEs [htr_closE])
|
|
43 |
end;
|
|
44 |
|
4091
|
45 |
claset_ref() := mk_htr_cs (claset());
|
915
|
46 |
|
|
47 |
(* Properties of the pointwise extension to environments *)
|
|
48 |
|
|
49 |
goalw ECR.thy [hastyenv_def]
|
|
50 |
"!!ve.[| ve:ValEnv; te:TyEnv; hastyenv(ve,te); <v,t>:HasTyRel |] ==> \
|
|
51 |
\ hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))";
|
4152
|
52 |
by Safe_tac;
|
2034
|
53 |
by (stac ve_dom_owr 1);
|
915
|
54 |
by (assume_tac 1);
|
|
55 |
by (etac (HasTyRel.dom_subset RS subsetD RS SigmaD1 RS ValNEE) 1);
|
2034
|
56 |
by (stac te_dom_owr 1);
|
2469
|
57 |
by (Asm_simp_tac 1);
|
915
|
58 |
by (rtac (excluded_middle RS disjE) 1);
|
2034
|
59 |
by (stac ve_app_owr2 1);
|
915
|
60 |
by (assume_tac 1);
|
|
61 |
by (assume_tac 1);
|
2034
|
62 |
by (stac te_app_owr2 1);
|
915
|
63 |
by (assume_tac 1);
|
|
64 |
by (dtac (ve_dom_owr RS subst) 1);
|
|
65 |
by (etac (HasTyRel.dom_subset RS subsetD RS SigmaD1 RS ValNEE) 1);
|
2469
|
66 |
by ((Fast_tac 1) THEN (Fast_tac 1));
|
4091
|
67 |
by (asm_simp_tac (simpset() addsimps [ve_app_owr1,te_app_owr1]) 1);
|
915
|
68 |
qed "hastyenv_owr";
|
|
69 |
|
|
70 |
goalw ECR.thy [isofenv_def,hastyenv_def]
|
|
71 |
"!!ve.[| ve:ValEnv; te:TyEnv; isofenv(ve,te) |] ==> hastyenv(ve,te)";
|
4152
|
72 |
by Safe_tac;
|
915
|
73 |
by (dtac bspec 1);
|
|
74 |
by (assume_tac 1);
|
4152
|
75 |
by Safe_tac;
|
915
|
76 |
by (dtac HasTyRel.htr_constI 1);
|
|
77 |
by (assume_tac 2);
|
|
78 |
by (etac te_appI 1);
|
|
79 |
by (etac ve_domI 1);
|
2469
|
80 |
by (ALLGOALS Asm_full_simp_tac);
|
915
|
81 |
qed "basic_consistency_lem";
|
|
82 |
|
|
83 |
|
|
84 |
|