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