author | paulson |
Wed, 15 Aug 2007 12:52:56 +0200 | |
changeset 24286 | 7619080e49f0 |
parent 24215 | 5458fbf18276 |
child 24287 | c857dac06da6 |
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:
16767
diff
changeset
|
8 |
signature RES_ATP = |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
9 |
sig |
17306 | 10 |
val prover: string ref |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
11 |
val destdir: string ref |
17849 | 12 |
val helper_path: string -> string -> string |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
13 |
val problem_name: string ref |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
14 |
val time_limit: int ref |
21224 | 15 |
val set_prover: string -> unit |
21588 | 16 |
|
19194 | 17 |
datatype mode = Auto | Fol | Hol |
19450
651d949776f8
exported linkup_logic_mode and changed the default setting
paulson
parents:
19445
diff
changeset
|
18 |
val linkup_logic_mode : mode ref |
19722 | 19 |
val write_subgoal_file: bool -> mode -> Proof.context -> thm list -> thm list -> int -> string |
20289 | 20 |
val atp_method: (Proof.context -> thm list -> int -> tactic) -> |
21 |
Method.src -> Proof.context -> Proof.method |
|
19194 | 22 |
val cond_rm_tmp: string -> unit |
20246
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
23 |
val include_all: bool ref |
19194 | 24 |
val run_relevance_filter: bool ref |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
25 |
val run_blacklist_filter: bool ref |
19894 | 26 |
val add_all : unit -> unit |
19227
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
27 |
val add_claset : unit -> unit |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
28 |
val add_simpset : unit -> unit |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
29 |
val add_clasimp : unit -> unit |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
30 |
val add_atpset : unit -> unit |
19894 | 31 |
val rm_all : unit -> unit |
19227
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
32 |
val rm_claset : unit -> unit |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
33 |
val rm_simpset : unit -> unit |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
34 |
val rm_atpset : unit -> unit |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
35 |
val rm_clasimp : unit -> unit |
21311 | 36 |
val is_fol_thms : thm list -> bool |
22989 | 37 |
val tvar_classes_of_terms : term list -> string list |
38 |
val tfree_classes_of_terms : term list -> string list |
|
39 |
val type_consts_of_terms : theory -> term list -> string list |
|
15347 | 40 |
end; |
41 |
||
22865 | 42 |
structure ResAtp: RES_ATP = |
15347 | 43 |
struct |
44 |
||
22130 | 45 |
fun timestamp s = Output.debug (fn () => ("At " ^ Time.toString (Time.now()) ^ ": " ^ s)); |
21070 | 46 |
|
19194 | 47 |
(********************************************************************) |
48 |
(* some settings for both background automatic ATP calling procedure*) |
|
49 |
(* and also explicit ATP invocation methods *) |
|
50 |
(********************************************************************) |
|
51 |
||
52 |
(*** background linkup ***) |
|
21224 | 53 |
val time_limit = ref 60; |
21588 | 54 |
val prover = ref ""; |
21224 | 55 |
|
21588 | 56 |
fun set_prover atp = |
21224 | 57 |
case String.map Char.toLower atp of |
21588 | 58 |
"e" => |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
59 |
(ReduceAxiomsN.max_new := 100; |
21224 | 60 |
ReduceAxiomsN.theory_const := false; |
61 |
prover := "E") |
|
21588 | 62 |
| "spass" => |
21224 | 63 |
(ReduceAxiomsN.max_new := 40; |
64 |
ReduceAxiomsN.theory_const := true; |
|
65 |
prover := "spass") |
|
21588 | 66 |
| "vampire" => |
21224 | 67 |
(ReduceAxiomsN.max_new := 60; |
68 |
ReduceAxiomsN.theory_const := false; |
|
69 |
prover := "vampire") |
|
70 |
| _ => error ("No such prover: " ^ atp); |
|
71 |
||
72 |
val _ = set_prover "E"; (* use E as the default prover *) |
|
73 |
||
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
74 |
val destdir = ref ""; (*Empty means write files to /tmp*) |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
75 |
val problem_name = ref "prob"; |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
76 |
|
17819 | 77 |
(*Return the path to a "helper" like SPASS or tptp2X, first checking that |
21588 | 78 |
it exists. FIXME: modify to use Path primitives and move to some central place.*) |
17819 | 79 |
fun helper_path evar base = |
80 |
case getenv evar of |
|
81 |
"" => error ("Isabelle environment variable " ^ evar ^ " not defined") |
|
21588 | 82 |
| home => |
17819 | 83 |
let val path = home ^ "/" ^ base |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21790
diff
changeset
|
84 |
in if File.exists (File.explode_platform_path path) then path |
21588 | 85 |
else error ("Could not find the file " ^ path) |
86 |
end; |
|
17819 | 87 |
|
21588 | 88 |
fun probfile_nosuffix _ = |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
89 |
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:
21790
diff
changeset
|
90 |
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:
17435
diff
changeset
|
91 |
then !destdir ^ "/" ^ !problem_name |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
92 |
else error ("No such directory: " ^ !destdir); |
15644 | 93 |
|
17717 | 94 |
fun prob_pathname n = probfile_nosuffix n ^ "_" ^ Int.toString n; |
95 |
||
19194 | 96 |
fun atp_input_file () = |
21588 | 97 |
let val file = !problem_name |
19194 | 98 |
in |
21588 | 99 |
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:
21790
diff
changeset
|
100 |
else if File.exists (File.explode_platform_path (!destdir)) |
21588 | 101 |
then !destdir ^ "/" ^ file |
102 |
else error ("No such directory: " ^ !destdir) |
|
19194 | 103 |
end; |
104 |
||
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
105 |
val include_all = ref true; |
19194 | 106 |
val include_simpset = ref false; |
21588 | 107 |
val include_claset = ref false; |
19194 | 108 |
val include_atpset = ref true; |
20246
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
109 |
|
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
110 |
(*Tests show that follow_defs gives VERY poor results with "include_all"*) |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
111 |
fun add_all() = (include_all:=true; ReduceAxiomsN.follow_defs := false); |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
112 |
fun rm_all() = include_all:=false; |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
113 |
|
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
114 |
fun add_simpset() = include_simpset:=true; |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
115 |
fun rm_simpset() = include_simpset:=false; |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
116 |
|
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
117 |
fun add_claset() = include_claset:=true; |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
118 |
fun rm_claset() = include_claset:=false; |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
119 |
|
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
120 |
fun add_clasimp() = (include_simpset:=true;include_claset:=true); |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
121 |
fun rm_clasimp() = (include_simpset:=false;include_claset:=false); |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
122 |
|
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
123 |
fun add_atpset() = include_atpset:=true; |
fdfe7399e057
"all theorems" mode forces definition-expansion off.
paulson
parents:
20224
diff
changeset
|
124 |
fun rm_atpset() = include_atpset:=false; |
19194 | 125 |
|
126 |
||
127 |
(**** relevance filter ****) |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
128 |
val run_relevance_filter = ReduceAxiomsN.run_relevance_filter; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
129 |
val run_blacklist_filter = ref true; |
19194 | 130 |
|
131 |
(******************************************************************) |
|
132 |
(* detect whether a given problem (clauses) is FOL/HOL/HOLC/HOLCS *) |
|
133 |
(******************************************************************) |
|
134 |
||
135 |
datatype logic = FOL | HOL | HOLC | HOLCS; |
|
136 |
||
137 |
fun string_of_logic FOL = "FOL" |
|
138 |
| string_of_logic HOL = "HOL" |
|
139 |
| string_of_logic HOLC = "HOLC" |
|
140 |
| string_of_logic HOLCS = "HOLCS"; |
|
141 |
||
142 |
fun is_fol_logic FOL = true |
|
143 |
| is_fol_logic _ = false |
|
144 |
||
145 |
(*HOLCS will not occur here*) |
|
146 |
fun upgrade_lg HOLC _ = HOLC |
|
147 |
| upgrade_lg HOL HOLC = HOLC |
|
148 |
| upgrade_lg HOL _ = HOL |
|
21588 | 149 |
| upgrade_lg FOL lg = lg; |
19194 | 150 |
|
151 |
(* check types *) |
|
19451 | 152 |
fun has_bool_hfn (Type("bool",_)) = true |
153 |
| has_bool_hfn (Type("fun",_)) = true |
|
154 |
| has_bool_hfn (Type(_, Ts)) = exists has_bool_hfn Ts |
|
155 |
| has_bool_hfn _ = false; |
|
19194 | 156 |
|
19451 | 157 |
fun is_hol_fn tp = |
19194 | 158 |
let val (targs,tr) = strip_type tp |
159 |
in |
|
21588 | 160 |
exists (has_bool_hfn) (tr::targs) |
19194 | 161 |
end; |
162 |
||
19451 | 163 |
fun is_hol_pred tp = |
164 |
let val (targs,tr) = strip_type tp |
|
165 |
in |
|
21588 | 166 |
exists (has_bool_hfn) targs |
19451 | 167 |
end; |
19194 | 168 |
|
169 |
exception FN_LG of term; |
|
170 |
||
21588 | 171 |
fun fn_lg (t as Const(f,tp)) (lg,seen) = |
172 |
if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) |
|
173 |
| fn_lg (t as Free(f,tp)) (lg,seen) = |
|
174 |
if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) |
|
19194 | 175 |
| fn_lg (t as Var(f,tp)) (lg,seen) = |
20854 | 176 |
if is_hol_fn tp then (upgrade_lg HOL lg,insert (op =) t seen) else (lg,insert (op =) t seen) |
177 |
| fn_lg (t as Abs(_,_,_)) (lg,seen) = (upgrade_lg HOLC lg,insert (op =) t seen) |
|
21588 | 178 |
| fn_lg f _ = raise FN_LG(f); |
19194 | 179 |
|
180 |
||
181 |
fun term_lg [] (lg,seen) = (lg,seen) |
|
182 |
| term_lg (tm::tms) (FOL,seen) = |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
183 |
let val (f,args) = strip_comb tm |
21588 | 184 |
val (lg',seen') = if f mem seen then (FOL,seen) |
185 |
else fn_lg f (FOL,seen) |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
186 |
in |
21588 | 187 |
if is_fol_logic lg' then () |
22130 | 188 |
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:
20446
diff
changeset
|
189 |
term_lg (args@tms) (lg',seen') |
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
190 |
end |
19194 | 191 |
| term_lg _ (lg,seen) = (lg,seen) |
192 |
||
193 |
exception PRED_LG of term; |
|
194 |
||
21588 | 195 |
fun pred_lg (t as Const(P,tp)) (lg,seen)= |
196 |
if is_hol_pred tp then (upgrade_lg HOL lg, insert (op =) t seen) |
|
197 |
else (lg,insert (op =) t seen) |
|
19194 | 198 |
| pred_lg (t as Free(P,tp)) (lg,seen) = |
21588 | 199 |
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:
21431
diff
changeset
|
200 |
else (lg,insert (op =) t seen) |
20854 | 201 |
| pred_lg (t as Var(_,_)) (lg,seen) = (upgrade_lg HOL lg, insert (op =) t seen) |
19194 | 202 |
| pred_lg P _ = raise PRED_LG(P); |
203 |
||
204 |
||
205 |
fun lit_lg (Const("Not",_) $ P) (lg,seen) = lit_lg P (lg,seen) |
|
206 |
| lit_lg P (lg,seen) = |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
207 |
let val (pred,args) = strip_comb P |
21588 | 208 |
val (lg',seen') = if pred mem seen then (lg,seen) |
209 |
else pred_lg pred (lg,seen) |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
210 |
in |
21588 | 211 |
if is_fol_logic lg' then () |
22130 | 212 |
else Output.debug (fn () => ("Found a HOL predicate: " ^ Display.raw_string_of_term pred)); |
21588 | 213 |
term_lg args (lg',seen') |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
214 |
end; |
19194 | 215 |
|
216 |
fun lits_lg [] (lg,seen) = (lg,seen) |
|
217 |
| lits_lg (lit::lits) (FOL,seen) = |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
218 |
let val (lg,seen') = lit_lg lit (FOL,seen) |
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
219 |
in |
21588 | 220 |
if is_fol_logic lg then () |
22130 | 221 |
else Output.debug (fn () => ("Found a HOL literal: " ^ Display.raw_string_of_term lit)); |
21588 | 222 |
lits_lg lits (lg,seen') |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
223 |
end |
19194 | 224 |
| lits_lg lits (lg,seen) = (lg,seen); |
225 |
||
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
226 |
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:
21431
diff
changeset
|
227 |
| 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:
19205
diff
changeset
|
228 |
| dest_disj_aux t disjs = t::disjs; |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
229 |
|
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
230 |
fun dest_disj t = dest_disj_aux t []; |
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
231 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
232 |
fun logic_of_clause tm = lits_lg (dest_disj tm); |
19194 | 233 |
|
234 |
fun logic_of_clauses [] (lg,seen) = (lg,seen) |
|
235 |
| logic_of_clauses (cls::clss) (FOL,seen) = |
|
19227
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
236 |
let val (lg,seen') = logic_of_clause cls (FOL,seen) |
21588 | 237 |
val _ = |
19641
f1de44e61ec1
replaced low-level Term.str_of by Display.raw_string_of_term (should actually use Display.string_of_term);
wenzelm
parents:
19617
diff
changeset
|
238 |
if is_fol_logic lg then () |
22130 | 239 |
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:
19205
diff
changeset
|
240 |
in |
21588 | 241 |
logic_of_clauses clss (lg,seen') |
19227
d15f2baa7ecc
Added more functions to the signature and tidied up some functions.
mengj
parents:
19205
diff
changeset
|
242 |
end |
19194 | 243 |
| logic_of_clauses (cls::clss) (lg,seen) = (lg,seen); |
244 |
||
245 |
fun problem_logic_goals_aux [] (lg,seen) = lg |
|
21588 | 246 |
| problem_logic_goals_aux (subgoal::subgoals) (lg,seen) = |
19194 | 247 |
problem_logic_goals_aux subgoals (logic_of_clauses subgoal (lg,seen)); |
21588 | 248 |
|
19194 | 249 |
fun problem_logic_goals subgoals = problem_logic_goals_aux subgoals (FOL,[]); |
250 |
||
21311 | 251 |
fun is_fol_thms ths = ((fst(logic_of_clauses (map prop_of ths) (FOL,[]))) = FOL); |
252 |
||
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
253 |
(***************************************************************) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
254 |
(* Retrieving and filtering lemmas *) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
255 |
(***************************************************************) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
256 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
257 |
(*** white list and black list of lemmas ***) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
258 |
|
24215 | 259 |
(*The rule subsetI is frequently omitted by the relevance filter. This could be theory data |
260 |
or identified with ATPset (which however is too big currently)*) |
|
261 |
val whitelist = [subsetI]; |
|
21588 | 262 |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
263 |
(*** retrieve lemmas from clasimpset and atpset, may filter them ***) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
264 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
265 |
(*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:
19746
diff
changeset
|
266 |
|
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
267 |
fun setinsert (x,s) = Symtab.update (x,()) s; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
268 |
|
20757
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
paulson
parents:
20661
diff
changeset
|
269 |
(*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:
21588
diff
changeset
|
270 |
"Accessible_Part.acc.defs", as these are definitions arising from packages.*) |
20757
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
paulson
parents:
20661
diff
changeset
|
271 |
fun is_package_def a = |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21790
diff
changeset
|
272 |
let val names = NameSpace.explode a |
21690
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
273 |
in |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
274 |
length names > 2 andalso |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
275 |
not (hd names = "local") andalso |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
276 |
String.isSuffix "_def" a orelse String.isSuffix "_defs" a |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
277 |
end; |
20757
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
paulson
parents:
20661
diff
changeset
|
278 |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
279 |
(** 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:
19746
diff
changeset
|
280 |
val xor_words = List.foldl Word.xorb 0w0; |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
281 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
282 |
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:
20643
diff
changeset
|
283 |
| 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:
19746
diff
changeset
|
284 |
| hashw_term ((Var(_,_)), w) = w |
20661
46832fee1215
Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
paulson
parents:
20643
diff
changeset
|
285 |
| 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:
19746
diff
changeset
|
286 |
| hashw_term ((Abs(_,_,t)), w) = hashw_term (t, w) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
287 |
| 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:
19746
diff
changeset
|
288 |
|
21070 | 289 |
fun hash_literal (Const("Not",_)$P) = Word.notb(hashw_term(P,0w0)) |
290 |
| hash_literal P = hashw_term(P,0w0); |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
291 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
292 |
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:
19746
diff
changeset
|
293 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
294 |
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:
20446
diff
changeset
|
295 |
|
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
296 |
exception HASH_CLAUSE; |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
297 |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
298 |
(*Create a hash table for clauses, of the given size*) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
299 |
fun mk_clause_table n = |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
300 |
Polyhash.mkTable (hash_term o prop_of, equal_thm) |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
301 |
(n, HASH_CLAUSE); |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
302 |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
303 |
(*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:
20854
diff
changeset
|
304 |
(thm * (string * int)) tuples. The theorems are hashed into the table. *) |
21588 | 305 |
fun make_unique xs = |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
306 |
let val ht = mk_clause_table 7000 |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
307 |
in |
22130 | 308 |
Output.debug (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses")); |
21588 | 309 |
app (ignore o Polyhash.peekInsert ht) xs; |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
310 |
Polyhash.listItems ht |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
311 |
end; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
312 |
|
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
313 |
(*Remove existing axiom clauses from the conjecture clauses, as this can dramatically |
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
314 |
boost an ATP's performance (for some reason)*) |
21588 | 315 |
fun subtract_cls c_clauses ax_clauses = |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
316 |
let val ht = mk_clause_table 2200 |
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
317 |
fun known x = isSome (Polyhash.peek ht x) |
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
318 |
in |
21588 | 319 |
app (ignore o Polyhash.peekInsert ht) ax_clauses; |
320 |
filter (not o known) c_clauses |
|
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
321 |
end; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
322 |
|
21588 | 323 |
(*Filter axiom clauses, but keep supplied clauses and clauses in whitelist. |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
324 |
Duplicates are removed later.*) |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
325 |
fun get_relevant_clauses thy cls_thms white_cls goals = |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
326 |
white_cls @ (ReduceAxiomsN.relevance_filter thy cls_thms goals); |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
327 |
|
21224 | 328 |
fun all_valid_thms ctxt = |
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
329 |
let |
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
330 |
fun blacklisted s = !run_blacklist_filter andalso is_package_def s |
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
331 |
|
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
332 |
fun extern_valid space (name, ths) = |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
333 |
let |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
334 |
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:
22217
diff
changeset
|
335 |
val xname = NameSpace.extern space name; |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
336 |
in |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
337 |
if blacklisted name then I |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
338 |
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:
22217
diff
changeset
|
339 |
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:
22217
diff
changeset
|
340 |
else I |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
341 |
end; |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
342 |
val thy = ProofContext.theory_of ctxt; |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
343 |
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:
22217
diff
changeset
|
344 |
|
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
345 |
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:
22217
diff
changeset
|
346 |
in |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
347 |
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:
22217
diff
changeset
|
348 |
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:
22217
diff
changeset
|
349 |
(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:
22217
diff
changeset
|
350 |
end; |
21224 | 351 |
|
21588 | 352 |
fun multi_name a (th, (n,pairs)) = |
21224 | 353 |
(n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs) |
354 |
||
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
355 |
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:
22217
diff
changeset
|
356 |
| 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:
22217
diff
changeset
|
357 |
| 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:
21397
diff
changeset
|
358 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
359 |
val multi_base_blacklist = |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
360 |
["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm"]; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
361 |
|
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
362 |
(*Ignore blacklisted basenames*) |
21588 | 363 |
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:
22217
diff
changeset
|
364 |
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:
22217
diff
changeset
|
365 |
else add_single_names ((a, ths), pairs); |
21224 | 366 |
|
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
367 |
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:
21243
diff
changeset
|
368 |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
369 |
(*The single theorems go BEFORE the multiple ones. Blacklist is applied to all.*) |
21588 | 370 |
fun name_thm_pairs ctxt = |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
371 |
let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt) |
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
372 |
val ht = mk_clause_table 900 (*ht for blacklisted theorems*) |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
373 |
fun blacklisted x = !run_blacklist_filter andalso isSome (Polyhash.peek ht x) |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
374 |
in |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
375 |
app (fn th => ignore (Polyhash.peekInsert ht (th,()))) (ResBlacklist.get ctxt); |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
376 |
filter (not o blacklisted o #2) |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
377 |
(foldl add_single_names (foldl add_multi_names [] mults) singles) |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
378 |
end; |
21224 | 379 |
|
380 |
fun check_named ("",th) = (warning ("No name for theorem " ^ string_of_thm th); false) |
|
381 |
| check_named (_,th) = true; |
|
19894 | 382 |
|
22193 | 383 |
fun display_thm (name,th) = Output.debug (fn () => name ^ ": " ^ string_of_thm th); |
384 |
||
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
385 |
(* get lemmas from claset, simpset, atpset and extra supplied rules *) |
21588 | 386 |
fun get_clasimp_atp_lemmas ctxt user_thms = |
19894 | 387 |
let val included_thms = |
21588 | 388 |
if !include_all |
389 |
then (tap (fn ths => Output.debug |
|
22130 | 390 |
(fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems"))) |
21588 | 391 |
(name_thm_pairs ctxt)) |
392 |
else |
|
393 |
let val claset_thms = |
|
394 |
if !include_claset then ResAxioms.claset_rules_of ctxt |
|
395 |
else [] |
|
396 |
val simpset_thms = |
|
397 |
if !include_simpset then ResAxioms.simpset_rules_of ctxt |
|
398 |
else [] |
|
399 |
val atpset_thms = |
|
400 |
if !include_atpset then ResAxioms.atpset_rules_of ctxt |
|
401 |
else [] |
|
22193 | 402 |
val _ = (Output.debug (fn () => "ATP theorems: "); app display_thm atpset_thms) |
21588 | 403 |
in claset_thms @ simpset_thms @ atpset_thms end |
404 |
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:
22217
diff
changeset
|
405 |
(map ResAxioms.pairname |
24215 | 406 |
(if null user_thms then whitelist else user_thms)) |
19894 | 407 |
in |
21224 | 408 |
(filter check_named included_thms, user_rules) |
19894 | 409 |
end; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
410 |
|
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
411 |
(***************************************************************) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
412 |
(* Type Classes Present in the Axiom or Conjecture Clauses *) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
413 |
(***************************************************************) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
414 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
415 |
fun add_classes (sorts, cset) = foldl setinsert cset (List.concat sorts); |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
416 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
417 |
(*Remove this trivial type class*) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
418 |
fun delete_type cset = Symtab.delete_safe "HOL.type" cset; |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
419 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
420 |
fun tvar_classes_of_terms ts = |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
421 |
let val sorts_list = map (map #2 o term_tvars) ts |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
422 |
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:
21243
diff
changeset
|
423 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
424 |
fun tfree_classes_of_terms ts = |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
425 |
let val sorts_list = map (map #2 o term_tfrees) ts |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
426 |
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:
20457
diff
changeset
|
427 |
|
21373
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
428 |
(*fold type constructors*) |
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
429 |
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:
21311
diff
changeset
|
430 |
| fold_type_consts f T x = x; |
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
431 |
|
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
432 |
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:
21311
diff
changeset
|
433 |
|
21397
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
434 |
(*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:
21373
diff
changeset
|
435 |
fun add_type_consts_in_term thy = |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
436 |
let val const_typargs = Sign.const_typargs thy |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
437 |
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:
21373
diff
changeset
|
438 |
| 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:
21373
diff
changeset
|
439 |
| 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:
21373
diff
changeset
|
440 |
| add_tcs _ x = x |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
441 |
in add_tcs end |
21373
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
442 |
|
21397
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
443 |
fun type_consts_of_terms thy ts = |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
444 |
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:
21311
diff
changeset
|
445 |
|
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
446 |
|
19194 | 447 |
(***************************************************************) |
448 |
(* ATP invocation methods setup *) |
|
449 |
(***************************************************************) |
|
450 |
||
21588 | 451 |
fun cnf_hyps_thms ctxt = |
20224
9c40a144ee0e
moved basic assumption operations from structure ProofContext to Assumption;
wenzelm
parents:
20131
diff
changeset
|
452 |
let val ths = Assumption.prems_of ctxt |
19617 | 453 |
in fold (fold (insert Thm.eq_thm) o ResAxioms.skolem_thm) ths [] end; |
19194 | 454 |
|
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
455 |
(*Translation mode can be auto-detected, or forced to be first-order or higher-order*) |
19194 | 456 |
datatype mode = Auto | Fol | Hol; |
457 |
||
19450
651d949776f8
exported linkup_logic_mode and changed the default setting
paulson
parents:
19445
diff
changeset
|
458 |
val linkup_logic_mode = ref Auto; |
19205
4ec788c69f82
Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents:
19194
diff
changeset
|
459 |
|
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
460 |
(*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:
22217
diff
changeset
|
461 |
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:
22217
diff
changeset
|
462 |
if is_fol_logic logic then filter (Meson.is_fol_term thy o prop_of o fst) cls |
21588 | 463 |
else cls; |
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
464 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
465 |
(**** Predicates to detect unwanted clauses (prolific or likely to cause unsoundness) ****) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
466 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
467 |
(** 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:
21431
diff
changeset
|
468 |
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:
21431
diff
changeset
|
469 |
inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **) |
21588 | 470 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
471 |
fun occurs ix = |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
472 |
let fun occ(Var (jx,_)) = (ix=jx) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
473 |
| occ(t1$t2) = occ t1 orelse occ t2 |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
474 |
| occ(Abs(_,_,t)) = occ t |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
475 |
| occ _ = false |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
476 |
in occ end; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
477 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
478 |
fun is_recordtype T = not (null (RecordPackage.dest_recTs T)); |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
479 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
480 |
(*Unwanted equalities include |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
481 |
(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:
21431
diff
changeset
|
482 |
(2) those between a variable and a record, since these seem to be prolific "cases" thms |
21588 | 483 |
*) |
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
484 |
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:
21431
diff
changeset
|
485 |
| too_general_eqterms _ = false; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
486 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
487 |
fun too_general_equality (Const ("op =", _) $ x $ y) = |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
488 |
too_general_eqterms (x,y) orelse too_general_eqterms(y,x) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
489 |
| too_general_equality _ = false; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
490 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
491 |
(* tautologous? *) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
492 |
fun is_taut (Const ("Trueprop", _) $ Const ("True", _)) = true |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
493 |
| is_taut _ = false; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
494 |
|
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
495 |
(*True if the term contains a variable whose (atomic) type is in the given list.*) |
21588 | 496 |
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:
21397
diff
changeset
|
497 |
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:
21397
diff
changeset
|
498 |
| var_tycon _ = false |
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
499 |
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:
21397
diff
changeset
|
500 |
|
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
501 |
(*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:
22193
diff
changeset
|
502 |
they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=(). |
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
503 |
The resulting clause will have no type constraint, yielding false proofs. Even "bool" |
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
504 |
leads to many unsound proofs, though (obviously) only for higher-order problems.*) |
24215 | 505 |
val unwanted_types = ["Product_Type.unit","bool"]; |
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
506 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
507 |
fun unwanted t = |
24215 | 508 |
is_taut t orelse has_typed_var unwanted_types t orelse |
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
509 |
forall too_general_equality (dest_disj t); |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
510 |
|
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
511 |
(*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:
21397
diff
changeset
|
512 |
likely to lead to unsound proofs.*) |
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
513 |
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:
21397
diff
changeset
|
514 |
|
23387 | 515 |
fun tptp_writer logic = ResHolClause.tptp_write_file (logic=FOL); |
19445
da75577642a9
tidying; ATP options including CASC mode for Vampire
paulson
parents:
19442
diff
changeset
|
516 |
|
23387 | 517 |
fun dfg_writer logic = ResHolClause.dfg_write_file (logic=FOL); |
19194 | 518 |
|
20022 | 519 |
(*Called by the oracle-based methods declared in res_atp_methods.ML*) |
19722 | 520 |
fun write_subgoal_file dfg mode ctxt conjectures user_thms n = |
21588 | 521 |
let val conj_cls = make_clauses conjectures |
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22724
diff
changeset
|
522 |
|> ResAxioms.assume_abstract_list "subgoal" |> Meson.finish_cnf |
21588 | 523 |
val hyp_cls = cnf_hyps_thms ctxt |
524 |
val goal_cls = conj_cls@hyp_cls |
|
525 |
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:
22217
diff
changeset
|
526 |
val thy = ProofContext.theory_of ctxt |
21588 | 527 |
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:
21397
diff
changeset
|
528 |
Auto => problem_logic_goals [goal_tms] |
21588 | 529 |
| Fol => FOL |
530 |
| Hol => HOL |
|
531 |
val (included_thms,user_rules) = get_clasimp_atp_lemmas ctxt user_thms |
|
22865 | 532 |
val cla_simp_atp_clauses = included_thms |
21588 | 533 |
|> 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:
22217
diff
changeset
|
534 |
|> restrict_to_logic thy logic |
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
535 |
|> remove_unwanted_clauses |
21588 | 536 |
val user_cls = ResAxioms.cnf_rules_pairs user_rules |
537 |
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:
21397
diff
changeset
|
538 |
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:
21311
diff
changeset
|
539 |
and axtms = map (prop_of o #1) axclauses |
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
540 |
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:
21397
diff
changeset
|
541 |
and tycons = type_consts_of_terms thy (goal_tms@axtms) |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
542 |
(*TFrees in conjecture clauses; TVars in axiom clauses*) |
22645 | 543 |
val (supers',arity_clauses) = ResClause.make_arity_clauses thy tycons supers |
544 |
val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' |
|
21588 | 545 |
val writer = if dfg then dfg_writer else tptp_writer |
546 |
and file = atp_input_file() |
|
547 |
and user_lemmas_names = map #1 user_rules |
|
19194 | 548 |
in |
21588 | 549 |
writer logic goal_cls file (axclauses,classrel_clauses,arity_clauses) user_lemmas_names; |
22130 | 550 |
Output.debug (fn () => "Writing to " ^ file); |
21588 | 551 |
file |
19194 | 552 |
end; |
553 |
||
554 |
||
555 |
(**** remove tmp files ****) |
|
21588 | 556 |
fun cond_rm_tmp file = |
22130 | 557 |
if !Output.debugging orelse !destdir <> "" |
558 |
then Output.debug (fn () => "ATP input kept...") |
|
20870 | 559 |
else OS.FileSys.remove file; |
19194 | 560 |
|
561 |
||
562 |
(****** setup ATPs as Isabelle methods ******) |
|
563 |
||
21588 | 564 |
fun atp_meth tac ths ctxt = |
19194 | 565 |
let val thy = ProofContext.theory_of ctxt |
21588 | 566 |
val _ = ResClause.init thy |
567 |
val _ = ResHolClause.init thy |
|
568 |
in Method.SIMPLE_METHOD' (tac ctxt ths) end; |
|
19194 | 569 |
|
570 |
fun atp_method tac = Method.thms_ctxt_args (atp_meth tac); |
|
571 |
||
572 |
(***************************************************************) |
|
573 |
(* automatic ATP invocation *) |
|
574 |
(***************************************************************) |
|
575 |
||
17306 | 576 |
(* call prover with settings and problem file for the current subgoal *) |
17764 | 577 |
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:
16767
diff
changeset
|
578 |
let |
17422 | 579 |
fun make_atp_list [] n = [] |
17717 | 580 |
| make_atp_list (sg_term::xs) n = |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
581 |
let |
17717 | 582 |
val probfile = prob_pathname n |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
583 |
val time = Int.toString (!time_limit) |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
584 |
in |
22130 | 585 |
Output.debug (fn () => "problem file in watcher_call_provers is " ^ probfile); |
17764 | 586 |
(*options are separated by Watcher.setting_sep, currently #"%"*) |
17306 | 587 |
if !prover = "spass" |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
588 |
then |
19445
da75577642a9
tidying; ATP options including CASC mode for Vampire
paulson
parents:
19442
diff
changeset
|
589 |
let val spass = helper_path "SPASS_HOME" "SPASS" |
da75577642a9
tidying; ATP options including CASC mode for Vampire
paulson
parents:
19442
diff
changeset
|
590 |
val sopts = |
da75577642a9
tidying; ATP options including CASC mode for Vampire
paulson
parents:
19442
diff
changeset
|
591 |
"-Auto%-SOS=1%-PGiven=0%-PProblem=0%-Splits=0%-FullRed=0%-DocProof%-TimeLimit=" ^ time |
21588 | 592 |
in |
19445
da75577642a9
tidying; ATP options including CASC mode for Vampire
paulson
parents:
19442
diff
changeset
|
593 |
("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:
16767
diff
changeset
|
594 |
end |
17306 | 595 |
else if !prover = "vampire" |
21588 | 596 |
then |
17819 | 597 |
let val vampire = helper_path "VAMPIRE_HOME" "vampire" |
21132 | 598 |
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:
16767
diff
changeset
|
599 |
in |
19445
da75577642a9
tidying; ATP options including CASC mode for Vampire
paulson
parents:
19442
diff
changeset
|
600 |
("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:
16767
diff
changeset
|
601 |
end |
21588 | 602 |
else if !prover = "E" |
603 |
then |
|
604 |
let val Eprover = helper_path "E_HOME" "eproof" |
|
605 |
in |
|
606 |
("E", Eprover, |
|
21900 | 607 |
"--tstp-in%--tstp-out%-l5%-xAutoDev%-tAutoDev%--silent%--cpu-limit=" ^ time, probfile) :: |
21588 | 608 |
make_atp_list xs (n+1) |
609 |
end |
|
610 |
else error ("Invalid prover name: " ^ !prover) |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
611 |
end |
15452 | 612 |
|
17422 | 613 |
val atp_list = make_atp_list sg_terms 1 |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
614 |
in |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
615 |
Watcher.callResProvers(childout,atp_list); |
22130 | 616 |
Output.debug (fn () => "Sent commands to watcher!") |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
617 |
end |
21588 | 618 |
|
22193 | 619 |
(*For debugging the generated set of theorem names*) |
21888 | 620 |
fun trace_vector fname = |
22193 | 621 |
let val path = File.explode_platform_path (fname ^ "_thm_names") |
21888 | 622 |
in Vector.app (File.append path o (fn s => s ^ "\n")) end; |
16357 | 623 |
|
20022 | 624 |
(*We write out problem files for each subgoal. Argument probfile generates filenames, |
18986 | 625 |
and allows the suppression of the suffix "_1" in problem-generation mode. |
626 |
FIXME: does not cope with &&, and it isn't easy because one could have multiple |
|
627 |
subgoals, each involving &&.*) |
|
20022 | 628 |
fun write_problem_files probfile (ctxt,th) = |
18753
aa82bd41555d
ResClasimp.get_clasimp_lemmas now takes all subgoals rather than only the first
paulson
parents:
18700
diff
changeset
|
629 |
let val goals = Thm.prems_of th |
22130 | 630 |
val _ = Output.debug (fn () => "number of subgoals = " ^ Int.toString (length goals)) |
17717 | 631 |
val thy = ProofContext.theory_of ctxt |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
632 |
fun get_neg_subgoals [] _ = [] |
22865 | 633 |
| 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:
21980
diff
changeset
|
634 |
get_neg_subgoals gls (n+1) |
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
635 |
val goal_cls = get_neg_subgoals goals 1 |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
636 |
val logic = case !linkup_logic_mode of |
23387 | 637 |
Auto => problem_logic_goals (map (map prop_of) goal_cls) |
21588 | 638 |
| Fol => FOL |
639 |
| Hol => HOL |
|
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
640 |
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:
22217
diff
changeset
|
641 |
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:
22217
diff
changeset
|
642 |
|> restrict_to_logic thy logic |
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
643 |
|> remove_unwanted_clauses |
22130 | 644 |
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:
20457
diff
changeset
|
645 |
val white_cls = ResAxioms.cnf_rules_pairs white_thms |
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
646 |
(*clauses relevant to goal gl*) |
21563
b4718f2c15f0
Goals in clause form are sent to the relevance filter.
mengj
parents:
21509
diff
changeset
|
647 |
val axcls_list = map (fn ngcls => get_relevant_clauses thy included_cls white_cls (map prop_of ngcls)) goal_cls |
22130 | 648 |
val _ = app (fn axcls => Output.debug (fn () => "filtered clauses = " ^ Int.toString(length axcls))) |
21070 | 649 |
axcls_list |
21588 | 650 |
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:
20457
diff
changeset
|
651 |
fun write_all [] [] _ = [] |
21588 | 652 |
| write_all (ccls::ccls_list) (axcls::axcls_list) k = |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
653 |
let val fname = probfile k |
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
654 |
val _ = Output.debug (fn () => "About to write file " ^ fname) |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
655 |
val axcls = make_unique axcls |
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
656 |
val _ = Output.debug (fn () => "Conjecture Clauses (before duplicate removal)") |
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
657 |
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:
20854
diff
changeset
|
658 |
val ccls = subtract_cls ccls axcls |
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
659 |
val _ = Output.debug (fn () => "Conjecture Clauses (AFTER duplicate removal)") |
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
660 |
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:
21397
diff
changeset
|
661 |
val ccltms = map prop_of ccls |
21373
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
662 |
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:
21397
diff
changeset
|
663 |
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:
21397
diff
changeset
|
664 |
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:
21397
diff
changeset
|
665 |
and tycons = type_consts_of_terms thy (ccltms@axtms) |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
666 |
(*TFrees in conjecture clauses; TVars in axiom clauses*) |
22645 | 667 |
val (supers',arity_clauses) = ResClause.make_arity_clauses thy tycons supers |
668 |
val _ = Output.debug (fn () => "arity clauses = " ^ Int.toString (length arity_clauses)) |
|
669 |
val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' |
|
22130 | 670 |
val _ = Output.debug (fn () => "classrel clauses = " ^ Int.toString (length classrel_clauses)) |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
671 |
val clnames = writer logic ccls fname (axcls,classrel_clauses,arity_clauses) [] |
21888 | 672 |
val thm_names = Vector.fromList clnames |
22193 | 673 |
val _ = if !Output.debugging then trace_vector fname thm_names else () |
20870 | 674 |
in (thm_names,fname) :: write_all ccls_list axcls_list (k+1) end |
675 |
val (thm_names_list, filenames) = ListPair.unzip (write_all goal_cls axcls_list 1) |
|
19194 | 676 |
in |
20870 | 677 |
(filenames, thm_names_list) |
19194 | 678 |
end; |
15644 | 679 |
|
21588 | 680 |
val last_watcher_pid = ref (NONE : (TextIO.instream * TextIO.outstream * |
17775 | 681 |
Posix.Process.pid * string list) option); |
682 |
||
683 |
fun kill_last_watcher () = |
|
21588 | 684 |
(case !last_watcher_pid of |
17775 | 685 |
NONE => () |
21588 | 686 |
| SOME (_, _, pid, files) => |
22130 | 687 |
(Output.debug (fn () => "Killing old watcher, pid = " ^ string_of_pid pid); |
21588 | 688 |
Watcher.killWatcher pid; |
689 |
ignore (map (try cond_rm_tmp) files))) |
|
22130 | 690 |
handle OS.SysErr _ => Output.debug (fn () => "Attempt to kill watcher failed"); |
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
691 |
|
21980
d22f7e3c5ad9
x-symbol is no longer switched off in the ATP linkup
paulson
parents:
21900
diff
changeset
|
692 |
(*writes out the current problems and calls ATPs*) |
d22f7e3c5ad9
x-symbol is no longer switched off in the ATP linkup
paulson
parents:
21900
diff
changeset
|
693 |
fun isar_atp (ctxt, th) = |
17717 | 694 |
if Thm.no_prems th then () |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
695 |
else |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
696 |
let |
17775 | 697 |
val _ = kill_last_watcher() |
20870 | 698 |
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:
21999
diff
changeset
|
699 |
val (childin, childout, pid) = Watcher.createWatcher (ctxt, th, thm_names_list) |
15608 | 700 |
in |
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
701 |
last_watcher_pid := SOME (childin, childout, pid, files); |
22130 | 702 |
Output.debug (fn () => "problem files: " ^ space_implode ", " files); |
703 |
Output.debug (fn () => "pid: " ^ string_of_pid pid); |
|
22578 | 704 |
watcher_call_provers (Thm.theory_of_thm th) (Thm.prems_of th) (childin, childout, pid) |
20954 | 705 |
end; |
706 |
||
707 |
(*For ML scripts, and primarily, for debugging*) |
|
21588 | 708 |
fun callatp () = |
20954 | 709 |
let val th = topthm() |
710 |
val ctxt = ProofContext.init (theory_of_thm th) |
|
21980
d22f7e3c5ad9
x-symbol is no longer switched off in the ATP linkup
paulson
parents:
21900
diff
changeset
|
711 |
in isar_atp (ctxt, th) end; |
15608 | 712 |
|
21588 | 713 |
val isar_atp_writeonly = setmp print_mode [] |
17717 | 714 |
(fn (ctxt,th) => |
715 |
if Thm.no_prems th then () |
|
21588 | 716 |
else |
717 |
let val probfile = if Thm.nprems_of th = 1 then probfile_nosuffix |
|
718 |
else prob_pathname |
|
20022 | 719 |
in ignore (write_problem_files probfile (ctxt,th)) end); |
15452 | 720 |
|
16357 | 721 |
|
22865 | 722 |
(** the Isar toplevel command **) |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
723 |
|
22865 | 724 |
fun sledgehammer state = |
725 |
let |
|
726 |
val (ctxt, (_, goal)) = Proof.get_goal (Toplevel.proof_of state); |
|
727 |
val thy = ProofContext.theory_of ctxt; |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
728 |
in |
22130 | 729 |
Output.debug (fn () => "subgoals in isar_atp:\n" ^ |
21588 | 730 |
Pretty.string_of (ProofContext.pretty_term ctxt |
731 |
(Logic.mk_conjunction_list (Thm.prems_of goal)))); |
|
22130 | 732 |
Output.debug (fn () => "current theory: " ^ Context.theory_name thy); |
16925
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16904
diff
changeset
|
733 |
ResClause.init thy; |
19194 | 734 |
ResHolClause.init thy; |
21132 | 735 |
if !time_limit > 0 then isar_atp (ctxt, goal) |
22865 | 736 |
else (warning ("Writing problem file only: " ^ !problem_name); |
22193 | 737 |
isar_atp_writeonly (ctxt, goal)) |
19205
4ec788c69f82
Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents:
19194
diff
changeset
|
738 |
end; |
4ec788c69f82
Tidying. New invoke_atp_ml for top-level debugging. Flag to force FOL mode.
paulson
parents:
19194
diff
changeset
|
739 |
|
22865 | 740 |
val _ = OuterSyntax.add_parsers |
741 |
[OuterSyntax.command "sledgehammer" "call automatic theorem provers" OuterKeyword.diag |
|
742 |
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_proof o Toplevel.keep sledgehammer))]; |
|
17091 | 743 |
|
15347 | 744 |
end; |