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