| author | wenzelm | 
| Mon, 02 Nov 2009 17:29:48 +0100 | |
| changeset 33379 | b834b42e4aa1 | 
| parent 33339 | d41f77196338 | 
| child 33522 | 737589bb9bb8 | 
| permissions | -rw-r--r-- | 
| 33311 | 1 | (* Title: HOL/Tools/res_axioms.ML | 
| 2 | Author: Jia Meng, Cambridge University Computer Laboratory | |
| 15347 | 3 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 4 | Transformation of axiom rules (elim/intro/etc) into CNF forms. | 
| 15347 | 5 | *) | 
| 6 | ||
| 15997 | 7 | signature RES_AXIOMS = | 
| 21505 | 8 | sig | 
| 32955 | 9 | val trace: bool Unsynchronized.ref | 
| 10 | val trace_msg: (unit -> string) -> unit | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 11 | val cnf_axiom: theory -> thm -> thm list | 
| 24669 | 12 | val pairname: thm -> string * thm | 
| 27184 | 13 | val multi_base_blacklist: string list | 
| 25243 | 14 | val bad_for_atp: thm -> bool | 
| 25761 | 15 | val type_has_empty_sort: typ -> bool | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 16 | val cnf_rules_pairs: theory -> (string * thm) list -> (thm * (string * int)) list | 
| 24669 | 17 | val neg_clausify: thm list -> thm list | 
| 18 | val expand_defs_tac: thm -> tactic | |
| 24827 | 19 | val combinators: thm -> thm | 
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 20 | val neg_conjecture_clauses: Proof.context -> thm -> int -> thm list * (string * typ) list | 
| 32740 | 21 | val suppress_endtheory: bool Unsynchronized.ref | 
| 22 | (*for emergency use where endtheory causes problems*) | |
| 24669 | 23 | val setup: theory -> theory | 
| 21505 | 24 | end; | 
| 19196 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 mengj parents: 
19175diff
changeset | 25 | |
| 33316 | 26 | structure Res_Axioms: RES_AXIOMS = | 
| 15997 | 27 | struct | 
| 15347 | 28 | |
| 32955 | 29 | val trace = Unsynchronized.ref false; | 
| 30 | fun trace_msg msg = if ! trace then tracing (msg ()) else (); | |
| 31 | ||
| 20902 | 32 | (* FIXME legacy *) | 
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 33 | fun freeze_thm th = #1 (Drule.freeze_thaw th); | 
| 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 34 | |
| 25761 | 35 | fun type_has_empty_sort (TFree (_, [])) = true | 
| 36 | | type_has_empty_sort (TVar (_, [])) = true | |
| 37 | | type_has_empty_sort (Type (_, Ts)) = exists type_has_empty_sort Ts | |
| 38 | | type_has_empty_sort _ = false; | |
| 27184 | 39 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 40 | |
| 15997 | 41 | (**** Transformation of Elimination Rules into First-Order Formulas****) | 
| 15347 | 42 | |
| 29064 | 43 | val cfalse = cterm_of @{theory HOL} HOLogic.false_const;
 | 
| 44 | val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop HOLogic.false_const);
 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 45 | |
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 46 | (*Converts an elim-rule into an equivalent theorem that does not have the | 
| 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 47 | predicate variable. Leaves other theorems unchanged. We simply instantiate the | 
| 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 48 | conclusion variable to False.*) | 
| 16009 | 49 | fun transform_elim th = | 
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 50 | case concl_of th of (*conclusion variable*) | 
| 24669 | 51 |        Const("Trueprop",_) $ (v as Var(_,Type("bool",[]))) =>
 | 
| 29064 | 52 |            Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th
 | 
| 24669 | 53 |     | v as Var(_, Type("prop",[])) =>
 | 
| 29064 | 54 |            Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th
 | 
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 55 | | _ => th; | 
| 15997 | 56 | |
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 57 | (*To enforce single-threading*) | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 58 | exception Clausify_failure of theory; | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 59 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 60 | |
| 16009 | 61 | (**** SKOLEMIZATION BY INFERENCE (lcp) ****) | 
| 62 | ||
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 63 | fun rhs_extra_types lhsT rhs = | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 64 | let val lhs_vars = Term.add_tfreesT lhsT [] | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 65 | fun add_new_TFrees (TFree v) = | 
| 24821 
cc55041ca8ba
skolem_cache: ignore internal theorems -- major speedup;
 wenzelm parents: 
