| author | haftmann |
| Tue, 11 May 2010 08:29:42 +0200 | |
| changeset 36808 | cbeb3484fa07 |
| parent 36603 | d5d6111761a6 |
| child 36945 | 9bec62c10714 |
| permissions | -rw-r--r-- |
| 35826 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML |
| 33311 | 2 |
Author: Jia Meng, Cambridge University Computer Laboratory |
|
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36228
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
| 15347 | 4 |
|
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
5 |
Transformation of axiom rules (elim/intro/etc) into CNF forms. |
| 15347 | 6 |
*) |
7 |
||
| 35826 | 8 |
signature SLEDGEHAMMER_FACT_PREPROCESSOR = |
| 21505 | 9 |
sig |
| 32955 | 10 |
val trace: bool Unsynchronized.ref |
11 |
val trace_msg: (unit -> string) -> unit |
|
| 35865 | 12 |
val skolem_prefix: string |
|
36492
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
13 |
val skolem_infix: string |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
14 |
val cnf_axiom: theory -> thm -> thm list |
| 27184 | 15 |
val multi_base_blacklist: string list |
| 25243 | 16 |
val bad_for_atp: thm -> bool |
|
35568
8fbbfc39508f
renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents:
33832
diff
changeset
|
17 |
val type_has_topsort: typ -> bool |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
18 |
val cnf_rules_pairs: theory -> (string * thm) list -> (thm * (string * int)) list |
| 32740 | 19 |
val suppress_endtheory: bool Unsynchronized.ref |
20 |
(*for emergency use where endtheory causes problems*) |
|
|
36478
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36398
diff
changeset
|
21 |
val strip_subgoal : thm -> int -> (string * typ) list * term list * term |
|
36398
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
22 |
val neg_clausify: thm -> thm list |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
23 |
val neg_conjecture_clauses: |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
24 |
Proof.context -> thm -> int -> thm list list * (string * typ) list |
|
36394
1a48d18449d8
move "neg_clausify" method and "clausify" attribute to "sledgehammer_isar.ML"
blanchet
parents:
36393
diff
changeset
|
25 |
val neg_clausify_tac: Proof.context -> int -> tactic |
| 24669 | 26 |
val setup: theory -> theory |
| 21505 | 27 |
end; |
|
19196
62ee8c10d796
Added functions to retrieve local and global atpset rules.
mengj
parents:
19175
diff
changeset
|
28 |
|
| 35826 | 29 |
structure Sledgehammer_Fact_Preprocessor : SLEDGEHAMMER_FACT_PREPROCESSOR = |
| 15997 | 30 |
struct |
| 15347 | 31 |
|
| 35865 | 32 |
open Sledgehammer_FOL_Clause |
33 |
||
| 32955 | 34 |
val trace = Unsynchronized.ref false; |
| 35865 | 35 |
fun trace_msg msg = if !trace then tracing (msg ()) else (); |
36 |
||
37 |
val skolem_prefix = "sko_" |
|
|
36492
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
38 |
val skolem_infix = "$" |
| 32955 | 39 |
|
| 33832 | 40 |
fun freeze_thm th = #1 (Drule.legacy_freeze_thaw th); |
|
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
41 |
|
|
35568
8fbbfc39508f
renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents:
33832
diff
changeset
|
42 |
val type_has_topsort = Term.exists_subtype |
|
8fbbfc39508f
renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents:
33832
diff
changeset
|
43 |
(fn TFree (_, []) => true |
|
8fbbfc39508f
renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents:
33832
diff
changeset
|
44 |
| TVar (_, []) => true |
|
8fbbfc39508f
renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents:
33832
diff
changeset
|
45 |
| _ => false); |
| 27184 | 46 |
|
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
47 |
|
| 15997 | 48 |
(**** Transformation of Elimination Rules into First-Order Formulas****) |
| 15347 | 49 |
|
| 29064 | 50 |
val cfalse = cterm_of @{theory HOL} HOLogic.false_const;
|
51 |
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:
20457
diff
changeset
|
52 |
|
|
21430
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents:
21290
diff
changeset
|
53 |
(*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:
21290
diff
changeset
|
54 |
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:
21290
diff
changeset
|
55 |
conclusion variable to False.*) |
| 16009 | 56 |
fun transform_elim th = |
|
21430
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents:
21290
diff
changeset
|
57 |
case concl_of th of (*conclusion variable*) |
| 35963 | 58 |
@{const Trueprop} $ (v as Var (_, @{typ bool})) =>
|
| 29064 | 59 |
Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th
|
| 35963 | 60 |
| v as Var(_, @{typ prop}) =>
|
| 29064 | 61 |
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:
21290
diff
changeset
|
62 |
| _ => th; |
| 15997 | 63 |
|
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
64 |
(*To enforce single-threading*) |
|
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
65 |
exception Clausify_failure of theory; |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
66 |
|
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
67 |
|
| 16009 | 68 |
(**** SKOLEMIZATION BY INFERENCE (lcp) ****) |
69 |
||
|
36492
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
70 |
(*Keep the full complexity of the original name*) |
|
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
71 |
fun flatten_name s = space_implode "_X" (Long_Name.explode s); |
|
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
72 |
|
|
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
73 |
fun skolem_name thm_name nref var_name = |
|
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
74 |
skolem_prefix ^ thm_name ^ "_" ^ Int.toString (Unsynchronized.inc nref) ^ |
|
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
75 |
skolem_infix ^ (if var_name = "" then "g" else flatten_name var_name) |
|
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
76 |
|
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
77 |
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:
24669
diff
changeset
|
78 |
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:
24669
diff
changeset
|
79 |
fun add_new_TFrees (TFree v) = |
|
24821
cc55041ca8ba
skolem_cache: ignore internal theorems -- major speedup;
wenzelm
parents:
24785
diff
changeset
|
80 |
if member (op =) lhs_vars v then I else insert (op =) (TFree v) |
|
cc55041ca8ba
skolem_cache: ignore internal theorems -- major speedup;
wenzelm
parents:
24785
diff
changeset
|
81 |
| add_new_TFrees _ = I |
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
82 |
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:
24669
diff
changeset
|
83 |
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:
24669
diff
changeset
|
84 |
|
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
85 |
(*Traverse a theorem, declaring Skolem function definitions. String s is the suggested |
|
27174
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
86 |
prefix for the Skolem constant.*) |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
87 |
fun declare_skofuns s th = |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
88 |
let |
| 33222 | 89 |
val nref = Unsynchronized.ref 0 (* FIXME ??? *) |
|
36492
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
90 |
fun dec_sko (Const (@{const_name Ex}, _) $ (xtp as Abs (s', T, p))) (axs, thy) =
|
|
27174
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
91 |
(*Existential: declare a Skolem function, then insert into body and continue*) |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
92 |
let |
|
36492
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
93 |
val cname = skolem_name s nref s' |
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
29064
diff
changeset
|
94 |
val args0 = OldTerm.term_frees xtp (*get the formal parameter list*) |
|
27174
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
95 |
val Ts = map type_of args0 |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
96 |
val extraTs = rhs_extra_types (Ts ---> T) xtp |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
97 |
val argsx = map (fn T => Free (gensym "vsk", T)) extraTs |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
98 |
val args = argsx @ args0 |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
99 |
val cT = extraTs ---> Ts ---> T |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
100 |
val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp) |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
101 |
(*Forms a lambda-abstraction over the formal parameters*) |
| 28110 | 102 |
val (c, thy') = |
|
33173
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents:
33158
diff
changeset
|
103 |
Sign.declare_const ((Binding.conceal (Binding.name cname), cT), NoSyn) thy |
|
27174
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
104 |
val cdef = cname ^ "_def" |
|
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35984
diff
changeset
|
105 |
val ((_, ax), thy'') = |
| 35984 | 106 |
Thm.add_def true false (Binding.name cdef, Logic.mk_equals (c, rhs)) thy' |
107 |
val ax' = Drule.export_without_context ax |
|
108 |
in dec_sko (subst_bound (list_comb (c, args), p)) (ax' :: axs, thy'') end |
|
| 35963 | 109 |
| dec_sko (Const (@{const_name All}, _) $ (Abs (a, T, p))) thx =
|
|
27174
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
110 |
(*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:
29265
diff
changeset
|
111 |
let val fname = Name.variant (OldTerm.add_term_names (p, [])) a |
|
27174
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
112 |
in dec_sko (subst_bound (Free (fname, T), p)) thx end |
| 35963 | 113 |
| dec_sko (@{const "op &"} $ p $ q) thx = dec_sko q (dec_sko p thx)
|
114 |
| dec_sko (@{const "op |"} $ p $ q) thx = dec_sko q (dec_sko p thx)
|
|
115 |
| dec_sko (@{const Trueprop} $ p) thx = dec_sko p thx
|
|
|
27174
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
116 |
| dec_sko t thx = thx (*Do nothing otherwise*) |
|
c2c484480f40
declare_skofuns/skolem: canonical argument order;
wenzelm
parents:
26939
diff
changeset
|
117 |
in fn thy => dec_sko (Thm.prop_of th) ([], thy) end; |
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
118 |
|
|
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
119 |
(*Traverse a theorem, accumulating Skolem function definitions.*) |
|
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22724
diff
changeset
|
120 |
fun assume_skofuns s th = |
| 33222 | 121 |
let val sko_count = Unsynchronized.ref 0 (* FIXME ??? *) |
|
36492
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
122 |
fun dec_sko (Const (@{const_name Ex}, _) $ (xtp as Abs (s', T, p))) defs =
|
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
123 |
(*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:
20457
diff
changeset
|
124 |
let val skos = map (#1 o Logic.dest_equals) defs (*existing sko fns*) |
| 33040 | 125 |
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:
20457
diff
changeset
|
126 |
val Ts = map type_of args |
|
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
127 |
val cT = Ts ---> T |
|
36492
60532b9bcd1c
save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents:
36478
diff
changeset
|
128 |
val id = skolem_name s sko_count s' |
|
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22724
diff
changeset
|
129 |
val c = Free (id, cT) |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
130 |
val rhs = list_abs_free (map dest_Free args, |
|
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
131 |
HOLogic.choice_const T $ xtp) |
|
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
132 |
(*Forms a lambda-abstraction over the formal parameters*) |
| 27330 | 133 |
val def = Logic.mk_equals (c, rhs) |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
134 |
in dec_sko (subst_bound (list_comb(c,args), p)) |
|
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
135 |
(def :: defs) |
|
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
136 |
end |
| 35963 | 137 |
| dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) defs =
|
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
138 |
(*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:
29265
diff
changeset
|
139 |
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:
20457
diff
changeset
|
140 |
in dec_sko (subst_bound (Free(fname,T), p)) defs end |
| 35963 | 141 |
| dec_sko (@{const "op &"} $ p $ q) defs = dec_sko q (dec_sko p defs)
|
142 |
| dec_sko (@{const "op |"} $ p $ q) defs = dec_sko q (dec_sko p defs)
|
|
143 |
| dec_sko (@{const Trueprop} $ p) defs = dec_sko p defs
|
|
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
144 |
| dec_sko t defs = defs (*Do nothing otherwise*) |
| 20419 | 145 |
in dec_sko (prop_of th) [] end; |
146 |
||
147 |
||
| 24827 | 148 |
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****) |
| 20419 | 149 |
|
150 |
(*Returns the vars of a theorem*) |
|
151 |
fun vars_of_thm th = |
|
| 22691 | 152 |
map (Thm.cterm_of (theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th []); |
| 20419 | 153 |
|
154 |
(*Make a version of fun_cong with a given variable name*) |
|
155 |
local |
|
156 |
val fun_cong' = fun_cong RS asm_rl; (*renumber f, g to prevent clashes with (a,0)*) |
|
157 |
val cx = hd (vars_of_thm fun_cong'); |
|
158 |
val ty = typ_of (ctyp_of_term cx); |
|
| 20445 | 159 |
val thy = theory_of_thm fun_cong; |
| 20419 | 160 |
fun mkvar a = cterm_of thy (Var((a,0),ty)); |
161 |
in |
|
162 |
fun xfun_cong x = Thm.instantiate ([], [(cx, mkvar x)]) fun_cong' |
|
163 |
end; |
|
164 |
||
|
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
165 |
(*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:
20789
diff
changeset
|
166 |
serves as an upper bound on how many to remove.*) |
|
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
167 |
fun strip_lambdas 0 th = th |
| 24669 | 168 |
| strip_lambdas n th = |
|
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
169 |
case prop_of th of |
| 35963 | 170 |
_ $ (Const (@{const_name "op ="}, _) $ _ $ Abs (x, _, _)) =>
|
| 24669 | 171 |
strip_lambdas (n-1) (freeze_thm (th RS xfun_cong x)) |
172 |
| _ => th; |
|
| 20419 | 173 |
|
| 24669 | 174 |
val lambda_free = not o Term.has_abs; |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
175 |
|
| 32010 | 176 |
val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B}));
|
177 |
val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C}));
|
|
178 |
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:
20789
diff
changeset
|
179 |
|
| 24827 | 180 |
(*FIXME: requires more use of cterm constructors*) |
181 |
fun abstract ct = |
|
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
182 |
let |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
183 |
val thy = theory_of_cterm ct |
|
25256
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
184 |
val Abs(x,_,body) = term_of ct |
| 35963 | 185 |
val Type(@{type_name fun}, [xT,bodyT]) = typ_of (ctyp_of_term ct)
|
| 24827 | 186 |
val cxT = ctyp_of thy xT and cbodyT = ctyp_of thy bodyT |
| 27184 | 187 |
fun makeK() = instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] @{thm abs_K}
|
| 24827 | 188 |
in |
189 |
case body of |
|
190 |
Const _ => makeK() |
|
191 |
| Free _ => makeK() |
|
192 |
| Var _ => makeK() (*though Var isn't expected*) |
|
| 27184 | 193 |
| Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
|
| 24827 | 194 |
| rator$rand => |
| 27184 | 195 |
if loose_bvar1 (rator,0) then (*C or S*) |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
196 |
if loose_bvar1 (rand,0) then (*S*) |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
197 |
let val crator = cterm_of thy (Abs(x,xT,rator)) |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
198 |
val crand = cterm_of thy (Abs(x,xT,rand)) |
| 27184 | 199 |
val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
|
200 |
val (_,rhs) = Thm.dest_equals (cprop_of abs_S') |
|
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
201 |
in |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
202 |
Thm.transitive abs_S' (Conv.binop_conv abstract rhs) |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
203 |
end |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
204 |
else (*C*) |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
205 |
let val crator = cterm_of thy (Abs(x,xT,rator)) |
| 27184 | 206 |
val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
|
207 |
val (_,rhs) = Thm.dest_equals (cprop_of abs_C') |
|
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
208 |
in |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
209 |
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:
27174
diff
changeset
|
210 |
end |
| 27184 | 211 |
else if loose_bvar1 (rand,0) then (*B or eta*) |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
212 |
if rand = Bound 0 then eta_conversion ct |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
213 |
else (*B*) |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
214 |
let val crand = cterm_of thy (Abs(x,xT,rand)) |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
215 |
val crator = cterm_of thy rator |
| 27184 | 216 |
val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
|
217 |
val (_,rhs) = Thm.dest_equals (cprop_of abs_B') |
|
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
218 |
in |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
219 |
Thm.transitive abs_B' (Conv.arg_conv abstract rhs) |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
220 |
end |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
221 |
else makeK() |
| 24827 | 222 |
| _ => error "abstract: Bad term" |
223 |
end; |
|
|
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
224 |
|
| 20419 | 225 |
(*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:
28262
diff
changeset
|
226 |
prefix for the constants.*) |
| 24827 | 227 |
fun combinators_aux ct = |
228 |
if lambda_free (term_of ct) then reflexive ct |
|
229 |
else |
|
230 |
case term_of ct of |
|
231 |
Abs _ => |
|
| 32994 | 232 |
let val (cv, cta) = Thm.dest_abs NONE ct |
233 |
val (v, _) = dest_Free (term_of cv) |
|
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
234 |
val u_th = combinators_aux cta |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
235 |
val cu = Thm.rhs_of u_th |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
236 |
val comb_eq = abstract (Thm.cabs cv cu) |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
237 |
in transitive (abstract_rule v cv u_th) comb_eq end |
| 32994 | 238 |
| _ $ _ => |
239 |
let val (ct1, ct2) = Thm.dest_comb ct |
|
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
240 |
in combination (combinators_aux ct1) (combinators_aux ct2) end; |
| 27184 | 241 |
|
| 24827 | 242 |
fun combinators th = |
| 27184 | 243 |
if lambda_free (prop_of th) then th |
| 24827 | 244 |
else |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
245 |
let val th = Drule.eta_contraction_rule th |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
246 |
val eqth = combinators_aux (cprop_of th) |
|
25256
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
247 |
in equal_elim eqth th end |
| 27184 | 248 |
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:
32010
diff
changeset
|
249 |
(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:
32010
diff
changeset
|
250 |
["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:
32010
diff
changeset
|
251 |
" Exception message: " ^ msg]); |
|
25256
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
252 |
TrueI); (*A type variable of sort {} will cause make abstraction fail.*)
|
| 16009 | 253 |
|
254 |
(*cterms are used throughout for efficiency*) |
|
| 29064 | 255 |
val cTrueprop = Thm.cterm_of @{theory HOL} HOLogic.Trueprop;
|
| 16009 | 256 |
|
257 |
(*cterm version of mk_cTrueprop*) |
|
258 |
fun c_mkTrueprop A = Thm.capply cTrueprop A; |
|
259 |
||
260 |
(*Given an abstraction over n variables, replace the bound variables by free |
|
261 |
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:
20457
diff
changeset
|
262 |
fun c_variant_abs_multi (ct0, vars) = |
| 16009 | 263 |
let val (cv,ct) = Thm.dest_abs NONE ct0 |
264 |
in c_variant_abs_multi (ct, cv::vars) end |
|
265 |
handle CTERM _ => (ct0, rev vars); |
|
266 |
||
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
267 |
(*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:
20457
diff
changeset
|
268 |
an existential formula by a use of that function. |
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
269 |
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:
20457
diff
changeset
|
270 |
fun skolem_of_def def = |
|
22902
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents:
22846
diff
changeset
|
271 |
let val (c,rhs) = Thm.dest_equals (cprop_of (freeze_thm def)) |
| 16009 | 272 |
val (ch, frees) = c_variant_abs_multi (rhs, []) |
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
273 |
val (chilbert,cabs) = Thm.dest_comb ch |
|
26627
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents:
26618
diff
changeset
|
274 |
val thy = Thm.theory_of_cterm chilbert |
|
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents:
26618
diff
changeset
|
275 |
val t = Thm.term_of chilbert |
| 35963 | 276 |
val T = case t of |
277 |
Const (@{const_name Eps}, Type (@{type_name fun},[_,T])) => T
|
|
278 |
| _ => raise THM ("skolem_of_def: expected Eps", 0, [def])
|
|
| 22596 | 279 |
val cex = Thm.cterm_of thy (HOLogic.exists_const T) |
| 16009 | 280 |
val ex_tm = c_mkTrueprop (Thm.capply cex cabs) |
281 |
and conc = c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees))); |
|
| 31454 | 282 |
fun tacf [prem] = rewrite_goals_tac [def] THEN rtac (prem RS @{thm someI_ex}) 1
|
| 23352 | 283 |
in Goal.prove_internal [ex_tm] conc tacf |
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
284 |
|> forall_intr_list frees |
| 26653 | 285 |
|> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) |
|
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35828
diff
changeset
|
286 |
|> Thm.varifyT_global |
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
287 |
end; |
| 16009 | 288 |
|
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
289 |
|
|
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
290 |
(*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:
24854
diff
changeset
|
291 |
fun to_nnf th ctxt0 = |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
292 |
let val th1 = th |> transform_elim |> zero_var_indexes |
| 32262 | 293 |
val ((_, [th2]), ctxt) = Variable.import true [th1] ctxt0 |
294 |
val th3 = th2 |
|
| 35625 | 295 |
|> Conv.fconv_rule Object_Logic.atomize |
| 32262 | 296 |
|> Meson.make_nnf ctxt |> strip_lambdas ~1 |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24854
diff
changeset
|
297 |
in (th3, ctxt) end; |
| 16009 | 298 |
|
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
299 |
(*Generate Skolem functions for a theorem supplied in nnf*) |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24854
diff
changeset
|
300 |
fun assume_skolem_of_def s th = |
|
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22724
diff
changeset
|
301 |
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:
18009
diff
changeset
|
302 |
|
| 25007 | 303 |
|
| 35963 | 304 |
(*** Blacklisting (FIXME: duplicated in "Sledgehammer_Fact_Filter"?) ***) |
| 25007 | 305 |
|
306 |
val max_lambda_nesting = 3; |
|
| 27184 | 307 |
|
| 25007 | 308 |
fun excessive_lambdas (f$t, k) = excessive_lambdas (f,k) orelse excessive_lambdas (t,k) |
309 |
| excessive_lambdas (Abs(_,_,t), k) = k=0 orelse excessive_lambdas (t,k-1) |
|
310 |
| excessive_lambdas _ = false; |
|
311 |
||
312 |
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT); |
|
313 |
||
314 |
(*Don't count nested lambdas at the level of formulas, as they are quantifiers*) |
|
315 |
fun excessive_lambdas_fm Ts (Abs(_,T,t)) = excessive_lambdas_fm (T::Ts) t |
|
316 |
| excessive_lambdas_fm Ts t = |
|
317 |
if is_formula_type (fastype_of1 (Ts, t)) |
|
318 |
then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t)) |
|
319 |
else excessive_lambdas (t, max_lambda_nesting); |
|
320 |
||
| 33027 | 321 |
(*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:
25243
diff
changeset
|
322 |
val max_apply_depth = 15; |
| 27184 | 323 |
|
|
25256
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
324 |
fun apply_depth (f$t) = Int.max (apply_depth f, apply_depth t + 1) |
|
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
325 |
| apply_depth (Abs(_,_,t)) = apply_depth t |
|
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
326 |
| apply_depth _ = 0; |
|
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
327 |
|
| 27184 | 328 |
fun too_complex t = |
329 |
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:
25761
diff
changeset
|
330 |
Meson.too_many_clauses NONE t orelse |
|
25256
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
331 |
excessive_lambdas_fm [] t; |
| 27184 | 332 |
|
| 25243 | 333 |
fun is_strange_thm th = |
334 |
case head_of (concl_of th) of |
|
| 35963 | 335 |
Const (a, _) => (a <> @{const_name Trueprop} andalso
|
336 |
a <> @{const_name "=="})
|
|
| 25243 | 337 |
| _ => false; |
338 |
||
| 27184 | 339 |
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:
33222
diff
changeset
|
340 |
too_complex (prop_of th) |
|
35568
8fbbfc39508f
renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents:
33832
diff
changeset
|
341 |
orelse exists_type type_has_topsort (prop_of th) |
| 25761 | 342 |
orelse is_strange_thm th; |
| 25243 | 343 |
|
| 35963 | 344 |
(* FIXME: put other record thms here, or declare as "no_atp" *) |
| 25007 | 345 |
val multi_base_blacklist = |
| 35963 | 346 |
["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits", |
347 |
"split_asm", "cases", "ext_cases"]; |
|
| 25007 | 348 |
|
|
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22724
diff
changeset
|
349 |
fun fake_name th = |
|
27865
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
wenzelm
parents:
27809
diff
changeset
|
350 |
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:
22724
diff
changeset
|
351 |
else gensym "unknown_thm_"; |
|
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22724
diff
changeset
|
352 |
|
| 27184 | 353 |
(*Skolemize a named theorem, with Skolem functions as additional premises.*) |
354 |
fun skolem_thm (s, th) = |
|
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30291
diff
changeset
|
355 |
if member (op =) multi_base_blacklist (Long_Name.base_name s) orelse bad_for_atp th then [] |
| 27184 | 356 |
else |
357 |
let |
|
|
36603
d5d6111761a6
renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
wenzelm
parents:
36492
diff
changeset
|
358 |
val ctxt0 = Variable.global_thm_context th |
| 27184 | 359 |
val (nnfth, ctxt1) = to_nnf th ctxt0 |
360 |
val (cnfs, ctxt2) = Meson.make_cnf (assume_skolem_of_def s nnfth) nnfth ctxt1 |
|
361 |
in cnfs |> map combinators |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf end |
|
362 |
handle THM _ => []; |
|
363 |
||
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
364 |
(*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:
24669
diff
changeset
|
365 |
Skolem functions.*) |
| 33522 | 366 |
structure ThmCache = Theory_Data |
| 22846 | 367 |
( |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
368 |
type T = thm list Thmtab.table * unit Symtab.table; |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
369 |
val empty = (Thmtab.empty, Symtab.empty); |
| 26618 | 370 |
val extend = I; |
| 33522 | 371 |
fun merge ((cache1, seen1), (cache2, seen2)) : T = |
| 27184 | 372 |
(Thmtab.merge (K true) (cache1, cache2), Symtab.merge (K true) (seen1, seen2)); |
| 22846 | 373 |
); |
| 22516 | 374 |
|
| 27184 | 375 |
val lookup_cache = Thmtab.lookup o #1 o ThmCache.get; |
376 |
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:
20457
diff
changeset
|
377 |
|
| 27184 | 378 |
val update_cache = ThmCache.map o apfst o Thmtab.update; |
379 |
fun mark_seen name = ThmCache.map (apsnd (Symtab.update (name, ()))); |
|
| 25007 | 380 |
|
|
36228
df47dc6c0e03
got rid of somewhat pointless "pairname" function in Sledgehammer
blanchet
parents:
36106
diff
changeset
|
381 |
(* Convert Isabelle theorems into axiom clauses. *) |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
382 |
fun cnf_axiom thy th0 = |
| 27184 | 383 |
let val th = Thm.transfer thy th0 in |
384 |
case lookup_cache thy th of |
|
385 |
NONE => map Thm.close_derivation (skolem_thm (fake_name th, th)) |
|
386 |
| SOME cls => cls |
|
| 22516 | 387 |
end; |
| 15347 | 388 |
|
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
389 |
|
| 22471 | 390 |
(**** Translate a set of theorems into CNF ****) |
| 15347 | 391 |
|
| 19894 | 392 |
fun pair_name_cls k (n, []) = [] |
393 |
| 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:
20457
diff
changeset
|
394 |
|
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
395 |
fun cnf_rules_pairs_aux _ pairs [] = pairs |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
396 |
| cnf_rules_pairs_aux thy pairs ((name,th)::ths) = |
|
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
397 |
let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs |
| 35826 | 398 |
handle THM _ => pairs | |
| 35865 | 399 |
CLAUSE _ => pairs |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
400 |
in cnf_rules_pairs_aux thy pairs' ths end; |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
401 |
|
|
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21254
diff
changeset
|
402 |
(*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:
27174
diff
changeset
|
403 |
fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l); |
| 19353 | 404 |
|
|
19196
62ee8c10d796
Added functions to retrieve local and global atpset rules.
mengj
parents:
19175
diff
changeset
|
405 |
|
| 35865 | 406 |
(**** Convert all facts of the theory into FOL or HOL clauses ****) |
| 15347 | 407 |
|
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
408 |
local |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
409 |
|
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
410 |
fun skolem_def (name, th) thy = |
|
36603
d5d6111761a6
renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
wenzelm
parents:
36492
diff
changeset
|
411 |
let val ctxt0 = Variable.global_thm_context th in |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
412 |
(case try (to_nnf th) ctxt0 of |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
413 |
NONE => (NONE, thy) |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
414 |
| SOME (nnfth, ctxt1) => |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
415 |
let val (defs, thy') = declare_skofuns (flatten_name name) nnfth thy |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
416 |
in (SOME (th, ctxt0, ctxt1, nnfth, defs), thy') end) |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
417 |
end; |
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
418 |
|
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
419 |
fun skolem_cnfs (th, ctxt0, ctxt1, nnfth, defs) = |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
420 |
let |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
421 |
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:
28262
diff
changeset
|
422 |
val cnfs' = cnfs |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
423 |
|> map combinators |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
424 |
|> Variable.export ctxt2 ctxt0 |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
425 |
|> Meson.finish_cnf |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
426 |
|> map Thm.close_derivation; |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
427 |
in (th, cnfs') end; |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
428 |
|
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
429 |
in |
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
430 |
|
| 27184 | 431 |
fun saturate_skolem_cache thy = |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
432 |
let |
|
33306
4138ba02b681
replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
wenzelm
parents:
33222
diff
changeset
|
433 |
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:
33222
diff
changeset
|
434 |
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:
33222
diff
changeset
|
435 |
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:
33222
diff
changeset
|
436 |
else cons (name, ths)); |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
437 |
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:
30291
diff
changeset
|
438 |
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:
28262
diff
changeset
|
439 |
else fold_index (fn (i, th) => |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
440 |
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:
28262
diff
changeset
|
441 |
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:
28262
diff
changeset
|
442 |
in |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
443 |
if null new_facts then NONE |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
444 |
else |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
445 |
let |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
446 |
val (defs, thy') = thy |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
447 |
|> fold (mark_seen o #1) new_facts |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
448 |
|> 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:
28262
diff
changeset
|
449 |
|>> map_filter I; |
| 29368 | 450 |
val cache_entries = Par_List.map skolem_cnfs defs; |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
451 |
in SOME (fold update_cache cache_entries thy') end |
|
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
452 |
end; |
| 27184 | 453 |
|
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
454 |
end; |
| 24854 | 455 |
|
| 32740 | 456 |
val suppress_endtheory = Unsynchronized.ref false; |
| 27184 | 457 |
|
458 |
fun clause_cache_endtheory thy = |
|
459 |
if ! suppress_endtheory then NONE |
|
460 |
else saturate_skolem_cache thy; |
|
461 |
||
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20445
diff
changeset
|
462 |
|
| 22516 | 463 |
(*The cache can be kept smaller by inspecting the prop of each thm. Can ignore all that are |
464 |
lambda_free, but then the individual theory caches become much bigger.*) |
|
| 21071 | 465 |
|
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
466 |
|
|
36398
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
467 |
fun strip_subgoal goal i = |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
468 |
let |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
469 |
val (t, frees) = Logic.goal_params (prop_of goal) i |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
470 |
val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees) |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
471 |
val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees |
|
36478
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
blanchet
parents:
36398
diff
changeset
|
472 |
in (rev (map dest_Free frees), hyp_ts, concl_t) end |
|
36398
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
473 |
|
|
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21900
diff
changeset
|
474 |
(*** Converting a subgoal into negated conjecture clauses. ***) |
|
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21900
diff
changeset
|
475 |
|
| 32262 | 476 |
fun neg_skolemize_tac ctxt = |
| 35625 | 477 |
EVERY' [rtac ccontr, Object_Logic.atomize_prems_tac, Meson.skolemize_tac ctxt]; |
| 22471 | 478 |
|
|
36398
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
479 |
fun neg_skolemize_tac ctxt = |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
480 |
EVERY' [rtac ccontr, Object_Logic.atomize_prems_tac, Meson.skolemize_tac ctxt]; |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
481 |
|
|
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35865
diff
changeset
|
482 |
val neg_clausify = |
|
36398
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
483 |
single #> Meson.make_clauses_unsorted #> map combinators #> Meson.finish_cnf |
|
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21900
diff
changeset
|
484 |
|
|
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
485 |
fun neg_conjecture_clauses ctxt st0 n = |
|
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
486 |
let |
| 32262 | 487 |
val st = Seq.hd (neg_skolemize_tac ctxt n st0) |
|
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
488 |
val ({params, prems, ...}, _) = Subgoal.focus (Variable.set_body false ctxt) n st
|
|
36398
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
489 |
in |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
490 |
(map neg_clausify prems, map (Term.dest_Free o Thm.term_of o #2) params) |
|
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
491 |
end |
|
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21900
diff
changeset
|
492 |
|
| 24669 | 493 |
(*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:
21900
diff
changeset
|
494 |
leading !!-bound universal variables, to express generality. *) |
|
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
495 |
fun neg_clausify_tac ctxt = |
| 32262 | 496 |
neg_skolemize_tac ctxt THEN' |
|
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
497 |
SUBGOAL (fn (prop, i) => |
|
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
498 |
let val ts = Logic.strip_assums_hyp prop in |
|
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
499 |
EVERY' |
| 32283 | 500 |
[Subgoal.FOCUS |
|
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
501 |
(fn {prems, ...} =>
|
|
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
502 |
(Method.insert_tac |
|
36398
de8522a5cbae
make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents:
36394
diff
changeset
|
503 |
(map forall_intr_vars (maps neg_clausify prems)) i)) ctxt, |
|
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32231
diff
changeset
|
504 |
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:
21900
diff
changeset
|
505 |
end); |
|
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21900
diff
changeset
|
506 |
|
| 27184 | 507 |
|
508 |
(** setup **) |
|
509 |
||
510 |
val setup = |
|
511 |
perhaps saturate_skolem_cache #> |
|
512 |
Theory.at_end clause_cache_endtheory; |
|
|
18510
0a6c24f549c3
the "skolem" attribute and better initialization of the clause database
paulson
parents:
18404
diff
changeset
|
513 |
|
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
514 |
end; |