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