author | obua |
Sun, 09 May 2004 23:04:36 +0200 | |
changeset 14722 | 8e739a6eaf11 |
parent 13105 | 3d1e7a199bdc |
child 15090 | 970c2668c694 |
permissions | -rw-r--r-- |
12189 | 1 |
(* Title: ZF/Tools/typechk.ML |
6049 | 2 |
ID: $Id$ |
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
6153 | 4 |
Copyright 1999 University of Cambridge |
6049 | 5 |
|
12189 | 6 |
Tactics for type checking -- from CTT. |
6049 | 7 |
*) |
8 |
||
6153 | 9 |
infix 4 addTCs delTCs; |
10 |
||
12189 | 11 |
signature BASIC_TYPE_CHECK = |
12 |
sig |
|
13 |
type tcset |
|
14 |
val addTCs: tcset * thm list -> tcset |
|
15 |
val delTCs: tcset * thm list -> tcset |
|
16 |
val typecheck_tac: tcset -> tactic |
|
17 |
val type_solver_tac: tcset -> thm list -> int -> tactic |
|
18 |
val print_tc: tcset -> unit |
|
19 |
val print_tcset: theory -> unit |
|
20 |
val tcset_ref_of: theory -> tcset ref |
|
21 |
val tcset_of: theory -> tcset |
|
22 |
val tcset: unit -> tcset |
|
12202 | 23 |
val TCSET: (tcset -> tactic) -> tactic |
24 |
val TCSET': (tcset -> 'a -> tactic) -> 'a -> tactic |
|
12189 | 25 |
val AddTCs: thm list -> unit |
26 |
val DelTCs: thm list -> unit |
|
27 |
val TC_add_global: theory attribute |
|
28 |
val TC_del_global: theory attribute |
|
29 |
val TC_add_local: Proof.context attribute |
|
30 |
val TC_del_local: Proof.context attribute |
|
31 |
val Typecheck_tac: tactic |
|
32 |
val Type_solver_tac: thm list -> int -> tactic |
|
33 |
end; |
|
34 |
||
35 |
signature TYPE_CHECK = |
|
36 |
sig |
|
37 |
include BASIC_TYPE_CHECK |
|
38 |
val setup: (theory -> theory) list |
|
39 |
end; |
|
40 |
||
41 |
structure TypeCheck: TYPE_CHECK = |
|
6153 | 42 |
struct |
12189 | 43 |
|
6153 | 44 |
datatype tcset = |
12189 | 45 |
TC of {rules: thm list, (*the type-checking rules*) |
46 |
net: thm Net.net}; (*discrimination net of the same rules*) |
|
6153 | 47 |
|
48 |
||
13105
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12311
diff
changeset
|
49 |
val mem_thm = gen_mem Drule.eq_thm_prop |
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12311
diff
changeset
|
50 |
and rem_thm = gen_rem Drule.eq_thm_prop; |
6153 | 51 |
|
52 |
fun addTC (cs as TC{rules, net}, th) = |
|
12189 | 53 |
if mem_thm (th, rules) then |
54 |
(warning ("Ignoring duplicate type-checking rule\n" ^ |
|
55 |
string_of_thm th); |
|
56 |
cs) |
|
6153 | 57 |
else |
12189 | 58 |
TC{rules = th::rules, |
59 |
net = Net.insert_term ((concl_of th, th), net, K false)}; |
|
6153 | 60 |
|
61 |
||
62 |
fun delTC (cs as TC{rules, net}, th) = |
|
63 |
if mem_thm (th, rules) then |
|
13105
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12311
diff
changeset
|
64 |
TC{net = Net.delete_term ((concl_of th, th), net, Drule.eq_thm_prop), |
12189 | 65 |
rules = rem_thm (rules,th)} |
66 |
else (warning ("No such type-checking rule\n" ^ (string_of_thm th)); |
|
67 |
cs); |
|
6153 | 68 |
|
69 |
val op addTCs = foldl addTC; |
|
70 |
val op delTCs = foldl delTC; |
|
71 |
||
72 |
||
73 |
(*resolution using a net rather than rules*) |
|
74 |
fun net_res_tac maxr net = |
|
75 |
SUBGOAL |
|
76 |
(fn (prem,i) => |
|
77 |
let val rls = Net.unify_term net (Logic.strip_assums_concl prem) |
|
12189 | 78 |
in |
79 |
if length rls <= maxr then resolve_tac rls i else no_tac |
|
6153 | 80 |
end); |
81 |
||
12189 | 82 |
fun is_rigid_elem (Const("Trueprop",_) $ (Const("op :",_) $ a $ _)) = |
6049 | 83 |
not (is_Var (head_of a)) |
84 |
| is_rigid_elem _ = false; |
|
85 |
||
12189 | 86 |
(*Try solving a:A by assumption provided a is rigid!*) |
6049 | 87 |
val test_assume_tac = SUBGOAL(fn (prem,i) => |
88 |
if is_rigid_elem (Logic.strip_assums_concl prem) |
|
89 |
then assume_tac i else eq_assume_tac i); |
|
90 |
||
12189 | 91 |
(*Type checking solves a:?A (a rigid, ?A maybe flexible). |
6049 | 92 |
match_tac is too strict; would refuse to instantiate ?A*) |
6153 | 93 |
fun typecheck_step_tac (TC{net,...}) = |
94 |
FIRSTGOAL (test_assume_tac ORELSE' net_res_tac 3 net); |
|
6049 | 95 |
|
6153 | 96 |
fun typecheck_tac tcset = REPEAT (typecheck_step_tac tcset); |
6049 | 97 |
|
6153 | 98 |
(*Compiles a term-net for speed*) |
99 |
val basic_res_tac = net_resolve_tac [TrueI,refl,reflexive_thm,iff_refl, |
|
12189 | 100 |
ballI,allI,conjI,impI]; |
6049 | 101 |
|
102 |
(*Instantiates variables in typing conditions. |
|
103 |
drawback: does not simplify conjunctions*) |
|
6153 | 104 |
fun type_solver_tac tcset hyps = SELECT_GOAL |
105 |
(DEPTH_SOLVE (etac FalseE 1 |
|
12189 | 106 |
ORELSE basic_res_tac 1 |
107 |
ORELSE (ares_tac hyps 1 |
|
108 |
APPEND typecheck_step_tac tcset))); |
|
6153 | 109 |
|
110 |
||
111 |
||
112 |
fun merge_tc (TC{rules,net}, TC{rules=rules',net=net'}) = |
|
13105
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12311
diff
changeset
|
113 |
TC {rules = gen_union Drule.eq_thm_prop (rules,rules'), |
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12311
diff
changeset
|
114 |
net = Net.merge (net, net', Drule.eq_thm_prop)}; |
6153 | 115 |
|
116 |
(*print tcsets*) |
|
117 |
fun print_tc (TC{rules,...}) = |
|
118 |
Pretty.writeln |
|
119 |
(Pretty.big_list "type-checking rules:" (map Display.pretty_thm rules)); |
|
120 |
||
121 |
||
12189 | 122 |
(** global tcset **) |
123 |
||
6153 | 124 |
structure TypecheckingArgs = |
125 |
struct |
|
126 |
val name = "ZF/type-checker"; |
|
127 |
type T = tcset ref; |
|
128 |
val empty = ref (TC{rules=[], net=Net.empty}); |
|
6556 | 129 |
fun copy (ref tc) = ref tc; |
130 |
val prep_ext = copy; |
|
6153 | 131 |
fun merge (ref tc1, ref tc2) = ref (merge_tc (tc1, tc2)); |
132 |
fun print _ (ref tc) = print_tc tc; |
|
133 |
end; |
|
134 |
||
135 |
structure TypecheckingData = TheoryDataFun(TypecheckingArgs); |
|
6049 | 136 |
|
6153 | 137 |
val print_tcset = TypecheckingData.print; |
138 |
val tcset_ref_of_sg = TypecheckingData.get_sg; |
|
139 |
val tcset_ref_of = TypecheckingData.get; |
|
140 |
||
12189 | 141 |
|
6153 | 142 |
(* access global tcset *) |
143 |
||
144 |
val tcset_of_sg = ! o tcset_ref_of_sg; |
|
145 |
val tcset_of = tcset_of_sg o sign_of; |
|
146 |
||
147 |
val tcset = tcset_of o Context.the_context; |
|
148 |
val tcset_ref = tcset_ref_of_sg o sign_of o Context.the_context; |
|
149 |
||
12202 | 150 |
fun TCSET tacf st = tacf (tcset_of_sg (Thm.sign_of_thm st)) st; |
151 |
fun TCSET' tacf i st = tacf (tcset_of_sg (Thm.sign_of_thm st)) i st; |
|
152 |
||
12189 | 153 |
|
6153 | 154 |
(* change global tcset *) |
155 |
||
156 |
fun change_tcset f x = tcset_ref () := (f (tcset (), x)); |
|
157 |
||
158 |
val AddTCs = change_tcset (op addTCs); |
|
159 |
val DelTCs = change_tcset (op delTCs); |
|
160 |
||
161 |
fun Typecheck_tac st = typecheck_tac (tcset()) st; |
|
162 |
||
163 |
fun Type_solver_tac hyps = type_solver_tac (tcset()) hyps; |
|
164 |
||
165 |
||
166 |
||
12189 | 167 |
(** local tcset **) |
168 |
||
169 |
structure LocalTypecheckingArgs = |
|
170 |
struct |
|
171 |
val name = TypecheckingArgs.name; |
|
172 |
type T = tcset; |
|
173 |
val init = tcset_of; |
|
174 |
fun print _ tcset = print_tc tcset; |
|
175 |
end; |
|
176 |
||
177 |
structure LocalTypecheckingData = ProofDataFun(LocalTypecheckingArgs); |
|
178 |
||
179 |
||
180 |
(* attributes *) |
|
181 |
||
182 |
fun global_att f (thy, th) = |
|
183 |
let val r = tcset_ref_of thy |
|
184 |
in r := f (! r, th); (thy, th) end; |
|
185 |
||
186 |
fun local_att f (ctxt, th) = (LocalTypecheckingData.map (fn tcset => f (tcset, th)) ctxt, th); |
|
187 |
||
188 |
val TC_add_global = global_att addTC; |
|
189 |
val TC_del_global = global_att delTC; |
|
190 |
val TC_add_local = local_att addTC; |
|
191 |
val TC_del_local = local_att delTC; |
|
192 |
||
193 |
val TC_attr = |
|
194 |
(Attrib.add_del_args TC_add_global TC_del_global, |
|
195 |
Attrib.add_del_args TC_add_local TC_del_local); |
|
196 |
||
197 |
||
198 |
(* methods *) |
|
199 |
||
200 |
fun TC_args x = Method.only_sectioned_args |
|
201 |
[Args.add -- Args.colon >> K (I, TC_add_local), |
|
202 |
Args.del -- Args.colon >> K (I, TC_del_local)] x; |
|
203 |
||
204 |
fun typecheck ctxt = |
|
205 |
Method.SIMPLE_METHOD (CHANGED (typecheck_tac (LocalTypecheckingData.get ctxt))); |
|
206 |
||
207 |
||
208 |
||
209 |
(** theory setup **) |
|
210 |
||
211 |
val setup = |
|
212 |
[TypecheckingData.init, LocalTypecheckingData.init, |
|
213 |
Attrib.add_attributes [("TC", TC_attr, "declaration of typecheck rule")], |
|
214 |
Method.add_methods [("typecheck", TC_args typecheck, "ZF typecheck")]]; |
|
215 |
||
216 |
||
217 |
(** outer syntax **) |
|
218 |
||
219 |
val print_tcsetP = |
|
220 |
OuterSyntax.improper_command "print_tcset" "print context of ZF type-checker" |
|
221 |
OuterSyntax.Keyword.diag |
|
222 |
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o (Toplevel.keep |
|
223 |
(Toplevel.node_case print_tcset (LocalTypecheckingData.print o Proof.context_of))))); |
|
224 |
||
225 |
val _ = OuterSyntax.add_parsers [print_tcsetP]; |
|
226 |
||
227 |
||
228 |
end; |
|
229 |
||
230 |
structure BasicTypeCheck: BASIC_TYPE_CHECK = TypeCheck; |
|
231 |
open BasicTypeCheck; |