author | huffman |
Sun, 09 May 2010 17:47:43 -0700 | |
changeset 36777 | be5461582d0f |
parent 36565 | 061475351a09 |
child 36909 | 7d5587f6d5f7 |
permissions | -rw-r--r-- |
36375
2482446a604c
move the minimizer to the Sledgehammer directory
blanchet
parents:
36237
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML |
33311 | 2 |
Author: Jia Meng, NICTA |
36375
2482446a604c
move the minimizer to the Sledgehammer directory
blanchet
parents:
36237
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
4 |
|
24311 | 5 |
FOL clauses translated from HOL formulae. |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
6 |
*) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
7 |
|
35826 | 8 |
signature SLEDGEHAMMER_HOL_CLAUSE = |
24311 | 9 |
sig |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
10 |
type name = Sledgehammer_FOL_Clause.name |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
11 |
type name_pool = Sledgehammer_FOL_Clause.name_pool |
35865 | 12 |
type kind = Sledgehammer_FOL_Clause.kind |
13 |
type fol_type = Sledgehammer_FOL_Clause.fol_type |
|
14 |
type classrel_clause = Sledgehammer_FOL_Clause.classrel_clause |
|
15 |
type arity_clause = Sledgehammer_FOL_Clause.arity_clause |
|
24311 | 16 |
type axiom_name = string |
17 |
type polarity = bool |
|
35865 | 18 |
type hol_clause_id = int |
19 |
||
24311 | 20 |
datatype combterm = |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
21 |
CombConst of name * fol_type * fol_type list (* Const and Free *) | |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
22 |
CombVar of name * fol_type | |
35865 | 23 |
CombApp of combterm * combterm |
24311 | 24 |
datatype literal = Literal of polarity * combterm |
35865 | 25 |
datatype hol_clause = |
26 |
HOLClause of {clause_id: hol_clause_id, axiom_name: axiom_name, th: thm, |
|
27 |
kind: kind, literals: literal list, ctypes_sorts: typ list} |
|
28 |
||
29 |
val type_of_combterm : combterm -> fol_type |
|
30 |
val strip_combterm_comb : combterm -> combterm * combterm list |
|
31 |
val literals_of_term : theory -> term -> literal list * typ list |
|
32 |
exception TRIVIAL |
|
33 |
val make_conjecture_clauses : bool -> theory -> thm list -> hol_clause list |
|
34 |
val make_axiom_clauses : bool -> theory -> |
|
35 |
(thm * (axiom_name * hol_clause_id)) list -> (axiom_name * hol_clause) list |
|
36 |
val get_helper_clauses : bool -> theory -> bool -> |
|
37 |
hol_clause list * (thm * (axiom_name * hol_clause_id)) list * string list -> |
|
38 |
hol_clause list |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
39 |
val write_tptp_file : bool -> bool -> bool -> Path.T -> |
35865 | 40 |
hol_clause list * hol_clause list * hol_clause list * hol_clause list * |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
41 |
classrel_clause list * arity_clause list -> name_pool option * int |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
42 |
val write_dfg_file : bool -> bool -> Path.T -> |
35865 | 43 |
hol_clause list * hol_clause list * hol_clause list * hol_clause list * |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
44 |
classrel_clause list * arity_clause list -> name_pool option * int |
24311 | 45 |
end |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
46 |
|
35826 | 47 |
structure Sledgehammer_HOL_Clause : SLEDGEHAMMER_HOL_CLAUSE = |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
48 |
struct |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
49 |
|
36167
c1a35be8e476
make Sledgehammer's output more debugging friendly
blanchet
parents:
35963
diff
changeset
|
50 |
open Sledgehammer_Util |
35865 | 51 |
open Sledgehammer_FOL_Clause |
52 |
open Sledgehammer_Fact_Preprocessor |
|
22825
bd774e01d6d5
Fixing bugs in the partial-typed and fully-typed translations
paulson
parents:
22130
diff
changeset
|
53 |
|
35963 | 54 |
(* Parameter "full_types" below indicates that full type information is to be |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
55 |
exported. |
22825
bd774e01d6d5
Fixing bugs in the partial-typed and fully-typed translations
paulson
parents:
22130
diff
changeset
|
56 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
57 |
If "explicit_apply" is false, each function will be directly applied to as |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
58 |
many arguments as possible, avoiding use of the "apply" operator. Use of |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
59 |
"hBOOL" is also minimized. |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
60 |
*) |
22064 | 61 |
|
33035 | 62 |
fun min_arity_of const_min_arity c = the_default 0 (Symtab.lookup const_min_arity c); |
22064 | 63 |
|
22825
bd774e01d6d5
Fixing bugs in the partial-typed and fully-typed translations
paulson
parents:
22130
diff
changeset
|
64 |
(*True if the constant ever appears outside of the top-level position in literals. |
bd774e01d6d5
Fixing bugs in the partial-typed and fully-typed translations
paulson
parents:
22130
diff
changeset
|
65 |
If false, the constant always receives all of its arguments and is used as a predicate.*) |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
66 |
fun needs_hBOOL explicit_apply const_needs_hBOOL c = |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
67 |
explicit_apply orelse |
33035 | 68 |
the_default false (Symtab.lookup const_needs_hBOOL c); |
22064 | 69 |
|
19198
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
70 |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
71 |
(******************************************************) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
72 |
(* data types for typed combinator expressions *) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
73 |
(******************************************************) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
74 |
|
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
75 |
type axiom_name = string; |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
76 |
type polarity = bool; |
35865 | 77 |
type hol_clause_id = int; |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
78 |
|
35865 | 79 |
datatype combterm = |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
80 |
CombConst of name * fol_type * fol_type list (* Const and Free *) | |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
81 |
CombVar of name * fol_type | |
35865 | 82 |
CombApp of combterm * combterm |
24311 | 83 |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
84 |
datatype literal = Literal of polarity * combterm; |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
85 |
|
35865 | 86 |
datatype hol_clause = |
87 |
HOLClause of {clause_id: hol_clause_id, axiom_name: axiom_name, th: thm, |
|
88 |
kind: kind, literals: literal list, ctypes_sorts: typ list}; |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
89 |
|
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
90 |
|
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
91 |
(*********************************************************************) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
92 |
(* convert a clause with type Term.term to a clause with type clause *) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
93 |
(*********************************************************************) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
94 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
95 |
fun isFalse (Literal (pol, CombConst ((c, _), _, _))) = |
35865 | 96 |
(pol andalso c = "c_False") orelse (not pol andalso c = "c_True") |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
97 |
| isFalse _ = false; |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
98 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
99 |
fun isTrue (Literal (pol, CombConst ((c, _), _, _))) = |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
100 |
(pol andalso c = "c_True") orelse |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
101 |
(not pol andalso c = "c_False") |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
102 |
| isTrue _ = false; |
24311 | 103 |
|
35865 | 104 |
fun isTaut (HOLClause {literals,...}) = exists isTrue literals; |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
105 |
|
30151 | 106 |
fun type_of dfg (Type (a, Ts)) = |
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
107 |
let val (folTypes,ts) = types_of dfg Ts in |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
108 |
(TyConstr (`(make_fixed_type_const dfg) a, folTypes), ts) |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
109 |
end |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
110 |
| type_of _ (tp as TFree (a, _)) = (TyFree (`make_fixed_type_var a), [tp]) |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
111 |
| type_of _ (tp as TVar (x, _)) = |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
112 |
(TyVar (make_schematic_type_var x, string_of_indexname x), [tp]) |
30151 | 113 |
and types_of dfg Ts = |
114 |
let val (folTyps,ts) = ListPair.unzip (map (type_of dfg) Ts) |
|
35865 | 115 |
in (folTyps, union_all ts) end; |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
116 |
|
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
117 |
(* same as above, but no gathering of sort information *) |
30151 | 118 |
fun simp_type_of dfg (Type (a, Ts)) = |
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
119 |
TyConstr (`(make_fixed_type_const dfg) a, map (simp_type_of dfg) Ts) |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
120 |
| simp_type_of _ (TFree (a, _)) = TyFree (`make_fixed_type_var a) |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
121 |
| simp_type_of _ (TVar (x, _)) = |
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
122 |
TyVar (make_schematic_type_var x, string_of_indexname x); |
18440
72ee07f4b56b
Literals in clauses are sorted now. Added functions for clause equivalence tests and hashing clauses to words.
mengj
parents:
18356
diff
changeset
|
123 |
|
18356
443717b3a9ad
Added new type embedding methods for translating HOL clauses.
mengj
parents:
18276
diff
changeset
|
124 |
|
30151 | 125 |
fun const_type_of dfg thy (c,t) = |
126 |
let val (tp,ts) = type_of dfg t |
|
127 |
in (tp, ts, map (simp_type_of dfg) (Sign.const_typargs thy (c,t))) end; |
|
18356
443717b3a9ad
Added new type embedding methods for translating HOL clauses.
mengj
parents:
18276
diff
changeset
|
128 |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
129 |
(* convert a Term.term (with combinators) into a combterm, also accummulate sort info *) |
30151 | 130 |
fun combterm_of dfg thy (Const(c,t)) = |
131 |
let val (tp,ts,tvar_list) = const_type_of dfg thy (c,t) |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
132 |
val c' = CombConst (`(make_fixed_const dfg) c, tp, tvar_list) |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
133 |
in (c',ts) end |
32994 | 134 |
| combterm_of dfg _ (Free(v,t)) = |
30151 | 135 |
let val (tp,ts) = type_of dfg t |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
136 |
val v' = CombConst (`make_fixed_var v, tp, []) |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
137 |
in (v',ts) end |
32994 | 138 |
| combterm_of dfg _ (Var(v,t)) = |
30151 | 139 |
let val (tp,ts) = type_of dfg t |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
140 |
val v' = CombVar ((make_schematic_var v, string_of_indexname v), tp) |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
141 |
in (v',ts) end |
30151 | 142 |
| combterm_of dfg thy (P $ Q) = |
143 |
let val (P',tsP) = combterm_of dfg thy P |
|
144 |
val (Q',tsQ) = combterm_of dfg thy Q |
|
33042 | 145 |
in (CombApp(P',Q'), union (op =) tsP tsQ) end |
35865 | 146 |
| combterm_of _ _ (t as Abs _) = raise CLAUSE ("HOL CLAUSE", t); |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
147 |
|
35865 | 148 |
fun predicate_of dfg thy ((@{const Not} $ P), polarity) = predicate_of dfg thy (P, not polarity) |
30151 | 149 |
| predicate_of dfg thy (t,polarity) = (combterm_of dfg thy (Envir.eta_contract t), polarity); |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
150 |
|
35865 | 151 |
fun literals_of_term1 dfg thy args (@{const Trueprop} $ P) = literals_of_term1 dfg thy args P |
152 |
| literals_of_term1 dfg thy args (@{const "op |"} $ P $ Q) = |
|
30151 | 153 |
literals_of_term1 dfg thy (literals_of_term1 dfg thy args P) Q |
154 |
| literals_of_term1 dfg thy (lits,ts) P = |
|
155 |
let val ((pred,ts'),pol) = predicate_of dfg thy (P,true) |
|
21509
6c5755ad9cae
ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents:
21470
diff
changeset
|
156 |
in |
33042 | 157 |
(Literal(pol,pred)::lits, union (op =) ts ts') |
21509
6c5755ad9cae
ATP linkup now generates "new TPTP" rather than "old TPTP"
paulson
parents:
21470
diff
changeset
|
158 |
end; |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
159 |
|
30151 | 160 |
fun literals_of_term_dfg dfg thy P = literals_of_term1 dfg thy ([],[]) P; |
161 |
val literals_of_term = literals_of_term_dfg false; |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
162 |
|
35865 | 163 |
(* Trivial problem, which resolution cannot handle (empty clause) *) |
164 |
exception TRIVIAL; |
|
28835 | 165 |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
166 |
(* making axiom and conjecture clauses *) |
35865 | 167 |
fun make_clause dfg thy (clause_id, axiom_name, kind, th) = |
30151 | 168 |
let val (lits,ctypes_sorts) = literals_of_term_dfg dfg thy (prop_of th) |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
169 |
in |
35865 | 170 |
if forall isFalse lits then |
171 |
raise TRIVIAL |
|
24311 | 172 |
else |
35865 | 173 |
HOLClause {clause_id = clause_id, axiom_name = axiom_name, th = th, |
174 |
kind = kind, literals = lits, ctypes_sorts = ctypes_sorts} |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
175 |
end; |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
176 |
|
20016 | 177 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
178 |
fun add_axiom_clause dfg thy (th, (name, id)) = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
179 |
let val cls = make_clause dfg thy (id, name, Axiom, th) in |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
180 |
not (isTaut cls) ? cons (name, cls) |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
181 |
end |
19354 | 182 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
183 |
fun make_axiom_clauses dfg thy clauses = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
184 |
fold (add_axiom_clause dfg thy) clauses [] |
19354 | 185 |
|
32994 | 186 |
fun make_conjecture_clauses_aux _ _ _ [] = [] |
30151 | 187 |
| make_conjecture_clauses_aux dfg thy n (th::ths) = |
35865 | 188 |
make_clause dfg thy (n,"conjecture", Conjecture, th) :: |
30151 | 189 |
make_conjecture_clauses_aux dfg thy (n+1) ths; |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
190 |
|
30151 | 191 |
fun make_conjecture_clauses dfg thy = make_conjecture_clauses_aux dfg thy 0; |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
192 |
|
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
193 |
|
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
194 |
(**********************************************************************) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
195 |
(* convert clause into ATP specific formats: *) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
196 |
(* TPTP used by Vampire and E *) |
19720 | 197 |
(* DFG used by SPASS *) |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
198 |
(**********************************************************************) |
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
199 |
|
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
200 |
(*Result of a function type; no need to check that the argument type matches.*) |
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36168
diff
changeset
|
201 |
fun result_type (TyConstr (("tc_fun", _), [_, tp2])) = tp2 |
36168
0a6ed065683d
give more sensible names to "fol_type" constructors, as requested by a FIXME comment
blanchet
parents:
36167
diff
changeset
|
202 |
| result_type _ = raise Fail "Non-function type" |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
203 |
|
32994 | 204 |
fun type_of_combterm (CombConst (_, tp, _)) = tp |
205 |
| type_of_combterm (CombVar (_, tp)) = tp |
|
206 |
| type_of_combterm (CombApp (t1, _)) = result_type (type_of_combterm t1); |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
207 |
|
22064 | 208 |
(*gets the head of a combinator application, along with the list of arguments*) |
35865 | 209 |
fun strip_combterm_comb u = |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
210 |
let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts) |
22064 | 211 |
| stripc x = x |
212 |
in stripc(u,[]) end; |
|
213 |
||
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
214 |
val type_wrapper = "ti"; |
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
215 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
216 |
fun head_needs_hBOOL explicit_apply const_needs_hBOOL |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
217 |
(CombConst ((c, _), _, _)) = |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
218 |
needs_hBOOL explicit_apply const_needs_hBOOL c |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
219 |
| head_needs_hBOOL _ _ _ = true |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
220 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
221 |
fun wrap_type full_types (s, ty) pool = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
222 |
if full_types then |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
223 |
let val (s', pool) = string_of_fol_type ty pool in |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
224 |
(type_wrapper ^ paren_pack [s, s'], pool) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
225 |
end |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
226 |
else |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
227 |
(s, pool) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
228 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
229 |
fun wrap_type_if full_types explicit_apply cnh (head, s, tp) = |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
230 |
if head_needs_hBOOL explicit_apply cnh head then |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
231 |
wrap_type full_types (s, tp) |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
232 |
else |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
233 |
pair s |
24311 | 234 |
|
35865 | 235 |
fun apply ss = "hAPP" ^ paren_pack ss; |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
236 |
|
22064 | 237 |
fun rev_apply (v, []) = v |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
238 |
| rev_apply (v, arg::args) = apply [rev_apply (v, args), arg]; |
22064 | 239 |
|
240 |
fun string_apply (v, args) = rev_apply (v, rev args); |
|
241 |
||
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
242 |
(* Apply an operator to the argument strings, using either the "apply" operator |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
243 |
or direct function application. *) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
244 |
fun string_of_application full_types cma (CombConst ((s, s'), _, tvars), args) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
245 |
pool = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
246 |
let |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
247 |
val s = if s = "equal" then "c_fequal" else s |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
248 |
val nargs = min_arity_of cma s |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
249 |
val args1 = List.take (args, nargs) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
250 |
handle Subscript => |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
251 |
raise Fail (quote s ^ " has arity " ^ Int.toString nargs ^ |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
252 |
" but is applied to " ^ commas (map quote args)) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
253 |
val args2 = List.drop (args, nargs) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
254 |
val (targs, pool) = if full_types then ([], pool) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
255 |
else pool_map string_of_fol_type tvars pool |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
256 |
val (s, pool) = nice_name (s, s') pool |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
257 |
in (string_apply (s ^ paren_pack (args1 @ targs), args2), pool) end |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
258 |
| string_of_application _ _ (CombVar (name, _), args) pool = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
259 |
nice_name name pool |>> (fn s => string_apply (s, args)) |
24311 | 260 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
261 |
fun string_of_combterm (params as (full_types, explicit_apply, cma, cnh)) t |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
262 |
pool = |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
263 |
let |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
264 |
val (head, args) = strip_combterm_comb t |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
265 |
val (ss, pool) = pool_map (string_of_combterm params) args pool |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
266 |
val (s, pool) = string_of_application full_types cma (head, ss) pool |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
267 |
in |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
268 |
wrap_type_if full_types explicit_apply cnh (head, s, type_of_combterm t) |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
269 |
pool |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
270 |
end |
18356
443717b3a9ad
Added new type embedding methods for translating HOL clauses.
mengj
parents:
18276
diff
changeset
|
271 |
|
22064 | 272 |
(*Boolean-valued terms are here converted to literals.*) |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
273 |
fun boolify params c = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
274 |
string_of_combterm params c #>> prefix "hBOOL" o paren_pack o single |
22064 | 275 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
276 |
fun string_of_predicate (params as (_, explicit_apply, _, cnh)) t = |
22064 | 277 |
case t of |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
278 |
(CombApp (CombApp (CombConst (("equal", _), _, _), t1), t2)) => |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
279 |
(* DFG only: new TPTP prefers infix equality *) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
280 |
pool_map (string_of_combterm params) [t1, t2] |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
281 |
#>> prefix "equal" o paren_pack |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
282 |
| _ => |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
283 |
case #1 (strip_combterm_comb t) of |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
284 |
CombConst ((s, _), _, _) => |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
285 |
(if needs_hBOOL explicit_apply cnh s then boolify else string_of_combterm) |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
286 |
params t |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
287 |
| _ => boolify params t |
18356
443717b3a9ad
Added new type embedding methods for translating HOL clauses.
mengj
parents:
18276
diff
changeset
|
288 |
|
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
289 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
290 |
(*** TPTP format ***) |
19720 | 291 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
292 |
fun tptp_of_equality params pos (t1, t2) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
293 |
pool_map (string_of_combterm params) [t1, t2] |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
294 |
#>> space_implode (if pos then " = " else " != ") |
24311 | 295 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
296 |
fun tptp_literal params |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
297 |
(Literal (pos, CombApp (CombApp (CombConst (("equal", _), _, _), t1), |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
298 |
t2))) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
299 |
tptp_of_equality params pos (t1, t2) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
300 |
| tptp_literal params (Literal (pos, pred)) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
301 |
string_of_predicate params pred #>> tptp_sign pos |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
302 |
|
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
303 |
(* Given a clause, returns its literals paired with a list of literals |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
304 |
concerning TFrees; the latter should only occur in conjecture clauses. *) |
36556
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
305 |
fun tptp_type_literals params pos (HOLClause {literals, ctypes_sorts, ...}) |
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
306 |
pool = |
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
307 |
let |
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
308 |
val (lits, pool) = pool_map (tptp_literal params) literals pool |
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
309 |
val (tylits, pool) = pool_map (tptp_of_type_literal pos) |
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
310 |
(add_type_literals ctypes_sorts) pool |
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
311 |
in ((lits, tylits), pool) end |
24311 | 312 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
313 |
fun tptp_clause params (cls as HOLClause {axiom_name, clause_id, kind, ...}) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
314 |
pool = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
315 |
let |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
316 |
val ((lits, tylits), pool) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
317 |
tptp_type_literals params (kind = Conjecture) cls pool |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
318 |
in |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
319 |
((gen_tptp_cls (clause_id, axiom_name, kind, lits, tylits), tylits), pool) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
320 |
end |
17998
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
321 |
|
0a869029ca58
A new file that defines datatypes representing FOL clauses that are derived from HOL formulae. This file also has functions that convert lambda terms to combinator expressions, and functions that convert clauses to TPTP format.
mengj
parents:
diff
changeset
|
322 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
323 |
(*** DFG format ***) |
21561 | 324 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
325 |
fun dfg_literal params (Literal (pos, pred)) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
326 |
string_of_predicate params pred #>> dfg_sign pos |
19720 | 327 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
328 |
fun dfg_type_literals params pos (HOLClause {literals, ctypes_sorts, ...}) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
329 |
pool_map (dfg_literal params) literals |
36556
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
blanchet
parents:
36481
diff
changeset
|
330 |
#>> rpair (map (dfg_of_type_literal pos) (add_type_literals ctypes_sorts)) |
19720 | 331 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
332 |
fun get_uvars (CombConst _) vars pool = (vars, pool) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
333 |
| get_uvars (CombVar (name, _)) vars pool = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
334 |
nice_name name pool |>> (fn var => var :: vars) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
335 |
| get_uvars (CombApp (P, Q)) vars pool = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
336 |
let val (vars, pool) = get_uvars P vars pool in get_uvars Q vars pool end |
19720 | 337 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
338 |
fun get_uvars_l (Literal (_, c)) = get_uvars c []; |
19720 | 339 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
340 |
fun dfg_vars (HOLClause {literals, ...}) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
341 |
pool_map get_uvars_l literals #>> union_all |
24311 | 342 |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
343 |
fun dfg_clause params (cls as HOLClause {axiom_name, clause_id, kind, |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
344 |
ctypes_sorts, ...}) pool = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
345 |
let |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
346 |
val ((lits, tylits), pool) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
347 |
dfg_type_literals params (kind = Conjecture) cls pool |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
348 |
val (vars, pool) = dfg_vars cls pool |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
349 |
val tvars = get_tvar_strs ctypes_sorts |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
350 |
in |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
351 |
((gen_dfg_cls (clause_id, axiom_name, kind, lits, tylits, tvars @ vars), |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
352 |
tylits), pool) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
353 |
end |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
354 |
|
19720 | 355 |
|
22064 | 356 |
(** For DFG format: accumulate function and predicate declarations **) |
19720 | 357 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
358 |
fun add_types tvars = fold add_fol_type_funcs tvars |
19720 | 359 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
360 |
fun add_decls (full_types, explicit_apply, cma, cnh) |
36565
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
361 |
(CombConst ((c, _), ctp, tvars)) (funcs, preds) = |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
362 |
(if c = "equal" then |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
363 |
(add_types tvars funcs, preds) |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
364 |
else |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
365 |
let val arity = min_arity_of cma c |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
366 |
val ntys = if not full_types then length tvars else 0 |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
367 |
val addit = Symtab.update(c, arity + ntys) |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
368 |
in |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
369 |
if needs_hBOOL explicit_apply cnh c then |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
370 |
(add_types tvars (addit funcs), preds) |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
371 |
else |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
372 |
(add_types tvars funcs, addit preds) |
061475351a09
fix more "undeclared symbol" errors related to SPASS's DFG format
blanchet
parents:
36556
diff
changeset
|
373 |
end) |>> full_types ? add_fol_type_funcs ctp |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
374 |
| add_decls _ (CombVar (_, ctp)) (funcs, preds) = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
375 |
(add_fol_type_funcs ctp funcs, preds) |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
376 |
| add_decls params (CombApp (P, Q)) decls = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
377 |
decls |> add_decls params P |> add_decls params Q |
19720 | 378 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
379 |
fun add_literal_decls params (Literal (_, c)) = add_decls params c |
19720 | 380 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
381 |
fun add_clause_decls params (HOLClause {literals, ...}) decls = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
382 |
fold (add_literal_decls params) literals decls |
27187 | 383 |
handle Symtab.DUP a => error ("function " ^ a ^ " has multiple arities") |
19720 | 384 |
|
31800 | 385 |
fun decls_of_clauses params clauses arity_clauses = |
36237
86e62a98deea
fix bug in SPASS's DFG output files, where "tc_bool" wasn't declared;
blanchet
parents:
36235
diff
changeset
|
386 |
let val functab = |
86e62a98deea
fix bug in SPASS's DFG output files, where "tc_bool" wasn't declared;
blanchet
parents:
36235
diff
changeset
|
387 |
init_functab |> fold Symtab.update [(type_wrapper, 2), ("hAPP", 2), |
86e62a98deea
fix bug in SPASS's DFG output files, where "tc_bool" wasn't declared;
blanchet
parents:
36235
diff
changeset
|
388 |
("tc_bool", 0)] |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
389 |
val predtab = Symtab.empty |> Symtab.update ("hBOOL", 1) |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
390 |
val (functab, predtab) = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
391 |
(functab, predtab) |> fold (add_clause_decls params) clauses |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
392 |
|>> fold add_arity_clause_funcs arity_clauses |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
393 |
in (Symtab.dest functab, Symtab.dest predtab) end |
19720 | 394 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
395 |
fun add_clause_preds (HOLClause {ctypes_sorts, ...}) preds = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
396 |
fold add_type_sort_preds ctypes_sorts preds |
27187 | 397 |
handle Symtab.DUP a => error ("predicate " ^ a ^ " has multiple arities") |
21398
11996e938dfe
Includes type:sort constraints in its code to collect predicates in axiom clauses.
paulson
parents:
21254
diff
changeset
|
398 |
|
11996e938dfe
Includes type:sort constraints in its code to collect predicates in axiom clauses.
paulson
parents:
21254
diff
changeset
|
399 |
(*Higher-order clauses have only the predicates hBOOL and type classes.*) |
24311 | 400 |
fun preds_of_clauses clauses clsrel_clauses arity_clauses = |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
401 |
Symtab.empty |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
402 |
|> fold add_clause_preds clauses |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
403 |
|> fold add_arity_clause_preds arity_clauses |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
404 |
|> fold add_classrel_clause_preds clsrel_clauses |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
405 |
|> Symtab.dest |
18440
72ee07f4b56b
Literals in clauses are sorted now. Added functions for clause equivalence tests and hashing clauses to words.
mengj
parents:
18356
diff
changeset
|
406 |
|
72ee07f4b56b
Literals in clauses are sorted now. Added functions for clause equivalence tests and hashing clauses to words.
mengj
parents:
18356
diff
changeset
|
407 |
(**********************************************************************) |
19198
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
408 |
(* write clauses to files *) |
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
409 |
(**********************************************************************) |
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
410 |
|
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
411 |
val init_counters = |
35865 | 412 |
Symtab.make [("c_COMBI", 0), ("c_COMBK", 0), ("c_COMBB", 0), ("c_COMBC", 0), |
413 |
("c_COMBS", 0)]; |
|
24311 | 414 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
415 |
fun count_combterm (CombConst ((c, _), _, _)) = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
416 |
Symtab.map_entry c (Integer.add 1) |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
417 |
| count_combterm (CombVar _) = I |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
418 |
| count_combterm (CombApp (t1, t2)) = count_combterm t1 #> count_combterm t2 |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
419 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
420 |
fun count_literal (Literal (_, t)) = count_combterm t |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
421 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
422 |
fun count_clause (HOLClause {literals, ...}) = fold count_literal literals |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
423 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
424 |
fun count_user_clause user_lemmas (HOLClause {axiom_name, literals, ...}) = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
425 |
member (op =) user_lemmas axiom_name ? fold count_literal literals |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
426 |
|
36228
df47dc6c0e03
got rid of somewhat pointless "pairname" function in Sledgehammer
blanchet
parents:
36222
diff
changeset
|
427 |
fun cnf_helper_thms thy = cnf_rules_pairs thy o map (`Thm.get_name_hint) |
20644 | 428 |
|
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
429 |
fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) = |
35865 | 430 |
if isFO then |
431 |
[] |
|
23386 | 432 |
else |
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
433 |
let |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
434 |
val axclauses = map #2 (make_axiom_clauses dfg thy axcls) |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
435 |
val ct = init_counters |> fold count_clause conjectures |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
436 |
|> fold (count_user_clause user_lemmas) axclauses |
33035 | 437 |
fun needed c = the (Symtab.lookup ct c) > 0 |
24311 | 438 |
val IK = if needed "c_COMBI" orelse needed "c_COMBK" |
35865 | 439 |
then cnf_helper_thms thy [@{thm COMBI_def}, @{thm COMBK_def}] |
24311 | 440 |
else [] |
441 |
val BC = if needed "c_COMBB" orelse needed "c_COMBC" |
|
35865 | 442 |
then cnf_helper_thms thy [@{thm COMBB_def}, @{thm COMBC_def}] |
21135 | 443 |
else [] |
35865 | 444 |
val S = if needed "c_COMBS" then cnf_helper_thms thy [@{thm COMBS_def}] |
24311 | 445 |
else [] |
35865 | 446 |
val other = cnf_helper_thms thy [@{thm fequal_imp_equal}, |
447 |
@{thm equal_imp_fequal}] |
|
20791
497e1c9d4a9f
Combinator axioms are now from Isabelle theorems, no need to use helper files.
mengj
parents:
20660
diff
changeset
|
448 |
in |
30151 | 449 |
map #2 (make_axiom_clauses dfg thy (other @ IK @ BC @ S)) |
23386 | 450 |
end; |
20791
497e1c9d4a9f
Combinator axioms are now from Isabelle theorems, no need to use helper files.
mengj
parents:
20660
diff
changeset
|
451 |
|
22064 | 452 |
(*Find the minimal arity of each function mentioned in the term. Also, note which uses |
453 |
are not at top level, to see if hBOOL is needed.*) |
|
30150 | 454 |
fun count_constants_term toplev t (const_min_arity, const_needs_hBOOL) = |
35865 | 455 |
let val (head, args) = strip_combterm_comb t |
22064 | 456 |
val n = length args |
36233 | 457 |
val (const_min_arity, const_needs_hBOOL) = |
458 |
(const_min_arity, const_needs_hBOOL) |
|
459 |
|> fold (count_constants_term false) args |
|
22064 | 460 |
in |
461 |
case head of |
|
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
462 |
CombConst ((a, _),_,_) => (*predicate or function version of "equal"?*) |
24311 | 463 |
let val a = if a="equal" andalso not toplev then "c_fequal" else a |
464 |
in |
|
36233 | 465 |
(const_min_arity |> Symtab.map_default (a, n) (Integer.min n), |
466 |
const_needs_hBOOL |> not toplev ? Symtab.update (a, true)) |
|
24311 | 467 |
end |
32994 | 468 |
| _ => (const_min_arity, const_needs_hBOOL) |
22064 | 469 |
end; |
470 |
||
471 |
(*A literal is a top-level term*) |
|
30150 | 472 |
fun count_constants_lit (Literal (_,t)) (const_min_arity, const_needs_hBOOL) = |
473 |
count_constants_term true t (const_min_arity, const_needs_hBOOL); |
|
22064 | 474 |
|
35865 | 475 |
fun count_constants_clause (HOLClause {literals, ...}) |
476 |
(const_min_arity, const_needs_hBOOL) = |
|
30150 | 477 |
fold count_constants_lit literals (const_min_arity, const_needs_hBOOL); |
22064 | 478 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
479 |
fun display_arity explicit_apply const_needs_hBOOL (c,n) = |
35865 | 480 |
trace_msg (fn () => "Constant: " ^ c ^ |
35826 | 481 |
" arity:\t" ^ Int.toString n ^ |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
482 |
(if needs_hBOOL explicit_apply const_needs_hBOOL c then |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
483 |
" needs hBOOL" |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
484 |
else |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
485 |
"")) |
22064 | 486 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
487 |
fun count_constants explicit_apply |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
488 |
(conjectures, _, extra_clauses, helper_clauses, _, _) = |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
489 |
if not explicit_apply then |
30150 | 490 |
let val (const_min_arity, const_needs_hBOOL) = |
491 |
fold count_constants_clause conjectures (Symtab.empty, Symtab.empty) |
|
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
492 |
|> fold count_constants_clause extra_clauses |
30149 | 493 |
|> fold count_constants_clause helper_clauses |
36481
af99c98121d6
make Sledgehammer more friendly if no subgoal is left
blanchet
parents:
36402
diff
changeset
|
494 |
val _ = app (display_arity explicit_apply const_needs_hBOOL) |
af99c98121d6
make Sledgehammer more friendly if no subgoal is left
blanchet
parents:
36402
diff
changeset
|
495 |
(Symtab.dest (const_min_arity)) |
30150 | 496 |
in (const_min_arity, const_needs_hBOOL) end |
497 |
else (Symtab.empty, Symtab.empty); |
|
22064 | 498 |
|
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
499 |
fun header () = |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
500 |
"% This file was generated by Isabelle (most likely Sledgehammer)\n" ^ |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
501 |
"% " ^ timestamp () ^ "\n" |
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
502 |
|
35865 | 503 |
(* TPTP format *) |
31749 | 504 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
505 |
fun write_tptp_file readable_names full_types explicit_apply file clauses = |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
506 |
let |
36167
c1a35be8e476
make Sledgehammer's output more debugging friendly
blanchet
parents:
35963
diff
changeset
|
507 |
fun section _ [] = [] |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
508 |
| section name ss = |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
509 |
"\n% " ^ name ^ plural_s (length ss) ^ " (" ^ Int.toString (length ss) ^ |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
510 |
")\n" :: ss |
36222
0e3e49bd658d
don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents:
36219
diff
changeset
|
511 |
val pool = empty_name_pool readable_names |
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
512 |
val (conjectures, axclauses, _, helper_clauses, |
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
513 |
classrel_clauses, arity_clauses) = clauses |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
514 |
val (cma, cnh) = count_constants explicit_apply clauses |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
515 |
val params = (full_types, explicit_apply, cma, cnh) |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
516 |
val ((conjecture_clss, tfree_litss), pool) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
517 |
pool_map (tptp_clause params) conjectures pool |>> ListPair.unzip |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
518 |
val tfree_clss = map tptp_tfree_clause (fold (union (op =)) tfree_litss []) |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
519 |
val (ax_clss, pool) = pool_map (apfst fst oo tptp_clause params) axclauses |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
520 |
pool |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
521 |
val classrel_clss = map tptp_classrel_clause classrel_clauses |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
522 |
val arity_clss = map tptp_arity_clause arity_clauses |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
523 |
val (helper_clss, pool) = pool_map (apfst fst oo tptp_clause params) |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
524 |
helper_clauses pool |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
525 |
val conjecture_offset = |
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
526 |
length ax_clss + length classrel_clss + length arity_clss |
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
527 |
+ length helper_clss |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
528 |
val _ = |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
529 |
File.write_list file |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
530 |
(header () :: |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
531 |
section "Relevant fact" ax_clss @ |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
532 |
section "Class relationship" classrel_clss @ |
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
533 |
section "Arity declaration" arity_clss @ |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
534 |
section "Helper fact" helper_clss @ |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
535 |
section "Conjecture" conjecture_clss @ |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
536 |
section "Type variable" tfree_clss) |
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
537 |
in (pool, conjecture_offset) end |
19720 | 538 |
|
35865 | 539 |
(* DFG format *) |
19720 | 540 |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
541 |
fun write_dfg_file full_types explicit_apply file clauses = |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
542 |
let |
36219
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents:
36218
diff
changeset
|
543 |
(* Some of the helper functions below are not name-pool-aware. However, |
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents:
36218
diff
changeset
|
544 |
there's no point in adding name pool support if the DFG format is on its |
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents:
36218
diff
changeset
|
545 |
way out anyway. *) |
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents:
36218
diff
changeset
|
546 |
val pool = NONE |
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
547 |
val (conjectures, axclauses, _, helper_clauses, |
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
548 |
classrel_clauses, arity_clauses) = clauses |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
549 |
val (cma, cnh) = count_constants explicit_apply clauses |
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36233
diff
changeset
|
550 |
val params = (full_types, explicit_apply, cma, cnh) |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
551 |
val ((conjecture_clss, tfree_litss), pool) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
552 |
pool_map (dfg_clause params) conjectures pool |>> ListPair.unzip |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
553 |
and problem_name = Path.implode (Path.base file) |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
554 |
val (axstrs, pool) = pool_map (apfst fst oo dfg_clause params) axclauses pool |
35865 | 555 |
val tfree_clss = map dfg_tfree_clause (union_all tfree_litss) |
36170
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
556 |
val (helper_clauses_strs, pool) = |
0cdb76723c88
added original constant names to Sledgehammer internal terms + output short names if "debug" is set (for increased readability)
blanchet
parents:
36169
diff
changeset
|
557 |
pool_map (apfst fst oo dfg_clause params) helper_clauses pool |
36218
0e4a01f3e7d3
get rid of "List.foldl" + add timestamp to SPASS
blanchet
parents:
36186
diff
changeset
|
558 |
val (funcs, cl_preds) = decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
559 |
and ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
560 |
val conjecture_offset = |
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
561 |
length axclauses + length classrel_clauses + length arity_clauses |
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
562 |
+ length helper_clauses |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
563 |
val _ = |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
564 |
File.write_list file |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
565 |
(header () :: |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
566 |
string_of_start problem_name :: |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
567 |
string_of_descrip problem_name :: |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
568 |
string_of_symbols (string_of_funcs funcs) |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
569 |
(string_of_preds (cl_preds @ ty_preds)) :: |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
570 |
"list_of_clauses(axioms, cnf).\n" :: |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
571 |
axstrs @ |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
572 |
map dfg_classrel_clause classrel_clauses @ |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
573 |
map dfg_arity_clause arity_clauses @ |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
574 |
helper_clauses_strs @ |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
575 |
["end_of_list.\n\nlist_of_clauses(conjectures, cnf).\n"] @ |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
576 |
conjecture_clss @ |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
577 |
tfree_clss @ |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
578 |
["end_of_list.\n\n", |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36375
diff
changeset
|
579 |
"end_problem.\n"]) |
36402
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
blanchet
parents:
36393
diff
changeset
|
580 |
in (pool, conjecture_offset) end |
19720 | 581 |
|
33311 | 582 |
end; |