| author | wenzelm | 
| Tue, 03 Jul 2007 22:27:08 +0200 | |
| changeset 23554 | 151d60fbfffe | 
| parent 23387 | 7cb8faa5d4d3 | 
| child 24215 | 5458fbf18276 | 
| permissions | -rw-r--r-- | 
| 19194 | 1 | (* Author: Jia Meng, Cambridge University Computer Laboratory, NICTA | 
| 15608 | 2 | ID: $Id$ | 
| 3 | Copyright 2004 University of Cambridge | |
| 15347 | 4 | |
| 5 | ATPs with TPTP format input. | |
| 6 | *) | |
| 15452 | 7 | |
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 8 | signature RES_ATP = | 
| 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 9 | sig | 
| 17306 | 10 | val prover: string ref | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 11 | val custom_spass: string list ref | 
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 12 | val destdir: string ref | 
| 17849 | 13 | val helper_path: string -> string -> string | 
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 14 | val problem_name: string ref | 
| 17690 
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
 paulson parents: 
17525diff
changeset | 15 | val time_limit: int ref | 
| 21224 | 16 | val set_prover: string -> unit | 
| 21588 | 17 | |
| 19194 | 18 | datatype mode = Auto | Fol | Hol | 
| 19450 
651d949776f8
exported linkup_logic_mode and changed the default setting
 paulson parents: 
19445diff
changeset | 19 | val linkup_logic_mode : mode ref | 
| 19722 | 20 | val write_subgoal_file: bool -> mode -> Proof.context -> thm list -> thm list -> int -> string | 
| 19194 | 21 | val vampire_time: int ref | 
| 22 | val eprover_time: int ref | |
| 19722 | 23 | val spass_time: int ref | 
| 19194 | 24 | val run_vampire: int -> unit | 
| 25 | val run_eprover: int -> unit | |
| 19722 | 26 | val run_spass: int -> unit | 
| 19194 | 27 | val vampireLimit: unit -> int | 
| 28 | val eproverLimit: unit -> int | |
| 19722 | 29 | val spassLimit: unit -> int | 
| 20289 | 30 | val atp_method: (Proof.context -> thm list -> int -> tactic) -> | 
| 31 | Method.src -> Proof.context -> Proof.method | |
| 19194 | 32 | val cond_rm_tmp: string -> unit | 
| 20246 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 33 | val include_all: bool ref | 
| 19194 | 34 | val run_relevance_filter: bool ref | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 35 | val run_blacklist_filter: bool ref | 
| 20372 
e42674e0486e
blacklist augmented to block some "unit" theorems that cause unsound resolution proofs
 paulson parents: 
20289diff
changeset | 36 | val blacklist: string list ref | 
| 19894 | 37 | val add_all : unit -> unit | 
| 19227 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 38 | val add_claset : unit -> unit | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 39 | val add_simpset : unit -> unit | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 40 | val add_clasimp : unit -> unit | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 41 | val add_atpset : unit -> unit | 
| 19894 | 42 | val rm_all : unit -> unit | 
| 19227 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 43 | val rm_claset : unit -> unit | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 44 | val rm_simpset : unit -> unit | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 45 | val rm_atpset : unit -> unit | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 46 | val rm_clasimp : unit -> unit | 
| 21311 | 47 | val is_fol_thms : thm list -> bool | 
| 22989 | 48 | val tvar_classes_of_terms : term list -> string list | 
| 49 | val tfree_classes_of_terms : term list -> string list | |
| 50 | val type_consts_of_terms : theory -> term list -> string list | |
| 15347 | 51 | end; | 
| 52 | ||
| 22865 | 53 | structure ResAtp: RES_ATP = | 
| 15347 | 54 | struct | 
| 55 | ||
| 22130 | 56 | fun timestamp s = Output.debug (fn () => ("At " ^ Time.toString (Time.now()) ^ ": " ^ s));
 | 
| 21070 | 57 | |
| 19194 | 58 | (********************************************************************) | 
| 59 | (* some settings for both background automatic ATP calling procedure*) | |
| 60 | (* and also explicit ATP invocation methods *) | |
| 61 | (********************************************************************) | |
| 62 | ||
| 63 | (*** background linkup ***) | |
| 21588 | 64 | val call_atp = ref false; | 
| 17150 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 quigley parents: 
17091diff
changeset | 65 | val hook_count = ref 0; | 
| 21224 | 66 | val time_limit = ref 60; | 
| 21588 | 67 | val prover = ref ""; | 
| 21224 | 68 | |
| 21588 | 69 | fun set_prover atp = | 
| 21224 | 70 | case String.map Char.toLower atp of | 
| 21588 | 71 | "e" => | 
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 72 | (ReduceAxiomsN.max_new := 100; | 
| 21224 | 73 | ReduceAxiomsN.theory_const := false; | 
| 74 | prover := "E") | |
| 21588 | 75 | | "spass" => | 
| 21224 | 76 | (ReduceAxiomsN.max_new := 40; | 
| 77 | ReduceAxiomsN.theory_const := true; | |
| 78 | prover := "spass") | |
| 21588 | 79 | | "vampire" => | 
| 21224 | 80 | (ReduceAxiomsN.max_new := 60; | 
| 81 | ReduceAxiomsN.theory_const := false; | |
| 82 | prover := "vampire") | |
| 83 |     | _ => error ("No such prover: " ^ atp);
 | |
| 84 | ||
| 85 | val _ = set_prover "E"; (* use E as the default prover *) | |
| 86 | ||
| 17305 
6cef3aedd661
axioms  now included in tptp files, no /bin/cat and various tidying
 paulson parents: 
17235diff
changeset | 87 | val custom_spass = (*specialized options for SPASS*) | 
| 17690 
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
 paulson parents: 
17525diff
changeset | 88 | ref ["-Auto=0","-FullRed=0","-IORe","-IOFc","-RTaut","-RFSub","-RBSub"]; | 
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 89 | val destdir = ref ""; (*Empty means write files to /tmp*) | 
| 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 90 | val problem_name = ref "prob"; | 
| 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 91 | |
| 17819 | 92 | (*Return the path to a "helper" like SPASS or tptp2X, first checking that | 
| 21588 | 93 | it exists. FIXME: modify to use Path primitives and move to some central place.*) | 
| 17819 | 94 | fun helper_path evar base = | 
| 95 | case getenv evar of | |
| 96 |       "" => error  ("Isabelle environment variable " ^ evar ^ " not defined")
 | |
| 21588 | 97 | | home => | 
| 17819 | 98 | let val path = home ^ "/" ^ base | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21790diff
changeset | 99 | in if File.exists (File.explode_platform_path path) then path | 
| 21588 | 100 |             else error ("Could not find the file " ^ path)
 | 
| 101 | end; | |
| 17819 | 102 | |
| 21588 | 103 | fun probfile_nosuffix _ = | 
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 104 | if !destdir = "" then File.platform_path (File.tmp_path (Path.basic (!problem_name))) | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21790diff
changeset | 105 | else if File.exists (File.explode_platform_path (!destdir)) | 
| 17484 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 106 | then !destdir ^ "/" ^ !problem_name | 
| 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 paulson parents: 
17435diff
changeset | 107 |   else error ("No such directory: " ^ !destdir);
 | 
