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