| author | haftmann | 
| Wed, 04 Oct 2006 14:17:38 +0200 | |
| changeset 20854 | f9cf9e62d11c | 
| parent 20664 | ffbc5a57191a | 
| child 21295 | 63552bc99cfb | 
| permissions | -rw-r--r-- | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1  | 
(* Title: Provers/blast.ML  | 
| 2894 | 2  | 
ID: $Id$  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 3083 | 4  | 
Copyright 1997 University of Cambridge  | 
| 2894 | 5  | 
|
6  | 
Generic tableau prover with proof reconstruction  | 
|
7  | 
||
| 2854 | 8  | 
SKOLEMIZES ReplaceI WRONGLY: allow new vars in prems, or forbid such rules??  | 
| 2894 | 9  | 
Needs explicit instantiation of assumptions?  | 
10  | 
||
| 18171 | 11  | 
Given the typeargs system, constructor Const could be eliminated, with  | 
12  | 
TConst replaced by a constructor that takes the typargs list as an argument.  | 
|
13  | 
However, Const is heavily used for logical connectives.  | 
|
| 2894 | 14  | 
|
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
15  | 
Blast_tac is often more powerful than fast_tac, but has some limitations.  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
16  | 
Blast_tac...  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
17  | 
* ignores wrappers (addss, addbefore, addafter, addWrapper, ...);  | 
| 4651 | 18  | 
this restriction is intrinsic  | 
| 2894 | 19  | 
* ignores elimination rules that don't have the correct format  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
20  | 
(conclusion must be a formula variable)  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
21  | 
* rules must not require higher-order unification, e.g. apply_type in ZF  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
22  | 
+ message "Function Var's argument not a bound variable" relates to this  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
23  | 
* its proof strategy is more general but can actually be slower  | 
| 2894 | 24  | 
|
25  | 
Known problems:  | 
|
| 3092 | 26  | 
"Recursive" chains of rules can sometimes exclude other unsafe formulae  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
27  | 
from expansion. This happens because newly-created formulae always  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
28  | 
have priority over existing ones. But obviously recursive rules  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
29  | 
such as transitivity are treated specially to prevent this. Sometimes  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
30  | 
the formulae get into the wrong order (see WRONG below).  | 
| 
3021
 
39806db47be9
Loop detection: before expanding a haz formula, see whether it is a duplicate
 
paulson 
parents: 
2999 
diff
changeset
 | 
31  | 
|
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
32  | 
With substition for equalities (hyp_subst_tac):  | 
| 3092 | 33  | 
When substitution affects a haz formula or literal, it is moved  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
34  | 
back to the list of safe formulae.  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
35  | 
But there's no way of putting it in the right place. A "moved" or  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
36  | 
"no DETERM" flag would prevent proofs failing here.  | 
| 2854 | 37  | 
*)  | 
38  | 
||
39  | 
(*Should be a type abbreviation?*)  | 
|
40  | 
type netpair = (int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net;  | 
|
41  | 
||
42  | 
signature BLAST_DATA =  | 
|
43  | 
sig  | 
|
44  | 
type claset  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
45  | 
val equality_name: string  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
46  | 
val not_name: string  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
47  | 
val notE : thm (* [| ~P; P |] ==> R *)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
48  | 
val ccontr : thm  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
49  | 
val contr_tac : int -> tactic  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
50  | 
val dup_intr : thm -> thm  | 
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
51  | 
val hyp_subst_tac : bool ref -> int -> tactic  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
52  | 
val claset : unit -> claset  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
53  | 
val rep_cs : (* dependent on classical.ML *)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
54  | 
      claset -> {safeIs: thm list, safeEs: thm list,
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
55  | 
hazIs: thm list, hazEs: thm list,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
56  | 
swrappers: (string * wrapper) list,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
57  | 
uwrappers: (string * wrapper) list,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
58  | 
safe0_netpair: netpair, safep_netpair: netpair,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
59  | 
haz_netpair: netpair, dup_netpair: netpair, xtra_netpair: ContextRules.netpair}  | 
| 7272 | 60  | 
val cla_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list  | 
| 7559 | 61  | 
val cla_meth': (claset -> int -> tactic) -> thm list -> Proof.context -> Proof.method  | 
| 2854 | 62  | 
end;  | 
63  | 
||
64  | 
||
65  | 
signature BLAST =  | 
|
66  | 
sig  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
67  | 
type claset  | 
| 
4233
 
85d004a96b98
Rationalized error handling: if low-level tactic (depth_tac) cannot accept the
 
paulson 
parents: 
4196 
diff
changeset
 | 
68  | 
exception TRANS of string (*reports translation errors*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
69  | 
datatype term =  | 
| 18177 | 70  | 
Const of string * term list  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
71  | 
| Skolem of string * term option ref list  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
72  | 
| Free of string  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
73  | 
| Var of term option ref  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
74  | 
| Bound of int  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
75  | 
| Abs of string*term  | 
| 17795 | 76  | 
| $ of term*term;  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
77  | 
type branch  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
78  | 
val depth_tac : claset -> int -> int -> tactic  | 
| 
15162
 
670ab8497818
depth limit (previously hard-coded with a value of 20) made a reference
 
webertj 
parents: 
14984 
diff
changeset
 | 
79  | 
val depth_limit : int ref  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
80  | 
val blast_tac : claset -> int -> tactic  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
81  | 
val Blast_tac : int -> tactic  | 
| 18708 | 82  | 
val setup : theory -> theory  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
83  | 
(*debugging tools*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
84  | 
val stats : bool ref  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
85  | 
val trace : bool ref  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
86  | 
val fullTrace : branch list list ref  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
87  | 
val fromType : (indexname * term) list ref -> Term.typ -> term  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
88  | 
val fromTerm : Term.term -> term  | 
| 4065 | 89  | 
val fromSubgoal : Term.term -> term  | 
90  | 
val instVars : term -> (unit -> unit)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
91  | 
val toTerm : int -> term -> Term.term  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
92  | 
val readGoal : theory -> string -> term  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
93  | 
val tryInThy : theory -> int -> string ->  | 
| 3083 | 94  | 
(int->tactic) list * branch list list * (int*int*exn) list  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
95  | 
val trygl : claset -> int -> int ->  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
96  | 
(int->tactic) list * branch list list * (int*int*exn) list  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
97  | 
val Trygl : int -> int ->  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
98  | 
(int->tactic) list * branch list list * (int*int*exn) list  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
99  | 
val normBr : branch -> branch  | 
| 2854 | 100  | 
end;  | 
101  | 
||
102  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
103  | 
functor BlastFun(Data: BLAST_DATA) : BLAST =  | 
| 2854 | 104  | 
struct  | 
105  | 
||
106  | 
type claset = Data.claset;  | 
|
107  | 
||
| 
4323
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
108  | 
val trace = ref false  | 
| 
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
109  | 
and stats = ref false; (*for runtime and search statistics*)  | 
| 2854 | 110  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
111  | 
datatype term =  | 
| 18177 | 112  | 
Const of string * term list (*typargs constant--as a terms!*)  | 
| 2854 | 113  | 
| Skolem of string * term option ref list  | 
| 
5343
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
114  | 
| Free of string  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
115  | 
| Var of term option ref  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
116  | 
| Bound of int  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
117  | 
| Abs of string*term  | 
| 5613 | 118  | 
| op $ of term*term;  | 
| 2854 | 119  | 
|
120  | 
||
121  | 
(** Basic syntactic operations **)  | 
|
122  | 
||
123  | 
fun is_Var (Var _) = true  | 
|
124  | 
| is_Var _ = false;  | 
|
125  | 
||
126  | 
fun dest_Var (Var x) = x;  | 
|
127  | 
||
128  | 
fun rand (f$x) = x;  | 
|
129  | 
||
130  | 
(* maps (f, [t1,...,tn]) to f(t1,...,tn) *)  | 
|
| 15570 | 131  | 
val list_comb : term * term list -> term = Library.foldl (op $);  | 
| 2854 | 132  | 
|
133  | 
(* maps f(t1,...,tn) to (f, [t1,...,tn]) ; naturally tail-recursive*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
134  | 
fun strip_comb u : term * term list =  | 
| 2854 | 135  | 
let fun stripc (f$t, ts) = stripc (f, t::ts)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
136  | 
| stripc x = x  | 
| 2854 | 137  | 
in stripc(u,[]) end;  | 
138  | 
||
139  | 
(* maps f(t1,...,tn) to f , which is never a combination *)  | 
|
140  | 
fun head_of (f$t) = head_of f  | 
|
141  | 
| head_of u = u;  | 
|
142  | 
||
143  | 
||
144  | 
(** Particular constants **)  | 
|
145  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
146  | 
fun negate P = Const (Data.not_name, []) $ P;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
147  | 
|
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
148  | 
fun isNot (Const (c, _) $ _) = c = Data.not_name  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
149  | 
| isNot _ = false;  | 
| 2854 | 150  | 
|
| 18177 | 151  | 
fun mkGoal P = Const ("*Goal*", []) $ P;
 | 
| 2854 | 152  | 
|
| 18177 | 153  | 
fun isGoal (Const ("*Goal*", _) $ _) = true
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
154  | 
| isGoal _ = false;  | 
| 2854 | 155  | 
|
| 18177 | 156  | 
val TruepropC = ObjectLogic.judgment_name (the_context ());  | 
157  | 
val TruepropT = Sign.the_const_type (the_context ()) TruepropC;  | 
|
| 18171 | 158  | 
|
| 18177 | 159  | 
fun mk_Trueprop t = Term.$ (Term.Const (TruepropC, TruepropT), t);  | 
| 2854 | 160  | 
|
| 18177 | 161  | 
fun strip_Trueprop (tm as Const (c, _) $ t) = if c = TruepropC then t else tm  | 
162  | 
| strip_Trueprop tm = tm;  | 
|
163  | 
||
| 2854 | 164  | 
|
165  | 
||
| 4065 | 166  | 
(*** Dealing with overloaded constants ***)  | 
| 2854 | 167  | 
|
| 4065 | 168  | 
(*alist is a map from TVar names to Vars. We need to unify the TVars  | 
169  | 
faithfully in order to track overloading*)  | 
|
| 18177 | 170  | 
fun fromType alist (Term.Type(a,Ts)) = list_comb (Const (a, []), map (fromType alist) Ts)  | 
| 4065 | 171  | 
| fromType alist (Term.TFree(a,_)) = Free a  | 
172  | 
| fromType alist (Term.TVar (ixn,_)) =  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
173  | 
(case (AList.lookup (op =) (!alist) ixn) of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
174  | 
NONE => let val t' = Var(ref NONE)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
175  | 
in alist := (ixn, t') :: !alist; t'  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
176  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
177  | 
| SOME v => v)  | 
| 2854 | 178  | 
|
| 17993 | 179  | 
(*refer to the theory in which blast is initialized*)  | 
| 18146 | 180  | 
val typargs = ref (fn ((_, T): string * typ) => [T]);  | 
| 2854 | 181  | 
|
| 18177 | 182  | 
fun fromConst alist (a, T) =  | 
183  | 
Const (a, map (fromType alist) (! typargs (a, T)));  | 
|
| 2854 | 184  | 
|
185  | 
||
186  | 
(*Tests whether 2 terms are alpha-convertible; chases instantiations*)  | 
|
| 18177 | 187  | 
fun (Const (a, ts)) aconv (Const (b, us)) = a=b andalso aconvs (ts, us)  | 
| 2854 | 188  | 
| (Skolem (a,_)) aconv (Skolem (b,_)) = a=b (*arglists must then be equal*)  | 
189  | 
| (Free a) aconv (Free b) = a=b  | 
|
| 15531 | 190  | 
| (Var(ref(SOME t))) aconv u = t aconv u  | 
191  | 
| t aconv (Var(ref(SOME u))) = t aconv u  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
192  | 
| (Var v) aconv (Var w) = v=w (*both Vars are un-assigned*)  | 
| 2854 | 193  | 
| (Bound i) aconv (Bound j) = i=j  | 
194  | 
| (Abs(_,t)) aconv (Abs(_,u)) = t aconv u  | 
|
195  | 
| (f$t) aconv (g$u) = (f aconv g) andalso (t aconv u)  | 
|
| 18177 | 196  | 
| _ aconv _ = false  | 
197  | 
and aconvs ([], []) = true  | 
|
198  | 
| aconvs (t :: ts, u :: us) = t aconv u andalso aconvs (ts, us)  | 
|
199  | 
| aconvs _ = false;  | 
|
| 2854 | 200  | 
|
201  | 
||
202  | 
fun mem_term (_, []) = false  | 
|
203  | 
| mem_term (t, t'::ts) = t aconv t' orelse mem_term(t,ts);  | 
|
204  | 
||
205  | 
fun ins_term(t,ts) = if mem_term(t,ts) then ts else t :: ts;  | 
|
206  | 
||
207  | 
fun mem_var (v: term option ref, []) = false  | 
|
208  | 
| mem_var (v, v'::vs) = v=v' orelse mem_var(v,vs);  | 
|
209  | 
||
210  | 
fun ins_var(v,vs) = if mem_var(v,vs) then vs else v :: vs;  | 
|
211  | 
||
212  | 
||
213  | 
(** Vars **)  | 
|
214  | 
||
215  | 
(*Accumulates the Vars in the term, suppressing duplicates*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
216  | 
fun add_term_vars (Skolem(a,args), vars) = add_vars_vars(args,vars)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
217  | 
| add_term_vars (Var (v as ref NONE), vars) = ins_var (v, vars)  | 
| 15531 | 218  | 
| add_term_vars (Var (ref (SOME u)), vars) = add_term_vars(u,vars)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
219  | 
| add_term_vars (Const (_,ts), vars) = add_terms_vars(ts,vars)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
220  | 
| add_term_vars (Abs (_,body), vars) = add_term_vars(body,vars)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
221  | 
| add_term_vars (f$t, vars) = add_term_vars (f, add_term_vars(t, vars))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
222  | 
| add_term_vars (_, vars) = vars  | 
| 2854 | 223  | 
(*Term list version. [The fold functionals are slow]*)  | 
224  | 
and add_terms_vars ([], vars) = vars  | 
|
225  | 
| add_terms_vars (t::ts, vars) = add_terms_vars (ts, add_term_vars(t,vars))  | 
|
226  | 
(*Var list version.*)  | 
|
227  | 
and add_vars_vars ([], vars) = vars  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
228  | 
| add_vars_vars (ref (SOME u) :: vs, vars) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
229  | 
add_vars_vars (vs, add_term_vars(u,vars))  | 
| 15531 | 230  | 
| add_vars_vars (v::vs, vars) = (*v must be a ref NONE*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
231  | 
add_vars_vars (vs, ins_var (v, vars));  | 
| 2854 | 232  | 
|
233  | 
||
234  | 
(*Chase assignments in "vars"; return a list of unassigned variables*)  | 
|
235  | 
fun vars_in_vars vars = add_vars_vars(vars,[]);  | 
|
236  | 
||
237  | 
||
238  | 
||
239  | 
(*increment a term's non-local bound variables  | 
|
240  | 
inc is increment for bound variables  | 
|
241  | 
lev is level at which a bound variable is considered 'loose'*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
242  | 
fun incr_bv (inc, lev, u as Bound i) = if i>=lev then Bound(i+inc) else u  | 
| 2854 | 243  | 
| incr_bv (inc, lev, Abs(a,body)) = Abs(a, incr_bv(inc,lev+1,body))  | 
244  | 
| incr_bv (inc, lev, f$t) = incr_bv(inc,lev,f) $ incr_bv(inc,lev,t)  | 
|
245  | 
| incr_bv (inc, lev, u) = u;  | 
|
246  | 
||
247  | 
fun incr_boundvars 0 t = t  | 
|
248  | 
| incr_boundvars inc t = incr_bv(inc,0,t);  | 
|
249  | 
||
250  | 
||
251  | 
(*Accumulate all 'loose' bound vars referring to level 'lev' or beyond.  | 
|
252  | 
(Bound 0) is loose at level 0 *)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
253  | 
fun add_loose_bnos (Bound i, lev, js) = if i<lev then js  | 
| 20854 | 254  | 
else insert (op =) (i - lev) js  | 
| 2854 | 255  | 
| add_loose_bnos (Abs (_,t), lev, js) = add_loose_bnos (t, lev+1, js)  | 
256  | 
| add_loose_bnos (f$t, lev, js) =  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
257  | 
add_loose_bnos (f, lev, add_loose_bnos (t, lev, js))  | 
| 2854 | 258  | 
| add_loose_bnos (_, _, js) = js;  | 
259  | 
||
260  | 
fun loose_bnos t = add_loose_bnos (t, 0, []);  | 
|
261  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
262  | 
fun subst_bound (arg, t) : term =  | 
| 2854 | 263  | 
let fun subst (t as Bound i, lev) =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
264  | 
if i<lev then t (*var is locally bound*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
265  | 
else if i=lev then incr_boundvars lev arg  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
266  | 
else Bound(i-1) (*loose: change it*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
267  | 
| subst (Abs(a,body), lev) = Abs(a, subst(body,lev+1))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
268  | 
| subst (f$t, lev) = subst(f,lev) $ subst(t,lev)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
269  | 
| subst (t,lev) = t  | 
| 2854 | 270  | 
in subst (t,0) end;  | 
271  | 
||
272  | 
||
| 
3101
 
e8a92f497295
Again "norm" DOES NOT normalize bodies of abstractions
 
paulson 
parents: 
3092 
diff
changeset
 | 
273  | 
(*Normalize...but not the bodies of ABSTRACTIONS*)  | 
| 2854 | 274  | 
fun norm t = case t of  | 
| 2952 | 275  | 
Skolem (a,args) => Skolem(a, vars_in_vars args)  | 
| 18177 | 276  | 
| Const(a,ts) => Const(a, map norm ts)  | 
| 15531 | 277  | 
| (Var (ref NONE)) => t  | 
278  | 
| (Var (ref (SOME u))) => norm u  | 
|
| 2854 | 279  | 
| (f $ u) => (case norm f of  | 
| 
3101
 
e8a92f497295
Again "norm" DOES NOT normalize bodies of abstractions
 
paulson 
parents: 
3092 
diff
changeset
 | 
280  | 
Abs(_,body) => norm (subst_bound (u, body))  | 
| 
 
e8a92f497295
Again "norm" DOES NOT normalize bodies of abstractions
 
paulson 
parents: 
3092 
diff
changeset
 | 
281  | 
| nf => nf $ norm u)  | 
| 2854 | 282  | 
| _ => t;  | 
283  | 
||
284  | 
||
285  | 
(*Weak (one-level) normalize for use in unification*)  | 
|
286  | 
fun wkNormAux t = case t of  | 
|
287  | 
(Var v) => (case !v of  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
288  | 
SOME u => wkNorm u  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
289  | 
| NONE => t)  | 
| 2854 | 290  | 
| (f $ u) => (case wkNormAux f of  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
291  | 
Abs(_,body) => wkNorm (subst_bound (u, body))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
292  | 
| nf => nf $ u)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
293  | 
| Abs (a,body) => (*eta-contract if possible*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
294  | 
(case wkNormAux body of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
295  | 
nb as (f $ t) =>  | 
| 20664 | 296  | 
if member (op =) (loose_bnos f) 0 orelse wkNorm t <> Bound 0  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
297  | 
then Abs(a,nb)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
298  | 
else wkNorm (incr_boundvars ~1 f)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
299  | 
| nb => Abs (a,nb))  | 
| 2854 | 300  | 
| _ => t  | 
301  | 
and wkNorm t = case head_of t of  | 
|
302  | 
Const _ => t  | 
|
303  | 
| Skolem(a,args) => t  | 
|
304  | 
| Free _ => t  | 
|
305  | 
| _ => wkNormAux t;  | 
|
306  | 
||
307  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
308  | 
(*Does variable v occur in u? For unification.  | 
| 
5734
 
bebd10cb7a8d
occurs check now handles Bound variables (for soundness)
 
paulson 
parents: 
5613 
diff
changeset
 | 
309  | 
Dangling bound vars are also forbidden.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
310  | 
fun varOccur v =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
311  | 
let fun occL lev [] = false (*same as (exists occ), but faster*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
312  | 
| occL lev (u::us) = occ lev u orelse occL lev us  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
313  | 
and occ lev (Var w) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
314  | 
v=w orelse  | 
| 15531 | 315  | 
(case !w of NONE => false  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
316  | 
| SOME u => occ lev u)  | 
| 
5734
 
bebd10cb7a8d
occurs check now handles Bound variables (for soundness)
 
paulson 
parents: 
5613 
diff
changeset
 | 
317  | 
| occ lev (Skolem(_,args)) = occL lev (map Var args)  | 
| 18177 | 318  | 
(*ignore Const, since term variables can't occur in types (?) *)  | 
| 
5734
 
bebd10cb7a8d
occurs check now handles Bound variables (for soundness)
 
paulson 
parents: 
5613 
diff
changeset
 | 
319  | 
| occ lev (Bound i) = lev <= i  | 
| 
 
bebd10cb7a8d
occurs check now handles Bound variables (for soundness)
 
paulson 
parents: 
5613 
diff
changeset
 | 
320  | 
| occ lev (Abs(_,u)) = occ (lev+1) u  | 
| 
 
bebd10cb7a8d
occurs check now handles Bound variables (for soundness)
 
paulson 
parents: 
5613 
diff
changeset
 | 
321  | 
| occ lev (f$u) = occ lev u orelse occ lev f  | 
| 
 
bebd10cb7a8d
occurs check now handles Bound variables (for soundness)
 
paulson 
parents: 
5613 
diff
changeset
 | 
322  | 
| occ lev _ = false;  | 
| 
 
bebd10cb7a8d
occurs check now handles Bound variables (for soundness)
 
paulson 
parents: 
5613 
diff
changeset
 | 
323  | 
in occ 0 end;  | 
| 2854 | 324  | 
|
325  | 
exception UNIFY;  | 
|
326  | 
||
327  | 
val trail = ref [] : term option ref list ref;  | 
|
328  | 
val ntrail = ref 0;  | 
|
329  | 
||
330  | 
||
331  | 
(*Restore the trail to some previous state: for backtracking*)  | 
|
332  | 
fun clearTo n =  | 
|
| 3083 | 333  | 
while !ntrail<>n do  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
334  | 
(hd(!trail) := NONE;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
335  | 
trail := tl (!trail);  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
336  | 
ntrail := !ntrail - 1);  | 
| 2854 | 337  | 
|
338  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
339  | 
(*First-order unification with bound variables.  | 
| 2854 | 340  | 
"vars" is a list of variables local to the rule and NOT to be put  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
341  | 
on the trail (no point in doing so)  | 
| 2854 | 342  | 
*)  | 
| 4065 | 343  | 
fun unify(vars,t,u) =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
344  | 
let val n = !ntrail  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
345  | 
fun update (t as Var v, u) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
346  | 
if t aconv u then ()  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
347  | 
else if varOccur v u then raise UNIFY  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
348  | 
else if mem_var(v, vars) then v := SOME u  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
349  | 
else (*avoid updating Vars in the branch if possible!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
350  | 
if is_Var u andalso mem_var(dest_Var u, vars)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
351  | 
then dest_Var u := SOME t  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
352  | 
else (v := SOME u;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
353  | 
trail := v :: !trail; ntrail := !ntrail + 1)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
354  | 
fun unifyAux (t,u) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
355  | 
case (wkNorm t, wkNorm u) of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
356  | 
(nt as Var v, nu) => update(nt,nu)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
357  | 
| (nu, nt as Var v) => update(nt,nu)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
358  | 
| (Const(a,ats), Const(b,bts)) => if a=b then unifysAux(ats,bts)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
359  | 
else raise UNIFY  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
360  | 
| (Abs(_,t'), Abs(_,u')) => unifyAux(t',u')  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
361  | 
(*NB: can yield unifiers having dangling Bound vars!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
362  | 
| (f$t', g$u') => (unifyAux(f,g); unifyAux(t',u'))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
363  | 
| (nt, nu) => if nt aconv nu then () else raise UNIFY  | 
| 18177 | 364  | 
and unifysAux ([], []) = ()  | 
365  | 
| unifysAux (t :: ts, u :: us) = (unifyAux (t, u); unifysAux (ts, us))  | 
|
366  | 
| unifysAux _ = raise UNIFY;  | 
|
| 3083 | 367  | 
in (unifyAux(t,u); true) handle UNIFY => (clearTo n; false)  | 
| 2854 | 368  | 
end;  | 
369  | 
||
370  | 
||
| 
16774
 
515b6020cf5d
experimental code to reduce the amount of type information in blast
 
paulson 
parents: 
15786 
diff
changeset
 | 
371  | 
(*Convert from "real" terms to prototerms; eta-contract.  | 
| 
 
515b6020cf5d
experimental code to reduce the amount of type information in blast
 
paulson 
parents: 
15786 
diff
changeset
 | 
372  | 
Code is similar to fromSubgoal.*)  | 
| 4065 | 373  | 
fun fromTerm t =  | 
374  | 
let val alistVar = ref []  | 
|
375  | 
and alistTVar = ref []  | 
|
376  | 
fun from (Term.Const aT) = fromConst alistTVar aT  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
377  | 
| from (Term.Free (a,_)) = Free a  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
378  | 
| from (Term.Bound i) = Bound i  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
379  | 
| from (Term.Var (ixn,T)) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
380  | 
(case (AList.lookup (op =) (!alistVar) ixn) of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
381  | 
NONE => let val t' = Var(ref NONE)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
382  | 
in alistVar := (ixn, t') :: !alistVar; t'  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
383  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
384  | 
| SOME v => v)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
385  | 
| from (Term.Abs (a,_,u)) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
386  | 
(case from u of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
387  | 
u' as (f $ Bound 0) =>  | 
| 20664 | 388  | 
if member (op =) (loose_bnos f) 0 then Abs(a,u')  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
389  | 
else incr_boundvars ~1 f  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
390  | 
| u' => Abs(a,u'))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
391  | 
| from (Term.$ (f,u)) = from f $ from u  | 
| 2854 | 392  | 
in from t end;  | 
393  | 
||
| 4065 | 394  | 
(*A debugging function: replaces all Vars by dummy Frees for visual inspection  | 
395  | 
of whether they are distinct. Function revert undoes the assignments.*)  | 
|
396  | 
fun instVars t =  | 
|
| 12902 | 397  | 
let val name = ref "a"  | 
| 4065 | 398  | 
val updated = ref []  | 
| 18177 | 399  | 
fun inst (Const(a,ts)) = List.app inst ts  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
400  | 
| inst (Var(v as ref NONE)) = (updated := v :: (!updated);  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
401  | 
                                       v       := SOME (Free ("?" ^ !name));
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
402  | 
name := Symbol.bump_string (!name))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
403  | 
| inst (Abs(a,t)) = inst t  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
404  | 
| inst (f $ u) = (inst f; inst u)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
405  | 
| inst _ = ()  | 
| 15570 | 406  | 
fun revert() = List.app (fn v => v:=NONE) (!updated)  | 
| 4065 | 407  | 
in inst t; revert end;  | 
408  | 
||
409  | 
||
| 2854 | 410  | 
(* A1==>...An==>B goes to [A1,...,An], where B is not an implication *)  | 
| 18177 | 411  | 
fun strip_imp_prems (Const ("==>", _) $ A $ B) = strip_Trueprop A :: strip_imp_prems B
 | 
| 2854 | 412  | 
| strip_imp_prems _ = [];  | 
413  | 
||
414  | 
(* A1==>...An==>B goes to B, where B is not an implication *)  | 
|
| 18177 | 415  | 
fun strip_imp_concl (Const ("==>", _) $ A $ B) = strip_imp_concl B
 | 
416  | 
| strip_imp_concl A = strip_Trueprop A;  | 
|
417  | 
||
| 2854 | 418  | 
|
419  | 
||
420  | 
(*** Conversion of Elimination Rules to Tableau Operations ***)  | 
|
421  | 
||
| 9170 | 422  | 
exception ElimBadConcl and ElimBadPrem;  | 
423  | 
||
424  | 
(*The conclusion becomes the goal/negated assumption *False*: delete it!  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
425  | 
If we don't find it then the premise is ill-formed and could cause  | 
| 9170 | 426  | 
PROOF FAILED*)  | 
427  | 
fun delete_concl [] = raise ElimBadPrem  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
428  | 
| delete_concl (P :: Ps) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
429  | 
(case P of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
430  | 
        Const (c, _) $ Var (ref (SOME (Const ("*False*", _)))) =>
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
431  | 
if c = "*Goal*" orelse c = Data.not_name then Ps  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
432  | 
else P :: delete_concl Ps  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
433  | 
| _ => P :: delete_concl Ps);  | 
| 2854 | 434  | 
|
| 18177 | 435  | 
fun skoPrem vars (Const ("all", _) $ Abs (_, P)) =
 | 
| 2854 | 436  | 
skoPrem vars (subst_bound (Skolem (gensym "S_", vars), P))  | 
437  | 
| skoPrem vars P = P;  | 
|
438  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
439  | 
fun convertPrem t =  | 
| 9170 | 440  | 
delete_concl (mkGoal (strip_imp_concl t) :: strip_imp_prems t);  | 
| 2854 | 441  | 
|
442  | 
(*Expects elimination rules to have a formula variable as conclusion*)  | 
|
443  | 
fun convertRule vars t =  | 
|
444  | 
let val (P::Ps) = strip_imp_prems t  | 
|
445  | 
val Var v = strip_imp_concl t  | 
|
| 18177 | 446  | 
  in  v := SOME (Const ("*False*", []));
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
447  | 
(P, map (convertPrem o skoPrem vars) Ps)  | 
| 9170 | 448  | 
end  | 
449  | 
handle Bind => raise ElimBadConcl;  | 
|
| 2854 | 450  | 
|
451  | 
||
452  | 
(*Like dup_elim, but puts the duplicated major premise FIRST*)  | 
|
| 
17257
 
0ab67cb765da
introduced binding priority 1 for linear combinators etc.
 
haftmann 
parents: 
16976 
diff
changeset
 | 
453  | 
fun rev_dup_elim th = (th RSN (2, revcut_rl)) |> assumption 2 |> Seq.hd;  | 
| 2854 | 454  | 
|
455  | 
||
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
456  | 
(*Rotate the assumptions in all new subgoals for the LIFO discipline*)  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
457  | 
local  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
458  | 
(*Count new hyps so that they can be rotated*)  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
459  | 
fun nNewHyps [] = 0  | 
| 18177 | 460  | 
    | nNewHyps (Const ("*Goal*", _) $ _ :: Ps) = nNewHyps Ps
 | 
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
461  | 
| nNewHyps (P::Ps) = 1 + nNewHyps Ps;  | 
| 2854 | 462  | 
|
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
463  | 
fun rot_tac [] i st = Seq.single st  | 
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
464  | 
| rot_tac (0::ks) i st = rot_tac ks (i+1) st  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
465  | 
| rot_tac (k::ks) i st = rot_tac ks (i+1) (rotate_rule (~k) i st);  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
466  | 
in  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
467  | 
fun rot_subgoals_tac (rot, rl) =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
468  | 
rot_tac (if rot then map nNewHyps rl else [])  | 
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
469  | 
end;  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
470  | 
|
| 2854 | 471  | 
|
| 
2999
 
fad7cc426242
Penalty for branching instantiations reduced from log3 to log4.
 
paulson 
parents: 
2952 
diff
changeset
 | 
472  | 
fun TRACE rl tac st i = if !trace then (prth rl; tac st i) else tac st i;  | 
| 2854 | 473  | 
|
| 
5343
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
474  | 
(*Resolution/matching tactics: if upd then the proof state may be updated.  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
475  | 
Matching makes the tactics more deterministic in the presence of Vars.*)  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
476  | 
fun emtac upd rl = TRACE rl (if upd then etac rl else ematch_tac [rl]);  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
477  | 
fun rmtac upd rl = TRACE rl (if upd then rtac rl else match_tac [rl]);  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
478  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
479  | 
(*Tableau rule from elimination rule.  | 
| 
5343
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
480  | 
Flag "upd" says that the inference updated the branch.  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
481  | 
Flag "dup" requests duplication of the affected formula.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
482  | 
fun fromRule vars rl =  | 
| 4065 | 483  | 
let val trl = rl |> rep_thm |> #prop |> fromTerm |> convertRule vars  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
484  | 
fun tac (upd, dup,rot) i =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
485  | 
emtac upd (if dup then rev_dup_elim rl else rl) i  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
486  | 
THEN  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
487  | 
rot_subgoals_tac (rot, #2 trl) i  | 
| 
3244
 
71b760618f30
Basis library version of type "option" now resides in its own structure Option
 
paulson 
parents: 
3101 
diff
changeset
 | 
488  | 
in Option.SOME (trl, tac) end  | 
| 9170 | 489  | 
handle ElimBadPrem => (*reject: prems don't preserve conclusion*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
490  | 
            (warning("Ignoring weak elimination rule\n" ^ string_of_thm rl);
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
491  | 
Option.NONE)  | 
| 9170 | 492  | 
| ElimBadConcl => (*ignore: conclusion is not just a variable*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
493  | 
           (if !trace then (warning("Ignoring ill-formed elimination rule:\n" ^
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
494  | 
"conclusion should be a variable\n" ^ string_of_thm rl))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
495  | 
else ();  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
496  | 
Option.NONE);  | 
| 2854 | 497  | 
|
498  | 
||
| 
3101
 
e8a92f497295
Again "norm" DOES NOT normalize bodies of abstractions
 
paulson 
parents: 
3092 
diff
changeset
 | 
499  | 
(*** Conversion of Introduction Rules ***)  | 
| 2854 | 500  | 
|
501  | 
fun convertIntrPrem t = mkGoal (strip_imp_concl t) :: strip_imp_prems t;  | 
|
502  | 
||
503  | 
fun convertIntrRule vars t =  | 
|
504  | 
let val Ps = strip_imp_prems t  | 
|
505  | 
val P = strip_imp_concl t  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
506  | 
in (mkGoal P, map (convertIntrPrem o skoPrem vars) Ps)  | 
| 2854 | 507  | 
end;  | 
508  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
509  | 
(*Tableau rule from introduction rule.  | 
| 
5343
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
510  | 
Flag "upd" says that the inference updated the branch.  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
511  | 
Flag "dup" requests duplication of the affected formula.  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
512  | 
Since haz rules are now delayed, "dup" is always FALSE for  | 
| 
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
513  | 
introduction rules.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
514  | 
fun fromIntrRule vars rl =  | 
| 4065 | 515  | 
let val trl = rl |> rep_thm |> #prop |> fromTerm |> convertIntrRule vars  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
516  | 
fun tac (upd,dup,rot) i =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
517  | 
rmtac upd (if dup then Data.dup_intr rl else rl) i  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
518  | 
THEN  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
519  | 
rot_subgoals_tac (rot, #2 trl) i  | 
| 2854 | 520  | 
in (trl, tac) end;  | 
521  | 
||
522  | 
||
| 3030 | 523  | 
val dummyVar = Term.Var (("etc",0), dummyT);
 | 
| 2854 | 524  | 
|
525  | 
(*Convert from prototerms to ordinary terms with dummy types  | 
|
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
526  | 
Ignore abstractions; identify all Vars; STOP at given depth*)  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
527  | 
fun toTerm 0 _ = dummyVar  | 
| 18177 | 528  | 
| toTerm d (Const(a,_)) = Term.Const (a,dummyT) (*no need to convert typargs*)  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
529  | 
| toTerm d (Skolem(a,_)) = Term.Const (a,dummyT)  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
530  | 
| toTerm d (Free a) = Term.Free (a,dummyT)  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
531  | 
| toTerm d (Bound i) = Term.Bound i  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
532  | 
| toTerm d (Var _) = dummyVar  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
533  | 
| toTerm d (Abs(a,_)) = dummyVar  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
534  | 
| toTerm d (f $ u) = Term.$ (toTerm d f, toTerm (d-1) u);  | 
| 2854 | 535  | 
|
536  | 
||
537  | 
fun netMkRules P vars (nps: netpair list) =  | 
|
538  | 
case P of  | 
|
| 18177 | 539  | 
      (Const ("*Goal*", _) $ G) =>
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
540  | 
let val pG = mk_Trueprop (toTerm 2 G)  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19037 
diff
changeset
 | 
541  | 
val intrs = maps (fn (inet,_) => Net.unify_term inet pG) nps  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
542  | 
in map (fromIntrRule vars o #2) (Tactic.orderlist intrs) end  | 
| 2854 | 543  | 
| _ =>  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
544  | 
let val pP = mk_Trueprop (toTerm 3 P)  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19037 
diff
changeset
 | 
545  | 
val elims = maps (fn (_,enet) => Net.unify_term enet pP) nps  | 
| 
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19037 
diff
changeset
 | 
546  | 
in map_filter (fromRule vars o #2) (Tactic.orderlist elims) end;  | 
| 2854 | 547  | 
|
| 3092 | 548  | 
|
549  | 
(*Pending formulae carry md (may duplicate) flags*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
550  | 
type branch =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
551  | 
    {pairs: ((term*bool) list * (*safe formulae on this level*)
 | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
552  | 
(term*bool) list) list, (*haz formulae on this level*)  | 
| 
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
553  | 
lits: term list, (*literals: irreducible formulae*)  | 
| 
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
554  | 
vars: term option ref list, (*variables occurring in branch*)  | 
| 
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
555  | 
lim: int}; (*resource limit*)  | 
| 3092 | 556  | 
|
557  | 
val fullTrace = ref[] : branch list list ref;  | 
|
558  | 
||
559  | 
(*Normalize a branch--for tracing*)  | 
|
560  | 
fun norm2 (G,md) = (norm G, md);  | 
|
561  | 
||
562  | 
fun normLev (Gs,Hs) = (map norm2 Gs, map norm2 Hs);  | 
|
563  | 
||
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
564  | 
fun normBr {pairs, lits, vars, lim} =
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
565  | 
     {pairs = map normLev pairs,
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
566  | 
lits = map norm lits,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
567  | 
vars = vars,  | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
568  | 
lim = lim};  | 
| 3092 | 569  | 
|
570  | 
||
| 4065 | 571  | 
val dummyTVar = Term.TVar(("a",0), []);
 | 
| 3092 | 572  | 
val dummyVar2 = Term.Var(("var",0), dummyT);
 | 
573  | 
||
| 4065 | 574  | 
(*convert Blast_tac's type representation to real types for tracing*)  | 
575  | 
fun showType (Free a) = Term.TFree (a,[])  | 
|
576  | 
| showType (Var _) = dummyTVar  | 
|
577  | 
| showType t =  | 
|
578  | 
(case strip_comb t of  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
579  | 
(Const (a, _), us) => Term.Type(a, map showType us)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
580  | 
| _ => dummyT);  | 
| 4065 | 581  | 
|
582  | 
(*Display top-level overloading if any*)  | 
|
| 18177 | 583  | 
fun topType thy (Const (c, ts)) = SOME (Sign.const_instance thy (c, map showType ts))  | 
584  | 
| topType thy (Abs(a,t)) = topType thy t  | 
|
585  | 
| topType thy (f $ u) = (case topType thy f of NONE => topType thy u | some => some)  | 
|
586  | 
| topType _ _ = NONE;  | 
|
| 4065 | 587  | 
|
588  | 
||
| 3092 | 589  | 
(*Convert from prototerms to ordinary terms with dummy types for tracing*)  | 
| 18177 | 590  | 
fun showTerm d (Const (a,_)) = Term.Const (a,dummyT)  | 
| 3092 | 591  | 
| showTerm d (Skolem(a,_)) = Term.Const (a,dummyT)  | 
592  | 
| showTerm d (Free a) = Term.Free (a,dummyT)  | 
|
593  | 
| showTerm d (Bound i) = Term.Bound i  | 
|
| 15531 | 594  | 
| showTerm d (Var(ref(SOME u))) = showTerm d u  | 
595  | 
| showTerm d (Var(ref NONE)) = dummyVar2  | 
|
| 3092 | 596  | 
| showTerm d (Abs(a,t)) = if d=0 then dummyVar  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
597  | 
else Term.Abs(a, dummyT, showTerm (d-1) t)  | 
| 3092 | 598  | 
| showTerm d (f $ u) = if d=0 then dummyVar  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
599  | 
else Term.$ (showTerm d f, showTerm (d-1) u);  | 
| 3092 | 600  | 
|
| 4065 | 601  | 
fun string_of sign d t = Sign.string_of_term sign (showTerm d t);  | 
| 3092 | 602  | 
|
| 19037 | 603  | 
(*Convert a Goal to an ordinary Not. Used also in dup_intr, where a goal like  | 
604  | 
Ex(P) is duplicated as the assumption ~Ex(P). *)  | 
|
605  | 
fun negOfGoal (Const ("*Goal*", _) $ G) = negate G
 | 
|
606  | 
| negOfGoal G = G;  | 
|
607  | 
||
608  | 
fun negOfGoal2 (G,md) = (negOfGoal G, md);  | 
|
609  | 
||
610  | 
(*Converts all Goals to Nots in the safe parts of a branch. They could  | 
|
611  | 
have been moved there from the literals list after substitution (equalSubst).  | 
|
612  | 
There can be at most one--this function could be made more efficient.*)  | 
|
613  | 
fun negOfGoals pairs = map (fn (Gs,haz) => (map negOfGoal2 Gs, haz)) pairs;  | 
|
614  | 
||
615  | 
(*Tactic. Convert *Goal* to negated assumption in FIRST position*)  | 
|
616  | 
fun negOfGoal_tac i = TRACE Data.ccontr (rtac Data.ccontr) i THEN  | 
|
617  | 
rotate_tac ~1 i;  | 
|
618  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
619  | 
fun traceTerm sign t =  | 
| 19037 | 620  | 
let val t' = norm (negOfGoal t)  | 
| 4065 | 621  | 
val stm = string_of sign 8 t'  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
622  | 
in  | 
| 18177 | 623  | 
case topType sign t' of  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
624  | 
NONE => stm (*no type to attach*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
625  | 
| SOME T => stm ^ "\t:: " ^ Sign.string_of_typ sign T  | 
| 4065 | 626  | 
end;  | 
| 3092 | 627  | 
|
628  | 
||
629  | 
(*Print tracing information at each iteration of prover*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
630  | 
fun tracing sign brs =  | 
| 14984 | 631  | 
let fun printPairs (((G,_)::_,_)::_) = immediate_output(traceTerm sign G)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
632  | 
| printPairs (([],(H,_)::_)::_) = immediate_output(traceTerm sign H ^ "\t (Unsafe)")  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
633  | 
| printPairs _ = ()  | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
634  | 
      fun printBrs (brs0 as {pairs, lits, lim, ...} :: brs) =
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
635  | 
(fullTrace := brs0 :: !fullTrace;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
636  | 
List.app (fn _ => immediate_output "+") brs;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
637  | 
             immediate_output (" [" ^ Int.toString lim ^ "] ");
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
638  | 
printPairs pairs;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
639  | 
writeln"")  | 
| 3092 | 640  | 
in if !trace then printBrs (map normBr brs) else ()  | 
641  | 
end;  | 
|
642  | 
||
| 
5343
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
643  | 
fun traceMsg s = if !trace then writeln s else ();  | 
| 4065 | 644  | 
|
| 3092 | 645  | 
(*Tracing: variables updated in the last branch operation?*)  | 
| 4065 | 646  | 
fun traceVars sign ntrl =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
647  | 
if !trace then  | 
| 4065 | 648  | 
(case !ntrail-ntrl of  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
649  | 
0 => ()  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
650  | 
| 1 => immediate_output"\t1 variable UPDATED:"  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
651  | 
          | n => immediate_output("\t" ^ Int.toString n ^ " variables UPDATED:");
 | 
| 4065 | 652  | 
(*display the instantiations themselves, though no variable names*)  | 
| 19502 | 653  | 
       List.app (fn v => immediate_output("   " ^ string_of sign 4 (the (!v))))
 | 
| 4065 | 654  | 
(List.take(!trail, !ntrail-ntrl));  | 
655  | 
writeln"")  | 
|
| 3092 | 656  | 
else ();  | 
657  | 
||
658  | 
(*Tracing: how many new branches are created?*)  | 
|
659  | 
fun traceNew prems =  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
660  | 
if !trace then  | 
| 3092 | 661  | 
case length prems of  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
662  | 
0 => immediate_output"branch closed by rule"  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
663  | 
| 1 => immediate_output"branch extended (1 new subgoal)"  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
664  | 
          | n => immediate_output("branch split: "^ Int.toString n ^ " new subgoals")
 | 
| 3092 | 665  | 
else ();  | 
666  | 
||
667  | 
||
668  | 
||
| 2854 | 669  | 
(*** Code for handling equality: naive substitution, like hyp_subst_tac ***)  | 
670  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
671  | 
(*Replace the ATOMIC term "old" by "new" in t*)  | 
| 2854 | 672  | 
fun subst_atomic (old,new) t =  | 
| 15531 | 673  | 
let fun subst (Var(ref(SOME u))) = subst u  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
674  | 
| subst (Abs(a,body)) = Abs(a, subst body)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
675  | 
| subst (f$t) = subst f $ subst t  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
676  | 
| subst t = if t aconv old then new else t  | 
| 2854 | 677  | 
in subst t end;  | 
678  | 
||
679  | 
(*Eta-contract a term from outside: just enough to reduce it to an atom*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
680  | 
fun eta_contract_atom (t0 as Abs(a, body)) =  | 
| 2854 | 681  | 
(case eta_contract2 body of  | 
| 20664 | 682  | 
f $ Bound 0 => if member (op =) (loose_bnos f) 0 then t0  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
683  | 
else eta_contract_atom (incr_boundvars ~1 f)  | 
| 2854 | 684  | 
| _ => t0)  | 
685  | 
| eta_contract_atom t = t  | 
|
686  | 
and eta_contract2 (f$t) = f $ eta_contract_atom t  | 
|
687  | 
| eta_contract2 t = eta_contract_atom t;  | 
|
688  | 
||
689  | 
||
690  | 
(*When can we safely delete the equality?  | 
|
691  | 
Not if it equates two constants; consider 0=1.  | 
|
692  | 
Not if it resembles x=t[x], since substitution does not eliminate x.  | 
|
693  | 
Not if it resembles ?x=0; another goal could instantiate ?x to Suc(i)  | 
|
694  | 
Prefer to eliminate Bound variables if possible.  | 
|
695  | 
Result: true = use as is, false = reorient first *)  | 
|
696  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
697  | 
(*Can t occur in u? For substitution.  | 
| 
4354
 
7f4da01bdf0e
Fixed the treatment of substitution for equations, restricting occurrences of
 
paulson 
parents: 
4323 
diff
changeset
 | 
698  | 
Does NOT examine the args of Skolem terms: substitution does not affect them.  | 
| 
4196
 
9953c0995b79
Now applies "map negOfGoal" to lits when expanding haz rules.
 
paulson 
parents: 
4149 
diff
changeset
 | 
699  | 
REFLEXIVE because hyp_subst_tac fails on x=x.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
700  | 
fun substOccur t =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
701  | 
let (*NO vars are permitted in u except the arguments of t, if it is  | 
| 
4354
 
7f4da01bdf0e
Fixed the treatment of substitution for equations, restricting occurrences of
 
paulson 
parents: 
4323 
diff
changeset
 | 
702  | 
a Skolem term. This ensures that no equations are deleted that could  | 
| 
 
7f4da01bdf0e
Fixed the treatment of substitution for equations, restricting occurrences of
 
paulson 
parents: 
4323 
diff
changeset
 | 
703  | 
be instantiated to a cycle. For example, x=?a is rejected because ?a  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
704  | 
could be instantiated to Suc(x).*)  | 
| 
4354
 
7f4da01bdf0e
Fixed the treatment of substitution for equations, restricting occurrences of
 
paulson 
parents: 
4323 
diff
changeset
 | 
705  | 
val vars = case t of  | 
| 
 
7f4da01bdf0e
Fixed the treatment of substitution for equations, restricting occurrences of
 
paulson 
parents: 
4323 
diff
changeset
 | 
706  | 
Skolem(_,vars) => vars  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
707  | 
| _ => []  | 
| 
4354
 
7f4da01bdf0e
Fixed the treatment of substitution for equations, restricting occurrences of
 
paulson 
parents: 
4323 
diff
changeset
 | 
708  | 
fun occEq u = (t aconv u) orelse occ u  | 
| 15531 | 709  | 
and occ (Var(ref(SOME u))) = occEq u  | 
| 
4354
 
7f4da01bdf0e
Fixed the treatment of substitution for equations, restricting occurrences of
 
paulson 
parents: 
4323 
diff
changeset
 | 
710  | 
| occ (Var v) = not (mem_var (v, vars))  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
711  | 
| occ (Abs(_,u)) = occEq u  | 
| 2854 | 712  | 
| occ (f$u) = occEq u orelse occEq f  | 
713  | 
| occ (_) = false;  | 
|
714  | 
in occEq end;  | 
|
715  | 
||
| 3092 | 716  | 
exception DEST_EQ;  | 
717  | 
||
| 18177 | 718  | 
(*Take apart an equality. NO constant Trueprop*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
719  | 
fun dest_eq (Const (c, _) $ t $ u) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
720  | 
if c = Data.equality_name then (eta_contract_atom t, eta_contract_atom u)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
721  | 
else raise DEST_EQ  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
722  | 
| dest_eq _ = raise DEST_EQ;  | 
| 3092 | 723  | 
|
| 
4196
 
9953c0995b79
Now applies "map negOfGoal" to lits when expanding haz rules.
 
paulson 
parents: 
4149 
diff
changeset
 | 
724  | 
(*Reject the equality if u occurs in (or equals!) t*)  | 
| 2854 | 725  | 
fun check (t,u,v) = if substOccur t u then raise DEST_EQ else v;  | 
726  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
727  | 
(*IF the goal is an equality with a substitutable variable  | 
| 2854 | 728  | 
THEN orient that equality ELSE raise exception DEST_EQ*)  | 
| 3092 | 729  | 
fun orientGoal (t,u) = case (t,u) of  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
730  | 
(Skolem _, _) => check(t,u,(t,u)) (*eliminates t*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
731  | 
| (_, Skolem _) => check(u,t,(u,t)) (*eliminates u*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
732  | 
| (Free _, _) => check(t,u,(t,u)) (*eliminates t*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
733  | 
| (_, Free _) => check(u,t,(u,t)) (*eliminates u*)  | 
| 2854 | 734  | 
| _ => raise DEST_EQ;  | 
735  | 
||
| 2894 | 736  | 
(*Substitute through the branch if an equality goal (else raise DEST_EQ).  | 
737  | 
Moves affected literals back into the branch, but it is not clear where  | 
|
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
738  | 
they should go: this could make proofs fail.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
739  | 
fun equalSubst sign (G, {pairs, lits, vars, lim}) =
 | 
| 3092 | 740  | 
let val (t,u) = orientGoal(dest_eq G)  | 
741  | 
val subst = subst_atomic (t,u)  | 
|
| 2854 | 742  | 
fun subst2(G,md) = (subst G, md)  | 
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
743  | 
(*substitute throughout list; extract affected formulae*)  | 
| 
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
744  | 
fun subForm ((G,md), (changed, pairs)) =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
745  | 
let val nG = subst G  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
746  | 
in if nG aconv G then (changed, (G,md)::pairs)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
747  | 
else ((nG,md)::changed, pairs)  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
748  | 
end  | 
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
749  | 
(*substitute throughout "stack frame"; extract affected formulae*)  | 
| 
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
750  | 
fun subFrame ((Gs,Hs), (changed, frames)) =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
751  | 
let val (changed', Gs') = foldr subForm (changed, []) Gs  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
752  | 
val (changed'', Hs') = foldr subForm (changed', []) Hs  | 
| 
4466
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
753  | 
in (changed'', (Gs',Hs')::frames) end  | 
| 
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
754  | 
(*substitute throughout literals; extract affected ones*)  | 
| 
 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 
paulson 
parents: 
4391 
diff
changeset
 | 
755  | 
fun subLit (lit, (changed, nlits)) =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
756  | 
let val nlit = subst lit  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
757  | 
in if nlit aconv lit then (changed, nlit::nlits)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
758  | 
else ((nlit,true)::changed, nlits)  | 
| 2854 | 759  | 
end  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
760  | 
val (changed, lits') = foldr subLit ([], []) lits  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
761  | 
val (changed', pairs') = foldr subFrame (changed, []) pairs  | 
| 3092 | 762  | 
  in  if !trace then writeln ("Substituting " ^ traceTerm sign u ^
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
763  | 
" for " ^ traceTerm sign t ^ " in branch" )  | 
| 3092 | 764  | 
else ();  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
765  | 
      {pairs = (changed',[])::pairs',   (*affected formulas, and others*)
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
766  | 
lits = lits', (*unaffected literals*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
767  | 
vars = vars,  | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
768  | 
lim = lim}  | 
| 2854 | 769  | 
end;  | 
770  | 
||
771  | 
||
772  | 
exception NEWBRANCHES and CLOSEF;  | 
|
773  | 
||
774  | 
exception PROVE;  | 
|
775  | 
||
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
776  | 
(*Trying eq_contr_tac first INCREASES the effort, slowing reconstruction*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
777  | 
val contr_tac = ematch_tac [Data.notE] THEN'  | 
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
778  | 
(eq_assume_tac ORELSE' assume_tac);  | 
| 2854 | 779  | 
|
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
780  | 
val eContr_tac = TRACE Data.notE contr_tac;  | 
| 2854 | 781  | 
val eAssume_tac = TRACE asm_rl (eq_assume_tac ORELSE' assume_tac);  | 
782  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
783  | 
(*Try to unify complementary literals and return the corresponding tactic. *)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
784  | 
fun tryClose (G, L) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
785  | 
let  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
786  | 
fun close t u tac = if unify ([], t, u) then SOME tac else NONE;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
787  | 
fun arg (_ $ t) = t;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
788  | 
in  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
789  | 
if isGoal G then close (arg G) L eAssume_tac  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
790  | 
else if isGoal L then close G (arg L) eAssume_tac  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
791  | 
else if isNot G then close (arg G) L eContr_tac  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
792  | 
else if isNot L then close G (arg L) eContr_tac  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
793  | 
else NONE  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
794  | 
end;  | 
| 2854 | 795  | 
|
796  | 
(*Were there Skolem terms in the premise? Must NOT chase Vars*)  | 
|
797  | 
fun hasSkolem (Skolem _) = true  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
798  | 
| hasSkolem (Abs (_,body)) = hasSkolem body  | 
| 2854 | 799  | 
| hasSkolem (f$t) = hasSkolem f orelse hasSkolem t  | 
800  | 
| hasSkolem _ = false;  | 
|
801  | 
||
802  | 
(*Attach the right "may duplicate" flag to new formulae: if they contain  | 
|
803  | 
Skolem terms then allow duplication.*)  | 
|
804  | 
fun joinMd md [] = []  | 
|
805  | 
| joinMd md (G::Gs) = (G, hasSkolem G orelse md) :: joinMd md Gs;  | 
|
806  | 
||
807  | 
||
808  | 
(** Backtracking and Pruning **)  | 
|
809  | 
||
810  | 
(*clashVar vars (n,trail) determines whether any of the last n elements  | 
|
811  | 
of "trail" occur in "vars" OR in their instantiations*)  | 
|
812  | 
fun clashVar [] = (fn _ => false)  | 
|
813  | 
| clashVar vars =  | 
|
814  | 
let fun clash (0, _) = false  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
815  | 
| clash (_, []) = false  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
816  | 
| clash (n, v::vs) = exists (varOccur v) vars orelse clash(n-1,vs)  | 
| 2854 | 817  | 
in clash end;  | 
818  | 
||
819  | 
||
820  | 
(*nbrs = # of branches just prior to closing this one. Delete choice points  | 
|
821  | 
for goals proved by the latest inference, provided NO variables in the  | 
|
822  | 
next branch have been updated.*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
823  | 
fun prune (1, nxtVars, choices) = choices (*DON'T prune at very end: allow  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
824  | 
backtracking over bad proofs*)  | 
| 17795 | 825  | 
| prune (nbrs: int, nxtVars, choices) =  | 
| 2854 | 826  | 
let fun traceIt last =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
827  | 
let val ll = length last  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
828  | 
and lc = length choices  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
829  | 
in if !trace andalso ll<lc then  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
830  | 
                    (writeln("Pruning " ^ Int.toString(lc-ll) ^ " levels");
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
831  | 
last)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
832  | 
else last  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
833  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
834  | 
fun pruneAux (last, _, _, []) = last  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
835  | 
| pruneAux (last, ntrl, trl, (ntrl',nbrs',exn) :: choices) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
836  | 
if nbrs' < nbrs  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
837  | 
then last (*don't backtrack beyond first solution of goal*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
838  | 
else if nbrs' > nbrs then pruneAux (last, ntrl, trl, choices)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
839  | 
else (* nbrs'=nbrs *)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
840  | 
if clashVar nxtVars (ntrl-ntrl', trl) then last  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
841  | 
else (*no clashes: can go back at least this far!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
842  | 
pruneAux(choices, ntrl', List.drop(trl, ntrl-ntrl'),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
843  | 
choices)  | 
| 2854 | 844  | 
in traceIt (pruneAux (choices, !ntrail, !trail, choices)) end;  | 
845  | 
||
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
846  | 
fun nextVars ({pairs, lits, vars, lim} :: _) = map Var vars
 | 
| 
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
847  | 
| nextVars [] = [];  | 
| 2854 | 848  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
849  | 
fun backtrack (choices as (ntrl, nbrs, exn)::_) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
850  | 
      (if !trace then (writeln ("Backtracking; now there are " ^
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
851  | 
Int.toString nbrs ^ " branches"))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
852  | 
else ();  | 
| 3083 | 853  | 
raise exn)  | 
854  | 
| backtrack _ = raise PROVE;  | 
|
| 2854 | 855  | 
|
| 2894 | 856  | 
(*Add the literal G, handling *Goal* and detecting duplicates.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
857  | 
fun addLit (Const ("*Goal*", _) $ G, lits) =
 | 
| 2894 | 858  | 
(*New literal is a *Goal*, so change all other Goals to Nots*)  | 
| 18177 | 859  | 
      let fun bad (Const ("*Goal*", _) $ _) = true
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
860  | 
| bad (Const (c, _) $ G') = c = Data.not_name andalso G aconv G'  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
861  | 
| bad _ = false;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
862  | 
fun change [] = []  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
863  | 
| change (lit :: lits) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
864  | 
(case lit of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
865  | 
Const (c, _) $ G' =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
866  | 
if c = "*Goal*" orelse c = Data.not_name then  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
867  | 
if G aconv G' then change lits  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
868  | 
else negate G' :: change lits  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
869  | 
else lit :: change lits  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
870  | 
| _ => lit :: change lits)  | 
| 2854 | 871  | 
in  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
872  | 
        Const ("*Goal*", []) $ G :: (if exists bad lits then change lits else lits)
 | 
| 2854 | 873  | 
end  | 
874  | 
| addLit (G,lits) = ins_term(G, lits)  | 
|
875  | 
||
876  | 
||
| 2952 | 877  | 
(*For calculating the "penalty" to assess on a branching factor of n  | 
878  | 
log2 seems a little too severe*)  | 
|
| 3083 | 879  | 
fun log n = if n<4 then 0 else 1 + log(n div 4);  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
880  | 
|
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
881  | 
|
| 
3021
 
39806db47be9
Loop detection: before expanding a haz formula, see whether it is a duplicate
 
paulson 
parents: 
2999 
diff
changeset
 | 
882  | 
(*match(t,u) says whether the term u might be an instance of the pattern t  | 
| 
 
39806db47be9
Loop detection: before expanding a haz formula, see whether it is a duplicate
 
paulson 
parents: 
2999 
diff
changeset
 | 
883  | 
Used to detect "recursive" rules such as transitivity*)  | 
| 
 
39806db47be9
Loop detection: before expanding a haz formula, see whether it is a duplicate
 
paulson 
parents: 
2999 
diff
changeset
 | 
884  | 
fun match (Var _) u = true  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
885  | 
| match (Const (a,tas)) (Const (b,tbs)) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
886  | 
a = "*Goal*" andalso b = Data.not_name orelse  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
887  | 
a = Data.not_name andalso b = "*Goal*" orelse  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
888  | 
a = b andalso matchs tas tbs  | 
| 4065 | 889  | 
| match (Free a) (Free b) = (a=b)  | 
890  | 
| match (Bound i) (Bound j) = (i=j)  | 
|
891  | 
| match (Abs(_,t)) (Abs(_,u)) = match t u  | 
|
892  | 
| match (f$t) (g$u) = match f g andalso match t u  | 
|
| 18177 | 893  | 
| match t u = false  | 
894  | 
and matchs [] [] = true  | 
|
895  | 
| matchs (t :: ts) (u :: us) = match t u andalso matchs ts us;  | 
|
| 
3021
 
39806db47be9
Loop detection: before expanding a haz formula, see whether it is a duplicate
 
paulson 
parents: 
2999 
diff
changeset
 | 
896  | 
|
| 
 
39806db47be9
Loop detection: before expanding a haz formula, see whether it is a duplicate
 
paulson 
parents: 
2999 
diff
changeset
 | 
897  | 
|
| 4300 | 898  | 
(*Branches closed: number of branches closed during the search  | 
899  | 
Branches tried: number of branches created by splitting (counting from 1)*)  | 
|
900  | 
val nclosed = ref 0  | 
|
901  | 
and ntried = ref 1;  | 
|
902  | 
||
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
903  | 
fun printStats (b, start, tacs) =  | 
| 
4323
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
904  | 
if b then  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
905  | 
writeln (endTiming start ^ " for search. Closed: "  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
906  | 
^ Int.toString (!nclosed) ^  | 
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
907  | 
" tried: " ^ Int.toString (!ntried) ^  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
908  | 
" tactics: " ^ Int.toString (length tacs))  | 
| 
4323
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
909  | 
else ();  | 
| 
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
910  | 
|
| 
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
911  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
912  | 
(*Tableau prover based on leanTaP. Argument is a list of branches. Each  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
913  | 
branch contains a list of unexpanded formulae, a list of literals, and a  | 
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
914  | 
bound on unsafe expansions.  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
915  | 
"start" is CPU time at start, for printing search time  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
916  | 
*)  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
917  | 
fun prove (sign, start, cs, brs, cont) =  | 
| 4653 | 918  | 
 let val {safe0_netpair, safep_netpair, haz_netpair, ...} = Data.rep_cs cs
 | 
| 2854 | 919  | 
val safeList = [safe0_netpair, safep_netpair]  | 
920  | 
and hazList = [haz_netpair]  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
921  | 
fun prv (tacs, trs, choices, []) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
922  | 
(printStats (!trace orelse !stats, start, tacs);  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
923  | 
cont (tacs, trs, choices)) (*all branches closed!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
924  | 
| prv (tacs, trs, choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
925  | 
              brs0 as {pairs = ((G,md)::br, haz)::pairs,
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
926  | 
lits, vars, lim} :: brs) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
927  | 
(*apply a safe rule only (possibly allowing instantiation);  | 
| 3917 | 928  | 
defer any haz formulae*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
929  | 
let exception PRV (*backtrack to precisely this recursion!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
930  | 
val ntrl = !ntrail  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
931  | 
val nbrs = length brs0  | 
| 2854 | 932  | 
val nxtVars = nextVars brs  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
933  | 
val G = norm G  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
934  | 
val rules = netMkRules G vars safeList  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
935  | 
(*Make a new branch, decrementing "lim" if instantiations occur*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
936  | 
fun newBr (vars',lim') prems =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
937  | 
map (fn prem =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
938  | 
if (exists isGoal prem)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
939  | 
                       then {pairs = ((joinMd md prem, []) ::
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
940  | 
negOfGoals ((br, haz)::pairs)),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
941  | 
lits = map negOfGoal lits,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
942  | 
vars = vars',  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
943  | 
lim = lim'}  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
944  | 
                       else {pairs = ((joinMd md prem, []) ::
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
945  | 
(br, haz) :: pairs),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
946  | 
lits = lits,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
947  | 
vars = vars',  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
948  | 
lim = lim'})  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
949  | 
prems @  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
950  | 
brs  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
951  | 
(*Seek a matching rule. If unifiable then add new premises  | 
| 2854 | 952  | 
to branch.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
953  | 
fun deeper [] = raise NEWBRANCHES  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
954  | 
| deeper (((P,prems),tac)::grls) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
955  | 
if unify(add_term_vars(P,[]), P, G)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
956  | 
then (*P comes from the rule; G comes from the branch.*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
957  | 
let val updated = ntrl < !ntrail (*branch updated*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
958  | 
val lim' = if updated  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
959  | 
then lim - (1+log(length rules))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
960  | 
else lim (*discourage branching updates*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
961  | 
val vars = vars_in_vars vars  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
962  | 
val vars' = foldr add_terms_vars vars prems  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
963  | 
val choices' = (ntrl, nbrs, PRV) :: choices  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
964  | 
val tacs' = (tac(updated,false,true))  | 
| 
5343
 
871b77df79a0
new version, more resistant to PROOF FAILED.  Now it distinguishes between
 
paulson 
parents: 
4653 
diff
changeset
 | 
965  | 
:: tacs (*no duplication; rotate*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
966  | 
in  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
967  | 
traceNew prems; traceVars sign ntrl;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
968  | 
(if null prems then (*closed the branch: prune!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
969  | 
(nclosed := !nclosed + 1;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
970  | 
prv(tacs', brs0::trs,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
971  | 
prune (nbrs, nxtVars, choices'),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
972  | 
brs))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
973  | 
else (*prems non-null*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
974  | 
if lim'<0 (*faster to kill ALL the alternatives*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
975  | 
then (traceMsg"Excessive branching: KILLED";  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
976  | 
clearTo ntrl; raise NEWBRANCHES)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
977  | 
else  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
978  | 
(ntried := !ntried + length prems - 1;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
979  | 
prv(tacs', brs0::trs, choices',  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
980  | 
newBr (vars',lim') prems)))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
981  | 
handle PRV =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
982  | 
if updated then  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
983  | 
(*Backtrack at this level.  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
984  | 
Reset Vars and try another rule*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
985  | 
(clearTo ntrl; deeper grls)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
986  | 
else (*backtrack to previous level*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
987  | 
backtrack choices  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
988  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
989  | 
else deeper grls  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
990  | 
(*Try to close branch by unifying with head goal*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
991  | 
fun closeF [] = raise CLOSEF  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
992  | 
| closeF (L::Ls) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
993  | 
case tryClose(G,L) of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
994  | 
NONE => closeF Ls  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
995  | 
| SOME tac =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
996  | 
let val choices' =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
997  | 
(if !trace then (immediate_output"branch closed";  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
998  | 
traceVars sign ntrl)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
999  | 
else ();  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1000  | 
prune (nbrs, nxtVars,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1001  | 
(ntrl, nbrs, PRV) :: choices))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1002  | 
in nclosed := !nclosed + 1;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1003  | 
prv (tac::tacs, brs0::trs, choices', brs)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1004  | 
handle PRV =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1005  | 
(*reset Vars and try another literal  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1006  | 
[this handler is pruned if possible!]*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1007  | 
(clearTo ntrl; closeF Ls)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1008  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1009  | 
(*Try to unify a queued formula (safe or haz) with head goal*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1010  | 
fun closeFl [] = raise CLOSEF  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1011  | 
| closeFl ((br, haz)::pairs) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1012  | 
closeF (map fst br)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1013  | 
handle CLOSEF => closeF (map fst haz)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1014  | 
handle CLOSEF => closeFl pairs  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1015  | 
in tracing sign brs0;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1016  | 
if lim<0 then (traceMsg "Limit reached. "; backtrack choices)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1017  | 
else  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1018  | 
prv (Data.hyp_subst_tac trace :: tacs,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1019  | 
brs0::trs, choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1020  | 
equalSubst sign  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1021  | 
                    (G, {pairs = (br,haz)::pairs,
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1022  | 
lits = lits, vars = vars, lim = lim})  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1023  | 
:: brs)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1024  | 
handle DEST_EQ => closeF lits  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1025  | 
handle CLOSEF => closeFl ((br,haz)::pairs)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1026  | 
handle CLOSEF => deeper rules  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1027  | 
handle NEWBRANCHES =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1028  | 
(case netMkRules G vars hazList of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1029  | 
[] => (*there are no plausible haz rules*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1030  | 
(traceMsg "moving formula to literals";  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1031  | 
prv (tacs, brs0::trs, choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1032  | 
                                   {pairs = (br,haz)::pairs,
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1033  | 
lits = addLit(G,lits),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1034  | 
vars = vars,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1035  | 
lim = lim} :: brs))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1036  | 
| _ => (*G admits some haz rules: try later*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1037  | 
(traceMsg "moving formula to haz list";  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1038  | 
prv (if isGoal G then negOfGoal_tac :: tacs  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1039  | 
else tacs,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1040  | 
brs0::trs,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1041  | 
choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1042  | 
                                 {pairs = (br, haz@[(negOfGoal G, md)])::pairs,
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1043  | 
lits = lits,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1044  | 
vars = vars,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1045  | 
lim = lim} :: brs)))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1046  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1047  | 
| prv (tacs, trs, choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1048  | 
              {pairs = ([],haz)::(Gs,haz')::pairs, lits, vars, lim} :: brs) =
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1049  | 
(*no more "safe" formulae: transfer haz down a level*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1050  | 
prv (tacs, trs, choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1051  | 
                {pairs = (Gs,haz@haz')::pairs,
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1052  | 
lits = lits,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1053  | 
vars = vars,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1054  | 
lim = lim} :: brs)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1055  | 
| prv (tacs, trs, choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1056  | 
              brs0 as {pairs = [([], (H,md)::Hs)],
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1057  | 
lits, vars, lim} :: brs) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1058  | 
(*no safe steps possible at any level: apply a haz rule*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1059  | 
let exception PRV (*backtrack to precisely this recursion!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1060  | 
val H = norm H  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1061  | 
val ntrl = !ntrail  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1062  | 
val rules = netMkRules H vars hazList  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1063  | 
(*new premises of haz rules may NOT be duplicated*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1064  | 
fun newPrem (vars,P,dup,lim') prem =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1065  | 
let val Gs' = map (fn Q => (Q,false)) prem  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1066  | 
and Hs' = if dup then Hs @ [(negOfGoal H, md)] else Hs  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1067  | 
and lits' = if (exists isGoal prem)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1068  | 
then map negOfGoal lits  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1069  | 
else lits  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1070  | 
                  in  {pairs = if exists (match P) prem then [(Gs',Hs')]
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1071  | 
(*Recursive in this premise. Don't make new  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1072  | 
"stack frame". New haz premises will end up  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1073  | 
at the BACK of the queue, preventing  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1074  | 
exclusion of others*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1075  | 
else [(Gs',[]), ([],Hs')],  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1076  | 
lits = lits',  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1077  | 
vars = vars,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1078  | 
lim = lim'}  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1079  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1080  | 
fun newBr x prems = map (newPrem x) prems @ brs  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1081  | 
(*Seek a matching rule. If unifiable then add new premises  | 
| 2854 | 1082  | 
to branch.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1083  | 
fun deeper [] = raise NEWBRANCHES  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1084  | 
| deeper (((P,prems),tac)::grls) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1085  | 
if unify(add_term_vars(P,[]), P, H)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1086  | 
then  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1087  | 
let val updated = ntrl < !ntrail (*branch updated*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1088  | 
val vars = vars_in_vars vars  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1089  | 
val vars' = foldr add_terms_vars vars prems  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1090  | 
(*duplicate H if md permits*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1091  | 
val dup = md (*earlier had "andalso vars' <> vars":  | 
| 
11152
 
32d002362005
Blast bug fix: now always duplicates when applying a haz rule,
 
paulson 
parents: 
11119 
diff
changeset
 | 
1092  | 
duplicate only if the subgoal has new vars*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1093  | 
(*any instances of P in the subgoals?  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1094  | 
NB: boolean "recur" affects tracing only!*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1095  | 
and recur = exists (exists (match P)) prems  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1096  | 
val lim' = (*Decrement "lim" extra if updates occur*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1097  | 
if updated then lim - (1+log(length rules))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1098  | 
else lim-1  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1099  | 
(*It is tempting to leave "lim" UNCHANGED if  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1100  | 
both dup and recur are false. Proofs are  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1101  | 
found at shallower depths, but looping  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1102  | 
occurs too often...*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1103  | 
val mayUndo =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1104  | 
(*Allowing backtracking from a rule application  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1105  | 
if other matching rules exist, if the rule  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1106  | 
updated variables, or if the rule did not  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1107  | 
introduce new variables. This latter condition  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1108  | 
means it is not a standard "gamma-rule" but  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1109  | 
some other form of unsafe rule. Aim is to  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1110  | 
emulate Fast_tac, which allows all unsafe steps  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1111  | 
to be undone.*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1112  | 
not(null grls) (*other rules to try?*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1113  | 
orelse updated  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1114  | 
orelse vars=vars' (*no new Vars?*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1115  | 
val tac' = tac(updated, dup, true)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1116  | 
(*if recur then perhaps shouldn't call rotate_tac: new  | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
1117  | 
formulae should be last, but that's WRONG if the new  | 
| 
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
1118  | 
formulae are Goals, since they remain in the first  | 
| 
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
1119  | 
position*)  | 
| 
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
1120  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1121  | 
in  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1122  | 
if lim'<0 andalso not (null prems)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1123  | 
then (*it's faster to kill ALL the alternatives*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1124  | 
(traceMsg"Excessive branching: KILLED";  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1125  | 
clearTo ntrl; raise NEWBRANCHES)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1126  | 
else  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1127  | 
traceNew prems;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1128  | 
if !trace andalso dup then immediate_output" (duplicating)"  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1129  | 
else ();  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1130  | 
if !trace andalso recur then immediate_output" (recursive)"  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1131  | 
else ();  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1132  | 
traceVars sign ntrl;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1133  | 
if null prems then nclosed := !nclosed + 1  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1134  | 
else ntried := !ntried + length prems - 1;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1135  | 
prv(tac' :: tacs,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1136  | 
brs0::trs,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1137  | 
(ntrl, length brs0, PRV) :: choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1138  | 
newBr (vars', P, dup, lim') prems)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1139  | 
handle PRV =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1140  | 
if mayUndo  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1141  | 
then (*reset Vars and try another rule*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1142  | 
(clearTo ntrl; deeper grls)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1143  | 
else (*backtrack to previous level*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1144  | 
backtrack choices  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1145  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1146  | 
else deeper grls  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1147  | 
in tracing sign brs0;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1148  | 
if lim<1 then (traceMsg "Limit reached. "; backtrack choices)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1149  | 
else deeper rules  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1150  | 
handle NEWBRANCHES =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1151  | 
(*cannot close branch: move H to literals*)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1152  | 
prv (tacs, brs0::trs, choices,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1153  | 
                      {pairs = [([], Hs)],
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1154  | 
lits = H::lits,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1155  | 
vars = vars,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1156  | 
lim = lim} :: brs)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1157  | 
end  | 
| 2854 | 1158  | 
| prv (tacs, trs, choices, _ :: brs) = backtrack choices  | 
| 12346 | 1159  | 
in prv ([], [], [(!ntrail, length brs, PROVE)], brs) end;  | 
| 2854 | 1160  | 
|
1161  | 
||
| 2883 | 1162  | 
(*Construct an initial branch.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1163  | 
fun initBranch (ts,lim) =  | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
1164  | 
    {pairs = [(map (fn t => (t,true)) ts, [])],
 | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1165  | 
lits = [],  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1166  | 
vars = add_terms_vars (ts,[]),  | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
1167  | 
lim = lim};  | 
| 2854 | 1168  | 
|
1169  | 
||
1170  | 
(*** Conversion & Skolemization of the Isabelle proof state ***)  | 
|
1171  | 
||
1172  | 
(*Make a list of all the parameters in a subgoal, even if nested*)  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1173  | 
local open Term  | 
| 2854 | 1174  | 
in  | 
1175  | 
fun discard_foralls (Const("all",_)$Abs(a,T,t)) = discard_foralls t
 | 
|
1176  | 
| discard_foralls t = t;  | 
|
1177  | 
end;  | 
|
1178  | 
||
1179  | 
(*List of variables not appearing as arguments to the given parameter*)  | 
|
1180  | 
fun getVars [] i = []  | 
|
| 20664 | 1181  | 
| getVars ((_,(v,is))::alist) (i: int) =  | 
1182  | 
if member (op =) is i then getVars alist i  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1183  | 
else v :: getVars alist i;  | 
| 2854 | 1184  | 
|
| 
4233
 
85d004a96b98
Rationalized error handling: if low-level tactic (depth_tac) cannot accept the
 
paulson 
parents: 
4196 
diff
changeset
 | 
1185  | 
exception TRANS of string;  | 
| 2854 | 1186  | 
|
| 
4233
 
85d004a96b98
Rationalized error handling: if low-level tactic (depth_tac) cannot accept the
 
paulson 
parents: 
4196 
diff
changeset
 | 
1187  | 
(*Translation of a subgoal: Skolemize all parameters*)  | 
| 4065 | 1188  | 
fun fromSubgoal t =  | 
1189  | 
let val alistVar = ref []  | 
|
1190  | 
and alistTVar = ref []  | 
|
| 2854 | 1191  | 
fun hdvar ((ix,(v,is))::_) = v  | 
1192  | 
fun from lev t =  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1193  | 
let val (ht,ts) = Term.strip_comb t  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1194  | 
fun apply u = list_comb (u, map (from lev) ts)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1195  | 
fun bounds [] = []  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1196  | 
| bounds (Term.Bound i::ts) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1197  | 
if i<lev then raise TRANS  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1198  | 
"Function unknown's argument not a parameter"  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1199  | 
else i-lev :: bounds ts  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1200  | 
| bounds ts = raise TRANS  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1201  | 
"Function unknown's argument not a bound variable"  | 
| 2854 | 1202  | 
in  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1203  | 
case ht of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1204  | 
Term.Const aT => apply (fromConst alistTVar aT)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1205  | 
| Term.Free (a,_) => apply (Free a)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1206  | 
| Term.Bound i => apply (Bound i)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1207  | 
| Term.Var (ix,_) =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1208  | 
(case (AList.lookup (op =) (!alistVar) ix) of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1209  | 
NONE => (alistVar := (ix, (ref NONE, bounds ts))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1210  | 
:: !alistVar;  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1211  | 
Var (hdvar(!alistVar)))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1212  | 
| SOME(v,is) => if is=bounds ts then Var v  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1213  | 
else raise TRANS  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1214  | 
                                ("Discrepancy among occurrences of "
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1215  | 
^ Syntax.string_of_vname ix))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1216  | 
| Term.Abs (a,_,body) =>  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1217  | 
if null ts then Abs(a, from (lev+1) body)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1218  | 
else raise TRANS "argument not in normal form"  | 
| 2854 | 1219  | 
end  | 
1220  | 
||
1221  | 
val npars = length (Logic.strip_params t)  | 
|
1222  | 
||
1223  | 
(*Skolemize a subgoal from a proof state*)  | 
|
1224  | 
fun skoSubgoal i t =  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1225  | 
if i<npars then  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1226  | 
skoSubgoal (i+1)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1227  | 
(subst_bound (Skolem (gensym "T_", getVars (!alistVar) i),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1228  | 
t))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1229  | 
else t  | 
| 2854 | 1230  | 
|
1231  | 
in skoSubgoal 0 (from 0 (discard_foralls t)) end;  | 
|
1232  | 
||
1233  | 
||
| 18533 | 1234  | 
fun reject_const thy c =  | 
1235  | 
if is_some (Sign.const_type thy c) then  | 
|
1236  | 
    error ("Blast: theory contains illegal constant " ^ quote c)
 | 
|
1237  | 
else ();  | 
|
1238  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1239  | 
fun initialize thy =  | 
| 18533 | 1240  | 
(fullTrace:=[]; trail := []; ntrail := 0;  | 
1241  | 
nclosed := 0; ntried := 1; typargs := Sign.const_typargs thy;  | 
|
1242  | 
reject_const thy "*Goal*"; reject_const thy "*False*");  | 
|
| 4300 | 1243  | 
|
1244  | 
||
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1245  | 
(*Tactic using tableau engine and proof reconstruction.  | 
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
1246  | 
"start" is CPU time at start, for printing SEARCH time  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1247  | 
(also prints reconstruction time)  | 
| 2854 | 1248  | 
"lim" is depth limit.*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1249  | 
fun timing_depth_tac start cs lim i st0 =  | 
| 
16774
 
515b6020cf5d
experimental code to reduce the amount of type information in blast
 
paulson 
parents: 
15786 
diff
changeset
 | 
1250  | 
let val st = (initialize (theory_of_thm st0); ObjectLogic.atomize_goal i st0);  | 
| 
9486
 
2df511ebb956
handle actual object-logic rules by atomizing the goal;
 
wenzelm 
parents: 
9170 
diff
changeset
 | 
1251  | 
      val {sign,...} = rep_thm st
 | 
| 
4323
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
1252  | 
val skoprem = fromSubgoal (List.nth(prems_of st, i-1))  | 
| 
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
1253  | 
val hyps = strip_imp_prems skoprem  | 
| 
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
1254  | 
and concl = strip_imp_concl skoprem  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1255  | 
fun cont (tacs,_,choices) =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1256  | 
let val start = startTiming()  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1257  | 
in  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1258  | 
case Seq.pull(EVERY' (rev tacs) i st) of  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1259  | 
              NONE => (writeln ("PROOF FAILED for depth " ^
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1260  | 
Int.toString lim);  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1261  | 
if !trace then error "************************\n"  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1262  | 
else ();  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1263  | 
backtrack choices)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1264  | 
| cell => (if (!trace orelse !stats)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1265  | 
then writeln (endTiming start ^ " for reconstruction")  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1266  | 
else ();  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1267  | 
Seq.make(fn()=> cell))  | 
| 
4323
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
1268  | 
end  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1269  | 
in prove (sign, start, cs, [initBranch (mkGoal concl :: hyps, lim)], cont)  | 
| 
4323
 
561242f8606b
Printing of statistics including time for search & reconstruction
 
paulson 
parents: 
4300 
diff
changeset
 | 
1270  | 
end  | 
| 
16774
 
515b6020cf5d
experimental code to reduce the amount of type information in blast
 
paulson 
parents: 
15786 
diff
changeset
 | 
1271  | 
handle PROVE => Seq.empty;  | 
| 2854 | 1272  | 
|
| 
4391
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
1273  | 
(*Public version with fixed depth*)  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
1274  | 
fun depth_tac cs lim i st = timing_depth_tac (startTiming()) cs lim i st;  | 
| 
 
cc3e8453d7f0
More deterministic and therefore faster (sometimes) proof reconstruction
 
paulson 
parents: 
4354 
diff
changeset
 | 
1275  | 
|
| 
15162
 
670ab8497818
depth limit (previously hard-coded with a value of 20) made a reference
 
webertj 
parents: 
14984 
diff
changeset
 | 
1276  | 
val depth_limit = ref 20;  | 
| 
 
670ab8497818
depth limit (previously hard-coded with a value of 20) made a reference
 
webertj 
parents: 
14984 
diff
changeset
 | 
1277  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1278  | 
fun blast_tac cs i st =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1279  | 
((DEEPEN (1, !depth_limit) (timing_depth_tac (startTiming()) cs) 0) i  | 
| 
5463
 
a5479f5cd482
Allows more backtracking in proof reconstruction, making it slower but more
 
paulson 
parents: 
5411 
diff
changeset
 | 
1280  | 
THEN flexflex_tac) st  | 
| 14466 | 1281  | 
handle TRANS s =>  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1282  | 
      ((if !trace then warning ("blast: " ^ s) else ());
 | 
| 14466 | 1283  | 
Seq.empty);  | 
| 2854 | 1284  | 
|
| 4078 | 1285  | 
fun Blast_tac i = blast_tac (Data.claset()) i;  | 
| 2854 | 1286  | 
|
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1287  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1288  | 
(*** For debugging: these apply the prover to a subgoal and return  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1289  | 
the resulting tactics, trace, etc. ***)  | 
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1290  | 
|
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1291  | 
(*Translate subgoal i from a proof state*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1292  | 
fun trygl cs lim i =  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1293  | 
let val st = topthm()  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1294  | 
                val {sign,...} = rep_thm st
 | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1295  | 
val skoprem = (initialize (theory_of_thm st);  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1296  | 
fromSubgoal (List.nth(prems_of st, i-1)))  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1297  | 
val hyps = strip_imp_prems skoprem  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1298  | 
and concl = strip_imp_concl skoprem  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1299  | 
in timeap prove (sign, startTiming(), cs,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1300  | 
[initBranch (mkGoal concl :: hyps, lim)], I)  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1301  | 
end  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1302  | 
        handle Subscript => error("There is no subgoal " ^ Int.toString i);
 | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1303  | 
|
| 4078 | 1304  | 
fun Trygl lim i = trygl (Data.claset()) lim i;  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1305  | 
|
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1306  | 
(*Read a string to make an initial, singleton branch*)  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1307  | 
fun readGoal thy s = Sign.read_prop thy s |> fromTerm |> rand |> mkGoal;  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1308  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1309  | 
fun tryInThy thy lim s =  | 
| 
16774
 
515b6020cf5d
experimental code to reduce the amount of type information in blast
 
paulson 
parents: 
15786 
diff
changeset
 | 
1310  | 
(initialize thy;  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1311  | 
timeap prove (thy,  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1312  | 
startTiming(),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1313  | 
Data.claset(),  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1314  | 
[initBranch ([readGoal thy s], lim)],  | 
| 
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1315  | 
I));  | 
| 
2924
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1316  | 
|
| 
 
af506c35b4ed
Control over excessive branching by applying a log2 penalty
 
paulson 
parents: 
2894 
diff
changeset
 | 
1317  | 
|
| 5926 | 1318  | 
(** method setup **)  | 
1319  | 
||
| 7559 | 1320  | 
fun blast_args m =  | 
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1321  | 
Method.bang_sectioned_args'  | 
| 
11152
 
32d002362005
Blast bug fix: now always duplicates when applying a haz rule,
 
paulson 
parents: 
11119 
diff
changeset
 | 
1322  | 
Data.cla_modifiers (Scan.lift (Scan.option Args.nat)) m;  | 
| 7155 | 1323  | 
|
| 15531 | 1324  | 
fun blast_meth NONE = Data.cla_meth' blast_tac  | 
1325  | 
| blast_meth (SOME lim) = Data.cla_meth' (fn cs => depth_tac cs lim);  | 
|
| 7155 | 1326  | 
|
| 
18525
 
ce1ae48c320f
avoid implicit assumptions about consts Not, op =, *Goal*, *False*;
 
wenzelm 
parents: 
18177 
diff
changeset
 | 
1327  | 
val setup =  | 
| 18708 | 1328  | 
  Method.add_methods [("blast", blast_args blast_meth, "tableau prover")];
 | 
| 5926 | 1329  | 
|
1330  | 
||
| 2854 | 1331  | 
end;  |