| 15644 | 108 | |
| 17717 | 109 | fun prob_pathname n = probfile_nosuffix n ^ "_" ^ Int.toString n; | 
| 110 | ||
| 19194 | 111 | |
| 112 | (*** ATP methods ***) | |
| 113 | val vampire_time = ref 60; | |
| 114 | val eprover_time = ref 60; | |
| 19722 | 115 | val spass_time = ref 60; | 
| 116 | ||
| 21588 | 117 | fun run_vampire time = | 
| 19194 | 118 | if (time >0) then vampire_time:= time | 
| 119 | else vampire_time:=60; | |
| 120 | ||
| 21588 | 121 | fun run_eprover time = | 
| 19194 | 122 | if (time > 0) then eprover_time:= time | 
| 123 | else eprover_time:=60; | |
| 124 | ||
| 21588 | 125 | fun run_spass time = | 
| 19722 | 126 | if (time > 0) then spass_time:=time | 
| 127 | else spass_time:=60; | |
| 128 | ||
| 129 | ||
| 19194 | 130 | fun vampireLimit () = !vampire_time; | 
| 131 | fun eproverLimit () = !eprover_time; | |
| 19722 | 132 | fun spassLimit () = !spass_time; | 
| 19194 | 133 | |
| 134 | fun atp_input_file () = | |
| 21588 | 135 | let val file = !problem_name | 
| 19194 | 136 | in | 
| 21588 | 137 | if !destdir = "" then File.platform_path (File.tmp_path (Path.basic file)) | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21790diff
changeset | 138 | else if File.exists (File.explode_platform_path (!destdir)) | 
| 21588 | 139 | then !destdir ^ "/" ^ file | 
| 140 |         else error ("No such directory: " ^ !destdir)
 | |
| 19194 | 141 | end; | 
| 142 | ||
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 143 | val include_all = ref true; | 
| 19194 | 144 | val include_simpset = ref false; | 
| 21588 | 145 | val include_claset = ref false; | 
| 19194 | 146 | val include_atpset = ref true; | 
| 20246 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 147 | |
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 148 | (*Tests show that follow_defs gives VERY poor results with "include_all"*) | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 149 | fun add_all() = (include_all:=true; ReduceAxiomsN.follow_defs := false); | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 150 | fun rm_all() = include_all:=false; | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 151 | |
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 152 | fun add_simpset() = include_simpset:=true; | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 153 | fun rm_simpset() = include_simpset:=false; | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 154 | |
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 155 | fun add_claset() = include_claset:=true; | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 156 | fun rm_claset() = include_claset:=false; | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 157 | |
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 158 | fun add_clasimp() = (include_simpset:=true;include_claset:=true); | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 159 | fun rm_clasimp() = (include_simpset:=false;include_claset:=false); | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 160 | |
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 161 | fun add_atpset() = include_atpset:=true; | 
| 
fdfe7399e057
"all theorems" mode forces definition-expansion off.
 paulson parents: 
20224diff
changeset | 162 | fun rm_atpset() = include_atpset:=false; | 
| 19194 | 163 | |
| 164 | ||
| 165 | (**** relevance filter ****) | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 166 | val run_relevance_filter = ReduceAxiomsN.run_relevance_filter; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 167 | val run_blacklist_filter = ref true; | 
| 19194 | 168 | |
| 169 | (******************************************************************) | |
| 170 | (* detect whether a given problem (clauses) is FOL/HOL/HOLC/HOLCS *) | |
| 171 | (******************************************************************) | |
| 172 | ||
| 173 | datatype logic = FOL | HOL | HOLC | HOLCS; | |
| 174 | ||
| 175 | fun string_of_logic FOL = "FOL" | |
| 176 | | string_of_logic HOL = "HOL" | |
| 177 | | string_of_logic HOLC = "HOLC" | |
| 178 | | string_of_logic HOLCS = "HOLCS"; | |
| 179 | ||
| 180 | fun is_fol_logic FOL = true | |
| 181 | | is_fol_logic _ = false | |
| 182 | ||
| 183 | (*HOLCS will not occur here*) | |
| 184 | fun upgrade_lg HOLC _ = HOLC | |
| 185 | | upgrade_lg HOL HOLC = HOLC | |
| 186 | | upgrade_lg HOL _ = HOL | |
| 21588 | 187 | | upgrade_lg FOL lg = lg; | 
| 19194 | 188 | |
| 189 | (* check types *) | |
| 19451 | 190 | fun has_bool_hfn (Type("bool",_)) = true
 | 
| 191 |   | has_bool_hfn (Type("fun",_)) = true
 | |
| 192 | | has_bool_hfn (Type(_, Ts)) = exists has_bool_hfn Ts | |
| 193 | | has_bool_hfn _ = false; | |
| 19194 | 194 | |
| 19451 | 195 | fun is_hol_fn tp = | 
| 19194 | 196 | let val (targs,tr) = strip_type tp | 
| 197 | in | |
| 21588 | 198 | exists (has_bool_hfn) (tr::targs) | 
| 19194 | 199 | end; | 
| 200 | ||
| 19451 | 201 | fun is_hol_pred tp = | 
| 202 | let val (targs,tr) = strip_type tp | |
| 203 | in | |
| 21588 | 204 | exists (has_bool_hfn) targs | 
| 19451 | 205 | end; | 
| 19194 | 206 | |
| 207 | exception FN_LG of term; | |
| 208 | ||
| 21588 | 209 | fun fn_lg (t as Const(f,tp)) (lg,seen) = | 
| 210 | if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) | |
| 211 | | fn_lg (t as Free(f,tp)) (lg,seen) = | |
| 212 | if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) | |
| 19194 | 213 | | fn_lg (t as Var(f,tp)) (lg,seen) = | 
| 20854 | 214 | if is_hol_fn tp then (upgrade_lg HOL lg,insert (op =) t seen) else (lg,insert (op =) t seen) | 
| 215 | | fn_lg (t as Abs(_,_,_)) (lg,seen) = (upgrade_lg HOLC lg,insert (op =) t seen) | |
| 21588 | 216 | | fn_lg f _ = raise FN_LG(f); | 
| 19194 | 217 | |
| 218 | ||
| 219 | fun term_lg [] (lg,seen) = (lg,seen) | |
| 220 | | term_lg (tm::tms) (FOL,seen) = | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 221 | let val (f,args) = strip_comb tm | 
| 21588 | 222 | val (lg',seen') = if f mem seen then (FOL,seen) | 
| 223 | else fn_lg f (FOL,seen) | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 224 | in | 
| 21588 | 225 | if is_fol_logic lg' then () | 
| 22130 | 226 |         else Output.debug (fn () => ("Found a HOL term: " ^ Display.raw_string_of_term f));
 | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 227 | term_lg (args@tms) (lg',seen') | 
| 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 228 | end | 
| 19194 | 229 | | term_lg _ (lg,seen) = (lg,seen) | 
| 230 | ||
| 231 | exception PRED_LG of term; | |
| 232 | ||
| 21588 | 233 | fun pred_lg (t as Const(P,tp)) (lg,seen)= | 
| 234 | if is_hol_pred tp then (upgrade_lg HOL lg, insert (op =) t seen) | |
| 235 | else (lg,insert (op =) t seen) | |
| 19194 | 236 | | pred_lg (t as Free(P,tp)) (lg,seen) = | 
| 21588 | 237 | if is_hol_pred tp then (upgrade_lg HOL lg, insert (op =) t seen) | 
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 238 | else (lg,insert (op =) t seen) | 
| 20854 | 239 | | pred_lg (t as Var(_,_)) (lg,seen) = (upgrade_lg HOL lg, insert (op =) t seen) | 
| 19194 | 240 | | pred_lg P _ = raise PRED_LG(P); | 
| 241 | ||
| 242 | ||
| 243 | fun lit_lg (Const("Not",_) $ P) (lg,seen) = lit_lg P (lg,seen)
 | |