24785diff
changeset | 66 | if member (op =) lhs_vars v then I else insert (op =) (TFree v) | 
| 
cc55041ca8ba
skolem_cache: ignore internal theorems -- major speedup;
 wenzelm parents: 
24785diff
changeset | 67 | | add_new_TFrees _ = I | 
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 68 | val rhs_consts = fold_aterms (fn Const c => insert (op =) c | _ => I) rhs [] | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 69 | in fold (#2 #> Term.fold_atyps add_new_TFrees) rhs_consts [] end; | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 70 | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 71 | (*Traverse a theorem, declaring Skolem function definitions. String s is the suggested | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 72 | prefix for the Skolem constant.*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 73 | fun declare_skofuns s th = | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 74 | let | 
| 33222 | 75 | val nref = Unsynchronized.ref 0 (* FIXME ??? *) | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 76 |     fun dec_sko (Const ("Ex",_) $ (xtp as Abs (_, T, p))) (axs, thy) =
 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 77 | (*Existential: declare a Skolem function, then insert into body and continue*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 78 | let | 
| 32740 | 79 | val cname = "sko_" ^ s ^ "_" ^ Int.toString (Unsynchronized.inc nref) | 
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
29064diff
changeset | 80 | val args0 = OldTerm.term_frees xtp (*get the formal parameter list*) | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 81 | val Ts = map type_of args0 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 82 | val extraTs = rhs_extra_types (Ts ---> T) xtp | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 83 | val argsx = map (fn T => Free (gensym "vsk", T)) extraTs | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 84 | val args = argsx @ args0 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 85 | val cT = extraTs ---> Ts ---> T | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 86 | val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 87 | (*Forms a lambda-abstraction over the formal parameters*) | 
| 28110 | 88 | val (c, thy') = | 
| 33173 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 wenzelm parents: 
33158diff
changeset | 89 | Sign.declare_const ((Binding.conceal (Binding.name cname), cT), NoSyn) thy | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 90 | val cdef = cname ^ "_def" | 
| 33222 | 91 | val thy'' = | 
| 92 | Theory.add_defs_i true false [(Binding.name cdef, Logic.mk_equals (c, rhs))] thy' | |
| 28965 | 93 | val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef) | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 94 | in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end | 
| 32994 | 95 |       | dec_sko (Const ("All", _) $ (Abs (a, T, p))) thx =
 | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 96 | (*Universal quant: insert a free variable into body and continue*) | 
| 29270 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 wenzelm parents: 
29265diff
changeset | 97 | let val fname = Name.variant (OldTerm.add_term_names (p, [])) a | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 98 | in dec_sko (subst_bound (Free (fname, T), p)) thx end | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 99 |       | dec_sko (Const ("op &", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 100 |       | dec_sko (Const ("op |", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 101 |       | dec_sko (Const ("Trueprop", _) $ p) thx = dec_sko p thx
 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 102 | | dec_sko t thx = thx (*Do nothing otherwise*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 103 | in fn thy => dec_sko (Thm.prop_of th) ([], thy) end; | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 104 | |
| 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 105 | (*Traverse a theorem, accumulating Skolem function definitions.*) | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 106 | fun assume_skofuns s th = | 
| 33222 | 107 | let val sko_count = Unsynchronized.ref 0 (* FIXME ??? *) | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 108 |       fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) defs =
 | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 109 | (*Existential: declare a Skolem function, then insert into body and continue*) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 110 | let val skos = map (#1 o Logic.dest_equals) defs (*existing sko fns*) | 
| 33040 | 111 | val args = subtract (op =) skos (OldTerm.term_frees xtp) (*the formal parameters*) | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 112 | val Ts = map type_of args | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 113 | val cT = Ts ---> T | 
| 32740 | 114 | val id = "sko_" ^ s ^ "_" ^ Int.toString (Unsynchronized.inc sko_count) | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 115 | val c = Free (id, cT) | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 116 | val rhs = list_abs_free (map dest_Free args, | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 117 | HOLogic.choice_const T $ xtp) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 118 | (*Forms a lambda-abstraction over the formal parameters*) | 
| 27330 | 119 | val def = Logic.mk_equals (c, rhs) | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 120 | in dec_sko (subst_bound (list_comb(c,args), p)) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 121 | (def :: defs) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 122 | end | 
| 32994 | 123 |         | dec_sko (Const ("All",_) $ Abs (a, T, p)) defs =
 | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 124 | (*Universal quant: insert a free variable into body and continue*) | 
| 29270 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 wenzelm parents: 
29265diff
changeset | 125 | let val fname = Name.variant (OldTerm.add_term_names (p,[])) a | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 126 | in dec_sko (subst_bound (Free(fname,T), p)) defs end | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 127 |         | dec_sko (Const ("op &", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
 | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 128 |         | dec_sko (Const ("op |", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
 | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 129 |         | dec_sko (Const ("Trueprop", _) $ p) defs = dec_sko p defs
 | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 130 | | dec_sko t defs = defs (*Do nothing otherwise*) | 
| 20419 | 131 | in dec_sko (prop_of th) [] end; | 
| 132 | ||
| 133 | ||
| 24827 | 134 | (**** REPLACING ABSTRACTIONS BY COMBINATORS ****) | 
| 20419 | 135 | |
| 136 | (*Returns the vars of a theorem*) | |
| 137 | fun vars_of_thm th = | |
| 22691 | 138 | map (Thm.cterm_of (theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th []); | 
| 20419 | 139 | |
| 140 | (*Make a version of fun_cong with a given variable name*) | |
| 141 | local | |
| 142 | val fun_cong' = fun_cong RS asm_rl; (*renumber f, g to prevent clashes with (a,0)*) | |
| 143 | val cx = hd (vars_of_thm fun_cong'); | |
| 144 | val ty = typ_of (ctyp_of_term cx); | |
| 20445 | 145 | val thy = theory_of_thm fun_cong; | 
| 20419 | 146 | fun mkvar a = cterm_of thy (Var((a,0),ty)); | 
| 147 | in | |
| 148 | fun xfun_cong x = Thm.instantiate ([], [(cx, mkvar x)]) fun_cong' | |
| 149 | end; | |
| 150 | ||
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 151 | (*Removes the lambdas from an equation of the form t = (%x. u). A non-negative n, | 
| 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 152 | serves as an upper bound on how many to remove.*) | 
| 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 153 | fun strip_lambdas 0 th = th | 
| 24669 | 154 | | strip_lambdas n th = | 
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 155 | case prop_of th of | 
| 24669 | 156 |           _ $ (Const ("op =", _) $ _ $ Abs (x,_,_)) =>
 | 
| 157 | strip_lambdas (n-1) (freeze_thm (th RS xfun_cong x)) | |
| 158 | | _ => th; | |
| 20419 | 159 | |
| 24669 | 160 | val lambda_free = not o Term.has_abs; | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 161 | |
| 32010 | 162 | val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B}));
 | 
| 163 | val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C}));
 | |
| 164 | val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S}));
 | |
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 165 | |
| 24827 | 166 | (*FIXME: requires more use of cterm constructors*) | 
| 167 | fun abstract ct = | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 168 | let | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 169 | val thy = theory_of_cterm ct | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 170 | val Abs(x,_,body) = term_of ct | 
| 24827 | 171 |       val Type("fun",[xT,bodyT]) = typ_of (ctyp_of_term ct)
 | 
| 172 | val cxT = ctyp_of thy xT and cbodyT = ctyp_of thy bodyT | |
| 27184 | 173 |       fun makeK() = instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] @{thm abs_K}
 | 
| 24827 | 174 | in | 
| 175 | case body of | |
| 176 | Const _ => makeK() | |
| 177 | | Free _ => makeK() | |
| 178 | | Var _ => makeK() (*though Var isn't expected*) | |
| 27184 | 179 |         | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
 | 
| 24827 | 180 | | rator$rand => | 
| 27184 | 181 | if loose_bvar1 (rator,0) then (*C or S*) | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 182 | if loose_bvar1 (rand,0) then (*S*) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 183 | let val crator = cterm_of thy (Abs(x,xT,rator)) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 184 | val crand = cterm_of thy (Abs(x,xT,rand)) | 
| 27184 | 185 |                      val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
 | 
| 186 | val (_,rhs) = Thm.dest_equals (cprop_of abs_S') | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 187 | in | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 188 | Thm.transitive abs_S' (Conv.binop_conv abstract rhs) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 189 | end | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 190 | else (*C*) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 191 | let val crator = cterm_of thy (Abs(x,xT,rator)) | 
| 27184 | 192 |                      val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
 | 
| 193 | val (_,rhs) = Thm.dest_equals (cprop_of abs_C') | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 194 | in | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 195 | Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 196 | end | 
| 27184 | 197 | else if loose_bvar1 (rand,0) then (*B or eta*) | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 198 | if rand = Bound 0 then eta_conversion ct | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 199 | else (*B*) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 200 | let val crand = cterm_of thy (Abs(x,xT,rand)) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 201 | val crator = cterm_of thy rator | 
| 27184 | 202 |                      val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
 | 
| 203 | val (_,rhs) = Thm.dest_equals (cprop_of abs_B') | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 204 | in | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 205 | Thm.transitive abs_B' (Conv.arg_conv abstract rhs) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 206 | end | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 207 | else makeK() | 
| 24827 | 208 | | _ => error "abstract: Bad term" | 
| 209 | end; | |
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 210 | |
| 20419 | 211 | (*Traverse a theorem, declaring abstraction function definitions. String s is the suggested | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 212 | prefix for the constants.*) | 
| 24827 | 213 | fun combinators_aux ct = | 
| 214 | if lambda_free (term_of ct) then reflexive ct | |
| 215 | else | |
| 216 | case term_of ct of | |
| 217 | Abs _ => | |
| 32994 | 218 | let val (cv, cta) = Thm.dest_abs NONE ct | 
| 219 | val (v, _) = dest_Free (term_of cv) | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 220 | val u_th = combinators_aux cta | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 221 | val cu = Thm.rhs_of u_th | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 222 | val comb_eq = abstract (Thm.cabs cv cu) | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 223 | in transitive (abstract_rule v cv u_th) comb_eq end | 
| 32994 | 224 | | _ $ _ => | 
| 225 | let val (ct1, ct2) = Thm.dest_comb ct | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 226 | in combination (combinators_aux ct1) (combinators_aux ct2) end; | 
| 27184 | 227 | |
| 24827 | 228 | fun combinators th = | 
| 27184 | 229 | if lambda_free (prop_of th) then th | 
| 24827 | 230 | else | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 231 | let val th = Drule.eta_contraction_rule th | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 232 | val eqth = combinators_aux (cprop_of th) | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 233 | in equal_elim eqth th end | 
| 27184 | 234 | handle THM (msg,_,_) => | 
| 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: 
32010diff
changeset | 235 | (warning (cat_lines | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32010diff
changeset | 236 | ["Error in the combinator translation of " ^ Display.string_of_thm_without_context th, | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32010diff
changeset | 237 | " Exception message: " ^ msg]); | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 238 |        TrueI);  (*A type variable of sort {} will cause make abstraction fail.*)
 | 
| 16009 | 239 | |
| 240 | (*cterms are used throughout for efficiency*) | |
| 29064 | 241 | val cTrueprop = Thm.cterm_of @{theory HOL} HOLogic.Trueprop;
 | 
| 16009 | 242 | |
| 243 | (*cterm version of mk_cTrueprop*) | |
| 244 | fun c_mkTrueprop A = Thm.capply cTrueprop A; | |
| 245 | ||
| 246 | (*Given an abstraction over n variables, replace the bound variables by free | |
| 247 | ones. Return the body, along with the list of free variables.*) | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 248 | fun c_variant_abs_multi (ct0, vars) = | 
| 16009 | 249 | let val (cv,ct) = Thm.dest_abs NONE ct0 | 
| 250 | in c_variant_abs_multi (ct, cv::vars) end | |
| 251 | handle CTERM _ => (ct0, rev vars); | |
| 252 | ||
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 253 | (*Given the definition of a Skolem function, return a theorem to replace | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 254 | an existential formula by a use of that function. | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 255 | Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 256 | fun skolem_of_def def = | 
| 22902 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 wenzelm parents: 
22846diff
changeset | 257 | let val (c,rhs) = Thm.dest_equals (cprop_of (freeze_thm def)) | 
| 16009 | 258 | val (ch, frees) = c_variant_abs_multi (rhs, []) | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 259 | val (chilbert,cabs) = Thm.dest_comb ch | 
| 26627 
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26618diff
changeset | 260 | val thy = Thm.theory_of_cterm chilbert | 
| 
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26618diff
changeset | 261 | val t = Thm.term_of chilbert | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 262 |       val T = case t of Const ("Hilbert_Choice.Eps", Type("fun",[_,T])) => T
 | 
| 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 263 |                       | _ => raise THM ("skolem_of_def: expected Eps", 0, [def])
 | 
| 22596 | 264 | val cex = Thm.cterm_of thy (HOLogic.exists_const T) | 
| 16009 | 265 | val ex_tm = c_mkTrueprop (Thm.capply cex cabs) | 
| 266 | and conc = c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees))); | |
| 31454 | 267 |       fun tacf [prem] = rewrite_goals_tac [def] THEN rtac (prem RS @{thm someI_ex}) 1
 | 
| 23352 | 268 | in Goal.prove_internal [ex_tm] conc tacf | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 269 | |> forall_intr_list frees | 
| 26653 | 270 | |> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 271 | |> Thm.varifyT | 
| 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 272 | end; | 
| 16009 | 273 | |
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 274 | |
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 275 | (*Converts an Isabelle theorem (intro, elim or simp format, even higher-order) into NNF.*) | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24854diff
changeset | 276 | fun to_nnf th ctxt0 = | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 277 | let val th1 = th |> transform_elim |> zero_var_indexes | 
| 32262 | 278 | val ((_, [th2]), ctxt) = Variable.import true [th1] ctxt0 | 
| 279 | val th3 = th2 | |
| 280 | |> Conv.fconv_rule ObjectLogic.atomize | |
| 281 | |> Meson.make_nnf ctxt |> strip_lambdas ~1 | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24854diff
changeset | 282 | in (th3, ctxt) end; | 
| 16009 | 283 | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 284 | (*Generate Skolem functions for a theorem supplied in nnf*) | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24854diff
changeset | 285 | fun assume_skolem_of_def s th = | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 286 | map (skolem_of_def o assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th); | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 287 | |
| 25007 | 288 | |
| 33316 | 289 | (*** Blacklisting (duplicated in Res_ATP?) ***) | 
| 25007 | 290 | |
| 291 | val max_lambda_nesting = 3; | |
| 27184 | 292 | |
| 25007 | 293 | fun excessive_lambdas (f$t, k) = excessive_lambdas (f,k) orelse excessive_lambdas (t,k) | 
| 294 | | excessive_lambdas (Abs(_,_,t), k) = k=0 orelse excessive_lambdas (t,k-1) | |
| 295 | | excessive_lambdas _ = false; | |
| 296 | ||
| 297 | fun is_formula_type T = (T = HOLogic.boolT orelse T = propT); | |
| 298 | ||
| 299 | (*Don't count nested lambdas at the level of formulas, as they are quantifiers*) | |
| 300 | fun excessive_lambdas_fm Ts (Abs(_,T,t)) = excessive_lambdas_fm (T::Ts) t | |
| 301 | | excessive_lambdas_fm Ts t = | |
| 302 | if is_formula_type (fastype_of1 (Ts, t)) | |
| 303 | then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t)) | |
| 304 | else excessive_lambdas (t, max_lambda_nesting); | |
| 305 | ||
| 33027 | 306 | (*The max apply_depth of any metis call in Metis_Examples (on 31-10-2007) was 11.*) | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 307 | val max_apply_depth = 15; | 
| 27184 | 308 | |
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 309 | fun apply_depth (f$t) = Int.max (apply_depth f, apply_depth t + 1) | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 310 | | apply_depth (Abs(_,_,t)) = apply_depth t | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 311 | | apply_depth _ = 0; | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 312 | |
| 27184 | 313 | fun too_complex t = | 
| 314 | apply_depth t > max_apply_depth orelse | |
| 26562 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 paulson parents: 
25761diff
changeset | 315 | Meson.too_many_clauses NONE t orelse | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 316 | excessive_lambdas_fm [] t; | 
| 27184 | 317 | |
| 25243 | 318 | fun is_strange_thm th = | 
| 319 | case head_of (concl_of th) of | |
| 33306 
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
 wenzelm parents: 
33222diff
changeset | 320 | Const (a, _) => (a <> "Trueprop" andalso a <> "==") | 
| 25243 | 321 | | _ => false; | 
| 322 | ||
| 27184 | 323 | fun bad_for_atp th = | 
| 33306 
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
 wenzelm parents: 
33222diff
changeset | 324 | too_complex (prop_of th) | 
| 27184 | 325 | orelse exists_type type_has_empty_sort (prop_of th) | 
| 25761 | 326 | orelse is_strange_thm th; | 
| 25243 | 327 | |
| 25007 | 328 | val multi_base_blacklist = | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 329 | ["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm", | 
| 33306 
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
 wenzelm parents: 
33222diff
changeset | 330 | "cases","ext_cases"]; (* FIXME put other record thms here, or declare as "noatp" *) | 
| 25007 | 331 | |
| 21071 | 332 | (*Keep the full complexity of the original name*) | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30291diff
changeset | 333 | fun flatten_name s = space_implode "_X" (Long_Name.explode s); | 
| 21071 | 334 | |
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 335 | fun fake_name th = | 
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
27809diff
changeset | 336 | if Thm.has_name_hint th then flatten_name (Thm.get_name_hint th) | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 337 | else gensym "unknown_thm_"; | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 338 | |
| 27184 | 339 | (*Skolemize a named theorem, with Skolem functions as additional premises.*) | 
| 340 | fun skolem_thm (s, th) = | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30291diff
changeset | 341 | if member (op =) multi_base_blacklist (Long_Name.base_name s) orelse bad_for_atp th then [] | 
| 27184 | 342 | else | 
| 343 | let | |
| 344 | val ctxt0 = Variable.thm_context th | |
| 345 | val (nnfth, ctxt1) = to_nnf th ctxt0 | |
| 346 | val (cnfs, ctxt2) = Meson.make_cnf (assume_skolem_of_def s nnfth) nnfth ctxt1 | |
| 347 | in cnfs |> map combinators |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf end | |
| 348 | handle THM _ => []; | |
| 349 | ||
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 350 | (*The cache prevents repeated clausification of a theorem, and also repeated declaration of | 
| 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 351 | Skolem functions.*) | 
| 22516 | 352 | structure ThmCache = TheoryDataFun | 
| 22846 | 353 | ( | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 354 | type T = thm list Thmtab.table * unit Symtab.table; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 355 | val empty = (Thmtab.empty, Symtab.empty); | 
| 26618 | 356 | val copy = I; | 
| 357 | val extend = I; | |
| 27184 | 358 | fun merge _ ((cache1, seen1), (cache2, seen2)) : T = | 
| 359 | (Thmtab.merge (K true) (cache1, cache2), Symtab.merge (K true) (seen1, seen2)); | |
| 22846 | 360 | ); | 
| 22516 | 361 | |
| 27184 | 362 | val lookup_cache = Thmtab.lookup o #1 o ThmCache.get; | 
| 363 | val already_seen = Symtab.defined o #2 o ThmCache.get; | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 364 | |
| 27184 | 365 | val update_cache = ThmCache.map o apfst o Thmtab.update; | 
| 366 | fun mark_seen name = ThmCache.map (apsnd (Symtab.update (name, ()))); | |
| 25007 | 367 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 368 | (*Exported function to convert Isabelle theorems into axiom clauses*) | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 369 | fun cnf_axiom thy th0 = | 
| 27184 | 370 | let val th = Thm.transfer thy th0 in | 
| 371 | case lookup_cache thy th of | |
| 372 | NONE => map Thm.close_derivation (skolem_thm (fake_name th, th)) | |
| 373 | | SOME cls => cls | |
| 22516 | 374 | end; | 
| 15347 | 375 | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 376 | |
| 30291 
a1c3abf57068
removed obsolete claset_rules_of, simpset_rules_of -- as proposed in the text;
 wenzelm parents: 
30280diff
changeset | 377 | (**** Rules from the context ****) | 
| 15347 | 378 | |
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
27809diff
changeset | 379 | fun pairname th = (Thm.get_name_hint th, th); | 
| 27184 | 380 | |
| 15347 | 381 | |
| 22471 | 382 | (**** Translate a set of theorems into CNF ****) | 
| 15347 | 383 | |
| 19894 | 384 | fun pair_name_cls k (n, []) = [] | 
| 385 | | pair_name_cls k (n, cls::clss) = (cls, (n,k)) :: pair_name_cls (k+1) (n, clss) | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 386 | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 387 | fun cnf_rules_pairs_aux _ pairs [] = pairs | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 388 | | cnf_rules_pairs_aux thy pairs ((name,th)::ths) = | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 389 | let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs | 
| 33316 | 390 | handle THM _ => pairs | Res_Clause.CLAUSE _ => pairs | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 391 | in cnf_rules_pairs_aux thy pairs' ths end; | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 392 | |
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21254diff
changeset | 393 | (*The combination of rev and tail recursion preserves the original order*) | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 394 | fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l); | 
| 19353 | 395 | |
| 19196 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 mengj parents: 
19175diff
changeset | 396 | |
| 33316 | 397 | (**** Convert all facts of the theory into clauses (Res_Clause.clause, or Res_HOL_Clause.clause) ****) | 
| 15347 | 398 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 399 | local | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 400 | |
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 401 | fun skolem_def (name, th) thy = | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 402 | let val ctxt0 = Variable.thm_context th in | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 403 | (case try (to_nnf th) ctxt0 of | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 404 | NONE => (NONE, thy) | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 405 | | SOME (nnfth, ctxt1) => | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 406 | let val (defs, thy') = declare_skofuns (flatten_name name) nnfth thy | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 407 | in (SOME (th, ctxt0, ctxt1, nnfth, defs), thy') end) | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 408 | end; | 
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 409 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 410 | fun skolem_cnfs (th, ctxt0, ctxt1, nnfth, defs) = | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 411 | let | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 412 | val (cnfs, ctxt2) = Meson.make_cnf (map skolem_of_def defs) nnfth ctxt1; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 413 | val cnfs' = cnfs | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 414 | |> map combinators | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 415 | |> Variable.export ctxt2 ctxt0 | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 416 | |> Meson.finish_cnf | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 417 | |> map Thm.close_derivation; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 418 | in (th, cnfs') end; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 419 | |
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 420 | in | 
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 421 | |
| 27184 | 422 | fun saturate_skolem_cache thy = | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 423 | let | 
| 33306 
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
 wenzelm parents: 
33222diff
changeset | 424 | val facts = PureThy.facts_of thy; | 
| 
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
 wenzelm parents: 
33222diff
changeset | 425 | val new_facts = (facts, []) |-> Facts.fold_static (fn (name, ths) => | 
| 
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
 wenzelm parents: 
33222diff
changeset | 426 | if Facts.is_concealed facts name orelse already_seen thy name then I | 
| 
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
 wenzelm parents: 
33222diff
changeset | 427 | else cons (name, ths)); | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 428 | val new_thms = (new_facts, []) |-> fold (fn (name, ths) => | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30291diff
changeset | 429 | if member (op =) multi_base_blacklist (Long_Name.base_name name) then I | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 430 | else fold_index (fn (i, th) => | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 431 | if bad_for_atp th orelse is_some (lookup_cache thy th) then I | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 432 | else cons (name ^ "_" ^ string_of_int (i + 1), Thm.transfer thy th)) ths); | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 433 | in | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 434 | if null new_facts then NONE | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 435 | else | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 436 | let | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 437 | val (defs, thy') = thy | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 438 | |> fold (mark_seen o #1) new_facts | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 439 | |> fold_map skolem_def (sort_distinct (Thm.thm_ord o pairself snd) new_thms) | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 440 | |>> map_filter I; | 
| 29368 | 441 | val cache_entries = Par_List.map skolem_cnfs defs; | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 442 | in SOME (fold update_cache cache_entries thy') end | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 443 | end; | 
| 27184 | 444 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 445 | end; | 
| 24854 | 446 | |
| 32740 | 447 | val suppress_endtheory = Unsynchronized.ref false; | 
| 27184 | 448 | |
| 449 | fun clause_cache_endtheory thy = | |
| 450 | if ! suppress_endtheory then NONE | |
| 451 | else saturate_skolem_cache thy; | |
| 452 | ||
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20445diff
changeset | 453 | |
| 22516 | 454 | (*The cache can be kept smaller by inspecting the prop of each thm. Can ignore all that are | 
| 455 | lambda_free, but then the individual theory caches become much bigger.*) | |
| 21071 | 456 | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 457 | |
| 16563 | 458 | (*** meson proof methods ***) | 
| 459 | ||
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 460 | (*Expand all new definitions of abstraction or Skolem functions in a proof state.*) | 
| 24827 | 461 | fun is_absko (Const ("==", _) $ Free (a,_) $ u) = String.isPrefix "sko_" a
 | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 462 | | is_absko _ = false; | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 463 | |
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 464 | fun is_okdef xs (Const ("==", _) $ t $ u) =   (*Definition of Free, not in certain terms*)
 | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 465 | is_Free t andalso not (member (op aconv) xs t) | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 466 | | is_okdef _ _ = false | 
| 22724 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 paulson parents: 
22691diff
changeset | 467 | |
| 24215 | 468 | (*This function tries to cope with open locales, which introduce hypotheses of the form | 
| 469 | Free == t, conjecture clauses, which introduce various hypotheses, and also definitions | |
| 24827 | 470 | of sko_ functions. *) | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 471 | fun expand_defs_tac st0 st = | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 472 | let val hyps0 = #hyps (rep_thm st0) | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 473 | val hyps = #hyps (crep_thm st) | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 474 | val newhyps = filter_out (member (op aconv) hyps0 o Thm.term_of) hyps | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 475 | val defs = filter (is_absko o Thm.term_of) newhyps | 
| 24669 | 476 | val remaining_hyps = filter_out (member (op aconv) (map Thm.term_of defs)) | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 477 | (map Thm.term_of hyps) | 
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
29064diff
changeset | 478 | val fixed = OldTerm.term_frees (concl_of st) @ | 
| 33339 | 479 | fold (union (op aconv)) (map OldTerm.term_frees remaining_hyps) [] | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 480 | in Seq.of_list [LocalDefs.expand (filter (is_okdef fixed o Thm.term_of) defs) st] end; | 
| 22724 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 paulson parents: 
22691diff
changeset | 481 | |
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 482 | |
| 32262 | 483 | fun meson_general_tac ctxt ths i st0 = | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 484 | let | 
| 32262 | 485 | val thy = ProofContext.theory_of ctxt | 
| 486 | val ctxt0 = Classical.put_claset HOL_cs ctxt | |
| 487 | in (Meson.meson_tac ctxt0 (maps (cnf_axiom thy) ths) i THEN expand_defs_tac st0) st0 end; | |
| 22724 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 paulson parents: 
22691diff
changeset | 488 | |
| 30515 | 489 | val meson_method_setup = | 
| 32262 | 490 |   Method.setup @{binding meson} (Attrib.thms >> (fn ths => fn ctxt =>
 | 
| 491 | SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ctxt ths))) | |
| 30515 | 492 | "MESON resolution proof procedure"; | 
| 15347 | 493 | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 494 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 495 | (*** Converting a subgoal into negated conjecture clauses. ***) | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 496 | |
| 32262 | 497 | fun neg_skolemize_tac ctxt = | 
| 498 | EVERY' [rtac ccontr, ObjectLogic.atomize_prems_tac, Meson.skolemize_tac ctxt]; | |
| 22471 | 499 | |
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 500 | val neg_clausify = Meson.make_clauses #> map combinators #> Meson.finish_cnf; | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 501 | |
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 502 | fun neg_conjecture_clauses ctxt st0 n = | 
| 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 503 | let | 
| 32262 | 504 | val st = Seq.hd (neg_skolemize_tac ctxt n st0) | 
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 505 |     val ({params, prems, ...}, _) = Subgoal.focus (Variable.set_body false ctxt) n st
 | 
| 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 506 | in (neg_clausify prems, map (Term.dest_Free o Thm.term_of o #2) params) end; | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 507 | |
| 24669 | 508 | (*Conversion of a subgoal to conjecture clauses. Each clause has | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 509 | leading !!-bound universal variables, to express generality. *) | 
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 510 | fun neg_clausify_tac ctxt = | 
| 32262 | 511 | neg_skolemize_tac ctxt THEN' | 
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 512 | SUBGOAL (fn (prop, i) => | 
| 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 513 | let val ts = Logic.strip_assums_hyp prop in | 
| 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 514 | EVERY' | 
| 32283 | 515 | [Subgoal.FOCUS | 
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 516 |          (fn {prems, ...} =>
 | 
| 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 517 | (Method.insert_tac | 
| 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 518 | (map forall_intr_vars (neg_clausify prems)) i)) ctxt, | 
| 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 519 | REPEAT_DETERM_N (length ts) o etac thin_rl] i | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 520 | end); | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 521 | |
| 30722 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 522 | val neg_clausify_setup = | 
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32231diff
changeset | 523 |   Method.setup @{binding neg_clausify} (Scan.succeed (SIMPLE_METHOD' o neg_clausify_tac))
 | 
| 30515 | 524 | "conversion of goal to conjecture clauses"; | 
| 24669 | 525 | |
| 27184 | 526 | |
| 527 | (** Attribute for converting a theorem into clauses **) | |
| 528 | ||
| 30722 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 529 | val clausify_setup = | 
| 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 530 |   Attrib.setup @{binding clausify}
 | 
| 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 531 | (Scan.lift OuterParse.nat >> | 
| 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 532 | (fn i => Thm.rule_attribute (fn context => fn th => | 
| 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 533 | Meson.make_meta_clause (nth (cnf_axiom (Context.theory_of context) th) i)))) | 
| 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 534 | "conversion of theorem to clauses"; | 
| 27184 | 535 | |
| 536 | ||
| 537 | ||
| 538 | (** setup **) | |
| 539 | ||
| 540 | val setup = | |
| 541 | meson_method_setup #> | |
| 30722 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 542 | neg_clausify_setup #> | 
| 
623d4831c8cf
simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
 wenzelm parents: 
30528diff
changeset | 543 | clausify_setup #> | 
| 27184 | 544 | perhaps saturate_skolem_cache #> | 
| 545 | Theory.at_end clause_cache_endtheory; | |
| 18510 
0a6c24f549c3
the "skolem" attribute and better initialization of the clause database
 paulson parents: 
18404diff
changeset | 546 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 547 | end; |