| author | paulson | 
| Thu, 11 Dec 2008 14:50:02 +0000 | |
| changeset 29679 | a624dc56e859 | 
| parent 28965 | 1de908189869 | 
| child 29064 | 70a61d58460e | 
| permissions | -rw-r--r-- | 
| 15347 | 1 | (* Author: Jia Meng, Cambridge University Computer Laboratory | 
| 2 | ID: $Id$ | |
| 3 | Copyright 2004 University of Cambridge | |
| 4 | ||
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 5 | Transformation of axiom rules (elim/intro/etc) into CNF forms. | 
| 15347 | 6 | *) | 
| 7 | ||
| 15997 | 8 | signature RES_AXIOMS = | 
| 21505 | 9 | sig | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 10 | val cnf_axiom: theory -> thm -> thm list | 
| 24669 | 11 | val pairname: thm -> string * thm | 
| 27184 | 12 | val multi_base_blacklist: string list | 
| 25243 | 13 | val bad_for_atp: thm -> bool | 
| 25761 | 14 | val type_has_empty_sort: typ -> bool | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 15 | val cnf_rules_pairs: theory -> (string * thm) list -> (thm * (string * int)) list | 
| 24669 | 16 | val neg_clausify: thm list -> thm list | 
| 17 | val expand_defs_tac: thm -> tactic | |
| 24827 | 18 | val combinators: thm -> thm | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 19 | val neg_conjecture_clauses: thm -> int -> thm list * (string * typ) list | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 20 | val claset_rules_of: Proof.context -> (string * thm) list (*FIXME DELETE*) | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 21 | val simpset_rules_of: Proof.context -> (string * thm) list (*FIXME DELETE*) | 
| 21505 | 22 | val atpset_rules_of: Proof.context -> (string * thm) list | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 23 | val suppress_endtheory: bool ref (*for emergency use where endtheory causes problems*) | 
| 24669 | 24 | val setup: theory -> theory | 
| 21505 | 25 | end; | 
| 19196 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 mengj parents: 
19175diff
changeset | 26 | |
| 24669 | 27 | structure ResAxioms: RES_AXIOMS = | 
| 15997 | 28 | struct | 
| 15347 | 29 | |
| 20902 | 30 | (* FIXME legacy *) | 
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 31 | fun freeze_thm th = #1 (Drule.freeze_thaw th); | 
| 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 32 | |
| 25761 | 33 | fun type_has_empty_sort (TFree (_, [])) = true | 
| 34 | | type_has_empty_sort (TVar (_, [])) = true | |
| 35 | | type_has_empty_sort (Type (_, Ts)) = exists type_has_empty_sort Ts | |
| 36 | | type_has_empty_sort _ = false; | |
| 27184 | 37 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 38 | |
| 15997 | 39 | (**** Transformation of Elimination Rules into First-Order Formulas****) | 
| 15347 | 40 | |
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 41 | val cfalse = cterm_of HOL.thy HOLogic.false_const; | 
| 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 42 | val ctp_false = cterm_of HOL.thy (HOLogic.mk_Trueprop HOLogic.false_const); | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 43 | |
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 44 | (*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 | 45 | 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 | 46 | conclusion variable to False.*) | 
| 16009 | 47 | fun transform_elim th = | 
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 48 | case concl_of th of (*conclusion variable*) | 
| 24669 | 49 |        Const("Trueprop",_) $ (v as Var(_,Type("bool",[]))) =>
 | 
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 50 | Thm.instantiate ([], [(cterm_of HOL.thy v, cfalse)]) th | 
| 24669 | 51 |     | v as Var(_, Type("prop",[])) =>
 | 