| 244 | | lit_lg P (lg,seen) = | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 245 | let val (pred,args) = strip_comb P | 
| 21588 | 246 | val (lg',seen') = if pred mem seen then (lg,seen) | 
| 247 | else pred_lg pred (lg,seen) | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 248 | in | 
| 21588 | 249 | if is_fol_logic lg' then () | 
| 22130 | 250 |         else Output.debug (fn () => ("Found a HOL predicate: " ^ Display.raw_string_of_term pred));
 | 
| 21588 | 251 | term_lg args (lg',seen') | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 252 | end; | 
| 19194 | 253 | |
| 254 | fun lits_lg [] (lg,seen) = (lg,seen) | |
| 255 | | lits_lg (lit::lits) (FOL,seen) = | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 256 | let val (lg,seen') = lit_lg lit (FOL,seen) | 
| 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 257 | in | 
| 21588 | 258 | if is_fol_logic lg then () | 
| 22130 | 259 |         else Output.debug (fn () => ("Found a HOL literal: " ^ Display.raw_string_of_term lit));
 | 
| 21588 | 260 | lits_lg lits (lg,seen') | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 261 | end | 
| 19194 | 262 | | lits_lg lits (lg,seen) = (lg,seen); | 
| 263 | ||
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 264 | fun dest_disj_aux (Const("Trueprop",_) $ t) disjs = dest_disj_aux t disjs
 | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 265 |   | dest_disj_aux (Const ("op |", _) $ t $ t') disjs = dest_disj_aux t (dest_disj_aux t' disjs)
 | 
| 19227 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 266 | | dest_disj_aux t disjs = t::disjs; | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 267 | |
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 268 | fun dest_disj t = dest_disj_aux t []; | 
| 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 269 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 270 | fun logic_of_clause tm = lits_lg (dest_disj tm); | 
| 19194 | 271 | |
| 272 | fun logic_of_clauses [] (lg,seen) = (lg,seen) | |
| 273 | | logic_of_clauses (cls::clss) (FOL,seen) = | |
| 19227 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 274 | let val (lg,seen') = logic_of_clause cls (FOL,seen) | 
| 21588 | 275 | val _ = | 
| 19641 
f1de44e61ec1
replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
 wenzelm parents: 
19617diff
changeset | 276 | if is_fol_logic lg then () | 
| 22130 | 277 |           else Output.debug (fn () => ("Found a HOL clause: " ^ Display.raw_string_of_term cls))
 | 
| 19227 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 278 | in | 
| 21588 | 279 | logic_of_clauses clss (lg,seen') | 
| 19227 
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
 mengj parents: 
19205diff
changeset | 280 | end | 
| 19194 | 281 | | logic_of_clauses (cls::clss) (lg,seen) = (lg,seen); | 
| 282 | ||
| 283 | fun problem_logic_goals_aux [] (lg,seen) = lg | |
| 21588 | 284 | | problem_logic_goals_aux (subgoal::subgoals) (lg,seen) = | 
| 19194 | 285 | problem_logic_goals_aux subgoals (logic_of_clauses subgoal (lg,seen)); | 
| 21588 | 286 | |
| 19194 | 287 | fun problem_logic_goals subgoals = problem_logic_goals_aux subgoals (FOL,[]); | 
| 288 | ||
| 21311 | 289 | fun is_fol_thms ths = ((fst(logic_of_clauses (map prop_of ths) (FOL,[]))) = FOL); | 
| 290 | ||
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 291 | (***************************************************************) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 292 | (* Retrieving and filtering lemmas *) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 293 | (***************************************************************) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 294 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 295 | (*** white list and black list of lemmas ***) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 296 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 297 | (*The rule subsetI is frequently omitted by the relevance filter.*) | 
| 21588 | 298 | val whitelist = ref [subsetI]; | 
| 299 | ||
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 300 | (*Names of theorems and theorem lists to be blacklisted. | 
| 20444 | 301 | |
| 302 | These theorems typically produce clauses that are prolific (match too many equality or | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 303 | membership literals) and relate to seldom-used facts. Some duplicate other rules. | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 304 | FIXME: this blacklist needs to be maintained using theory data and added to using | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 305 | an attribute.*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 306 | val blacklist = ref | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 307 | ["Datatype.prod.size", | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 308 | "Divides.dvd_0_left_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 309 | "Finite_Set.card_0_eq", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 310 | "Finite_Set.card_infinite", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 311 | "Finite_Set.Max_ge", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 312 | "Finite_Set.Max_in", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 313 | "Finite_Set.Max_le_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 314 | "Finite_Set.Max_less_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 315 | "Finite_Set.max.f_below_strict_below.below_f_conv", (*duplicates in Orderings.*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 316 | "Finite_Set.max.f_below_strict_below.strict_below_f_conv", (*duplicates in Orderings.*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 317 | "Finite_Set.Min_ge_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 318 | "Finite_Set.Min_gr_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 319 | "Finite_Set.Min_in", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 320 | "Finite_Set.Min_le", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 321 | "Finite_Set.min.f_below_strict_below.below_f_conv", (*duplicates in Orderings.*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 322 | "Finite_Set.min.f_below_strict_below.strict_below_f_conv", (*duplicates in Orderings.*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 323 | "Fun.vimage_image_eq", (*involves an existential quantifier*) | 
| 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 324 | "HOL.split_if_asm", (*splitting theorem*) | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 325 | "HOL.split_if", (*splitting theorem*) | 
| 22193 | 326 | "HOL.All_def", (*far worse than useless!!*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 327 | "IntDef.abs_split", | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 328 | "IntDef.Integ.Abs_Integ_inject", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 329 | "IntDef.Integ.Abs_Integ_inverse", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 330 | "IntDiv.zdvd_0_left", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 331 | "List.append_eq_append_conv", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 332 | "List.hd_Cons_tl", (*Says everything is [] or Cons. Probably prolific.*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 333 | "List.in_listsD", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 334 | "List.in_listsI", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 335 | "List.lists.Cons", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 336 | "List.listsE", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 337 | "Nat.less_one", (*not directional? obscure*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 338 | "Nat.not_gr0", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 339 | "Nat.one_eq_mult_iff", (*duplicate by symmetry*) | 
| 21243 | 340 | "Nat.of_nat_0_eq_iff", | 
| 341 | "Nat.of_nat_eq_0_iff", | |
| 342 | "Nat.of_nat_le_0_iff", | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 343 | "NatSimprocs.divide_le_0_iff_number_of", (*too many clauses*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 344 | "NatSimprocs.divide_less_0_iff_number_of", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 345 | "NatSimprocs.equation_minus_iff_1", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 346 | "NatSimprocs.equation_minus_iff_number_of", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 347 | "NatSimprocs.le_minus_iff_1", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 348 | "NatSimprocs.le_minus_iff_number_of", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 349 | "NatSimprocs.less_minus_iff_1", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 350 | "NatSimprocs.less_minus_iff_number_of", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 351 | "NatSimprocs.minus_equation_iff_number_of", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 352 | "NatSimprocs.minus_le_iff_1", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 353 | "NatSimprocs.minus_le_iff_number_of", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 354 | "NatSimprocs.minus_less_iff_1", (*not directional*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 355 | "NatSimprocs.mult_le_cancel_left_number_of", (*excessive case analysis*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 356 | "NatSimprocs.mult_le_cancel_right_number_of", (*excessive case analysis*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 357 | "NatSimprocs.mult_less_cancel_left_number_of", (*excessive case analysis*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 358 | "NatSimprocs.mult_less_cancel_right_number_of", (*excessive case analysis*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 359 | "NatSimprocs.zero_le_divide_iff_number_of", (*excessive case analysis*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 360 | "NatSimprocs.zero_less_divide_iff_number_of", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 361 | "OrderedGroup.abs_0_eq", (*duplicate by symmetry*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 362 | "OrderedGroup.diff_eq_0_iff_eq", (*prolific?*) | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22382diff
changeset | 363 | "OrderedGroup.sup_0_eq_0", | 
| 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22382diff
changeset | 364 | "OrderedGroup.inf_0_eq_0", | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 365 | "OrderedGroup.pprt_eq_0", (*obscure*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 366 | "OrderedGroup.pprt_eq_id", (*obscure*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 367 | "OrderedGroup.pprt_mono", (*obscure*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 368 | "Orderings.split_max", (*splitting theorem*) | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 369 | "Orderings.split_min", (*splitting theorem*) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 370 | "Power.zero_less_power_abs_iff", | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 371 | "Product_Type.split_eta_SetCompr", (*involves an existential quantifier*) | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 372 | "Product_Type.split_paired_Ball_Sigma", (*splitting theorem*) | 
| 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 373 | "Product_Type.split_paired_Bex_Sigma", (*splitting theorem*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 374 | "Product_Type.split_split_asm", (*splitting theorem*) | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 375 | "Product_Type.split_split", (*splitting theorem*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 376 | "Product_Type.unit_abs_eta_conv", | 
| 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 377 | "Product_Type.unit_induct", | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 378 | "Relation.diagI", | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 379 | "Relation.Domain_def", (*involves an existential quantifier*) | 
| 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 380 | "Relation.Image_def", (*involves an existential quantifier*) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 381 | "Relation.ImageI", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 382 | "Ring_and_Field.divide_cancel_left", (*fields are seldom used & often prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 383 | "Ring_and_Field.divide_cancel_right", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 384 | "Ring_and_Field.divide_divide_eq_left", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 385 | "Ring_and_Field.divide_divide_eq_right", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 386 | "Ring_and_Field.divide_eq_0_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 387 | "Ring_and_Field.divide_eq_1_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 388 | "Ring_and_Field.divide_eq_eq_1", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 389 | "Ring_and_Field.divide_le_0_1_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 390 | "Ring_and_Field.divide_le_eq_1_neg", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 391 | "Ring_and_Field.divide_le_eq_1_pos", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 392 | "Ring_and_Field.divide_less_0_1_iff", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 393 | "Ring_and_Field.divide_less_eq_1_neg", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 394 | "Ring_and_Field.divide_less_eq_1_pos", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 395 | "Ring_and_Field.eq_divide_eq_1", (*duplicate by symmetry*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 396 | "Ring_and_Field.field_mult_cancel_left", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 397 | "Ring_and_Field.field_mult_cancel_right", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 398 | "Ring_and_Field.inverse_le_iff_le_neg", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 399 | "Ring_and_Field.inverse_le_iff_le", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 400 | "Ring_and_Field.inverse_less_iff_less_neg", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 401 | "Ring_and_Field.inverse_less_iff_less", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 402 | "Ring_and_Field.le_divide_eq_1_neg", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 403 | "Ring_and_Field.le_divide_eq_1_pos", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 404 | "Ring_and_Field.less_divide_eq_1_neg", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 405 | "Ring_and_Field.less_divide_eq_1_pos", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 406 | "Ring_and_Field.one_eq_divide_iff", (*duplicate by symmetry*) | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 407 | "Set.ball_simps", "Set.bex_simps", (*quantifier rewriting: useless*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 408 | "Set.Collect_bex_eq", (*involves an existential quantifier*) | 
| 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 409 | "Set.Collect_ex_eq", (*involves an existential quantifier*) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 410 | "Set.Diff_eq_empty_iff", (*redundant with paramodulation*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 411 | "Set.Diff_insert0", | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 412 | "Set.empty_Union_conv", (*redundant with paramodulation*) | 
| 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 413 | "Set.full_SetCompr_eq", (*involves an existential quantifier*) | 
| 22193 | 414 | "Set.image_Collect", (*involves Collect and a boolean variable...*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 415 | "Set.image_def", (*involves an existential quantifier*) | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 416 | "Set.disjoint_insert", "Set.insert_disjoint", | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 417 | "Set.Int_UNIV", (*redundant with paramodulation*) | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 418 | "Set.Inter_UNIV_conv", | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 419 | "Set.Inter_iff", (*We already have InterI, InterE*) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 420 | "Set.psubsetE", (*too prolific and obscure*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 421 | "Set.psubsetI", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 422 | "Set.singleton_insert_inj_eq'", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 423 | "Set.singleton_insert_inj_eq", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 424 | "Set.singletonD", (*these two duplicate some "insert" lemmas*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 425 | "Set.singletonI", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 426 | "Set.Un_empty", (*redundant with paramodulation*) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 427 | "Set.UNION_def", (*involves an existential quantifier*) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 428 | "Set.Union_empty_conv", (*redundant with paramodulation*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 429 | "Set.Union_iff", (*We already have UnionI, UnionE*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 430 | "SetInterval.atLeastAtMost_iff", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 431 | "SetInterval.atLeastLessThan_iff", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 432 | "SetInterval.greaterThanAtMost_iff", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 433 | "SetInterval.greaterThanLessThan_iff", (*obscure and prolific*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 434 | "SetInterval.ivl_subset"]; (*excessive case analysis*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 435 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 436 | (*These might be prolific but are probably OK, and min and max are basic. | 
| 21588 | 437 | "Orderings.max_less_iff_conj", | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 438 | "Orderings.min_less_iff_conj", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 439 | "Orderings.min_max.below_inf.below_inf_conv", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 440 | "Orderings.min_max.below_sup.above_sup_conv", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 441 | Very prolific and somewhat obscure: | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 442 | "Set.InterD", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 443 | "Set.UnionI", | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 444 | *) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 445 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 446 | (*** retrieve lemmas from clasimpset and atpset, may filter them ***) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 447 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 448 | (*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 449 | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 450 | fun setinsert (x,s) = Symtab.update (x,()) s; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 451 | |
| 20757 
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
 paulson parents: 
20661diff
changeset | 452 | (*Reject theorems with names like "List.filter.filter_list_def" or | 
| 21690 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 453 | "Accessible_Part.acc.defs", as these are definitions arising from packages.*) | 
| 20757 
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
 paulson parents: 
20661diff
changeset | 454 | fun is_package_def a = | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21790diff
changeset | 455 | let val names = NameSpace.explode a | 
| 21690 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 456 | in | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 457 | length names > 2 andalso | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 458 | not (hd names = "local") andalso | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 459 | String.isSuffix "_def" a orelse String.isSuffix "_defs" a | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 460 | end; | 
| 20757 
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
 paulson parents: 
20661diff
changeset | 461 | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 462 | (** a hash function from Term.term to int, and also a hash table **) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 463 | val xor_words = List.foldl Word.xorb 0w0; | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 464 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 465 | fun hashw_term ((Const(c,_)), w) = Polyhash.hashw_string (c,w) | 
| 20661 
46832fee1215
Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
 paulson parents: 
20643diff
changeset | 466 | | hashw_term ((Free(a,_)), w) = Polyhash.hashw_string (a,w) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 467 | | hashw_term ((Var(_,_)), w) = w | 
| 20661 
46832fee1215
Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
 paulson parents: 
20643diff
changeset | 468 | | hashw_term ((Bound i), w) = Polyhash.hashw_int(i,w) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 469 | | hashw_term ((Abs(_,_,t)), w) = hashw_term (t, w) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 470 | | hashw_term ((P$Q), w) = hashw_term (Q, (hashw_term (P, w))); | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 471 | |
| 21070 | 472 | fun hash_literal (Const("Not",_)$P) = Word.notb(hashw_term(P,0w0))
 | 
| 473 | | hash_literal P = hashw_term(P,0w0); | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 474 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 475 | fun hash_term t = Word.toIntX (xor_words (map hash_literal (dest_disj t))); | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 476 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 477 | fun equal_thm (thm1,thm2) = Term.aconv(prop_of thm1, prop_of thm2); | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 478 | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 479 | exception HASH_CLAUSE; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 480 | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 481 | (*Create a hash table for clauses, of the given size*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 482 | fun mk_clause_table n = | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 483 | Polyhash.mkTable (hash_term o prop_of, equal_thm) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 484 | (n, HASH_CLAUSE); | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 485 | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 486 | (*Use a hash table to eliminate duplicates from xs. Argument is a list of | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 487 | (thm * (string * int)) tuples. The theorems are hashed into the table. *) | 
| 21588 | 488 | fun make_unique xs = | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 489 | let val ht = mk_clause_table 7000 | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 490 | in | 
| 22130 | 491 |       Output.debug (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses"));
 | 
| 21588 | 492 | app (ignore o Polyhash.peekInsert ht) xs; | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 493 | Polyhash.listItems ht | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 494 | end; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 495 | |
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 496 | (*Remove existing axiom clauses from the conjecture clauses, as this can dramatically | 
| 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 497 | boost an ATP's performance (for some reason)*) | 
| 21588 | 498 | fun subtract_cls c_clauses ax_clauses = | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 499 | let val ht = mk_clause_table 2200 | 
| 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 500 | fun known x = isSome (Polyhash.peek ht x) | 
| 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 501 | in | 
| 21588 | 502 | app (ignore o Polyhash.peekInsert ht) ax_clauses; | 
| 503 | filter (not o known) c_clauses | |
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 504 | end; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 505 | |
| 21588 | 506 | (*Filter axiom clauses, but keep supplied clauses and clauses in whitelist. | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 507 | Duplicates are removed later.*) | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 508 | fun get_relevant_clauses thy cls_thms white_cls goals = | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 509 | white_cls @ (ReduceAxiomsN.relevance_filter thy cls_thms goals); | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 510 | |
| 21224 | 511 | fun all_valid_thms ctxt = | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 512 | let | 
| 22865 | 513 | val banned_tab = foldl setinsert Symtab.empty (!blacklist) | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 514 | fun blacklisted s = !run_blacklist_filter andalso | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 515 | (isSome (Symtab.lookup banned_tab s) orelse is_package_def s) | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 516 | |
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 517 | fun extern_valid space (name, ths) = | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 518 | let | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 519 | val is_valid = ProofContext.valid_thms ctxt; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 520 | val xname = NameSpace.extern space name; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 521 | in | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 522 | if blacklisted name then I | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 523 | else if is_valid (xname, ths) then cons (xname, ths) | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 524 | else if is_valid (name, ths) then cons (name, ths) | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 525 | else I | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 526 | end; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 527 | val thy = ProofContext.theory_of ctxt; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 528 | val all_thys = thy :: Theory.ancestors_of thy; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 529 | |
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 530 | fun dest_valid (space, tab) = Symtab.fold (extern_valid space) tab []; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 531 | in | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 532 | maps (dest_valid o PureThy.theorems_of) all_thys @ | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 533 | fold (extern_valid (#1 (ProofContext.theorems_of ctxt))) | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 534 | (FactIndex.find (ProofContext.fact_index_of ctxt) ([], [])) [] | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 535 | end; | 
| 21224 | 536 | |
| 21588 | 537 | fun multi_name a (th, (n,pairs)) = | 
| 21224 | 538 |   (n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs)
 | 
| 539 | ||
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 540 | fun add_single_names ((a, []), pairs) = pairs | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 541 | | add_single_names ((a, [th]), pairs) = (a,th)::pairs | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 542 | | add_single_names ((a, ths), pairs) = #2 (foldl (multi_name a) (1,pairs) ths); | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 543 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 544 | val multi_base_blacklist = | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 545 | ["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm"]; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 546 | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 547 | (*Ignore blacklisted basenames*) | 
| 21588 | 548 | fun add_multi_names ((a, ths), pairs) = | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 549 | if (Sign.base_name a) mem_string multi_base_blacklist then pairs | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 550 | else add_single_names ((a, ths), pairs); | 
| 21224 | 551 | |
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 552 | fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a; | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 553 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 554 | (*The single theorems go BEFORE the multiple ones*) | 
| 21588 | 555 | fun name_thm_pairs ctxt = | 
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 556 | let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt) | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 557 | in foldl add_single_names (foldl add_multi_names [] mults) singles end; | 
| 21224 | 558 | |
| 559 | fun check_named ("",th) = (warning ("No name for theorem " ^ string_of_thm th); false)
 | |
| 560 | | check_named (_,th) = true; | |
| 19894 | 561 | |
| 22193 | 562 | fun display_thm (name,th) = Output.debug (fn () => name ^ ": " ^ string_of_thm th); | 
| 563 | ||
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 564 | (* get lemmas from claset, simpset, atpset and extra supplied rules *) | 
| 21588 | 565 | fun get_clasimp_atp_lemmas ctxt user_thms = | 
| 19894 | 566 | let val included_thms = | 
| 21588 | 567 | if !include_all | 
| 568 | then (tap (fn ths => Output.debug | |
| 22130 | 569 |                      (fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
 | 
| 21588 | 570 | (name_thm_pairs ctxt)) | 
| 571 | else | |
| 572 | let val claset_thms = | |
| 573 | if !include_claset then ResAxioms.claset_rules_of ctxt | |
| 574 | else [] | |
| 575 | val simpset_thms = | |
| 576 | if !include_simpset then ResAxioms.simpset_rules_of ctxt | |
| 577 | else [] | |
| 578 | val atpset_thms = | |
| 579 | if !include_atpset then ResAxioms.atpset_rules_of ctxt | |
| 580 | else [] | |
| 22193 | 581 | val _ = (Output.debug (fn () => "ATP theorems: "); app display_thm atpset_thms) | 
| 21588 | 582 | in claset_thms @ simpset_thms @ atpset_thms end | 
| 583 | val user_rules = filter check_named | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 584 | (map ResAxioms.pairname | 
| 21588 | 585 | (if null user_thms then !whitelist else user_thms)) | 
| 19894 | 586 | in | 
| 21224 | 587 | (filter check_named included_thms, user_rules) | 
| 19894 | 588 | end; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 589 | |
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 590 | (***************************************************************) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 591 | (* Type Classes Present in the Axiom or Conjecture Clauses *) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 592 | (***************************************************************) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 593 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 594 | fun add_classes (sorts, cset) = foldl setinsert cset (List.concat sorts); | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 595 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 596 | (*Remove this trivial type class*) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 597 | fun delete_type cset = Symtab.delete_safe "HOL.type" cset; | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 598 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 599 | fun tvar_classes_of_terms ts = | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 600 | let val sorts_list = map (map #2 o term_tvars) ts | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 601 | in Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list)) end; | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 602 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 603 | fun tfree_classes_of_terms ts = | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 604 | let val sorts_list = map (map #2 o term_tfrees) ts | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 605 | in Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list)) end; | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 606 | |
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 607 | (*fold type constructors*) | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 608 | fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x)) | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 609 | | fold_type_consts f T x = x; | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 610 | |
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 611 | val add_type_consts_in_type = fold_type_consts setinsert; | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 612 | |
| 21397 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 613 | (*Type constructors used to instantiate overloaded constants are the only ones needed.*) | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 614 | fun add_type_consts_in_term thy = | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 615 | let val const_typargs = Sign.const_typargs thy | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 616 | fun add_tcs (Const cT) x = fold add_type_consts_in_type (const_typargs cT) x | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 617 | | add_tcs (Abs (_, T, u)) x = add_tcs u x | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 618 | | add_tcs (t $ u) x = add_tcs t (add_tcs u x) | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 619 | | add_tcs _ x = x | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 620 | in add_tcs end | 
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 621 | |
| 21397 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 622 | fun type_consts_of_terms thy ts = | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 623 | Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty); | 
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 624 | |
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 625 | |
| 19194 | 626 | (***************************************************************) | 
| 627 | (* ATP invocation methods setup *) | |
| 628 | (***************************************************************) | |
| 629 | ||
| 21588 | 630 | fun cnf_hyps_thms ctxt = | 
| 20224 
9c40a144ee0e
moved basic assumption operations from structure ProofContext to Assumption;
 wenzelm parents: 
20131diff
changeset | 631 | let val ths = Assumption.prems_of ctxt | 
| 19617 | 632 | in fold (fold (insert Thm.eq_thm) o ResAxioms.skolem_thm) ths [] end; | 
| 19194 | 633 | |
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 634 | (*Translation mode can be auto-detected, or forced to be first-order or higher-order*) | 
| 19194 | 635 | datatype mode = Auto | Fol | Hol; | 
| 636 | ||
| 19450 
651d949776f8
exported linkup_logic_mode and changed the default setting
 paulson parents: 
19445diff
changeset | 637 | val linkup_logic_mode = ref Auto; | 
| 19205 
4ec788c69f82
Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
 paulson parents: 
19194diff
changeset | 638 | |
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 639 | (*Ensures that no higher-order theorems "leak out"*) | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 640 | fun restrict_to_logic thy logic cls = | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 641 | if is_fol_logic logic then filter (Meson.is_fol_term thy o prop_of o fst) cls | 
| 21588 | 642 | else cls; | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 643 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 644 | (**** Predicates to detect unwanted clauses (prolific or likely to cause unsoundness) ****) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 645 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 646 | (** Too general means, positive equality literal with a variable X as one operand, | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 647 | when X does not occur properly in the other operand. This rules out clearly | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 648 | inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **) | 
| 21588 | 649 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 650 | fun occurs ix = | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 651 | let fun occ(Var (jx,_)) = (ix=jx) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 652 | | occ(t1$t2) = occ t1 orelse occ t2 | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 653 | | occ(Abs(_,_,t)) = occ t | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 654 | | occ _ = false | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 655 | in occ end; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 656 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 657 | fun is_recordtype T = not (null (RecordPackage.dest_recTs T)); | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 658 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 659 | (*Unwanted equalities include | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 660 | (1) those between a variable that does not properly occur in the second operand, | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 661 | (2) those between a variable and a record, since these seem to be prolific "cases" thms | 
| 21588 | 662 | *) | 
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 663 | fun too_general_eqterms (Var (ix,T), t) = not (occurs ix t) orelse is_recordtype T | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 664 | | too_general_eqterms _ = false; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 665 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 666 | fun too_general_equality (Const ("op =", _) $ x $ y) =
 | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 667 | too_general_eqterms (x,y) orelse too_general_eqterms(y,x) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 668 | | too_general_equality _ = false; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 669 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 670 | (* tautologous? *) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 671 | fun is_taut (Const ("Trueprop", _) $ Const ("True", _)) = true
 | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 672 | | is_taut _ = false; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 673 | |
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 674 | (*True if the term contains a variable whose (atomic) type is in the given list.*) | 
| 21588 | 675 | fun has_typed_var tycons = | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 676 | let fun var_tycon (Var (_, Type(a,_))) = a mem_string tycons | 
| 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 677 | | var_tycon _ = false | 
| 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 678 | in exists var_tycon o term_vars end; | 
| 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 679 | |
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 680 | (*Clauses are forbidden to contain variables of these types. The typical reason is that | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 681 | they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=(). | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 682 | The resulting clause will have no type constraint, yielding false proofs. Even "bool" | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 683 | leads to many unsound proofs, though (obviously) only for higher-order problems.*) | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 684 | val unwanted_types = ref ["Product_Type.unit","bool"]; | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 685 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 686 | fun unwanted t = | 
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 687 | is_taut t orelse has_typed_var (!unwanted_types) t orelse | 
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 688 | forall too_general_equality (dest_disj t); | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 689 | |
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 690 | (*Clauses containing variables of type "unit" or "bool" are unlikely to be useful and | 
| 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 691 | likely to lead to unsound proofs.*) | 
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 692 | fun remove_unwanted_clauses cls = filter (not o unwanted o prop_of o fst) cls; | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 693 | |
| 23387 | 694 | fun tptp_writer logic = ResHolClause.tptp_write_file (logic=FOL); | 
| 19445 
da75577642a9
tidying; ATP options including CASC mode for Vampire
 paulson parents: 
19442diff
changeset | 695 | |
| 23387 | 696 | fun dfg_writer logic = ResHolClause.dfg_write_file (logic=FOL); | 
| 19194 | 697 | |
| 20022 | 698 | (*Called by the oracle-based methods declared in res_atp_methods.ML*) | 
| 19722 | 699 | fun write_subgoal_file dfg mode ctxt conjectures user_thms n = | 
| 21588 | 700 | let val conj_cls = make_clauses conjectures | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22724diff
changeset | 701 | |> ResAxioms.assume_abstract_list "subgoal" |> Meson.finish_cnf | 
| 21588 | 702 | val hyp_cls = cnf_hyps_thms ctxt | 
| 703 | val goal_cls = conj_cls@hyp_cls | |
| 704 | val goal_tms = map prop_of goal_cls | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 705 | val thy = ProofContext.theory_of ctxt | 
| 21588 | 706 | val logic = case mode of | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 707 | Auto => problem_logic_goals [goal_tms] | 
| 21588 | 708 | | Fol => FOL | 
| 709 | | Hol => HOL | |
| 710 | val (included_thms,user_rules) = get_clasimp_atp_lemmas ctxt user_thms | |
| 22865 | 711 | val cla_simp_atp_clauses = included_thms | 
| 21588 | 712 | |> ResAxioms.cnf_rules_pairs |> make_unique | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 713 | |> restrict_to_logic thy logic | 
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 714 | |> remove_unwanted_clauses | 
| 21588 | 715 | val user_cls = ResAxioms.cnf_rules_pairs user_rules | 
| 716 | val axclauses = make_unique (get_relevant_clauses thy cla_simp_atp_clauses user_cls goal_tms) | |
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 717 | val subs = tfree_classes_of_terms goal_tms | 
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 718 | and axtms = map (prop_of o #1) axclauses | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 719 | val supers = tvar_classes_of_terms axtms | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 720 | and tycons = type_consts_of_terms thy (goal_tms@axtms) | 
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 721 | (*TFrees in conjecture clauses; TVars in axiom clauses*) | 
| 22645 | 722 | val (supers',arity_clauses) = ResClause.make_arity_clauses thy tycons supers | 
| 723 | val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' | |
| 21588 | 724 | val writer = if dfg then dfg_writer else tptp_writer | 
| 725 | and file = atp_input_file() | |
| 726 | and user_lemmas_names = map #1 user_rules | |
| 19194 | 727 | in | 
| 21588 | 728 | writer logic goal_cls file (axclauses,classrel_clauses,arity_clauses) user_lemmas_names; | 
| 22130 | 729 | Output.debug (fn () => "Writing to " ^ file); | 
| 21588 | 730 | file | 
| 19194 | 731 | end; | 
| 732 | ||
| 733 | ||
| 734 | (**** remove tmp files ****) | |
| 21588 | 735 | fun cond_rm_tmp file = | 
| 22130 | 736 | if !Output.debugging orelse !destdir <> "" | 
| 737 | then Output.debug (fn () => "ATP input kept...") | |
| 20870 | 738 | else OS.FileSys.remove file; | 
| 19194 | 739 | |
| 740 | ||
| 741 | (****** setup ATPs as Isabelle methods ******) | |
| 742 | ||
| 21588 | 743 | fun atp_meth tac ths ctxt = | 
| 19194 | 744 | let val thy = ProofContext.theory_of ctxt | 
| 21588 | 745 | val _ = ResClause.init thy | 
| 746 | val _ = ResHolClause.init thy | |
| 747 | in Method.SIMPLE_METHOD' (tac ctxt ths) end; | |
| 19194 | 748 | |
| 749 | fun atp_method tac = Method.thms_ctxt_args (atp_meth tac); | |
| 750 | ||
| 751 | (***************************************************************) | |
| 752 | (* automatic ATP invocation *) | |
| 753 | (***************************************************************) | |
| 754 | ||
| 17306 | 755 | (* call prover with settings and problem file for the current subgoal *) | 
| 17764 | 756 | fun watcher_call_provers sign sg_terms (childin, childout, pid) = | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 757 | let | 
| 17422 | 758 | fun make_atp_list [] n = [] | 
| 17717 | 759 | | make_atp_list (sg_term::xs) n = | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 760 | let | 
| 17717 | 761 | val probfile = prob_pathname n | 
| 17690 
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
 paulson parents: 
17525diff
changeset | 762 | val time = Int.toString (!time_limit) | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 763 | in | 
| 22130 | 764 | Output.debug (fn () => "problem file in watcher_call_provers is " ^ probfile); | 
| 17764 | 765 | (*options are separated by Watcher.setting_sep, currently #"%"*) | 
| 17306 | 766 | if !prover = "spass" | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 767 | then | 
| 19445 
da75577642a9
tidying; ATP options including CASC mode for Vampire
 paulson parents: 
19442diff
changeset | 768 | let val spass = helper_path "SPASS_HOME" "SPASS" | 
| 
da75577642a9
tidying; ATP options including CASC mode for Vampire
 paulson parents: 
19442diff
changeset | 769 | val sopts = | 
| 
da75577642a9
tidying; ATP options including CASC mode for Vampire
 paulson parents: 
19442diff
changeset | 770 | "-Auto%-SOS=1%-PGiven=0%-PProblem=0%-Splits=0%-FullRed=0%-DocProof%-TimeLimit=" ^ time | 
| 21588 | 771 | in | 
| 19445 
da75577642a9
tidying; ATP options including CASC mode for Vampire
 paulson parents: 
19442diff
changeset | 772 |                   ("spass", spass, sopts, probfile) :: make_atp_list xs (n+1)
 | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 773 | end | 
| 17306 | 774 | else if !prover = "vampire" | 
| 21588 | 775 | then | 
| 17819 | 776 | let val vampire = helper_path "VAMPIRE_HOME" "vampire" | 
| 21132 | 777 | val vopts = "--mode casc%-t " ^ time (*what about -m 100000?*) | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 778 | in | 
| 19445 
da75577642a9
tidying; ATP options including CASC mode for Vampire
 paulson parents: 
19442diff
changeset | 779 |                   ("vampire", vampire, vopts, probfile) :: make_atp_list xs (n+1)
 | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 780 | end | 
| 21588 | 781 | else if !prover = "E" | 
| 782 | then | |
| 783 | let val Eprover = helper_path "E_HOME" "eproof" | |
| 784 | in | |
| 785 |                   ("E", Eprover,
 | |
| 21900 | 786 | "--tstp-in%--tstp-out%-l5%-xAutoDev%-tAutoDev%--silent%--cpu-limit=" ^ time, probfile) :: | 
| 21588 | 787 | make_atp_list xs (n+1) | 
| 788 | end | |
| 789 |              else error ("Invalid prover name: " ^ !prover)
 | |
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 790 | end | 
| 15452 | 791 | |
| 17422 | 792 | val atp_list = make_atp_list sg_terms 1 | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 793 | in | 
| 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 794 | Watcher.callResProvers(childout,atp_list); | 
| 22130 | 795 | Output.debug (fn () => "Sent commands to watcher!") | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 796 | end | 
| 21588 | 797 | |
| 22193 | 798 | (*For debugging the generated set of theorem names*) | 
| 21888 | 799 | fun trace_vector fname = | 
| 22193 | 800 | let val path = File.explode_platform_path (fname ^ "_thm_names") | 
| 21888 | 801 | in Vector.app (File.append path o (fn s => s ^ "\n")) end; | 
| 16357 | 802 | |
| 20022 | 803 | (*We write out problem files for each subgoal. Argument probfile generates filenames, | 
| 18986 | 804 | and allows the suppression of the suffix "_1" in problem-generation mode. | 
| 805 | FIXME: does not cope with &&, and it isn't easy because one could have multiple | |
| 806 | subgoals, each involving &&.*) | |
| 20022 | 807 | fun write_problem_files probfile (ctxt,th) = | 
| 18753 
aa82bd41555d
ResClasimp.get_clasimp_lemmas  now takes all subgoals rather than only the first
 paulson parents: 
18700diff
changeset | 808 | let val goals = Thm.prems_of th | 
| 22130 | 809 | val _ = Output.debug (fn () => "number of subgoals = " ^ Int.toString (length goals)) | 
| 17717 | 810 | val thy = ProofContext.theory_of ctxt | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 811 | fun get_neg_subgoals [] _ = [] | 
| 22865 | 812 | | get_neg_subgoals (gl::gls) n = #1 (ResAxioms.neg_conjecture_clauses th n) :: | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21980diff
changeset | 813 | get_neg_subgoals gls (n+1) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 814 | val goal_cls = get_neg_subgoals goals 1 | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 815 | val logic = case !linkup_logic_mode of | 
| 23387 | 816 | Auto => problem_logic_goals (map (map prop_of) goal_cls) | 
| 21588 | 817 | | Fol => FOL | 
| 818 | | Hol => HOL | |
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 819 | val (included_thms,white_thms) = get_clasimp_atp_lemmas ctxt [] | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 820 | val included_cls = included_thms |> ResAxioms.cnf_rules_pairs |> make_unique | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 821 | |> restrict_to_logic thy logic | 
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 822 | |> remove_unwanted_clauses | 
| 22130 | 823 | val _ = Output.debug (fn () => "included clauses = " ^ Int.toString(length included_cls)) | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 824 | val white_cls = ResAxioms.cnf_rules_pairs white_thms | 
| 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 825 | (*clauses relevant to goal gl*) | 
| 21563 
b4718f2c15f0
Goals in clause form are sent to the relevance filter.
 mengj parents: 
21509diff
changeset | 826 | val axcls_list = map (fn ngcls => get_relevant_clauses thy included_cls white_cls (map prop_of ngcls)) goal_cls | 
| 22130 | 827 | val _ = app (fn axcls => Output.debug (fn () => "filtered clauses = " ^ Int.toString(length axcls))) | 
| 21070 | 828 | axcls_list | 
| 21588 | 829 | val writer = if !prover = "spass" then dfg_writer else tptp_writer | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 830 | fun write_all [] [] _ = [] | 
| 21588 | 831 | | write_all (ccls::ccls_list) (axcls::axcls_list) k = | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 832 | let val fname = probfile k | 
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 833 | val _ = Output.debug (fn () => "About to write file " ^ fname) | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 834 | val axcls = make_unique axcls | 
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 835 | val _ = Output.debug (fn () => "Conjecture Clauses (before duplicate removal)") | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 836 | val _ = app (fn th => Output.debug (fn _ => string_of_thm th)) ccls | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 837 | val ccls = subtract_cls ccls axcls | 
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 838 | val _ = Output.debug (fn () => "Conjecture Clauses (AFTER duplicate removal)") | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 839 | val _ = app (fn th => Output.debug (fn _ => string_of_thm th)) ccls | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 840 | val ccltms = map prop_of ccls | 
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 841 | and axtms = map (prop_of o #1) axcls | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 842 | val subs = tfree_classes_of_terms ccltms | 
| 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 843 | and supers = tvar_classes_of_terms axtms | 
| 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 844 | and tycons = type_consts_of_terms thy (ccltms@axtms) | 
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 845 | (*TFrees in conjecture clauses; TVars in axiom clauses*) | 
| 22645 | 846 | val (supers',arity_clauses) = ResClause.make_arity_clauses thy tycons supers | 
| 847 | val _ = Output.debug (fn () => "arity clauses = " ^ Int.toString (length arity_clauses)) | |
| 848 | val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' | |
| 22130 | 849 | val _ = Output.debug (fn () => "classrel clauses = " ^ Int.toString (length classrel_clauses)) | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 850 | val clnames = writer logic ccls fname (axcls,classrel_clauses,arity_clauses) [] | 
| 21888 | 851 | val thm_names = Vector.fromList clnames | 
| 22193 | 852 | val _ = if !Output.debugging then trace_vector fname thm_names else () | 
| 20870 | 853 | in (thm_names,fname) :: write_all ccls_list axcls_list (k+1) end | 
| 854 | val (thm_names_list, filenames) = ListPair.unzip (write_all goal_cls axcls_list 1) | |
| 19194 | 855 | in | 
| 20870 | 856 | (filenames, thm_names_list) | 
| 19194 | 857 | end; | 
| 15644 | 858 | |
| 21588 | 859 | val last_watcher_pid = ref (NONE : (TextIO.instream * TextIO.outstream * | 
| 17775 | 860 | Posix.Process.pid * string list) option); | 
| 861 | ||
| 862 | fun kill_last_watcher () = | |
| 21588 | 863 | (case !last_watcher_pid of | 
| 17775 | 864 | NONE => () | 
| 21588 | 865 | | SOME (_, _, pid, files) => | 
| 22130 | 866 | (Output.debug (fn () => "Killing old watcher, pid = " ^ string_of_pid pid); | 
| 21588 | 867 | Watcher.killWatcher pid; | 
| 868 | ignore (map (try cond_rm_tmp) files))) | |
| 22130 | 869 | handle OS.SysErr _ => Output.debug (fn () => "Attempt to kill watcher failed"); | 
| 17525 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 paulson parents: 
17502diff
changeset | 870 | |
| 21980 
d22f7e3c5ad9
x-symbol is no longer switched off in the ATP linkup
 paulson parents: 
21900diff
changeset | 871 | (*writes out the current problems and calls ATPs*) | 
| 
d22f7e3c5ad9
x-symbol is no longer switched off in the ATP linkup
 paulson parents: 
21900diff
changeset | 872 | fun isar_atp (ctxt, th) = | 
| 17717 | 873 | if Thm.no_prems th then () | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 874 | else | 
| 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 875 | let | 
| 17775 | 876 | val _ = kill_last_watcher() | 
| 20870 | 877 | val (files,thm_names_list) = write_problem_files prob_pathname (ctxt,th) | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 878 | val (childin, childout, pid) = Watcher.createWatcher (ctxt, th, thm_names_list) | 
| 15608 | 879 | in | 
| 17772 
818cec5f82a4
major simplification: removal of the goalstring argument
 paulson parents: 
17764diff
changeset | 880 | last_watcher_pid := SOME (childin, childout, pid, files); | 
| 22130 | 881 | Output.debug (fn () => "problem files: " ^ space_implode ", " files); | 
| 882 | Output.debug (fn () => "pid: " ^ string_of_pid pid); | |
| 22578 | 883 | watcher_call_provers (Thm.theory_of_thm th) (Thm.prems_of th) (childin, childout, pid) | 
| 20954 | 884 | end; | 
| 885 | ||
| 886 | (*For ML scripts, and primarily, for debugging*) | |
| 21588 | 887 | fun callatp () = | 
| 20954 | 888 | let val th = topthm() | 
| 889 | val ctxt = ProofContext.init (theory_of_thm th) | |
| 21980 
d22f7e3c5ad9
x-symbol is no longer switched off in the ATP linkup
 paulson parents: 
21900diff
changeset | 890 | in isar_atp (ctxt, th) end; | 
| 15608 | 891 | |
| 21588 | 892 | val isar_atp_writeonly = setmp print_mode [] | 
| 17717 | 893 | (fn (ctxt,th) => | 
| 894 | if Thm.no_prems th then () | |
| 21588 | 895 | else | 
| 896 | let val probfile = if Thm.nprems_of th = 1 then probfile_nosuffix | |
| 897 | else prob_pathname | |
| 20022 | 898 | in ignore (write_problem_files probfile (ctxt,th)) end); | 
| 15452 | 899 | |
| 16357 | 900 | |
| 22865 | 901 | (** the Isar toplevel command **) | 
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 902 | |
| 22865 | 903 | fun sledgehammer state = | 
| 904 | let | |
| 905 | val (ctxt, (_, goal)) = Proof.get_goal (Toplevel.proof_of state); | |
| 906 | val thy = ProofContext.theory_of ctxt; | |
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 907 | in | 
| 22130 | 908 | Output.debug (fn () => "subgoals in isar_atp:\n" ^ | 
| 21588 | 909 | Pretty.string_of (ProofContext.pretty_term ctxt | 
| 910 | (Logic.mk_conjunction_list (Thm.prems_of goal)))); | |
| 22130 | 911 | Output.debug (fn () => "current theory: " ^ Context.theory_name thy); | 
| 20419 | 912 | inc hook_count; | 
| 22130 | 913 | Output.debug (fn () => "in hook for time: " ^ Int.toString (!hook_count)); | 
| 16925 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 paulson parents: 
16904diff
changeset | 914 | ResClause.init thy; | 
| 19194 | 915 | ResHolClause.init thy; | 
| 21132 | 916 | if !time_limit > 0 then isar_atp (ctxt, goal) | 
| 22865 | 917 |     else (warning ("Writing problem file only: " ^ !problem_name);
 | 
| 22193 | 918 | isar_atp_writeonly (ctxt, goal)) | 
| 19205 
4ec788c69f82
Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
 paulson parents: 
19194diff
changeset | 919 | end; | 
| 
4ec788c69f82
Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
 paulson parents: 
19194diff
changeset | 920 | |
| 22865 | 921 | val _ = OuterSyntax.add_parsers | 
| 922 | [OuterSyntax.command "sledgehammer" "call automatic theorem provers" OuterKeyword.diag | |
| 923 | (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_proof o Toplevel.keep sledgehammer))]; | |
| 17091 | 924 | |
| 15347 | 925 | end; |