author | blanchet |
Thu, 31 Mar 2011 11:16:52 +0200 | |
changeset 42180 | a6c141925a8a |
parent 41990 | 7f2793d51efc |
child 42227 | 662b50b7126f |
permissions | -rw-r--r-- |
40114 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
2 |
Author: Fabian Immler, TU Muenchen |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
3 |
Author: Makarius |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
5 |
|
39494
bf7dd4902321
rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
blanchet
parents:
39452
diff
changeset
|
6 |
Translation of HOL to FOL for Sledgehammer. |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
7 |
*) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
8 |
|
40068 | 9 |
signature SLEDGEHAMMER_ATP_TRANSLATE = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
10 |
sig |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
11 |
type 'a problem = 'a ATP_Problem.problem |
40114 | 12 |
type translated_formula |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
13 |
|
41134 | 14 |
datatype type_system = |
15 |
Tags of bool | |
|
16 |
Preds of bool | |
|
17 |
Const_Args | |
|
18 |
No_Types |
|
19 |
||
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
20 |
val precise_overloaded_args : bool Unsynchronized.ref |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
21 |
val fact_prefix : string |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
22 |
val conjecture_prefix : string |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
23 |
val types_dangerous_types : type_system -> bool |
41136
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
24 |
val num_atp_type_args : theory -> type_system -> string -> int |
41088 | 25 |
val translate_atp_fact : |
41990
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
26 |
Proof.context -> bool -> (string * 'a) * thm |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
27 |
-> translated_formula option * ((string * 'a) * thm) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39975
diff
changeset
|
28 |
val prepare_atp_problem : |
41134 | 29 |
Proof.context -> bool -> bool -> type_system -> bool -> term list -> term |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
30 |
-> (translated_formula option * ((string * 'a) * thm)) list |
38818
61cf050f8b2e
improve SPASS hack, when a clause comes from several facts
blanchet
parents:
38752
diff
changeset
|
31 |
-> string problem * string Symtab.table * int * (string * 'a) list vector |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
32 |
val atp_problem_weights : string problem -> (string * real) list |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
33 |
end; |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
34 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
35 |
structure Sledgehammer_ATP_Translate : SLEDGEHAMMER_ATP_TRANSLATE = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
36 |
struct |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
37 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
38 |
open ATP_Problem |
39494
bf7dd4902321
rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
blanchet
parents:
39452
diff
changeset
|
39 |
open Metis_Translate |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
40 |
open Sledgehammer_Util |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
41 |
|
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
42 |
(* FIXME: Remove references once appropriate defaults have been determined |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
43 |
empirically. *) |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
44 |
val precise_overloaded_args = Unsynchronized.ref false |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
45 |
|
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
46 |
val fact_prefix = "fact_" |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
47 |
val conjecture_prefix = "conj_" |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
48 |
val helper_prefix = "help_" |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
49 |
val class_rel_clause_prefix = "clrel_"; |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
50 |
val arity_clause_prefix = "arity_" |
39975
7c50d5ca5c04
avoid generating several formulas with the same name ("tfrees")
blanchet
parents:
39954
diff
changeset
|
51 |
val tfree_prefix = "tfree_" |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
52 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
53 |
(* Freshness almost guaranteed! *) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
54 |
val sledgehammer_weak_prefix = "Sledgehammer:" |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
55 |
|
40114 | 56 |
type translated_formula = |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38748
diff
changeset
|
57 |
{name: string, |
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38748
diff
changeset
|
58 |
kind: kind, |
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38748
diff
changeset
|
59 |
combformula: (name, combterm) formula, |
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38748
diff
changeset
|
60 |
ctypes_sorts: typ list} |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
61 |
|
41134 | 62 |
datatype type_system = |
63 |
Tags of bool | |
|
64 |
Preds of bool | |
|
65 |
Const_Args | |
|
66 |
No_Types |
|
67 |
||
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
68 |
fun types_dangerous_types (Tags _) = true |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
69 |
| types_dangerous_types (Preds _) = true |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
70 |
| types_dangerous_types _ = false |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
71 |
|
41136
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
72 |
(* This is an approximation. If it returns "true" for a constant that isn't |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
73 |
overloaded (i.e., that has one uniform definition), needless clutter is |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
74 |
generated; if it returns "false" for an overloaded constant, the ATP gets a |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
75 |
license to do unsound reasoning if the type system is "overloaded_args". *) |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
76 |
fun is_overloaded thy s = |
41748
657712cc8847
fix handling of "fequal" in generated ATP problems -- the bug was visible if "explicit_apply" was on and "singleton_conv" is one of the facts being translated, as it resulted in an arity error in the ATP (e.g., E)
blanchet
parents:
41726
diff
changeset
|
77 |
not (s = @{const_name HOL.eq}) andalso |
657712cc8847
fix handling of "fequal" in generated ATP problems -- the bug was visible if "explicit_apply" was on and "singleton_conv" is one of the facts being translated, as it resulted in an arity error in the ATP (e.g., E)
blanchet
parents:
41726
diff
changeset
|
78 |
not (s = @{const_name Metis.fequal}) andalso |
657712cc8847
fix handling of "fequal" in generated ATP problems -- the bug was visible if "explicit_apply" was on and "singleton_conv" is one of the facts being translated, as it resulted in an arity error in the ATP (e.g., E)
blanchet
parents:
41726
diff
changeset
|
79 |
(not (!precise_overloaded_args) orelse s = @{const_name finite} orelse |
41147
0e1903273712
fix translation of higher-order equality ("fequal") if "precise_overloaded_args" is "true"
blanchet
parents:
41145
diff
changeset
|
80 |
length (Defs.specifications_of (Theory.defs_of thy) s) > 1) |
41136
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
81 |
|
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
82 |
fun needs_type_args thy type_sys s = |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
83 |
case type_sys of |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
84 |
Tags full_types => not full_types andalso is_overloaded thy s |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
85 |
| Preds _ => is_overloaded thy s (* FIXME: could be more precise *) |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
86 |
| Const_Args => is_overloaded thy s |
41136
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
87 |
| No_Types => false |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
88 |
|
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
89 |
fun num_atp_type_args thy type_sys s = |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
90 |
if needs_type_args thy type_sys s then num_type_args thy s else 0 |
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
91 |
|
41137
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
92 |
fun atp_type_literals_for_types type_sys Ts = |
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
93 |
if type_sys = No_Types then [] else type_literals_for_types Ts |
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
94 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
95 |
fun mk_anot phi = AConn (ANot, [phi]) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
96 |
fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2]) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
97 |
fun mk_ahorn [] phi = phi |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
98 |
| mk_ahorn (phi :: phis) psi = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
99 |
AConn (AImplies, [fold (mk_aconn AAnd) phis phi, psi]) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
100 |
|
41145
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
101 |
fun close_universally phi = |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
102 |
let |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
103 |
fun term_vars bounds (ATerm (name as (s, _), tms)) = |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
104 |
(is_atp_variable s andalso not (member (op =) bounds name)) |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
105 |
? insert (op =) name |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
106 |
#> fold (term_vars bounds) tms |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
107 |
fun formula_vars bounds (AQuant (_, xs, phi)) = |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
108 |
formula_vars (xs @ bounds) phi |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
109 |
| formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
110 |
| formula_vars bounds (AAtom tm) = term_vars bounds tm |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
111 |
in |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
112 |
case formula_vars [] phi [] of [] => phi | xs => AQuant (AForall, xs, phi) |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
113 |
end |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
114 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
115 |
fun combformula_for_prop thy eq_as_iff = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
116 |
let |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
117 |
fun do_term bs t ts = |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
118 |
combterm_from_term thy bs (Envir.eta_contract t) |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
119 |
|>> AAtom ||> union (op =) ts |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
120 |
fun do_quant bs q s T t' = |
38518
54727b44e277
handle bound name conflicts gracefully in FOF translation
blanchet
parents:
38496
diff
changeset
|
121 |
let val s = Name.variant (map fst bs) s in |
54727b44e277
handle bound name conflicts gracefully in FOF translation
blanchet
parents:
38496
diff
changeset
|
122 |
do_formula ((s, T) :: bs) t' |
54727b44e277
handle bound name conflicts gracefully in FOF translation
blanchet
parents:
38496
diff
changeset
|
123 |
#>> (fn phi => AQuant (q, [`make_bound_var s], phi)) |
54727b44e277
handle bound name conflicts gracefully in FOF translation
blanchet
parents:
38496
diff
changeset
|
124 |
end |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
125 |
and do_conn bs c t1 t2 = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
126 |
do_formula bs t1 ##>> do_formula bs t2 |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
127 |
#>> (fn (phi1, phi2) => AConn (c, [phi1, phi2])) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
128 |
and do_formula bs t = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
129 |
case t of |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
130 |
@{const Not} $ t1 => |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
131 |
do_formula bs t1 #>> (fn phi => AConn (ANot, [phi])) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
132 |
| Const (@{const_name All}, _) $ Abs (s, T, t') => |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
133 |
do_quant bs AForall s T t' |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
134 |
| Const (@{const_name Ex}, _) $ Abs (s, T, t') => |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
135 |
do_quant bs AExists s T t' |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
136 |
| @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2 |
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
137 |
| @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2 |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38752
diff
changeset
|
138 |
| @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2 |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38829
diff
changeset
|
139 |
| Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 => |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
140 |
if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
141 |
| _ => do_term bs t |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
142 |
in do_formula [] end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
143 |
|
38618 | 144 |
val presimplify_term = prop_of o Meson.presimplify oo Skip_Proof.make_thm |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
145 |
|
41491 | 146 |
fun concealed_bound_name j = sledgehammer_weak_prefix ^ string_of_int j |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
147 |
fun conceal_bounds Ts t = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
148 |
subst_bounds (map (Free o apfst concealed_bound_name) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
149 |
(0 upto length Ts - 1 ~~ Ts), t) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
150 |
fun reveal_bounds Ts = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
151 |
subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j)) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
152 |
(0 upto length Ts - 1 ~~ Ts)) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
153 |
|
38608
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
154 |
(* Removes the lambdas from an equation of the form "t = (%x. u)". |
39890 | 155 |
(Cf. "extensionalize_theorem" in "Meson_Clausify".) *) |
38608
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
156 |
fun extensionalize_term t = |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
157 |
let |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
158 |
fun aux j (@{const Trueprop} $ t') = @{const Trueprop} $ aux j t' |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
159 |
| aux j (t as Const (s, Type (_, [Type (_, [_, T']), |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
160 |
Type (_, [_, res_T])])) |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
161 |
$ t2 $ Abs (var_s, var_T, t')) = |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38829
diff
changeset
|
162 |
if s = @{const_name HOL.eq} orelse s = @{const_name "=="} then |
38608
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
163 |
let val var_t = Var ((var_s, j), var_T) in |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
164 |
Const (s, T' --> T' --> res_T) |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
165 |
$ betapply (t2, var_t) $ subst_bound (var_t, t') |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
166 |
|> aux (j + 1) |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
167 |
end |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
168 |
else |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
169 |
t |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
170 |
| aux _ t = t |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
171 |
in aux (maxidx_of_term t + 1) t end |
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
172 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
173 |
fun introduce_combinators_in_term ctxt kind t = |
38491
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
174 |
let val thy = ProofContext.theory_of ctxt in |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
175 |
if Meson.is_fol_term thy t then |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
176 |
t |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
177 |
else |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
178 |
let |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
179 |
fun aux Ts t = |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
180 |
case t of |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
181 |
@{const Not} $ t1 => @{const Not} $ aux Ts t1 |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
182 |
| (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') => |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
183 |
t0 $ Abs (s, T, aux (T :: Ts) t') |
38652
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38618
diff
changeset
|
184 |
| (t0 as Const (@{const_name All}, _)) $ t1 => |
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38618
diff
changeset
|
185 |
aux Ts (t0 $ eta_expand Ts t1 1) |
38491
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
186 |
| (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') => |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
187 |
t0 $ Abs (s, T, aux (T :: Ts) t') |
38652
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38618
diff
changeset
|
188 |
| (t0 as Const (@{const_name Ex}, _)) $ t1 => |
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38618
diff
changeset
|
189 |
aux Ts (t0 $ eta_expand Ts t1 1) |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
190 |
| (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
191 |
| (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38752
diff
changeset
|
192 |
| (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38829
diff
changeset
|
193 |
| (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _]))) |
38491
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
194 |
$ t1 $ t2 => |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
195 |
t0 $ aux Ts t1 $ aux Ts t2 |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
196 |
| _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
197 |
t |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
198 |
else |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
199 |
t |> conceal_bounds Ts |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
200 |
|> Envir.eta_contract |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
201 |
|> cterm_of thy |
39890 | 202 |
|> Meson_Clausify.introduce_combinators_in_cterm |
38491
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
203 |
|> prop_of |> Logic.dest_equals |> snd |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
204 |
|> reveal_bounds Ts |
39370
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
blanchet
parents:
39005
diff
changeset
|
205 |
val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single |
38491
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
206 |
in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
207 |
handle THM _ => |
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
208 |
(* A type variable of sort "{}" will make abstraction fail. *) |
38613
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
209 |
if kind = Conjecture then HOLogic.false_const |
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
210 |
else HOLogic.true_const |
38491
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
blanchet
parents:
38282
diff
changeset
|
211 |
end |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
212 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
213 |
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
214 |
same in Sledgehammer to prevent the discovery of unreplable proofs. *) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
215 |
fun freeze_term t = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
216 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
217 |
fun aux (t $ u) = aux t $ aux u |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
218 |
| aux (Abs (s, T, t)) = Abs (s, T, aux t) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
219 |
| aux (Var ((s, i), T)) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
220 |
Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
221 |
| aux t = t |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
222 |
in t |> exists_subterm is_Var t ? aux end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
223 |
|
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
224 |
(* making fact and conjecture formulas *) |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
225 |
fun make_formula ctxt eq_as_iff presimp name kind t = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
226 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
227 |
val thy = ProofContext.theory_of ctxt |
38608
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38606
diff
changeset
|
228 |
val t = t |> Envir.beta_eta_contract |
38652
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38618
diff
changeset
|
229 |
|> transform_elim_term |
41211
1e2e16bc0077
no need to do a super-duper atomization if Metis fails afterwards anyway
blanchet
parents:
41199
diff
changeset
|
230 |
|> Object_Logic.atomize_term thy |
38652
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38618
diff
changeset
|
231 |
val need_trueprop = (fastype_of t = HOLogic.boolT) |
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38618
diff
changeset
|
232 |
val t = t |> need_trueprop ? HOLogic.mk_Trueprop |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
233 |
|> extensionalize_term |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
234 |
|> presimp ? presimplify_term thy |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
235 |
|> perhaps (try (HOLogic.dest_Trueprop)) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
236 |
|> introduce_combinators_in_term ctxt kind |
38613
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
237 |
|> kind <> Axiom ? freeze_term |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
238 |
val (combformula, ctypes_sorts) = combformula_for_prop thy eq_as_iff t [] |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
239 |
in |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38748
diff
changeset
|
240 |
{name = name, combformula = combformula, kind = kind, |
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38748
diff
changeset
|
241 |
ctypes_sorts = ctypes_sorts} |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
242 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
243 |
|
41990
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
244 |
fun make_fact ctxt keep_trivial eq_as_iff presimp ((name, _), th) = |
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
245 |
case (keep_trivial, |
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
246 |
make_formula ctxt eq_as_iff presimp name Axiom (prop_of th)) of |
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
247 |
(false, {combformula = AAtom (CombConst (("c_True", _), _, _)), ...}) => |
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
248 |
NONE |
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
249 |
| (_, formula) => SOME formula |
38613
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
250 |
fun make_conjecture ctxt ts = |
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
251 |
let val last = length ts - 1 in |
41491 | 252 |
map2 (fn j => make_formula ctxt true true (string_of_int j) |
38613
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
253 |
(if j = last then Conjecture else Hypothesis)) |
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
254 |
(0 upto last) ts |
4ca2cae2653f
use "hypothesis" rather than "conjecture" for hypotheses in TPTP format;
blanchet
parents:
38610
diff
changeset
|
255 |
end |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
256 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
257 |
(** Helper facts **) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
258 |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
259 |
fun fold_formula f (AQuant (_, _, phi)) = fold_formula f phi |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
260 |
| fold_formula f (AConn (_, phis)) = fold (fold_formula f) phis |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
261 |
| fold_formula f (AAtom tm) = f tm |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
262 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
263 |
fun count_term (ATerm ((s, _), tms)) = |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
264 |
(if is_atp_variable s then I |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
265 |
else Symtab.map_entry s (Integer.add 1)) |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
266 |
#> fold count_term tms |
41406 | 267 |
fun count_formula x = fold_formula count_term x |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
268 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
269 |
val init_counters = |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
270 |
metis_helpers |> map fst |> sort_distinct string_ord |> map (rpair 0) |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
271 |
|> Symtab.make |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
272 |
|
41145
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
273 |
fun get_helper_facts ctxt explicit_forall type_sys formulas = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
274 |
let |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
275 |
val no_dangerous_types = types_dangerous_types type_sys |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
276 |
val ct = init_counters |> fold count_formula formulas |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
277 |
fun is_used s = the (Symtab.lookup ct s) > 0 |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
278 |
fun dub c needs_full_types (th, j) = |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
279 |
((c ^ "_" ^ string_of_int j ^ (if needs_full_types then "ft" else ""), |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
280 |
false), th) |
41990
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
281 |
fun make_facts eq_as_iff = map_filter (make_fact ctxt false eq_as_iff false) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
282 |
in |
41145
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
283 |
(metis_helpers |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
284 |
|> filter (is_used o fst) |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
285 |
|> maps (fn (c, (needs_full_types, ths)) => |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
286 |
if needs_full_types andalso not no_dangerous_types then |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
287 |
[] |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
288 |
else |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
289 |
ths ~~ (1 upto length ths) |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
290 |
|> map (dub c needs_full_types) |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
291 |
|> make_facts (not needs_full_types)), |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
292 |
if type_sys = Tags false then |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
293 |
let |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
294 |
fun var s = ATerm (`I s, []) |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
295 |
fun tag tm = ATerm (`I type_tag_name, [var "X", tm]) |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
296 |
in |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
297 |
[Fof (helper_prefix ^ ascii_of "ti_ti", Axiom, |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
298 |
AAtom (ATerm (`I "equal", [tag (tag (var "Y")), tag (var "Y")])) |
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
299 |
|> explicit_forall ? close_universally, NONE)] |
41145
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
300 |
end |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
301 |
else |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
302 |
[]) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
303 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
304 |
|
41990
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
305 |
fun translate_atp_fact ctxt keep_trivial = |
7f2793d51efc
add option to function to keep trivial ATP formulas, needed for some experiments
blanchet
parents:
41770
diff
changeset
|
306 |
`(make_fact ctxt keep_trivial true true) |
39004
f1b465f889b5
translate the axioms to FOF once and for all ATPs
blanchet
parents:
39003
diff
changeset
|
307 |
|
41134 | 308 |
fun translate_formulas ctxt type_sys hyp_ts concl_t rich_facts = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
309 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
310 |
val thy = ProofContext.theory_of ctxt |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
311 |
val fact_ts = map (prop_of o snd o snd) rich_facts |
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
312 |
val (facts, fact_names) = |
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
313 |
rich_facts |
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
314 |
|> map_filter (fn (NONE, _) => NONE |
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
315 |
| (SOME fact, (name, _)) => SOME (fact, name)) |
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41088
diff
changeset
|
316 |
|> ListPair.unzip |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
317 |
(* Remove existing facts from the conjecture, as this can dramatically |
39005 | 318 |
boost an ATP's performance (for some reason). *) |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
319 |
val hyp_ts = hyp_ts |> filter_out (member (op aconv) fact_ts) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
320 |
val goal_t = Logic.list_implies (hyp_ts, concl_t) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
321 |
val subs = tfree_classes_of_terms [goal_t] |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
322 |
val supers = tvar_classes_of_terms fact_ts |
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
323 |
val tycons = type_consts_of_terms thy (goal_t :: fact_ts) |
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
324 |
(* TFrees in the conjecture; TVars in the facts *) |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
325 |
val conjs = make_conjecture ctxt (hyp_ts @ [concl_t]) |
41137
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
326 |
val (supers', arity_clauses) = |
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
327 |
if type_sys = No_Types then ([], []) |
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
328 |
else make_arity_clauses thy tycons supers |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
329 |
val class_rel_clauses = make_class_rel_clauses thy subs supers' |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
330 |
in |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
331 |
(fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses)) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
332 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
333 |
|
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
334 |
fun tag_with_type ty t = ATerm (`I type_tag_name, [ty, t]) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
335 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
336 |
fun fo_term_for_combtyp (CombTVar name) = ATerm (name, []) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
337 |
| fo_term_for_combtyp (CombTFree name) = ATerm (name, []) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
338 |
| fo_term_for_combtyp (CombType (name, tys)) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
339 |
ATerm (name, map fo_term_for_combtyp tys) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
340 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
341 |
fun fo_literal_for_type_literal (TyLitVar (class, name)) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
342 |
(true, ATerm (class, [ATerm (name, [])])) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
343 |
| fo_literal_for_type_literal (TyLitFree (class, name)) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
344 |
(true, ATerm (class, [ATerm (name, [])])) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
345 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
346 |
fun formula_for_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
347 |
|
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
348 |
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
349 |
considered dangerous because their "exhaust" properties can easily lead to |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
350 |
unsound ATP proofs. The checks below are an (unsound) approximation of |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
351 |
finiteness. *) |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
352 |
|
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
353 |
fun is_dtyp_dangerous _ (Datatype_Aux.DtTFree _) = true |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
354 |
| is_dtyp_dangerous ctxt (Datatype_Aux.DtType (s, Us)) = |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
355 |
is_type_constr_dangerous ctxt s andalso forall (is_dtyp_dangerous ctxt) Us |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
356 |
| is_dtyp_dangerous _ (Datatype_Aux.DtRec _) = false |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
357 |
and is_type_dangerous ctxt (Type (s, Ts)) = |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
358 |
is_type_constr_dangerous ctxt s andalso forall (is_type_dangerous ctxt) Ts |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
359 |
| is_type_dangerous _ _ = false |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
360 |
and is_type_constr_dangerous ctxt s = |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
361 |
let val thy = ProofContext.theory_of ctxt in |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
362 |
case Datatype_Data.get_info thy s of |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
363 |
SOME {descr, ...} => |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
364 |
forall (fn (_, (_, _, constrs)) => |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
365 |
forall (forall (is_dtyp_dangerous ctxt) o snd) constrs) descr |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
366 |
| NONE => |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
367 |
case Typedef.get_info ctxt s of |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
368 |
({rep_type, ...}, _) :: _ => is_type_dangerous ctxt rep_type |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
369 |
| [] => true |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
370 |
end |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
371 |
|
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
372 |
fun is_combtyp_dangerous ctxt (CombType ((s, _), tys)) = |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
373 |
(case strip_prefix_and_unascii type_const_prefix s of |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
374 |
SOME s' => forall (is_combtyp_dangerous ctxt) tys andalso |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
375 |
is_type_constr_dangerous ctxt (invert_const s') |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
376 |
| NONE => false) |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
377 |
| is_combtyp_dangerous _ _ = false |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
378 |
|
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
379 |
fun should_tag_with_type ctxt (Tags full_types) ty = |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
380 |
full_types orelse is_combtyp_dangerous ctxt ty |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
381 |
| should_tag_with_type _ _ _ = false |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
382 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
383 |
val fname_table = |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
384 |
[("c_False", (0, ("c_fFalse", @{const_name Metis.fFalse}))), |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
385 |
("c_True", (0, ("c_fTrue", @{const_name Metis.fTrue}))), |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
386 |
("c_Not", (1, ("c_fNot", @{const_name Metis.fNot}))), |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
387 |
("c_conj", (2, ("c_fconj", @{const_name Metis.fconj}))), |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
388 |
("c_disj", (2, ("c_fdisj", @{const_name Metis.fdisj}))), |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
389 |
("c_implies", (2, ("c_fimplies", @{const_name Metis.fimplies}))), |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
390 |
("equal", (2, ("c_fequal", @{const_name Metis.fequal})))] |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
391 |
|
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
392 |
fun fo_term_for_combterm ctxt type_sys = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
393 |
let |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
394 |
val thy = ProofContext.theory_of ctxt |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
395 |
fun aux top_level u = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
396 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
397 |
val (head, args) = strip_combterm_comb u |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
398 |
val (x, ty_args) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
399 |
case head of |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
400 |
CombConst (name as (s, s'), _, ty_args) => |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
401 |
(case AList.lookup (op =) fname_table s of |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
402 |
SOME (n, fname) => |
41150
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
403 |
(if top_level andalso length args = n then |
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
404 |
case s of |
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
405 |
"c_False" => ("$false", s') |
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
406 |
| "c_True" => ("$true", s') |
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
407 |
| _ => name |
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
408 |
else |
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
409 |
fname, []) |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
410 |
| NONE => |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
411 |
case strip_prefix_and_unascii const_prefix s of |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
412 |
NONE => (name, ty_args) |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
413 |
| SOME s'' => |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
414 |
let |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
415 |
val s'' = invert_const s'' |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
416 |
val ty_args = |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
417 |
if needs_type_args thy type_sys s'' then ty_args else [] |
41150
54b3c9c05664
make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
blanchet
parents:
41149
diff
changeset
|
418 |
in (name, ty_args) end) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
419 |
| CombVar (name, _) => (name, []) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
420 |
| CombApp _ => raise Fail "impossible \"CombApp\"" |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
421 |
val t = |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
422 |
ATerm (x, map fo_term_for_combtyp ty_args @ map (aux false) args) |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
423 |
val ty = combtyp_of u |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
424 |
in |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
425 |
t |> (if should_tag_with_type ctxt type_sys ty then |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
426 |
tag_with_type (fo_term_for_combtyp ty) |
41134 | 427 |
else |
428 |
I) |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
429 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
430 |
in aux true end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
431 |
|
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
432 |
fun formula_for_combformula ctxt type_sys = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
433 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
434 |
fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
435 |
| aux (AConn (c, phis)) = AConn (c, map aux phis) |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
436 |
| aux (AAtom tm) = AAtom (fo_term_for_combterm ctxt type_sys tm) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
437 |
in aux end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
438 |
|
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
439 |
fun formula_for_fact ctxt type_sys |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
440 |
({combformula, ctypes_sorts, ...} : translated_formula) = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
441 |
mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal) |
41137
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
442 |
(atp_type_literals_for_types type_sys ctypes_sorts)) |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
443 |
(formula_for_combformula ctxt type_sys combformula) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
444 |
|
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
445 |
(* Each fact is given a unique fact number to avoid name clashes (e.g., because |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
446 |
of monomorphization). The TPTP explicitly forbids name clashes, and some of |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
447 |
the remote provers might care. *) |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
448 |
fun problem_line_for_fact ctxt prefix type_sys |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
449 |
(j, formula as {name, kind, ...}) = |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
450 |
Fof (prefix ^ string_of_int j ^ "_" ^ ascii_of name, kind, |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
451 |
formula_for_fact ctxt type_sys formula, NONE) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
452 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
453 |
fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass, |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
454 |
superclass, ...}) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
455 |
let val ty_arg = ATerm (("T", "T"), []) in |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
456 |
Fof (class_rel_clause_prefix ^ ascii_of name, Axiom, |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
457 |
AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])), |
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
458 |
AAtom (ATerm (superclass, [ty_arg]))]), NONE) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
459 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
460 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
461 |
fun fo_literal_for_arity_literal (TConsLit (c, t, args)) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
462 |
(true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)])) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
463 |
| fo_literal_for_arity_literal (TVarLit (c, sort)) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
464 |
(false, ATerm (c, [ATerm (sort, [])])) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
465 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
466 |
fun problem_line_for_arity_clause (ArityClause {name, conclLit, premLits, |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
467 |
...}) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
468 |
Fof (arity_clause_prefix ^ ascii_of name, Axiom, |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
469 |
mk_ahorn (map (formula_for_fo_literal o apfst not |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
470 |
o fo_literal_for_arity_literal) premLits) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
471 |
(formula_for_fo_literal |
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
472 |
(fo_literal_for_arity_literal conclLit)), NONE) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
473 |
|
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
474 |
fun problem_line_for_conjecture ctxt type_sys |
40114 | 475 |
({name, kind, combformula, ...} : translated_formula) = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
476 |
Fof (conjecture_prefix ^ name, kind, |
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
477 |
formula_for_combformula ctxt type_sys combformula, NONE) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
478 |
|
41137
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
479 |
fun free_type_literals_for_conjecture type_sys |
40114 | 480 |
({ctypes_sorts, ...} : translated_formula) = |
41137
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
481 |
ctypes_sorts |> atp_type_literals_for_types type_sys |
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
blanchet
parents:
41136
diff
changeset
|
482 |
|> map fo_literal_for_type_literal |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
483 |
|
39975
7c50d5ca5c04
avoid generating several formulas with the same name ("tfrees")
blanchet
parents:
39954
diff
changeset
|
484 |
fun problem_line_for_free_type j lit = |
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
485 |
Fof (tfree_prefix ^ string_of_int j, Hypothesis, formula_for_fo_literal lit, |
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
486 |
NONE) |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
487 |
fun problem_lines_for_free_types type_sys conjs = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
488 |
let |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
489 |
val litss = map (free_type_literals_for_conjecture type_sys) conjs |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
490 |
val lits = fold (union (op =)) litss [] |
39975
7c50d5ca5c04
avoid generating several formulas with the same name ("tfrees")
blanchet
parents:
39954
diff
changeset
|
491 |
in map2 problem_line_for_free_type (0 upto length lits - 1) lits end |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
492 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
493 |
(** "hBOOL" and "hAPP" **) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
494 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
495 |
type const_info = {min_arity: int, max_arity: int, sub_level: bool} |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
496 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
497 |
fun consider_term top_level (ATerm ((s, _), ts)) = |
39452 | 498 |
(if is_atp_variable s then |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
499 |
I |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
500 |
else |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
501 |
let val n = length ts in |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
502 |
Symtab.map_default |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
503 |
(s, {min_arity = n, max_arity = 0, sub_level = false}) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
504 |
(fn {min_arity, max_arity, sub_level} => |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
505 |
{min_arity = Int.min (n, min_arity), |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
506 |
max_arity = Int.max (n, max_arity), |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
507 |
sub_level = sub_level orelse not top_level}) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
508 |
end) |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
509 |
#> fold (consider_term (top_level andalso s = type_tag_name)) ts |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
510 |
fun consider_formula (AQuant (_, _, phi)) = consider_formula phi |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
511 |
| consider_formula (AConn (_, phis)) = fold consider_formula phis |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
512 |
| consider_formula (AAtom tm) = consider_term true tm |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
513 |
|
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
514 |
fun consider_problem_line (Fof (_, _, phi, _)) = consider_formula phi |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
515 |
fun consider_problem problem = fold (fold consider_problem_line o snd) problem |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
516 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
517 |
(* needed for helper facts if the problem otherwise does not involve equality *) |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
518 |
val equal_entry = ("equal", {min_arity = 2, max_arity = 2, sub_level = false}) |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
519 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
520 |
fun const_table_for_problem explicit_apply problem = |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
521 |
if explicit_apply then |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
522 |
NONE |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
523 |
else |
41147
0e1903273712
fix translation of higher-order equality ("fequal") if "precise_overloaded_args" is "true"
blanchet
parents:
41145
diff
changeset
|
524 |
SOME (Symtab.empty |> Symtab.default equal_entry |> consider_problem problem) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
525 |
|
41134 | 526 |
fun min_arity_of thy type_sys NONE s = |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
527 |
(if s = "equal" orelse s = type_tag_name orelse |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
528 |
String.isPrefix type_const_prefix s orelse |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
529 |
String.isPrefix class_prefix s then |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
530 |
16383 (* large number *) |
38748 | 531 |
else case strip_prefix_and_unascii const_prefix s of |
41136
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
blanchet
parents:
41134
diff
changeset
|
532 |
SOME s' => num_atp_type_args thy type_sys (invert_const s') |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
533 |
| NONE => 0) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
534 |
| min_arity_of _ _ (SOME the_const_tab) s = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
535 |
case Symtab.lookup the_const_tab s of |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
536 |
SOME ({min_arity, ...} : const_info) => min_arity |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
537 |
| NONE => 0 |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
538 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
539 |
fun full_type_of (ATerm ((s, _), [ty, _])) = |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
540 |
if s = type_tag_name then SOME ty else NONE |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
541 |
| full_type_of _ = NONE |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
542 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
543 |
fun list_hAPP_rev _ t1 [] = t1 |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
544 |
| list_hAPP_rev NONE t1 (t2 :: ts2) = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
545 |
ATerm (`I "hAPP", [list_hAPP_rev NONE t1 ts2, t2]) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
546 |
| list_hAPP_rev (SOME ty) t1 (t2 :: ts2) = |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
547 |
case full_type_of t2 of |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
548 |
SOME ty2 => |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
549 |
let val ty' = ATerm (`make_fixed_type_const @{type_name fun}, |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
550 |
[ty2, ty]) in |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
551 |
ATerm (`I "hAPP", |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
552 |
[tag_with_type ty' (list_hAPP_rev (SOME ty') t1 ts2), t2]) |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
553 |
end |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
554 |
| NONE => list_hAPP_rev NONE t1 (t2 :: ts2) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
555 |
|
41134 | 556 |
fun repair_applications_in_term thy type_sys const_tab = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
557 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
558 |
fun aux opt_ty (ATerm (name as (s, _), ts)) = |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
559 |
if s = type_tag_name then |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
560 |
case ts of |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
561 |
[t1, t2] => ATerm (name, [aux NONE t1, aux (SOME t1) t2]) |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
562 |
| _ => raise Fail "malformed type tag" |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
563 |
else |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
564 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
565 |
val ts = map (aux NONE) ts |
41134 | 566 |
val (ts1, ts2) = chop (min_arity_of thy type_sys const_tab s) ts |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
567 |
in list_hAPP_rev opt_ty (ATerm (name, ts1)) (rev ts2) end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
568 |
in aux NONE end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
569 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
570 |
fun boolify t = ATerm (`I "hBOOL", [t]) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
571 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
572 |
(* True if the constant ever appears outside of the top-level position in |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
573 |
literals, or if it appears with different arities (e.g., because of different |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
574 |
type instantiations). If false, the constant always receives all of its |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
575 |
arguments and is used as a predicate. *) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
576 |
fun is_predicate NONE s = |
38589
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents:
38518
diff
changeset
|
577 |
s = "equal" orelse s = "$false" orelse s = "$true" orelse |
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents:
38518
diff
changeset
|
578 |
String.isPrefix type_const_prefix s orelse String.isPrefix class_prefix s |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
579 |
| is_predicate (SOME the_const_tab) s = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
580 |
case Symtab.lookup the_const_tab s of |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
581 |
SOME {min_arity, max_arity, sub_level} => |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
582 |
not sub_level andalso min_arity = max_arity |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
583 |
| NONE => false |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
584 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
585 |
fun repair_predicates_in_term pred_const_tab (t as ATerm ((s, _), ts)) = |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
586 |
if s = type_tag_name then |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
587 |
case ts of |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
588 |
[_, t' as ATerm ((s', _), _)] => |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
589 |
if is_predicate pred_const_tab s' then t' else boolify t |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41137
diff
changeset
|
590 |
| _ => raise Fail "malformed type tag" |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
591 |
else |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
592 |
t |> not (is_predicate pred_const_tab s) ? boolify |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
593 |
|
41134 | 594 |
fun repair_formula thy explicit_forall type_sys const_tab = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
595 |
let |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
596 |
val pred_const_tab = case type_sys of Tags _ => NONE | _ => const_tab |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
597 |
fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
598 |
| aux (AConn (c, phis)) = AConn (c, map aux phis) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
599 |
| aux (AAtom tm) = |
41134 | 600 |
AAtom (tm |> repair_applications_in_term thy type_sys const_tab |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
601 |
|> repair_predicates_in_term pred_const_tab) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
602 |
in aux #> explicit_forall ? close_universally end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
603 |
|
41134 | 604 |
fun repair_problem_line thy explicit_forall type_sys const_tab |
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
605 |
(Fof (ident, kind, phi, source)) = |
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
606 |
Fof (ident, kind, repair_formula thy explicit_forall type_sys const_tab phi, |
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
607 |
source) |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
608 |
fun repair_problem thy = map o apsnd o map ooo repair_problem_line thy |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
609 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
610 |
fun dest_Fof (Fof z) = z |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
611 |
|
41157 | 612 |
val factsN = "Relevant facts" |
613 |
val class_relsN = "Class relationships" |
|
614 |
val aritiesN = "Arity declarations" |
|
615 |
val helpersN = "Helper facts" |
|
616 |
val conjsN = "Conjectures" |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
617 |
val free_typesN = "Type variables" |
41157 | 618 |
|
619 |
fun offset_of_heading_in_problem _ [] j = j |
|
620 |
| offset_of_heading_in_problem needle ((heading, lines) :: problem) j = |
|
621 |
if heading = needle then j |
|
622 |
else offset_of_heading_in_problem needle problem (j + length lines) |
|
623 |
||
41134 | 624 |
fun prepare_atp_problem ctxt readable_names explicit_forall type_sys |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40145
diff
changeset
|
625 |
explicit_apply hyp_ts concl_t facts = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
626 |
let |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
627 |
val thy = ProofContext.theory_of ctxt |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
628 |
val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) = |
41134 | 629 |
translate_formulas ctxt type_sys hyp_ts concl_t facts |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
630 |
(* Reordering these might or might not confuse the proof reconstruction |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
631 |
code or the SPASS Flotter hack. *) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
632 |
val problem = |
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
633 |
[(factsN, map (problem_line_for_fact ctxt fact_prefix type_sys) |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
634 |
(0 upto length facts - 1 ~~ facts)), |
41157 | 635 |
(class_relsN, map problem_line_for_class_rel_clause class_rel_clauses), |
636 |
(aritiesN, map problem_line_for_arity_clause arity_clauses), |
|
637 |
(helpersN, []), |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
638 |
(conjsN, map (problem_line_for_conjecture ctxt type_sys) conjs), |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
639 |
(free_typesN, problem_lines_for_free_types type_sys conjs)] |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
640 |
val const_tab = const_table_for_problem explicit_apply problem |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
641 |
val problem = |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
642 |
problem |> repair_problem thy explicit_forall type_sys const_tab |
41157 | 643 |
val helper_lines = |
41145
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
644 |
get_helper_facts ctxt explicit_forall type_sys |
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
blanchet
parents:
41140
diff
changeset
|
645 |
(maps (map (#3 o dest_Fof) o snd) problem) |
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
646 |
|>> map (pair 0 |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
647 |
#> problem_line_for_fact ctxt helper_prefix type_sys |
41157 | 648 |
#> repair_problem_line thy explicit_forall type_sys const_tab) |
649 |
|> op @ |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
650 |
val (problem, pool) = |
41157 | 651 |
problem |> AList.update (op =) (helpersN, helper_lines) |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
652 |
|> nice_atp_problem readable_names |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
653 |
in |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
654 |
(problem, |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
655 |
case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty, |
41157 | 656 |
offset_of_heading_in_problem conjsN problem 0, |
657 |
fact_names |> Vector.fromList) |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
658 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
659 |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
660 |
(* FUDGE *) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
661 |
val conj_weight = 0.0 |
41770 | 662 |
val hyp_weight = 0.1 |
663 |
val fact_min_weight = 0.2 |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
664 |
val fact_max_weight = 1.0 |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
665 |
|
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
666 |
fun add_term_weights weight (ATerm (s, tms)) = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
667 |
(not (is_atp_variable s) andalso s <> "equal") ? Symtab.default (s, weight) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
668 |
#> fold (add_term_weights weight) tms |
41769
eb2e39555f98
extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
blanchet
parents:
41748
diff
changeset
|
669 |
fun add_problem_line_weights weight (Fof (_, _, phi, _)) = |
41384 | 670 |
fold_formula (add_term_weights weight) phi |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
671 |
|
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
672 |
fun add_conjectures_weights [] = I |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
673 |
| add_conjectures_weights conjs = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
674 |
let val (hyps, conj) = split_last conjs in |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
675 |
add_problem_line_weights conj_weight conj |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
676 |
#> fold (add_problem_line_weights hyp_weight) hyps |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
677 |
end |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
678 |
|
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
679 |
fun add_facts_weights facts = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
680 |
let |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
681 |
val num_facts = length facts |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
682 |
fun weight_of j = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
683 |
fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
684 |
/ Real.fromInt num_facts |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
685 |
in |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
686 |
map weight_of (0 upto num_facts - 1) ~~ facts |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
687 |
|> fold (uncurry add_problem_line_weights) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
688 |
end |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
689 |
|
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
690 |
(* Weights are from 0.0 (most important) to 1.0 (least important). *) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
691 |
fun atp_problem_weights problem = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
692 |
Symtab.empty |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
693 |
|> add_conjectures_weights (these (AList.lookup (op =) problem conjsN)) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
694 |
|> add_facts_weights (these (AList.lookup (op =) problem factsN)) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
695 |
|> Symtab.dest |
41726 | 696 |
|> sort (prod_ord Real.compare string_ord o pairself swap) |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
697 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
698 |
end; |