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