author | berghofe |
Fri, 07 Apr 2006 11:17:44 +0200 | |
changeset 19357 | dade85a75c9f |
parent 17456 | bcf7544875b2 |
permissions | -rw-r--r-- |
17456 | 1 |
(* Title: CCL/typecheck.ML |
0 | 2 |
ID: $Id$ |
1459 | 3 |
Author: Martin Coen, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1993 University of Cambridge |
5 |
*) |
|
6 |
||
7 |
(*** Lemmas for constructors and subtypes ***) |
|
8 |
||
9 |
(* 0-ary constructors do not need additional rules as they are handled *) |
|
10 |
(* correctly by applying SubtypeI *) |
|
11 |
(* |
|
17456 | 12 |
val Subtype_canTs = |
0 | 13 |
let fun tac prems = cut_facts_tac prems 1 THEN |
14 |
REPEAT (ares_tac (SubtypeI::canTs@icanTs) 1 ORELSE |
|
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
15 |
etac SubtypeE 1) |
17456 | 16 |
fun solve s = prove_goal (the_context ()) s (fn prems => [tac prems]) |
0 | 17 |
in map solve |
18 |
["P(one) ==> one : {x:Unit.P(x)}", |
|
19 |
"P(true) ==> true : {x:Bool.P(x)}", |
|
20 |
"P(false) ==> false : {x:Bool.P(x)}", |
|
21 |
"a : {x:A. b:{y:B(a).P(<x,y>)}} ==> <a,b> : {x:Sigma(A,B).P(x)}", |
|
22 |
"a : {x:A.P(inl(x))} ==> inl(a) : {x:A+B.P(x)}", |
|
23 |
"b : {x:B.P(inr(x))} ==> inr(b) : {x:A+B.P(x)}", |
|
24 |
"P(zero) ==> zero : {x:Nat.P(x)}", |
|
25 |
"a : {x:Nat.P(succ(x))} ==> succ(a) : {x:Nat.P(x)}", |
|
26 |
"P([]) ==> [] : {x:List(A).P(x)}", |
|
289 | 27 |
"h : {x:A. t : {y:List(A).P(x$y)}} ==> h$t : {x:List(A).P(x)}" |
0 | 28 |
] end; |
29 |
*) |
|
17456 | 30 |
val Subtype_canTs = |
0 | 31 |
let fun tac prems = cut_facts_tac prems 1 THEN |
32 |
REPEAT (ares_tac (SubtypeI::canTs@icanTs) 1 ORELSE |
|
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
33 |
etac SubtypeE 1) |
17456 | 34 |
fun solve s = prove_goal (the_context ()) s (fn prems => [tac prems]) |
0 | 35 |
in map solve |
36 |
["a : {x:A. b:{y:B(a).P(<x,y>)}} ==> <a,b> : {x:Sigma(A,B).P(x)}", |
|
3837 | 37 |
"a : {x:A. P(inl(x))} ==> inl(a) : {x:A+B. P(x)}", |
38 |
"b : {x:B. P(inr(x))} ==> inr(b) : {x:A+B. P(x)}", |
|
39 |
"a : {x:Nat. P(succ(x))} ==> succ(a) : {x:Nat. P(x)}", |
|
289 | 40 |
"h : {x:A. t : {y:List(A).P(x$y)}} ==> h$t : {x:List(A).P(x)}" |
0 | 41 |
] end; |
42 |
||
17456 | 43 |
val prems = goal (the_context ()) |
0 | 44 |
"[| f(t):B; ~t=bot |] ==> let x be t in f(x) : B"; |
45 |
by (cut_facts_tac prems 1); |
|
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
46 |
by (etac (letB RS ssubst) 1); |
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
47 |
by (assume_tac 1); |
757 | 48 |
qed "letT"; |
0 | 49 |
|
17456 | 50 |
val prems = goal (the_context ()) |
0 | 51 |
"[| a:A; f : Pi(A,B) |] ==> f ` a : B(a)"; |
52 |
by (REPEAT (ares_tac (applyT::prems) 1)); |
|
757 | 53 |
qed "applyT2"; |
0 | 54 |
|
17456 | 55 |
val prems = goal (the_context ()) |
3837 | 56 |
"[| a:A; a:A ==> P(a) |] ==> a : {x:A. P(x)}"; |
0 | 57 |
by (fast_tac (type_cs addSIs prems) 1); |
757 | 58 |
qed "rcall_lemma1"; |
0 | 59 |
|
17456 | 60 |
val prems = goal (the_context ()) |
3837 | 61 |
"[| a:{x:A. Q(x)}; [| a:A; Q(a) |] ==> P(a) |] ==> a : {x:A. P(x)}"; |
0 | 62 |
by (cut_facts_tac prems 1); |
63 |
by (fast_tac (type_cs addSIs prems) 1); |
|
757 | 64 |
qed "rcall_lemma2"; |
0 | 65 |
|
66 |
val rcall_lemmas = [asm_rl,rcall_lemma1,SubtypeD1,rcall_lemma2]; |
|
67 |
||
68 |
(***********************************TYPECHECKING*************************************) |
|
69 |
||
70 |
fun bvars (Const("all",_) $ Abs(s,_,t)) l = bvars t (s::l) |
|
71 |
| bvars _ l = l; |
|
72 |
||
17456 | 73 |
fun get_bno l n (Const("all",_) $ Abs(s,_,t)) = get_bno (s::l) n t |
0 | 74 |
| get_bno l n (Const("Trueprop",_) $ t) = get_bno l n t |
75 |
| get_bno l n (Const("Ball",_) $ _ $ Abs(s,_,t)) = get_bno (s::l) (n+1) t |
|
76 |
| get_bno l n (Const("op :",_) $ t $ _) = get_bno l n t |
|
77 |
| get_bno l n (t $ s) = get_bno l n t |
|
78 |
| get_bno l n (Bound m) = (m-length(l),n); |
|
79 |
||
80 |
(* Not a great way of identifying induction hypothesis! *) |
|
81 |
fun could_IH x = could_unify(x,hd (prems_of rcallT)) orelse |
|
82 |
could_unify(x,hd (prems_of rcall2T)) orelse |
|
83 |
could_unify(x,hd (prems_of rcall3T)); |
|
84 |
||
3537 | 85 |
fun IHinst tac rls = SUBGOAL (fn (Bi,i) => |
86 |
let val bvs = bvars Bi []; |
|
15570 | 87 |
val ihs = List.filter could_IH (Logic.strip_assums_hyp Bi); |
0 | 88 |
val rnames = map (fn x=> |
89 |
let val (a,b) = get_bno [] 0 x |
|
15570 | 90 |
in (List.nth(bvs,a),b) end) ihs; |
0 | 91 |
fun try_IHs [] = no_tac |
15570 | 92 |
| try_IHs ((x,y)::xs) = tac [("g",x)] (List.nth(rls,y-1)) i ORELSE (try_IHs xs); |
0 | 93 |
in try_IHs rnames end); |
94 |
||
95 |
(*****) |
|
96 |
||
97 |
val type_rls = canTs@icanTs@(applyT2::ncanTs)@incanTs@ |
|
98 |
precTs@letrecTs@[letT]@Subtype_canTs; |
|
99 |
||
17456 | 100 |
fun is_rigid_prog t = |
101 |
case (Logic.strip_assums_concl t) of |
|
0 | 102 |
(Const("Trueprop",_) $ (Const("op :",_) $ a $ _)) => (term_vars a = []) |
103 |
| _ => false; |
|
104 |
||
105 |
fun rcall_tac i = let fun tac ps rl i = res_inst_tac ps rl i THEN atac i; |
|
106 |
in IHinst tac rcallTs i end THEN |
|
107 |
eresolve_tac rcall_lemmas i; |
|
108 |
||
109 |
fun raw_step_tac prems i = ares_tac (prems@type_rls) i ORELSE |
|
110 |
rcall_tac i ORELSE |
|
111 |
ematch_tac [SubtypeE] i ORELSE |
|
112 |
match_tac [SubtypeI] i; |
|
113 |
||
3537 | 114 |
fun tc_step_tac prems = SUBGOAL (fn (Bi,i) => |
115 |
if is_rigid_prog Bi then raw_step_tac prems i else no_tac); |
|
0 | 116 |
|
117 |
fun typechk_tac rls i = SELECT_GOAL (REPEAT_FIRST (tc_step_tac rls)) i; |
|
118 |
||
119 |
val tac = typechk_tac [] 1; |
|
120 |
||
121 |
||
122 |
(*** Clean up Correctness Condictions ***) |
|
123 |
||
124 |
val clean_ccs_tac = REPEAT_FIRST (eresolve_tac ([SubtypeE]@rmIHs) ORELSE' |
|
125 |
hyp_subst_tac); |
|
126 |
||
17456 | 127 |
val clean_ccs_tac = |
0 | 128 |
let fun tac ps rl i = eres_inst_tac ps rl i THEN atac i; |
129 |
in TRY (REPEAT_FIRST (IHinst tac hyprcallTs ORELSE' |
|
130 |
eresolve_tac ([asm_rl,SubtypeE]@rmIHs) ORELSE' |
|
131 |
hyp_subst_tac)) end; |
|
132 |
||
133 |
fun gen_ccs_tac rls i = SELECT_GOAL (REPEAT_FIRST (tc_step_tac rls) THEN |
|
134 |
clean_ccs_tac) i; |