1 (* Title: Pure/logic.ML
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory
4 Copyright Cambridge University 1992
6 Supporting code for defining the abstract type "thm"
13 val assum_pairs : term -> (term*term)list
14 val auto_rename : bool ref
15 val close_form : term -> term
16 val count_prems : term * int -> int
17 val dest_equals : term -> term * term
18 val dest_flexpair : term -> term * term
19 val dest_implies : term -> term * term
20 val dest_inclass : term -> typ * class
21 val dest_type : term -> typ
22 val flatten_params : int -> term -> term
23 val freeze_vars : term -> term
24 val incr_indexes : typ list * int -> term -> term
25 val lift_fns : term * int -> (term -> term) * (term -> term)
26 val list_flexpairs : (term*term)list * term -> term
27 val list_implies : term list * term -> term
28 val list_rename_params: string list * term -> term
29 val mk_equals : term * term -> term
30 val mk_flexpair : term * term -> term
31 val mk_implies : term * term -> term
32 val mk_inclass : typ * class -> term
33 val mk_type : typ -> term
34 val occs : term * term -> bool
35 val rule_of : (term*term)list * term list * term -> term
36 val set_rename_prefix : string -> unit
37 val skip_flexpairs : term -> term
38 val strip_assums_concl: term -> term
39 val strip_assums_hyp : term -> term list
40 val strip_flexpairs : term -> (term*term)list * term
41 val strip_horn : term -> (term*term)list * term list * term
42 val strip_imp_concl : term -> term
43 val strip_imp_prems : term -> term list
44 val strip_params : term -> (string * typ) list
45 val strip_prems : int * term list * term -> term list * term
46 val thaw_vars : term -> term
47 val unvarify : term -> term
48 val varify : term -> term
51 functor LogicFun (structure Unify: UNIFY and Net:NET): LOGIC =
54 structure Sign = Unify.Sign;
55 structure Type = Sign.Type;
57 (*** Abstract syntax operations on the meta-connectives ***)
62 fun mk_equals(t,u) = equals(fastype_of t) $ t $ u;
64 fun dest_equals (Const("==",_) $ t $ u) = (t,u)
65 | dest_equals t = raise TERM("dest_equals", [t]);
69 fun mk_implies(A,B) = implies $ A $ B;
71 fun dest_implies (Const("==>",_) $ A $ B) = (A,B)
72 | dest_implies A = raise TERM("dest_implies", [A]);
74 (** nested implications **)
76 (* [A1,...,An], B goes to A1==>...An==>B *)
77 fun list_implies ([], B) = B : term
78 | list_implies (A::AS, B) = implies $ A $ list_implies(AS,B);
80 (* A1==>...An==>B goes to [A1,...,An], where B is not an implication *)
81 fun strip_imp_prems (Const("==>", _) $ A $ B) = A :: strip_imp_prems B
82 | strip_imp_prems _ = [];
84 (* A1==>...An==>B goes to B, where B is not an implication *)
85 fun strip_imp_concl (Const("==>", _) $ A $ B) = strip_imp_concl B
86 | strip_imp_concl A = A : term;
88 (*Strip and return premises: (i, [], A1==>...Ai==>B)
89 goes to ([Ai, A(i-1),...,A1] , B) (REVERSED)
90 if i<0 or else i too big then raises TERM*)
91 fun strip_prems (0, As, B) = (As, B)
92 | strip_prems (i, As, Const("==>", _) $ A $ B) =
93 strip_prems (i-1, A::As, B)
94 | strip_prems (_, As, A) = raise TERM("strip_prems", A::As);
96 (*Count premises -- quicker than (length ostrip_prems) *)
97 fun count_prems (Const("==>", _) $ A $ B, n) = count_prems (B,n+1)
98 | count_prems (_,n) = n;
100 (** flex-flex constraints **)
102 (*Make a constraint.*)
103 fun mk_flexpair(t,u) = flexpair(fastype_of t) $ t $ u;
105 fun dest_flexpair (Const("=?=",_) $ t $ u) = (t,u)
106 | dest_flexpair t = raise TERM("dest_flexpair", [t]);
108 (*make flexflex antecedents: ( [(a1,b1),...,(an,bn)] , C )
109 goes to (a1=?=b1) ==>...(an=?=bn)==>C *)
110 fun list_flexpairs ([], A) = A
111 | list_flexpairs ((t,u)::pairs, A) =
112 implies $ (mk_flexpair(t,u)) $ list_flexpairs(pairs,A);
114 (*Make the object-rule tpairs==>As==>B *)
115 fun rule_of (tpairs, As, B) = list_flexpairs(tpairs, list_implies(As, B));
117 (*Remove and return flexflex pairs:
118 (a1=?=b1)==>...(an=?=bn)==>C to ( [(a1,b1),...,(an,bn)] , C )
119 [Tail recursive in order to return a pair of results] *)
120 fun strip_flex_aux (pairs, Const("==>", _) $ (Const("=?=",_)$t$u) $ C) =
121 strip_flex_aux ((t,u)::pairs, C)
122 | strip_flex_aux (pairs,C) = (rev pairs, C);
124 fun strip_flexpairs A = strip_flex_aux([], A);
126 (*Discard flexflex pairs*)
127 fun skip_flexpairs (Const("==>", _) $ (Const("=?=",_)$_$_) $ C) =
129 | skip_flexpairs C = C;
131 (*strip a proof state (Horn clause):
132 (a1==b1)==>...(am==bm)==>B1==>...Bn==>C
133 goes to ( [(a1,b1),...,(am,bm)] , [B1,...,Bn] , C) *)
135 let val (tpairs,horn) = strip_flexpairs A
136 in (tpairs, strip_imp_prems horn, strip_imp_concl horn) end;
138 (** types as terms **)
140 fun mk_type ty = Const ("TYPE", itselfT ty);
142 fun dest_type (Const ("TYPE", Type ("itself", [ty]))) = ty
143 | dest_type t = raise TERM ("dest_type", [t]);
145 (** class constraints **)
147 fun mk_inclass (ty, c) =
148 Const (Sign.const_of_class c, itselfT ty --> propT) $ mk_type ty;
150 fun dest_inclass (t as Const (c_class, _) $ ty) =
151 ((dest_type ty, Sign.class_of_const c_class)
152 handle TERM _ => raise TERM ("dest_inclass", [t]))
153 | dest_inclass t = raise TERM ("dest_inclass", [t]);
156 (*** Low-level term operations ***)
158 (*Does t occur in u? Or is alpha-convertible to u?
159 The term t must contain no loose bound variables*)
160 fun t occs u = (t aconv u) orelse
162 Abs(_,_,body) => t occs body
163 | f$t' => t occs f orelse t occs t'
166 (*Close up a formula over all free variables by quantification*)
168 list_all_free (map dest_Free (sort atless (term_frees A)),
172 (*Freeze all (T)Vars by turning them into (T)Frees*)
173 fun freeze_vars(Var(ixn,T)) = Free(Syntax.string_of_vname ixn,
175 | freeze_vars(Const(a,T)) = Const(a,Type.freeze_vars T)
176 | freeze_vars(Free(a,T)) = Free(a,Type.freeze_vars T)
177 | freeze_vars(s$t) = freeze_vars s $ freeze_vars t
178 | freeze_vars(Abs(a,T,t)) = Abs(a,Type.freeze_vars T,freeze_vars t)
179 | freeze_vars(b) = b;
181 (*Reverse the effect of freeze_vars*)
182 fun thaw_vars(Const(a,T)) = Const(a,Type.thaw_vars T)
183 | thaw_vars(Free(a,T)) =
184 let val T' = Type.thaw_vars T
186 "?"::vn => let val (ixn,_) = Syntax.scan_varname vn
190 | thaw_vars(Abs(a,T,t)) = Abs(a,Type.thaw_vars T, thaw_vars t)
191 | thaw_vars(s$t) = thaw_vars s $ thaw_vars t
195 (*** Specialized operations for resolution... ***)
197 (*For all variables in the term, increment indexnames and lift over the Us
198 result is ?Gidx(B.(lev+n-1),...,B.lev) where lev is abstraction level *)
199 fun incr_indexes (Us: typ list, inc:int) t =
200 let fun incr (Var ((a,i), T), lev) =
201 Unify.combound (Var((a, i+inc), Us---> incr_tvar inc T),
203 | incr (Abs (a,T,body), lev) =
204 Abs (a, incr_tvar inc T, incr(body,lev+1))
205 | incr (Const(a,T),_) = Const(a, incr_tvar inc T)
206 | incr (Free(a,T),_) = Free(a, incr_tvar inc T)
207 | incr (f$t, lev) = incr(f,lev) $ incr(t,lev)
211 (*Make lifting functions from subgoal and increment.
212 lift_abs operates on tpairs (unification constraints)
213 lift_all operates on propositions *)
214 fun lift_fns (B,inc) =
215 let fun lift_abs (Us, Const("==>", _) $ _ $ B) u = lift_abs (Us,B) u
216 | lift_abs (Us, Const("all",_)$Abs(a,T,t)) u =
217 Abs(a, T, lift_abs (T::Us, t) u)
218 | lift_abs (Us, _) u = incr_indexes(rev Us, inc) u
219 fun lift_all (Us, Const("==>", _) $ A $ B) u =
220 implies $ A $ lift_all (Us,B) u
221 | lift_all (Us, Const("all",_)$Abs(a,T,t)) u =
222 all T $ Abs(a, T, lift_all (T::Us,t) u)
223 | lift_all (Us, _) u = incr_indexes(rev Us, inc) u;
224 in (lift_abs([],B), lift_all([],B)) end;
226 (*Strips assumptions in goal, yielding list of hypotheses. *)
227 fun strip_assums_hyp (Const("==>", _) $ H $ B) = H :: strip_assums_hyp B
228 | strip_assums_hyp (Const("all",_)$Abs(a,T,t)) = strip_assums_hyp t
229 | strip_assums_hyp B = [];
231 (*Strips assumptions in goal, yielding conclusion. *)
232 fun strip_assums_concl (Const("==>", _) $ H $ B) = strip_assums_concl B
233 | strip_assums_concl (Const("all",_)$Abs(a,T,t)) = strip_assums_concl t
234 | strip_assums_concl B = B;
236 (*Make a list of all the parameters in a subgoal, even if nested*)
237 fun strip_params (Const("==>", _) $ H $ B) = strip_params B
238 | strip_params (Const("all",_)$Abs(a,T,t)) = (a,T) :: strip_params t
239 | strip_params B = [];
241 (*Removes the parameters from a subgoal and renumber bvars in hypotheses,
242 where j is the total number of parameters (precomputed)
243 If n>0 then deletes assumption n. *)
244 fun remove_params j n A =
245 if j=0 andalso n<=0 then A (*nothing left to do...*)
247 Const("==>", _) $ H $ B =>
248 if n=1 then (remove_params j (n-1) B)
249 else implies $ (incr_boundvars j H) $ (remove_params j (n-1) B)
250 | Const("all",_)$Abs(a,T,t) => remove_params (j-1) n t
251 | _ => if n>0 then raise TERM("remove_params", [A])
254 (** Auto-renaming of parameters in subgoals **)
256 val auto_rename = ref false
257 and rename_prefix = ref "ka";
259 (*rename_prefix is not exported; it is set by this function.*)
260 fun set_rename_prefix a =
261 if a<>"" andalso forall is_letter (explode a)
262 then (rename_prefix := a; auto_rename := true)
263 else error"rename prefix must be nonempty and consist of letters";
265 (*Makes parameters in a goal have distinctive names (not guaranteed unique!)
266 A name clash could cause the printer to rename bound vars;
267 then res_inst_tac would not work properly.*)
268 fun rename_vars (a, []) = []
269 | rename_vars (a, (_,T)::vars) =
270 (a,T) :: rename_vars (bump_string a, vars);
272 (*Move all parameters to the front of the subgoal, renaming them apart;
273 if n>0 then deletes assumption n. *)
274 fun flatten_params n A =
275 let val params = strip_params A;
276 val vars = if !auto_rename
277 then rename_vars (!rename_prefix, params)
278 else variantlist(map #1 params,[]) ~~ map #2 params
279 in list_all (vars, remove_params (length vars) n A)
282 (*Makes parameters in a goal have the names supplied by the list cs.*)
283 fun list_rename_params (cs, Const("==>", _) $ A $ B) =
284 implies $ A $ list_rename_params (cs, B)
285 | list_rename_params (c::cs, Const("all",_)$Abs(_,T,t)) =
286 all T $ Abs(c, T, list_rename_params (cs, t))
287 | list_rename_params (cs, B) = B;
289 (*Strips assumptions in goal yielding ( [Hn,...,H1], [xm,...,x1], B )
290 where H1,...,Hn are the hypotheses and x1...xm are the parameters. *)
291 fun strip_assums_aux (Hs, params, Const("==>", _) $ H $ B) =
292 strip_assums_aux (H::Hs, params, B)
293 | strip_assums_aux (Hs, params, Const("all",_)$Abs(a,T,t)) =
294 strip_assums_aux (Hs, (a,T)::params, t)
295 | strip_assums_aux (Hs, params, B) = (Hs, params, B);
297 fun strip_assums A = strip_assums_aux ([],[],A);
300 (*Produces disagreement pairs, one for each assumption proof, in order.
301 A is the first premise of the lifted rule, and thus has the form
302 H1 ==> ... Hk ==> B and the pairs are (H1,B),...,(Hk,B) *)
304 let val (Hs, params, B) = strip_assums A
305 val D = Unify.rlist_abs(params, B)
306 fun pairrev ([],pairs) = pairs
307 | pairrev (H::Hs,pairs) =
308 pairrev(Hs, (Unify.rlist_abs(params,H), D) :: pairs)
309 in pairrev (Hs,[]) (*WAS: map pair (rev Hs) *)
313 (*Converts Frees to Vars and TFrees to TVars so that axioms can be written
314 without (?) everywhere*)
315 fun varify (Const(a,T)) = Const(a, Type.varifyT T)
316 | varify (Free(a,T)) = Var((a,0), Type.varifyT T)
317 | varify (Var(ixn,T)) = Var(ixn, Type.varifyT T)
318 | varify (Abs (a,T,body)) = Abs (a, Type.varifyT T, varify body)
319 | varify (f$t) = varify f $ varify t
322 (*Inverse of varify. Converts axioms back to their original form.*)
323 fun unvarify (Const(a,T)) = Const(a, Type.unvarifyT T)
324 | unvarify (Var((a,0), T)) = Free(a, Type.unvarifyT T)
325 | unvarify (Var(ixn,T)) = Var(ixn, Type.unvarifyT T) (*non-0 index!*)
326 | unvarify (Abs (a,T,body)) = Abs (a, Type.unvarifyT T, unvarify body)
327 | unvarify (f$t) = unvarify f $ unvarify t