| author | urbanc |
| Wed, 02 Nov 2005 15:31:12 +0100 | |
| changeset 18067 | 8b9848d150ba |
| parent 18009 | df077e122064 |
| child 18141 | 89e2e8bed08f |
| permissions | -rw-r--r-- |
| 15347 | 1 |
(* Author: Jia Meng, Cambridge University Computer Laboratory |
2 |
ID: $Id$ |
|
3 |
Copyright 2004 University of Cambridge |
|
4 |
||
5 |
Transformation of axiom rules (elim/intro/etc) into CNF forms. |
|
6 |
*) |
|
7 |
||
| 15997 | 8 |
signature RES_AXIOMS = |
9 |
sig |
|
10 |
exception ELIMR2FOL of string |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
11 |
val tagging_enabled : bool |
| 15997 | 12 |
val elimRule_tac : thm -> Tactical.tactic |
| 16012 | 13 |
val elimR2Fol : thm -> term |
| 15997 | 14 |
val transform_elim : thm -> thm |
|
16039
dfe264950511
Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents:
16012
diff
changeset
|
15 |
val clausify_axiom_pairs : (string*thm) -> (ResClause.clause*thm) list |
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
16 |
val clausify_axiom_pairsH : (string*thm) -> (ResHolClause.clause*thm) list |
| 15997 | 17 |
val cnf_axiom : (string * thm) -> thm list |
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
18 |
val cnf_axiomH : (string * thm) -> thm list |
| 15997 | 19 |
val meta_cnf_axiom : thm -> thm list |
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
20 |
val meta_cnf_axiomH : thm -> thm list |
| 16012 | 21 |
val rm_Eps : (term * term) list -> thm list -> term list |
| 15997 | 22 |
val claset_rules_of_thy : theory -> (string * thm) list |
23 |
val simpset_rules_of_thy : theory -> (string * thm) list |
|
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
24 |
val claset_rules_of_ctxt: Proof.context -> (string * thm) list |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
25 |
val simpset_rules_of_ctxt : Proof.context -> (string * thm) list |
| 17829 | 26 |
val clausify_rules_pairs : (string*thm) list -> (ResClause.clause*thm) list list |
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
27 |
val clausify_rules_pairsH : (string*thm) list -> (ResHolClause.clause*thm) list list |
| 16563 | 28 |
val clause_setup : (theory -> theory) list |
29 |
val meson_method_setup : (theory -> theory) list |
|
|
17905
1574533861b1
Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents:
17829
diff
changeset
|
30 |
val pairname : thm -> (string * thm) |
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
31 |
val repeat_RS : thm -> thm -> thm |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
32 |
|
| 15997 | 33 |
end; |
| 15347 | 34 |
|
| 15997 | 35 |
structure ResAxioms : RES_AXIOMS = |
36 |
||
37 |
struct |
|
| 15347 | 38 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
39 |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
40 |
val tagging_enabled = false (*compile_time option*) |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
41 |
|
| 15997 | 42 |
(**** Transformation of Elimination Rules into First-Order Formulas****) |
| 15347 | 43 |
|
| 15390 | 44 |
(* a tactic used to prove an elim-rule. *) |
| 16009 | 45 |
fun elimRule_tac th = |
46 |
((rtac impI 1) ORELSE (rtac notI 1)) THEN (etac th 1) THEN |
|
| 16588 | 47 |
REPEAT(fast_tac HOL_cs 1); |
| 15347 | 48 |
|
49 |
exception ELIMR2FOL of string; |
|
50 |
||
| 15390 | 51 |
(* functions used to construct a formula *) |
52 |
||
| 15347 | 53 |
fun make_disjs [x] = x |
| 15956 | 54 |
| make_disjs (x :: xs) = HOLogic.mk_disj(x, make_disjs xs) |
| 15347 | 55 |
|
56 |
fun make_conjs [x] = x |
|
| 15956 | 57 |
| make_conjs (x :: xs) = HOLogic.mk_conj(x, make_conjs xs) |
58 |
||
59 |
fun add_EX tm [] = tm |
|
60 |
| add_EX tm ((x,xtp)::xs) = add_EX (HOLogic.exists_const xtp $ Abs(x,xtp,tm)) xs; |
|
| 15347 | 61 |
|
62 |
||
63 |
||
| 15956 | 64 |
fun is_neg (Const("Trueprop",_) $ (Const("Not",_) $ Free(p,_))) (Const("Trueprop",_) $ Free(q,_)) = (p = q)
|
| 15371 | 65 |
| is_neg _ _ = false; |
66 |
||
| 15347 | 67 |
|
68 |
exception STRIP_CONCL; |
|
69 |
||
70 |
||
| 15371 | 71 |
fun strip_concl' prems bvs (Const ("==>",_) $ P $ Q) =
|
| 15956 | 72 |
let val P' = HOLogic.dest_Trueprop P |
73 |
val prems' = P'::prems |
|
74 |
in |
|
| 15371 | 75 |
strip_concl' prems' bvs Q |
| 15956 | 76 |
end |
| 15371 | 77 |
| strip_concl' prems bvs P = |
| 15956 | 78 |
let val P' = HOLogic.Not $ (HOLogic.dest_Trueprop P) |
79 |
in |
|
| 15371 | 80 |
add_EX (make_conjs (P'::prems)) bvs |
| 15956 | 81 |
end; |
| 15371 | 82 |
|
83 |
||
84 |
fun strip_concl prems bvs concl (Const ("all", _) $ Abs (x,xtp,body)) = strip_concl prems ((x,xtp)::bvs) concl body
|
|
85 |
| strip_concl prems bvs concl (Const ("==>",_) $ P $ Q) =
|
|
86 |
if (is_neg P concl) then (strip_concl' prems bvs Q) |
|
87 |
else |
|
| 15956 | 88 |
(let val P' = HOLogic.dest_Trueprop P |
| 15371 | 89 |
val prems' = P'::prems |
90 |
in |
|
91 |
strip_concl prems' bvs concl Q |
|
92 |
end) |
|
93 |
| strip_concl prems bvs concl _ = add_EX (make_conjs prems) bvs; |
|
| 15347 | 94 |
|
95 |
||
96 |
||
| 15371 | 97 |
fun trans_elim (main,others,concl) = |
98 |
let val others' = map (strip_concl [] [] concl) others |
|
| 15347 | 99 |
val disjs = make_disjs others' |
100 |
in |
|
| 15956 | 101 |
HOLogic.mk_imp (HOLogic.dest_Trueprop main, disjs) |
| 15347 | 102 |
end; |
103 |
||
104 |
||
| 15390 | 105 |
(* aux function of elim2Fol, take away predicate variable. *) |
| 15371 | 106 |
fun elimR2Fol_aux prems concl = |
| 15347 | 107 |
let val nprems = length prems |
108 |
val main = hd prems |
|
109 |
in |
|
| 15956 | 110 |
if (nprems = 1) then HOLogic.Not $ (HOLogic.dest_Trueprop main) |
| 15371 | 111 |
else trans_elim (main, tl prems, concl) |
| 15347 | 112 |
end; |
113 |
||
| 15956 | 114 |
|
| 16012 | 115 |
(* convert an elim rule into an equivalent formula, of type term. *) |
| 15347 | 116 |
fun elimR2Fol elimR = |
117 |
let val elimR' = Drule.freeze_all elimR |
|
118 |
val (prems,concl) = (prems_of elimR', concl_of elimR') |
|
119 |
in |
|
120 |
case concl of Const("Trueprop",_) $ Free(_,Type("bool",[]))
|
|
| 15956 | 121 |
=> HOLogic.mk_Trueprop (elimR2Fol_aux prems concl) |
122 |
| Free(x,Type("prop",[])) => HOLogic.mk_Trueprop(elimR2Fol_aux prems concl)
|
|
| 15347 | 123 |
| _ => raise ELIMR2FOL("Not an elimination rule!")
|
124 |
end; |
|
125 |
||
126 |
||
| 15390 | 127 |
(* check if a rule is an elim rule *) |
| 16009 | 128 |
fun is_elimR th = |
129 |
case (concl_of th) of (Const ("Trueprop", _) $ Var (idx,_)) => true
|
|
| 15347 | 130 |
| Var(indx,Type("prop",[])) => true
|
131 |
| _ => false; |
|
132 |
||
| 15997 | 133 |
(* convert an elim-rule into an equivalent theorem that does not have the |
134 |
predicate variable. Leave other theorems unchanged.*) |
|
| 16009 | 135 |
fun transform_elim th = |
136 |
if is_elimR th then |
|
137 |
let val tm = elimR2Fol th |
|
138 |
val ctm = cterm_of (sign_of_thm th) tm |
|
| 18009 | 139 |
in Goal.prove_raw [] ctm (fn _ => elimRule_tac th) end |
| 16563 | 140 |
else th; |
| 15997 | 141 |
|
142 |
||
143 |
(**** Transformation of Clasets and Simpsets into First-Order Axioms ****) |
|
144 |
||
| 15390 | 145 |
(* repeated resolution *) |
| 15347 | 146 |
fun repeat_RS thm1 thm2 = |
147 |
let val thm1' = thm1 RS thm2 handle THM _ => thm1 |
|
148 |
in |
|
149 |
if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2) |
|
150 |
end; |
|
151 |
||
152 |
||
| 16009 | 153 |
(*Convert a theorem into NNF and also skolemize it. Original version, using |
154 |
Hilbert's epsilon in the resulting clauses.*) |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
155 |
fun skolem_axiom_g mk_nnf th = |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
156 |
let val th' = (skolemize o mk_nnf o ObjectLogic.atomize_thm o Drule.freeze_all) th |
| 16588 | 157 |
in repeat_RS th' someI_ex |
158 |
end; |
|
| 15347 | 159 |
|
160 |
||
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
161 |
val skolem_axiom = skolem_axiom_g make_nnf; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
162 |
val skolem_axiomH = skolem_axiom_g make_nnf1; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
163 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
164 |
|
| 16009 | 165 |
fun cnf_rule th = make_clauses [skolem_axiom (transform_elim th)]; |
| 15347 | 166 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
167 |
fun cnf_ruleH th = make_clauses [skolem_axiomH (transform_elim th)]; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
168 |
|
| 16563 | 169 |
(*Transfer a theorem into theory Reconstruction.thy if it is not already |
|
15359
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15347
diff
changeset
|
170 |
inside that theory -- because it's needed for Skolemization *) |
|
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15347
diff
changeset
|
171 |
|
| 16563 | 172 |
(*This will refer to the final version of theory Reconstruction.*) |
173 |
val recon_thy_ref = Theory.self_ref (the_context ()); |
|
|
15359
8bad1f42fec0
new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents:
15347
diff
changeset
|
174 |
|
| 16563 | 175 |
(*If called while Reconstruction is being created, it will transfer to the |
176 |
current version. If called afterward, it will transfer to the final version.*) |
|
| 16009 | 177 |
fun transfer_to_Reconstruction th = |
| 16563 | 178 |
transfer (Theory.deref recon_thy_ref) th handle THM _ => th; |
| 15347 | 179 |
|
|
15955
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
180 |
fun is_taut th = |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
181 |
case (prop_of th) of |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
182 |
(Const ("Trueprop", _) $ Const ("True", _)) => true
|
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
183 |
| _ => false; |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
184 |
|
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
185 |
(* remove tautologous clauses *) |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
186 |
val rm_redundant_cls = List.filter (not o is_taut); |
| 15347 | 187 |
|
| 17829 | 188 |
(* transform an Isabelle theorem into CNF *) |
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
189 |
fun cnf_axiom_aux_g cnf_rule th = |
| 16173 | 190 |
map zero_var_indexes |
| 16009 | 191 |
(rm_redundant_cls (cnf_rule (transfer_to_Reconstruction th))); |
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
192 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
193 |
val cnf_axiom_aux = cnf_axiom_aux_g cnf_rule; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
194 |
val cnf_axiom_auxH = cnf_axiom_aux_g cnf_ruleH; |
| 15997 | 195 |
|
196 |
||
| 16009 | 197 |
(**** SKOLEMIZATION BY INFERENCE (lcp) ****) |
198 |
||
199 |
(*Traverse a term, accumulating Skolem function definitions.*) |
|
200 |
fun declare_skofuns s t thy = |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
201 |
let fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) (n, (thy, axs)) =
|
| 16009 | 202 |
(*Existential: declare a Skolem function, then insert into body and continue*) |
203 |
let val cname = s ^ "_" ^ Int.toString n |
|
| 16012 | 204 |
val args = term_frees xtp (*get the formal parameter list*) |
| 16009 | 205 |
val Ts = map type_of args |
206 |
val cT = Ts ---> T |
|
| 16125 | 207 |
val c = Const (Sign.full_name (Theory.sign_of thy) cname, cT) |
| 16009 | 208 |
val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp) |
| 16012 | 209 |
(*Forms a lambda-abstraction over the formal parameters*) |
| 16009 | 210 |
val def = equals cT $ c $ rhs |
211 |
val thy' = Theory.add_consts_i [(cname, cT, NoSyn)] thy |
|
| 16012 | 212 |
(*Theory is augmented with the constant, then its def*) |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
213 |
val cdef = cname ^ "_def" |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
214 |
val thy'' = Theory.add_defs_i false [(cdef, def)] thy' |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
215 |
in dec_sko (subst_bound (list_comb(c,args), p)) |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
216 |
(n+1, (thy'', get_axiom thy'' cdef :: axs)) |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
217 |
end |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
218 |
| dec_sko (Const ("All",_) $ (xtp as Abs(a,T,p))) (n, thx) =
|
| 16012 | 219 |
(*Universal quant: insert a free variable into body and continue*) |
| 16009 | 220 |
let val fname = variant (add_term_names (p,[])) a |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
221 |
in dec_sko (subst_bound (Free(fname,T), p)) (n, thx) end |
| 16009 | 222 |
| dec_sko (Const ("op &", _) $ p $ q) nthy =
|
223 |
dec_sko q (dec_sko p nthy) |
|
224 |
| dec_sko (Const ("op |", _) $ p $ q) nthy =
|
|
225 |
dec_sko q (dec_sko p nthy) |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
226 |
| dec_sko (Const ("HOL.tag", _) $ p) nthy =
|
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
227 |
dec_sko p nthy |
| 16009 | 228 |
| dec_sko (Const ("Trueprop", _) $ p) nthy =
|
229 |
dec_sko p nthy |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
230 |
| dec_sko t nthx = nthx (*Do nothing otherwise*) |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
231 |
in #2 (dec_sko t (1, (thy,[]))) end; |
| 16009 | 232 |
|
233 |
(*cterms are used throughout for efficiency*) |
|
234 |
val cTrueprop = Thm.cterm_of (Theory.sign_of HOL.thy) HOLogic.Trueprop; |
|
235 |
||
236 |
(*cterm version of mk_cTrueprop*) |
|
237 |
fun c_mkTrueprop A = Thm.capply cTrueprop A; |
|
238 |
||
239 |
(*Given an abstraction over n variables, replace the bound variables by free |
|
240 |
ones. Return the body, along with the list of free variables.*) |
|
241 |
fun c_variant_abs_multi (ct0, vars) = |
|
242 |
let val (cv,ct) = Thm.dest_abs NONE ct0 |
|
243 |
in c_variant_abs_multi (ct, cv::vars) end |
|
244 |
handle CTERM _ => (ct0, rev vars); |
|
245 |
||
246 |
(*Given the definition of a Skolem function, return a theorem to replace |
|
247 |
an existential formula by a use of that function.*) |
|
| 16588 | 248 |
fun skolem_of_def def = |
| 16009 | 249 |
let val (c,rhs) = Drule.dest_equals (cprop_of (Drule.freeze_all def)) |
250 |
val (ch, frees) = c_variant_abs_multi (rhs, []) |
|
251 |
val (chil,cabs) = Thm.dest_comb ch |
|
| 16588 | 252 |
val {sign,t, ...} = rep_cterm chil
|
| 16009 | 253 |
val (Const ("Hilbert_Choice.Eps", Type("fun",[_,T]))) = t
|
254 |
val cex = Thm.cterm_of sign (HOLogic.exists_const T) |
|
255 |
val ex_tm = c_mkTrueprop (Thm.capply cex cabs) |
|
256 |
and conc = c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees))); |
|
| 17959 | 257 |
in OldGoals.prove_goalw_cterm [def] (Drule.mk_implies (ex_tm, conc)) |
| 16009 | 258 |
(fn [prem] => [ rtac (prem RS someI_ex) 1 ]) |
259 |
end; |
|
260 |
||
261 |
||
262 |
(*Converts an Isabelle theorem (intro, elim or simp format) into nnf.*) |
|
263 |
fun to_nnf thy th = |
|
| 16588 | 264 |
th |> Thm.transfer thy |
265 |
|> transform_elim |> Drule.freeze_all |
|
266 |
|> ObjectLogic.atomize_thm |> make_nnf; |
|
| 16009 | 267 |
|
268 |
(*The cache prevents repeated clausification of a theorem, |
|
| 16800 | 269 |
and also repeated declaration of Skolem functions*) (* FIXME better use Termtab!? *) |
|
15955
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
270 |
val clause_cache = ref (Symtab.empty : (thm * thm list) Symtab.table) |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
271 |
|
| 16009 | 272 |
(*Declare Skolem functions for a theorem, supplied in nnf and with its name*) |
273 |
fun skolem thy (name,th) = |
|
| 16588 | 274 |
let val cname = (case name of "" => gensym "sko" | s => Sign.base_name s) |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
275 |
val (thy',axs) = declare_skofuns cname (#prop (rep_thm th)) thy |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
276 |
in (map skolem_of_def axs, thy') end; |
| 16009 | 277 |
|
278 |
(*Populate the clause cache using the supplied theorems*) |
|
279 |
fun skolemlist [] thy = thy |
|
280 |
| skolemlist ((name,th)::nths) thy = |
|
| 17412 | 281 |
(case Symtab.lookup (!clause_cache) name of |
| 16009 | 282 |
NONE => |
| 16588 | 283 |
let val (nnfth,ok) = (to_nnf thy th, true) |
284 |
handle THM _ => (asm_rl, false) |
|
285 |
in |
|
286 |
if ok then |
|
287 |
let val (skoths,thy') = skolem thy (name, nnfth) |
|
288 |
val cls = Meson.make_cnf skoths nnfth |
|
| 17412 | 289 |
in change clause_cache (Symtab.update (name, (th, cls))); |
| 16588 | 290 |
skolemlist nths thy' |
291 |
end |
|
292 |
else skolemlist nths thy |
|
293 |
end |
|
| 16009 | 294 |
| SOME _ => skolemlist nths thy) (*FIXME: check for duplicate names?*) |
295 |
||
296 |
(*Exported function to convert Isabelle theorems into axiom clauses*) |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
297 |
fun cnf_axiom_g cnf_axiom_aux (name,th) = |
| 15956 | 298 |
case name of |
|
15955
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
299 |
"" => cnf_axiom_aux th (*no name, so can't cache*) |
| 17412 | 300 |
| s => case Symtab.lookup (!clause_cache) s of |
|
15955
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
301 |
NONE => |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
302 |
let val cls = cnf_axiom_aux th |
| 17412 | 303 |
in change clause_cache (Symtab.update (s, (th, cls))); cls end |
|
15955
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
304 |
| SOME(th',cls) => |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
305 |
if eq_thm(th,th') then cls |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
306 |
else (*New theorem stored under the same name? Possible??*) |
|
87cf2ce8ede8
memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents:
15872
diff
changeset
|
307 |
let val cls = cnf_axiom_aux th |
| 17412 | 308 |
in change clause_cache (Symtab.update (s, (th, cls))); cls end; |
| 15347 | 309 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
310 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
311 |
val cnf_axiom = cnf_axiom_g cnf_axiom_aux; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
312 |
val cnf_axiomH = cnf_axiom_g cnf_axiom_auxH; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
313 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
314 |
|
| 15956 | 315 |
fun pairname th = (Thm.name_of_thm th, th); |
316 |
||
317 |
fun meta_cnf_axiom th = |
|
318 |
map Meson.make_meta_clause (cnf_axiom (pairname th)); |
|
| 15499 | 319 |
|
| 15347 | 320 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
321 |
fun meta_cnf_axiomH th = |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
322 |
map Meson.make_meta_clause (cnf_axiomH (pairname th)); |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
323 |
|
| 15347 | 324 |
(* changed: with one extra case added *) |
| 15956 | 325 |
fun univ_vars_of_aux (Const ("Hilbert_Choice.Eps",_) $ Abs(_,_,body)) vars =
|
326 |
univ_vars_of_aux body vars |
|
327 |
| univ_vars_of_aux (Const ("Ex",_) $ Abs(_,_,body)) vars =
|
|
328 |
univ_vars_of_aux body vars (* EX x. body *) |
|
| 15347 | 329 |
| univ_vars_of_aux (P $ Q) vars = |
| 15956 | 330 |
univ_vars_of_aux Q (univ_vars_of_aux P vars) |
| 15347 | 331 |
| univ_vars_of_aux (t as Var(_,_)) vars = |
| 15956 | 332 |
if (t mem vars) then vars else (t::vars) |
| 15347 | 333 |
| univ_vars_of_aux _ vars = vars; |
334 |
||
335 |
fun univ_vars_of t = univ_vars_of_aux t []; |
|
336 |
||
337 |
||
| 17819 | 338 |
(**** Generating Skoklem Functions ****) |
339 |
||
340 |
val skolem_prefix = "sk_"; |
|
341 |
val skolem_id = ref 0; |
|
342 |
||
343 |
fun gen_skolem_name () = |
|
344 |
let val sk_fun = skolem_prefix ^ string_of_int (! skolem_id); |
|
345 |
in inc skolem_id; sk_fun end; |
|
346 |
||
347 |
fun gen_skolem univ_vars tp = |
|
348 |
let |
|
349 |
val skolem_type = map Term.fastype_of univ_vars ---> tp |
|
350 |
val skolem_term = Const (gen_skolem_name (), skolem_type) |
|
351 |
in Term.list_comb (skolem_term, univ_vars) end; |
|
352 |
||
| 15347 | 353 |
fun get_new_skolem epss (t as (Const ("Hilbert_Choice.Eps",_) $ Abs(_,tp,_))) =
|
354 |
let val all_vars = univ_vars_of t |
|
| 17819 | 355 |
val sk_term = gen_skolem all_vars tp |
| 15347 | 356 |
in |
357 |
(sk_term,(t,sk_term)::epss) |
|
358 |
end; |
|
359 |
||
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
360 |
(*FIXME: use a-list lookup!!*) |
| 15531 | 361 |
fun sk_lookup [] t = NONE |
362 |
| sk_lookup ((tm,sk_tm)::tms) t = if (t = tm) then SOME (sk_tm) else (sk_lookup tms t); |
|
| 15347 | 363 |
|
| 15390 | 364 |
(* get the proper skolem term to replace epsilon term *) |
| 15347 | 365 |
fun get_skolem epss t = |
| 15956 | 366 |
case (sk_lookup epss t) of NONE => get_new_skolem epss t |
367 |
| SOME sk => (sk,epss); |
|
| 15347 | 368 |
|
| 16009 | 369 |
fun rm_Eps_cls_aux epss (t as (Const ("Hilbert_Choice.Eps",_) $ Abs(_,_,_))) =
|
370 |
get_skolem epss t |
|
| 15347 | 371 |
| rm_Eps_cls_aux epss (P $ Q) = |
| 16009 | 372 |
let val (P',epss') = rm_Eps_cls_aux epss P |
373 |
val (Q',epss'') = rm_Eps_cls_aux epss' Q |
|
374 |
in (P' $ Q',epss'') end |
|
| 15347 | 375 |
| rm_Eps_cls_aux epss t = (t,epss); |
376 |
||
| 16009 | 377 |
fun rm_Eps_cls epss th = rm_Eps_cls_aux epss (prop_of th); |
| 15347 | 378 |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
379 |
(* replace the epsilon terms in a formula by skolem terms. *) |
| 15347 | 380 |
fun rm_Eps _ [] = [] |
| 17829 | 381 |
| rm_Eps epss (th::ths) = |
| 16009 | 382 |
let val (th',epss') = rm_Eps_cls epss th |
| 17829 | 383 |
in th' :: (rm_Eps epss' ths) end; |
| 15347 | 384 |
|
385 |
||
386 |
||
| 15872 | 387 |
(**** Extract and Clausify theorems from a theory's claset and simpset ****) |
| 15347 | 388 |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
389 |
(*Preserve the name of "th" after the transformation "f"*) |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
390 |
fun preserve_name f th = Thm.name_thm (Thm.name_of_thm th, f th); |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
391 |
|
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
392 |
(*Tags identify the major premise or conclusion, as hints to resolution provers. |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
393 |
However, they don't appear to help in recent tests, and they complicate the code.*) |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
394 |
val tagI = thm "tagI"; |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
395 |
val tagD = thm "tagD"; |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
396 |
|
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
397 |
val tag_intro = preserve_name (fn th => th RS tagI); |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
398 |
val tag_elim = preserve_name (fn th => tagD RS th); |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
399 |
|
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
400 |
fun rules_of_claset cs = |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
401 |
let val {safeIs,safeEs,hazIs,hazEs,...} = rep_cs cs
|
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
402 |
val intros = safeIs @ hazIs |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
403 |
val elims = safeEs @ hazEs |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
404 |
in |
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
405 |
debug ("rules_of_claset intros: " ^ Int.toString(length intros) ^
|
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
406 |
" elims: " ^ Int.toString(length elims)); |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
407 |
if tagging_enabled |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
408 |
then map pairname (map tag_intro intros @ map tag_elim elims) |
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
409 |
else map pairname (intros @ elims) |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
410 |
end; |
| 15347 | 411 |
|
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
412 |
fun rules_of_simpset ss = |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
413 |
let val ({rules,...}, _) = rep_ss ss
|
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
414 |
val simps = Net.entries rules |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
415 |
in |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
416 |
debug ("rules_of_simpset: " ^ Int.toString(length simps));
|
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
417 |
map (fn r => (#name r, #thm r)) simps |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
418 |
end; |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
419 |
|
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
420 |
fun claset_rules_of_thy thy = rules_of_claset (claset_of thy); |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
421 |
fun simpset_rules_of_thy thy = rules_of_simpset (simpset_of thy); |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
422 |
|
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
423 |
fun claset_rules_of_ctxt ctxt = rules_of_claset (local_claset_of ctxt); |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17412
diff
changeset
|
424 |
fun simpset_rules_of_ctxt ctxt = rules_of_simpset (local_simpset_of ctxt); |
| 15347 | 425 |
|
426 |
||
| 15872 | 427 |
(**** Translate a set of classical/simplifier rules into CNF (still as type "thm") ****) |
| 15347 | 428 |
|
429 |
(* classical rules *) |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
430 |
fun cnf_rules_g cnf_axiom [] err_list = ([],err_list) |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
431 |
| cnf_rules_g cnf_axiom ((name,th) :: ths) err_list = |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
432 |
let val (ts,es) = cnf_rules_g cnf_axiom ths err_list |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
433 |
in (cnf_axiom (name,th) :: ts,es) handle _ => (ts, (th::es)) end; |
| 15347 | 434 |
|
435 |
||
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
436 |
val cnf_rules = cnf_rules_g cnf_axiom; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
437 |
val cnf_rulesH = cnf_rules_g cnf_axiomH; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
438 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
439 |
|
| 15872 | 440 |
(**** Convert all theorems of a claset/simpset into clauses (ResClause.clause) ****) |
| 15347 | 441 |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
442 |
fun addclause (c,th) l = |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
443 |
if ResClause.isTaut c then l else (c,th) :: l; |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
444 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
445 |
fun addclauseH (c,th) l = |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
446 |
if ResHolClause.isTaut c then l else (c,th)::l; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
447 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
448 |
|
| 17829 | 449 |
(* outputs a list of (clause,theorem) pairs *) |
450 |
fun clausify_axiom_pairs (thm_name,th) = |
|
451 |
let val isa_clauses = cnf_axiom (thm_name,th) |
|
|
16039
dfe264950511
Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents:
16012
diff
changeset
|
452 |
val isa_clauses' = rm_Eps [] isa_clauses |
|
dfe264950511
Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents:
16012
diff
changeset
|
453 |
val clauses_n = length isa_clauses |
|
dfe264950511
Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents:
16012
diff
changeset
|
454 |
fun make_axiom_clauses _ [] []= [] |
| 16897 | 455 |
| make_axiom_clauses i (cls::clss) (cls'::clss') = |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
456 |
addclause (ResClause.make_axiom_clause cls (thm_name,i), cls') |
|
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
457 |
(make_axiom_clauses (i+1) clss clss') |
| 15347 | 458 |
in |
|
16039
dfe264950511
Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents:
16012
diff
changeset
|
459 |
make_axiom_clauses 0 isa_clauses' isa_clauses |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
460 |
end |
| 15347 | 461 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
462 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
463 |
fun clausify_axiom_pairsH (thm_name,th) = |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
464 |
let val isa_clauses = cnf_axiomH (thm_name,th) |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
465 |
val isa_clauses' = rm_Eps [] isa_clauses |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
466 |
val clauses_n = length isa_clauses |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
467 |
fun make_axiom_clauses _ [] []= [] |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
468 |
| make_axiom_clauses i (cls::clss) (cls'::clss') = |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
469 |
addclauseH (ResHolClause.make_axiom_clause cls (thm_name,i), cls') |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
470 |
(make_axiom_clauses (i+1) clss clss') |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
471 |
in |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
472 |
make_axiom_clauses 0 isa_clauses' isa_clauses |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
473 |
end |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
474 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
475 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
476 |
|
| 17829 | 477 |
fun clausify_rules_pairs_aux result [] = result |
478 |
| clausify_rules_pairs_aux result ((name,th)::ths) = |
|
479 |
clausify_rules_pairs_aux (clausify_axiom_pairs (name,th) :: result) ths |
|
480 |
handle THM (msg,_,_) => |
|
481 |
(debug ("Cannot clausify " ^ name ^ ": " ^ msg);
|
|
482 |
clausify_rules_pairs_aux result ths) |
|
483 |
| ResClause.CLAUSE (msg,t) => |
|
484 |
(debug ("Cannot clausify " ^ name ^ ": " ^ msg ^
|
|
485 |
": " ^ TermLib.string_of_term t); |
|
486 |
clausify_rules_pairs_aux result ths) |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17279
diff
changeset
|
487 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
488 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
489 |
fun clausify_rules_pairs_auxH result [] = result |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
490 |
| clausify_rules_pairs_auxH result ((name,th)::ths) = |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
491 |
clausify_rules_pairs_auxH (clausify_axiom_pairsH (name,th) :: result) ths |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
492 |
handle THM (msg,_,_) => |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
493 |
(debug ("Cannot clausify " ^ name ^ ": " ^ msg);
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
494 |
clausify_rules_pairs_auxH result ths) |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
495 |
| ResHolClause.LAM2COMB (t) => |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
496 |
(debug ("Cannot clausify " ^ TermLib.string_of_term t);
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
497 |
clausify_rules_pairs_auxH result ths) |
|
16039
dfe264950511
Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents:
16012
diff
changeset
|
498 |
|
| 15347 | 499 |
|
|
18000
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
500 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
501 |
val clausify_rules_pairs = clausify_rules_pairs_aux []; |
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
502 |
|
|
ac059afd6b86
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents:
17959
diff
changeset
|
503 |
val clausify_rules_pairsH = clausify_rules_pairs_auxH []; |
| 16009 | 504 |
(*Setup function: takes a theory and installs ALL simprules and claset rules |
505 |
into the clause cache*) |
|
506 |
fun clause_cache_setup thy = |
|
507 |
let val simps = simpset_rules_of_thy thy |
|
508 |
and clas = claset_rules_of_thy thy |
|
509 |
in skolemlist clas (skolemlist simps thy) end; |
|
510 |
||
| 16563 | 511 |
val clause_setup = [clause_cache_setup]; |
512 |
||
513 |
||
514 |
(*** meson proof methods ***) |
|
515 |
||
516 |
fun cnf_rules_of_ths ths = List.concat (#1 (cnf_rules (map pairname ths) [])); |
|
517 |
||
518 |
fun meson_meth ths ctxt = |
|
519 |
Method.SIMPLE_METHOD' HEADGOAL |
|
520 |
(CHANGED_PROP o Meson.meson_claset_tac (cnf_rules_of_ths ths) (local_claset_of ctxt)); |
|
521 |
||
522 |
val meson_method_setup = |
|
523 |
[Method.add_methods |
|
524 |
[("meson", Method.thms_ctxt_args meson_meth,
|
|
525 |
"The MESON resolution proof procedure")]]; |
|
| 15347 | 526 |
|
527 |
end; |