| 21430 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 52 | Thm.instantiate ([], [(cterm_of HOL.thy v, ctp_false)]) th | 
| 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 paulson parents: 
21290diff
changeset | 53 | | _ => th; | 
| 15997 | 54 | |
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 55 | (*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 | 56 | exception Clausify_failure of theory; | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 57 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 58 | |
| 16009 | 59 | (**** SKOLEMIZATION BY INFERENCE (lcp) ****) | 
| 60 | ||
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 61 | 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 | 62 | 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 | 63 | fun add_new_TFrees (TFree v) = | 
| 24821 
cc55041ca8ba
skolem_cache: ignore internal theorems -- major speedup;
 wenzelm parents: 
24785diff
changeset | 64 | 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 | 65 | | 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 | 66 | 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 | 67 | 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 | 68 | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 69 | (*Traverse a theorem, declaring Skolem function definitions. String s is the suggested | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 70 | prefix for the Skolem constant.*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 71 | fun declare_skofuns s th = | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 72 | let | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 73 | val nref = ref 0 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 74 |     fun dec_sko (Const ("Ex",_) $ (xtp as Abs (_, T, p))) (axs, thy) =
 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 75 | (*Existential: declare a Skolem function, then insert into body and continue*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 76 | let | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 77 | val cname = "sko_" ^ s ^ "_" ^ Int.toString (inc nref) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 78 | val args0 = term_frees xtp (*get the formal parameter list*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 79 | val Ts = map type_of args0 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 80 | val extraTs = rhs_extra_types (Ts ---> T) xtp | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 81 | val argsx = map (fn T => Free (gensym "vsk", T)) extraTs | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 82 | val args = argsx @ args0 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 83 | val cT = extraTs ---> Ts ---> T | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 84 | 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 | 85 | (*Forms a lambda-abstraction over the formal parameters*) | 
| 28110 | 86 | val (c, thy') = | 
| 28965 | 87 | Sign.declare_const [Markup.property_internal] ((Binding.name cname, cT), NoSyn) thy | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 88 | val cdef = cname ^ "_def" | 
| 27330 | 89 | val thy'' = Theory.add_defs_i true false [(cdef, Logic.mk_equals (c, rhs))] thy' | 
| 28965 | 90 | val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef) | 
| 27174 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 91 | in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 92 |       | dec_sko (Const ("All", _) $ (xtp as Abs (a, T, p))) thx =
 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 93 | (*Universal quant: insert a free variable into body and continue*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 94 | let val fname = Name.variant (add_term_names (p, [])) a | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 95 | in dec_sko (subst_bound (Free (fname, T), p)) thx end | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 96 |       | 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 | 97 |       | 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 | 98 |       | dec_sko (Const ("Trueprop", _) $ p) thx = dec_sko p thx
 | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 99 | | dec_sko t thx = thx (*Do nothing otherwise*) | 
| 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 wenzelm parents: 
26939diff
changeset | 100 | in fn thy => dec_sko (Thm.prop_of th) ([], thy) end; | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 101 | |
| 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 102 | (*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 | 103 | fun assume_skofuns s th = | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 104 | let val sko_count = ref 0 | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 105 |       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 | 106 | (*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 | 107 | let val skos = map (#1 o Logic.dest_equals) defs (*existing sko fns*) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 108 | val args = term_frees xtp \\ skos (*the formal parameters*) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 109 | val Ts = map type_of args | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 110 | val cT = Ts ---> T | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 111 | val id = "sko_" ^ s ^ "_" ^ Int.toString (inc sko_count) | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 112 | val c = Free (id, cT) | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 113 | 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 | 114 | HOLogic.choice_const T $ xtp) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 115 | (*Forms a lambda-abstraction over the formal parameters*) | 
| 27330 | 116 | val def = Logic.mk_equals (c, rhs) | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 117 | 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 | 118 | (def :: defs) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 119 | end | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 120 |         | dec_sko (Const ("All",_) $ (xtp as Abs(a,T,p))) defs =
 | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 121 | (*Universal quant: insert a free variable into body and continue*) | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 122 | let val fname = Name.variant (add_term_names (p,[])) a | 
| 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 123 | 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 | 124 |         | 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 | 125 |         | 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 | 126 |         | 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 | 127 | | dec_sko t defs = defs (*Do nothing otherwise*) | 
| 20419 | 128 | in dec_sko (prop_of th) [] end; | 
| 129 | ||
| 130 | ||
| 24827 | 131 | (**** REPLACING ABSTRACTIONS BY COMBINATORS ****) | 
| 20419 | 132 | |
| 133 | (*Returns the vars of a theorem*) | |
| 134 | fun vars_of_thm th = | |
| 22691 | 135 | map (Thm.cterm_of (theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th []); | 
| 20419 | 136 | |
| 137 | (*Make a version of fun_cong with a given variable name*) | |
| 138 | local | |
| 139 | val fun_cong' = fun_cong RS asm_rl; (*renumber f, g to prevent clashes with (a,0)*) | |
| 140 | val cx = hd (vars_of_thm fun_cong'); | |
| 141 | val ty = typ_of (ctyp_of_term cx); | |
| 20445 | 142 | val thy = theory_of_thm fun_cong; | 
| 20419 | 143 | fun mkvar a = cterm_of thy (Var((a,0),ty)); | 
| 144 | in | |
| 145 | fun xfun_cong x = Thm.instantiate ([], [(cx, mkvar x)]) fun_cong' | |
| 146 | end; | |
| 147 | ||
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 148 | (*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 | 149 | 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 | 150 | fun strip_lambdas 0 th = th | 
| 24669 | 151 | | strip_lambdas n th = | 
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 152 | case prop_of th of | 
| 24669 | 153 |           _ $ (Const ("op =", _) $ _ $ Abs (x,_,_)) =>
 | 
| 154 | strip_lambdas (n-1) (freeze_thm (th RS xfun_cong x)) | |
| 155 | | _ => th; | |
| 20419 | 156 | |
| 24669 | 157 | 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 | 158 | |
| 24669 | 159 | val monomorphic = not o Term.exists_type (Term.exists_subtype Term.is_TVar); | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 160 | |
| 28262 
aa7ca36d67fd
back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
 wenzelm parents: 
28110diff
changeset | 161 | val [f_B,g_B] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_B}));
 | 
| 
aa7ca36d67fd
back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
 wenzelm parents: 
28110diff
changeset | 162 | val [g_C,f_C] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_C}));
 | 
| 
aa7ca36d67fd
back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
 wenzelm parents: 
28110diff
changeset | 163 | val [f_S,g_S] = map (cterm_of (the_context ())) (term_vars (prop_of @{thm abs_S}));
 | 
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 164 | |
| 24827 | 165 | (*FIXME: requires more use of cterm constructors*) | 
| 166 | fun abstract ct = | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 167 | let | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 168 | val thy = theory_of_cterm ct | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 169 | val Abs(x,_,body) = term_of ct | 
| 24827 | 170 |       val Type("fun",[xT,bodyT]) = typ_of (ctyp_of_term ct)
 | 
| 171 | val cxT = ctyp_of thy xT and cbodyT = ctyp_of thy bodyT | |
| 27184 | 172 |       fun makeK() = instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] @{thm abs_K}
 | 
| 24827 | 173 | in | 
| 174 | case body of | |
| 175 | Const _ => makeK() | |
| 176 | | Free _ => makeK() | |
| 177 | | Var _ => makeK() (*though Var isn't expected*) | |
| 27184 | 178 |         | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
 | 
| 24827 | 179 | | rator$rand => | 
| 27184 | 180 | 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 | 181 | if loose_bvar1 (rand,0) then (*S*) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 182 | 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 | 183 | val crand = cterm_of thy (Abs(x,xT,rand)) | 
| 27184 | 184 |                      val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
 | 
| 185 | 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 | 186 | in | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 187 | Thm.transitive abs_S' (Conv.binop_conv abstract rhs) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 188 | end | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 189 | else (*C*) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 190 | let val crator = cterm_of thy (Abs(x,xT,rator)) | 
| 27184 | 191 |                      val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
 | 
| 192 | 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 | 193 | in | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 194 | 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 | 195 | end | 
| 27184 | 196 | 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 | 197 | if rand = Bound 0 then eta_conversion ct | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 198 | else (*B*) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 199 | 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 | 200 | val crator = cterm_of thy rator | 
| 27184 | 201 |                      val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
 | 
| 202 | 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 | 203 | in | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 204 | Thm.transitive abs_B' (Conv.arg_conv abstract rhs) | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 205 | end | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 206 | else makeK() | 
| 24827 | 207 | | _ => error "abstract: Bad term" | 
| 208 | end; | |
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 209 | |
| 20419 | 210 | (*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 | 211 | prefix for the constants.*) | 
| 24827 | 212 | fun combinators_aux ct = | 
| 213 | if lambda_free (term_of ct) then reflexive ct | |
| 214 | else | |
| 215 | case term_of ct of | |
| 216 | Abs _ => | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 217 | let val (cv,cta) = Thm.dest_abs NONE ct | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 218 | val (v,Tv) = (dest_Free o term_of) cv | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 219 | val u_th = combinators_aux cta | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 220 | val cu = Thm.rhs_of u_th | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 221 | val comb_eq = abstract (Thm.cabs cv cu) | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 222 | in transitive (abstract_rule v cv u_th) comb_eq end | 
| 24827 | 223 | | t1 $ t2 => | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 224 | let val (ct1,ct2) = Thm.dest_comb ct | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 225 | in combination (combinators_aux ct1) (combinators_aux ct2) end; | 
| 27184 | 226 | |
| 24827 | 227 | fun combinators th = | 
| 27184 | 228 | if lambda_free (prop_of th) then th | 
| 24827 | 229 | else | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 230 | let val th = Drule.eta_contraction_rule th | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 231 | val eqth = combinators_aux (cprop_of th) | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 232 | in equal_elim eqth th end | 
| 27184 | 233 | handle THM (msg,_,_) => | 
| 26928 | 234 |       (warning ("Error in the combinator translation of " ^ Display.string_of_thm th);
 | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 235 |        warning ("  Exception message: " ^ msg);
 | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 236 |        TrueI);  (*A type variable of sort {} will cause make abstraction fail.*)
 | 
| 16009 | 237 | |
| 238 | (*cterms are used throughout for efficiency*) | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 239 | val cTrueprop = Thm.cterm_of HOL.thy HOLogic.Trueprop; | 
| 16009 | 240 | |
| 241 | (*cterm version of mk_cTrueprop*) | |
| 242 | fun c_mkTrueprop A = Thm.capply cTrueprop A; | |
| 243 | ||
| 244 | (*Given an abstraction over n variables, replace the bound variables by free | |
| 245 | 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 | 246 | fun c_variant_abs_multi (ct0, vars) = | 
| 16009 | 247 | let val (cv,ct) = Thm.dest_abs NONE ct0 | 
| 248 | in c_variant_abs_multi (ct, cv::vars) end | |
| 249 | handle CTERM _ => (ct0, rev vars); | |
| 250 | ||
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 251 | (*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 | 252 | an existential formula by a use of that function. | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 253 | 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 | 254 | fun skolem_of_def def = | 
| 22902 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 wenzelm parents: 
22846diff
changeset | 255 | let val (c,rhs) = Thm.dest_equals (cprop_of (freeze_thm def)) | 
| 16009 | 256 | val (ch, frees) = c_variant_abs_multi (rhs, []) | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 257 | val (chilbert,cabs) = Thm.dest_comb ch | 
| 26627 
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26618diff
changeset | 258 | val thy = Thm.theory_of_cterm chilbert | 
| 
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26618diff
changeset | 259 | val t = Thm.term_of chilbert | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 260 |       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 | 261 |                       | _ => raise THM ("skolem_of_def: expected Eps", 0, [def])
 | 
| 22596 | 262 | val cex = Thm.cterm_of thy (HOLogic.exists_const T) | 
| 16009 | 263 | val ex_tm = c_mkTrueprop (Thm.capply cex cabs) | 
| 264 | and conc = c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees))); | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 265 | fun tacf [prem] = rewrite_goals_tac [def] THEN rtac (prem RS someI_ex) 1 | 
| 23352 | 266 | in Goal.prove_internal [ex_tm] conc tacf | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 267 | |> forall_intr_list frees | 
| 26653 | 268 | |> 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 | 269 | |> Thm.varifyT | 
| 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 270 | end; | 
| 16009 | 271 | |
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 272 | |
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 273 | (*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 | 274 | fun to_nnf th ctxt0 = | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 275 | let val th1 = th |> transform_elim |> zero_var_indexes | 
| 27184 | 276 | val ((_,[th2]),ctxt) = Variable.import_thms true [th1] ctxt0 | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24854diff
changeset | 277 | val th3 = th2 |> Conv.fconv_rule ObjectLogic.atomize |> Meson.make_nnf |> strip_lambdas ~1 | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24854diff
changeset | 278 | in (th3, ctxt) end; | 
| 16009 | 279 | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 280 | (*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 | 281 | 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 | 282 | 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 | 283 | |
| 24669 | 284 | fun assert_lambda_free ths msg = | 
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 285 | case filter (not o lambda_free o prop_of) ths of | 
| 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 286 | [] => () | 
| 26928 | 287 | | ths' => error (msg ^ "\n" ^ cat_lines (map Display.string_of_thm ths')); | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20445diff
changeset | 288 | |
| 25007 | 289 | |
| 27184 | 290 | (*** Blacklisting (duplicated in ResAtp?) ***) | 
| 25007 | 291 | |
| 292 | val max_lambda_nesting = 3; | |
| 27184 | 293 | |
| 25007 | 294 | fun excessive_lambdas (f$t, k) = excessive_lambdas (f,k) orelse excessive_lambdas (t,k) | 
| 295 | | excessive_lambdas (Abs(_,_,t), k) = k=0 orelse excessive_lambdas (t,k-1) | |
| 296 | | excessive_lambdas _ = false; | |
| 297 | ||
| 298 | fun is_formula_type T = (T = HOLogic.boolT orelse T = propT); | |
| 299 | ||
| 300 | (*Don't count nested lambdas at the level of formulas, as they are quantifiers*) | |
| 301 | fun excessive_lambdas_fm Ts (Abs(_,T,t)) = excessive_lambdas_fm (T::Ts) t | |
| 302 | | excessive_lambdas_fm Ts t = | |
| 303 | if is_formula_type (fastype_of1 (Ts, t)) | |
| 304 | then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t)) | |
| 305 | else excessive_lambdas (t, max_lambda_nesting); | |
| 306 | ||
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 307 | (*The max apply_depth of any metis call in MetisExamples (on 31-10-2007) was 11.*) | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 308 | val max_apply_depth = 15; | 
| 27184 | 309 | |
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 310 | 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 | 311 | | apply_depth (Abs(_,_,t)) = apply_depth t | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 312 | | apply_depth _ = 0; | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 313 | |
| 27184 | 314 | fun too_complex t = | 
| 315 | 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 | 316 | Meson.too_many_clauses NONE t orelse | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 317 | excessive_lambdas_fm [] t; | 
| 27184 | 318 | |
| 25243 | 319 | fun is_strange_thm th = | 
| 320 | case head_of (concl_of th) of | |
| 321 | Const (a,_) => (a <> "Trueprop" andalso a <> "==") | |
| 322 | | _ => false; | |
| 323 | ||
| 27184 | 324 | fun bad_for_atp th = | 
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
27809diff
changeset | 325 | Thm.is_internal th | 
| 27184 | 326 | orelse too_complex (prop_of th) | 
| 327 | orelse exists_type type_has_empty_sort (prop_of th) | |
| 25761 | 328 | orelse is_strange_thm th; | 
| 25243 | 329 | |
| 25007 | 330 | val multi_base_blacklist = | 
| 25256 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 331 | ["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm", | 
| 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 paulson parents: 
25243diff
changeset | 332 | "cases","ext_cases"]; (*FIXME: put other record thms here, or use the "Internal" marker*) | 
| 25007 | 333 | |
| 21071 | 334 | (*Keep the full complexity of the original name*) | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 335 | fun flatten_name s = space_implode "_X" (NameSpace.explode s); | 
| 21071 | 336 | |
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 337 | fun fake_name th = | 
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
27809diff
changeset | 338 | 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 | 339 | else gensym "unknown_thm_"; | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 340 | |
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 341 | fun name_or_string th = | 
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
27809diff
changeset | 342 | if Thm.has_name_hint th then Thm.get_name_hint th | 
| 26928 | 343 | else Display.string_of_thm th; | 
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 344 | |
| 27184 | 345 | (*Skolemize a named theorem, with Skolem functions as additional premises.*) | 
| 346 | fun skolem_thm (s, th) = | |
| 347 | if member (op =) multi_base_blacklist (Sign.base_name s) orelse bad_for_atp th then [] | |
| 348 | else | |
| 349 | let | |
| 350 | val ctxt0 = Variable.thm_context th | |
| 351 | val (nnfth, ctxt1) = to_nnf th ctxt0 | |
| 352 | val (cnfs, ctxt2) = Meson.make_cnf (assume_skolem_of_def s nnfth) nnfth ctxt1 | |
| 353 | in cnfs |> map combinators |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf end | |
| 354 | handle THM _ => []; | |
| 355 | ||
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 356 | (*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 | 357 | Skolem functions.*) | 
| 22516 | 358 | structure ThmCache = TheoryDataFun | 
| 22846 | 359 | ( | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 360 | type T = thm list Thmtab.table * unit Symtab.table; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 361 | val empty = (Thmtab.empty, Symtab.empty); | 
| 26618 | 362 | val copy = I; | 
| 363 | val extend = I; | |
| 27184 | 364 | fun merge _ ((cache1, seen1), (cache2, seen2)) : T = | 
| 365 | (Thmtab.merge (K true) (cache1, cache2), Symtab.merge (K true) (seen1, seen2)); | |
| 22846 | 366 | ); | 
| 22516 | 367 | |
| 27184 | 368 | val lookup_cache = Thmtab.lookup o #1 o ThmCache.get; | 
| 369 | 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 | 370 | |
| 27184 | 371 | val update_cache = ThmCache.map o apfst o Thmtab.update; | 
| 372 | fun mark_seen name = ThmCache.map (apsnd (Symtab.update (name, ()))); | |
| 25007 | 373 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 374 | (*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 | 375 | fun cnf_axiom thy th0 = | 
| 27184 | 376 | let val th = Thm.transfer thy th0 in | 
| 377 | case lookup_cache thy th of | |
| 378 | NONE => map Thm.close_derivation (skolem_thm (fake_name th, th)) | |
| 379 | | SOME cls => cls | |
| 22516 | 380 | end; | 
| 15347 | 381 | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 382 | |
| 15872 | 383 | (**** Extract and Clausify theorems from a theory's claset and simpset ****) | 
| 15347 | 384 | |
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
27809diff
changeset | 385 | fun pairname th = (Thm.get_name_hint th, th); | 
| 27184 | 386 | |
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17412diff
changeset | 387 | fun rules_of_claset cs = | 
| 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17412diff
changeset | 388 |   let val {safeIs,safeEs,hazIs,hazEs,...} = rep_cs cs
 | 
| 19175 | 389 | val intros = safeIs @ hazIs | 
| 18532 | 390 | val elims = map Classical.classical_rule (safeEs @ hazEs) | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 391 | in map pairname (intros @ elims) end; | 
| 15347 | 392 | |
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17412diff
changeset | 393 | fun rules_of_simpset ss = | 
| 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17412diff
changeset | 394 |   let val ({rules,...}, _) = rep_ss ss
 | 
| 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17412diff
changeset | 395 | val simps = Net.entries rules | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 396 | in map (fn r => (#name r, #thm r)) simps end; | 
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17412diff
changeset | 397 | |
| 21505 | 398 | fun claset_rules_of ctxt = rules_of_claset (local_claset_of ctxt); | 
| 399 | fun simpset_rules_of ctxt = rules_of_simpset (local_simpset_of ctxt); | |
| 19196 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 mengj parents: 
19175diff
changeset | 400 | |
| 24042 | 401 | fun atpset_rules_of ctxt = map pairname (ResAtpset.get ctxt); | 
| 20774 | 402 | |
| 15347 | 403 | |
| 22471 | 404 | (**** Translate a set of theorems into CNF ****) | 
| 15347 | 405 | |
| 19894 | 406 | fun pair_name_cls k (n, []) = [] | 
| 407 | | 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 | 408 | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 409 | fun cnf_rules_pairs_aux _ pairs [] = pairs | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 410 | | cnf_rules_pairs_aux thy pairs ((name,th)::ths) = | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 411 | let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 412 | handle THM _ => pairs | ResClause.CLAUSE _ => pairs | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 413 | 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 | 414 | |
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21254diff
changeset | 415 | (*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 | 416 | fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l); | 
| 19353 | 417 | |
| 19196 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 mengj parents: 
19175diff
changeset | 418 | |
| 27184 | 419 | (**** Convert all facts of the theory into clauses (ResClause.clause, or ResHolClause.clause) ****) | 
| 15347 | 420 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 421 | local | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 422 | |
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 423 | fun skolem_def (name, th) thy = | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 424 | let val ctxt0 = Variable.thm_context th in | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 425 | (case try (to_nnf th) ctxt0 of | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 426 | NONE => (NONE, thy) | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 427 | | SOME (nnfth, ctxt1) => | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 428 | 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 | 429 | in (SOME (th, ctxt0, ctxt1, nnfth, defs), thy') end) | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 430 | end; | 
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 431 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 432 | fun skolem_cnfs (th, ctxt0, ctxt1, nnfth, defs) = | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 433 | let | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 434 | 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 | 435 | val cnfs' = cnfs | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 436 | |> map combinators | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 437 | |> Variable.export ctxt2 ctxt0 | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 438 | |> Meson.finish_cnf | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 439 | |> map Thm.close_derivation; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 440 | in (th, cnfs') end; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 441 | |
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 442 | in | 
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 443 | |
| 27184 | 444 | fun saturate_skolem_cache thy = | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 445 | let | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 446 | val new_facts = (PureThy.facts_of thy, []) |-> Facts.fold_static (fn (name, ths) => | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 447 | if already_seen thy name then I else cons (name, ths)); | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 448 | val new_thms = (new_facts, []) |-> fold (fn (name, ths) => | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 449 | if member (op =) multi_base_blacklist (Sign.base_name name) then I | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 450 | else fold_index (fn (i, th) => | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 451 | 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 | 452 | 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 | 453 | in | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 454 | if null new_facts then NONE | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 455 | else | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 456 | let | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 457 | val (defs, thy') = thy | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 458 | |> fold (mark_seen o #1) new_facts | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 459 | |> 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 | 460 | |>> map_filter I; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 461 | val cache_entries = ParList.map skolem_cnfs defs; | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 462 | in SOME (fold update_cache cache_entries thy') end | 
| 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 463 | end; | 
| 27184 | 464 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 465 | end; | 
| 24854 | 466 | |
| 27184 | 467 | val suppress_endtheory = ref false; | 
| 468 | ||
| 469 | fun clause_cache_endtheory thy = | |
| 470 | if ! suppress_endtheory then NONE | |
| 471 | else saturate_skolem_cache thy; | |
| 472 | ||
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20445diff
changeset | 473 | |
| 22516 | 474 | (*The cache can be kept smaller by inspecting the prop of each thm. Can ignore all that are | 
| 475 | lambda_free, but then the individual theory caches become much bigger.*) | |
| 21071 | 476 | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 477 | |
| 16563 | 478 | (*** meson proof methods ***) | 
| 479 | ||
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 480 | (*Expand all new definitions of abstraction or Skolem functions in a proof state.*) | 
| 24827 | 481 | 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 | 482 | | is_absko _ = false; | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 483 | |
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 484 | 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 | 485 | 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 | 486 | | is_okdef _ _ = false | 
| 22724 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 paulson parents: 
22691diff
changeset | 487 | |
| 24215 | 488 | (*This function tries to cope with open locales, which introduce hypotheses of the form | 
| 489 | Free == t, conjecture clauses, which introduce various hypotheses, and also definitions | |
| 24827 | 490 | of sko_ functions. *) | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 491 | fun expand_defs_tac st0 st = | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 492 | let val hyps0 = #hyps (rep_thm st0) | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 493 | val hyps = #hyps (crep_thm st) | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 494 | 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 | 495 | val defs = filter (is_absko o Thm.term_of) newhyps | 
| 24669 | 496 | 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 | 497 | (map Thm.term_of hyps) | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 498 | val fixed = term_frees (concl_of st) @ | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 499 | foldl (gen_union (op aconv)) [] (map term_frees remaining_hyps) | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 500 | 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 | 501 | |
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 502 | |
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 503 | fun meson_general_tac ths i st0 = | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 504 | let | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 505 | val thy = Thm.theory_of_thm st0 | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 506 | in (Meson.meson_claset_tac (maps (cnf_axiom thy) ths) HOL_cs i THEN expand_defs_tac st0) st0 end; | 
| 22724 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 paulson parents: 
22691diff
changeset | 507 | |
| 21588 | 508 | val meson_method_setup = Method.add_methods | 
| 509 |   [("meson", Method.thms_args (fn ths =>
 | |
| 22724 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 paulson parents: 
22691diff
changeset | 510 | Method.SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ths)), | 
| 21588 | 511 | "MESON resolution proof procedure")]; | 
| 15347 | 512 | |
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 513 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 514 | (*** Converting a subgoal into negated conjecture clauses. ***) | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 515 | |
| 24300 | 516 | val neg_skolemize_tac = EVERY' [rtac ccontr, ObjectLogic.atomize_prems_tac, Meson.skolemize_tac]; | 
| 22471 | 517 | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24854diff
changeset | 518 | fun neg_clausify sts = | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24854diff
changeset | 519 | sts |> 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 | 520 | |
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 521 | fun neg_conjecture_clauses st0 n = | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 522 | let val st = Seq.hd (neg_skolemize_tac n st0) | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 523 | val (params,_,_) = strip_context (Logic.nth_prem (n, Thm.prop_of st)) | 
| 27187 | 524 | in (neg_clausify (the (metahyps_thms n st)), params) end | 
| 525 | handle Option => error "unable to Skolemize subgoal"; | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 526 | |
| 24669 | 527 | (*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 | 528 | leading !!-bound universal variables, to express generality. *) | 
| 24669 | 529 | val neg_clausify_tac = | 
| 530 | neg_skolemize_tac THEN' | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 531 | SUBGOAL | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 532 | (fn (prop,_) => | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 533 | let val ts = Logic.strip_assums_hyp prop | 
| 24669 | 534 | in EVERY1 | 
| 535 | [METAHYPS | |
| 536 | (fn hyps => | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 537 | (Method.insert_tac | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 538 | (map forall_intr_vars (neg_clausify hyps)) 1)), | 
| 24669 | 539 | REPEAT_DETERM_N (length ts) o (etac thin_rl)] | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 540 | end); | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 541 | |
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 542 | val setup_methods = Method.add_methods | 
| 24669 | 543 |   [("neg_clausify", Method.no_args (Method.SIMPLE_METHOD' neg_clausify_tac),
 | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21900diff
changeset | 544 | "conversion of goal to conjecture clauses")]; | 
| 24669 | 545 | |
| 27184 | 546 | |
| 547 | (** Attribute for converting a theorem into clauses **) | |
| 548 | ||
| 27809 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 wenzelm parents: 
27330diff
changeset | 549 | val clausify = Attrib.syntax (Scan.lift OuterParse.nat | 
| 27184 | 550 | >> (fn i => Thm.rule_attribute (fn context => fn th => | 
| 551 | Meson.make_meta_clause (nth (cnf_axiom (Context.theory_of context) th) i)))); | |
| 552 | ||
| 553 | val setup_attrs = Attrib.add_attributes | |
| 554 |   [("clausify", clausify, "conversion of theorem to clauses")];
 | |
| 555 | ||
| 556 | ||
| 557 | ||
| 558 | (** setup **) | |
| 559 | ||
| 560 | val setup = | |
| 561 | meson_method_setup #> | |
| 562 | setup_methods #> | |
| 563 | setup_attrs #> | |
| 564 | perhaps saturate_skolem_cache #> | |
| 565 | Theory.at_end clause_cache_endtheory; | |
| 18510 
0a6c24f549c3
the "skolem" attribute and better initialization of the clause database
 paulson parents: 
18404diff
changeset | 566 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 567 | end; |