author | blanchet |
Wed, 17 Mar 2010 19:37:44 +0100 | |
changeset 35827 | f552152d7747 |
parent 35826 | 1590abc3d42a |
child 35865 | 2f8fb5242799 |
permissions | -rw-r--r-- |
35826 | 1 |
(* Title: HOL/Sledgehammer/sledgehammer_hol_clause.ML |
33311 | 2 |
Author: Jia Meng, NICTA |
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
|
3 |
|
24311 | 4 |
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
|
5 |
*) |
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 |
|
35826 | 7 |
signature SLEDGEHAMMER_HOL_CLAUSE = |
24311 | 8 |
sig |
9 |
val ext: thm |
|
10 |
val comb_I: thm |
|
11 |
val comb_K: thm |
|
12 |
val comb_B: thm |
|
13 |
val comb_C: thm |
|
14 |
val comb_S: thm |
|
30153
051d3825a15d
turned "read-only refs" typ_level and minimize_applies into constant values;
wenzelm
parents:
30151
diff
changeset
|
15 |
val minimize_applies: bool |
24311 | 16 |
type axiom_name = string |
17 |
type polarity = bool |
|
18 |
type clause_id = int |
|
19 |
datatype combterm = |
|
35826 | 20 |
CombConst of string * Sledgehammer_FOL_Clause.fol_type * Sledgehammer_FOL_Clause.fol_type list (*Const and Free*) |
21 |
| CombVar of string * Sledgehammer_FOL_Clause.fol_type |
|
24311 | 22 |
| CombApp of combterm * combterm |
23 |
datatype literal = Literal of polarity * combterm |
|
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
24 |
datatype clause = Clause of {clause_id: clause_id, axiom_name: axiom_name, th: thm, |
35826 | 25 |
kind: Sledgehammer_FOL_Clause.kind,literals: literal list, ctypes_sorts: typ list} |
26 |
val type_of_combterm: combterm -> Sledgehammer_FOL_Clause.fol_type |
|
24311 | 27 |
val strip_comb: combterm -> combterm * combterm list |
24940 | 28 |
val literals_of_term: theory -> term -> literal list * typ list |
28835 | 29 |
exception TOO_TRIVIAL |
31838 | 30 |
val make_conjecture_clauses: bool -> theory -> thm list -> clause list |
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
31 |
val make_axiom_clauses: bool -> |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
32 |
theory -> |
31838 | 33 |
(thm * (axiom_name * clause_id)) list -> (axiom_name * clause) list |
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
34 |
val get_helper_clauses: bool -> |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
35 |
theory -> |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
36 |
bool -> |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
37 |
clause list * (thm * (axiom_name * clause_id)) list * string list -> |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
38 |
clause list |
31838 | 39 |
val tptp_write_file: bool -> Path.T -> |
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
40 |
clause list * clause list * clause list * clause list * |
35826 | 41 |
Sledgehammer_FOL_Clause.classrelClause list * Sledgehammer_FOL_Clause.arityClause list -> |
31839
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
42 |
int * int |
31838 | 43 |
val dfg_write_file: bool -> Path.T -> |
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
44 |
clause list * clause list * clause list * clause list * |
35826 | 45 |
Sledgehammer_FOL_Clause.classrelClause list * Sledgehammer_FOL_Clause.arityClause list -> |
31839
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
46 |
int * int |
24311 | 47 |
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
|
48 |
|
35826 | 49 |
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
|
50 |
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
|
51 |
|
35826 | 52 |
structure SFC = Sledgehammer_FOL_Clause; |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
53 |
|
20791
497e1c9d4a9f
Combinator axioms are now from Isabelle theorems, no need to use helper files.
mengj
parents:
20660
diff
changeset
|
54 |
(* theorems for combinators and function extensionality *) |
497e1c9d4a9f
Combinator axioms are now from Isabelle theorems, no need to use helper files.
mengj
parents:
20660
diff
changeset
|
55 |
val ext = thm "HOL.ext"; |
35827 | 56 |
val comb_I = thm "Sledgehammer.COMBI_def"; |
57 |
val comb_K = thm "Sledgehammer.COMBK_def"; |
|
58 |
val comb_B = thm "Sledgehammer.COMBB_def"; |
|
59 |
val comb_C = thm "Sledgehammer.COMBC_def"; |
|
60 |
val comb_S = thm "Sledgehammer.COMBS_def"; |
|
61 |
val fequal_imp_equal = thm "Sledgehammer.fequal_imp_equal"; |
|
62 |
val equal_imp_fequal = thm "Sledgehammer.equal_imp_fequal"; |
|
20791
497e1c9d4a9f
Combinator axioms are now from Isabelle theorems, no need to use helper files.
mengj
parents:
20660
diff
changeset
|
63 |
|
22825
bd774e01d6d5
Fixing bugs in the partial-typed and fully-typed translations
paulson
parents:
22130
diff
changeset
|
64 |
|
31800 | 65 |
(* Parameter t_full below indicates that full type information is to be |
66 |
exported *) |
|
22825
bd774e01d6d5
Fixing bugs in the partial-typed and fully-typed translations
paulson
parents:
22130
diff
changeset
|
67 |
|
22064 | 68 |
(*If true, each function will be directly applied to as many arguments as possible, avoiding |
24385
ab62948281a9
Deleted the partially-typed translation and the combinator code.
paulson
parents:
24323
diff
changeset
|
69 |
use of the "apply" operator. Use of hBOOL is also minimized.*) |
30153
051d3825a15d
turned "read-only refs" typ_level and minimize_applies into constant values;
wenzelm
parents:
30151
diff
changeset
|
70 |
val minimize_applies = true; |
22064 | 71 |
|
33035 | 72 |
fun min_arity_of const_min_arity c = the_default 0 (Symtab.lookup const_min_arity c); |
22064 | 73 |
|
22825
bd774e01d6d5
Fixing bugs in the partial-typed and fully-typed translations
paulson
parents:
22130
diff
changeset
|
74 |
(*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
|
75 |
If false, the constant always receives all of its arguments and is used as a predicate.*) |
33035 | 76 |
fun needs_hBOOL const_needs_hBOOL c = |
77 |
not minimize_applies orelse |
|
78 |
the_default false (Symtab.lookup const_needs_hBOOL c); |
|
22064 | 79 |
|
19198
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
80 |
|
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
|
81 |
(******************************************************) |
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
|
82 |
(* 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
|
83 |
(******************************************************) |
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 |
|
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 |
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
|
86 |
type polarity = bool; |
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
|
87 |
type clause_id = int; |
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
|
88 |
|
35826 | 89 |
datatype combterm = CombConst of string * SFC.fol_type * SFC.fol_type list (*Const and Free*) |
90 |
| CombVar of string * SFC.fol_type |
|
24311 | 91 |
| CombApp of combterm * combterm |
92 |
||
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
|
93 |
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
|
94 |
|
24311 | 95 |
datatype clause = |
96 |
Clause of {clause_id: clause_id, |
|
97 |
axiom_name: axiom_name, |
|
98 |
th: thm, |
|
35826 | 99 |
kind: SFC.kind, |
24311 | 100 |
literals: literal list, |
24940 | 101 |
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
|
102 |
|
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
|
103 |
|
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
|
104 |
(*********************************************************************) |
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 |
(* 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
|
106 |
(*********************************************************************) |
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
|
107 |
|
21561 | 108 |
fun isFalse (Literal(pol, CombConst(c,_,_))) = |
20360 | 109 |
(pol andalso c = "c_False") orelse |
110 |
(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
|
111 |
| 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
|
112 |
|
21561 | 113 |
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
|
114 |
(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
|
115 |
(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
|
116 |
| isTrue _ = false; |
24311 | 117 |
|
118 |
fun isTaut (Clause {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
|
119 |
|
30151 | 120 |
fun type_of dfg (Type (a, Ts)) = |
121 |
let val (folTypes,ts) = types_of dfg Ts |
|
35826 | 122 |
in (SFC.Comp(SFC.make_fixed_type_const dfg a, folTypes), ts) end |
32994 | 123 |
| type_of _ (tp as TFree (a, _)) = |
35826 | 124 |
(SFC.AtomF (SFC.make_fixed_type_var a), [tp]) |
32994 | 125 |
| type_of _ (tp as TVar (v, _)) = |
35826 | 126 |
(SFC.AtomV (SFC.make_schematic_type_var v), [tp]) |
30151 | 127 |
and types_of dfg Ts = |
128 |
let val (folTyps,ts) = ListPair.unzip (map (type_of dfg) Ts) |
|
35826 | 129 |
in (folTyps, SFC.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
|
130 |
|
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
|
131 |
(* same as above, but no gathering of sort information *) |
30151 | 132 |
fun simp_type_of dfg (Type (a, Ts)) = |
35826 | 133 |
SFC.Comp(SFC.make_fixed_type_const dfg a, map (simp_type_of dfg) Ts) |
134 |
| simp_type_of _ (TFree (a, _)) = SFC.AtomF(SFC.make_fixed_type_var a) |
|
135 |
| simp_type_of _ (TVar (v, _)) = SFC.AtomV(SFC.make_schematic_type_var v); |
|
18440
72ee07f4b56b
Literals in clauses are sorted now. Added functions for clause equivalence tests and hashing clauses to words.
mengj
parents:
18356
diff
changeset
|
136 |
|
18356
443717b3a9ad
Added new type embedding methods for translating HOL clauses.
mengj
parents:
18276
diff
changeset
|
137 |
|
30151 | 138 |
fun const_type_of dfg thy (c,t) = |
139 |
let val (tp,ts) = type_of dfg t |
|
140 |
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
|
141 |
|
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
|
142 |
(* convert a Term.term (with combinators) into a combterm, also accummulate sort info *) |
30151 | 143 |
fun combterm_of dfg thy (Const(c,t)) = |
144 |
let val (tp,ts,tvar_list) = const_type_of dfg thy (c,t) |
|
35826 | 145 |
val c' = CombConst(SFC.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
|
146 |
in (c',ts) end |
32994 | 147 |
| combterm_of dfg _ (Free(v,t)) = |
30151 | 148 |
let val (tp,ts) = type_of dfg t |
35826 | 149 |
val v' = CombConst(SFC.make_fixed_var v, tp, []) |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
150 |
in (v',ts) end |
32994 | 151 |
| combterm_of dfg _ (Var(v,t)) = |
30151 | 152 |
let val (tp,ts) = type_of dfg t |
35826 | 153 |
val v' = CombVar(SFC.make_schematic_var v,tp) |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
154 |
in (v',ts) end |
30151 | 155 |
| combterm_of dfg thy (P $ Q) = |
156 |
let val (P',tsP) = combterm_of dfg thy P |
|
157 |
val (Q',tsQ) = combterm_of dfg thy Q |
|
33042 | 158 |
in (CombApp(P',Q'), union (op =) tsP tsQ) end |
35826 | 159 |
| combterm_of _ _ (t as Abs _) = raise SFC.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
|
160 |
|
30151 | 161 |
fun predicate_of dfg thy ((Const("Not",_) $ P), polarity) = predicate_of dfg thy (P, not polarity) |
162 |
| 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
|
163 |
|
30151 | 164 |
fun literals_of_term1 dfg thy args (Const("Trueprop",_) $ P) = literals_of_term1 dfg thy args P |
165 |
| literals_of_term1 dfg thy args (Const("op |",_) $ P $ Q) = |
|
166 |
literals_of_term1 dfg thy (literals_of_term1 dfg thy args P) Q |
|
167 |
| literals_of_term1 dfg thy (lits,ts) P = |
|
168 |
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
|
169 |
in |
33042 | 170 |
(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
|
171 |
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
|
172 |
|
30151 | 173 |
fun literals_of_term_dfg dfg thy P = literals_of_term1 dfg thy ([],[]) P; |
174 |
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
|
175 |
|
28835 | 176 |
(* Problem too trivial for resolution (empty clause) *) |
177 |
exception TOO_TRIVIAL; |
|
178 |
||
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
|
179 |
(* making axiom and conjecture clauses *) |
30151 | 180 |
fun make_clause dfg thy (clause_id,axiom_name,kind,th) = |
181 |
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
|
182 |
in |
24311 | 183 |
if forall isFalse lits |
28835 | 184 |
then raise TOO_TRIVIAL |
24311 | 185 |
else |
186 |
Clause {clause_id = clause_id, axiom_name = axiom_name, th = th, kind = kind, |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
187 |
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
|
188 |
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
|
189 |
|
20016 | 190 |
|
30151 | 191 |
fun add_axiom_clause dfg thy ((th,(name,id)), pairs) = |
35826 | 192 |
let val cls = make_clause dfg thy (id, name, SFC.Axiom, th) |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
193 |
in |
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
194 |
if isTaut cls then pairs else (name,cls)::pairs |
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
195 |
end; |
19354 | 196 |
|
30190 | 197 |
fun make_axiom_clauses dfg thy = List.foldl (add_axiom_clause dfg thy) []; |
19354 | 198 |
|
32994 | 199 |
fun make_conjecture_clauses_aux _ _ _ [] = [] |
30151 | 200 |
| make_conjecture_clauses_aux dfg thy n (th::ths) = |
35826 | 201 |
make_clause dfg thy (n,"conjecture", SFC.Conjecture, th) :: |
30151 | 202 |
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
|
203 |
|
30151 | 204 |
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
|
205 |
|
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
|
206 |
|
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 |
(**********************************************************************) |
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
|
208 |
(* 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
|
209 |
(* TPTP used by Vampire and E *) |
19720 | 210 |
(* 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
|
211 |
(**********************************************************************) |
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
|
212 |
|
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
213 |
(*Result of a function type; no need to check that the argument type matches.*) |
35826 | 214 |
fun result_type (SFC.Comp ("tc_fun", [_, tp2])) = tp2 |
27187 | 215 |
| result_type _ = error "result_type" |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
216 |
|
32994 | 217 |
fun type_of_combterm (CombConst (_, tp, _)) = tp |
218 |
| type_of_combterm (CombVar (_, tp)) = tp |
|
219 |
| 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
|
220 |
|
22064 | 221 |
(*gets the head of a combinator application, along with the list of arguments*) |
222 |
fun strip_comb u = |
|
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
223 |
let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts) |
22064 | 224 |
| stripc x = x |
225 |
in stripc(u,[]) end; |
|
226 |
||
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
227 |
val type_wrapper = "ti"; |
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
228 |
|
30150 | 229 |
fun head_needs_hBOOL const_needs_hBOOL (CombConst(c,_,_)) = needs_hBOOL const_needs_hBOOL c |
32994 | 230 |
| head_needs_hBOOL _ _ = true; |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
231 |
|
31791 | 232 |
fun wrap_type t_full (s, tp) = |
233 |
if t_full then |
|
35826 | 234 |
type_wrapper ^ SFC.paren_pack [s, SFC.string_of_fol_type tp] |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
235 |
else s; |
24311 | 236 |
|
35826 | 237 |
fun apply ss = "hAPP" ^ SFC.paren_pack ss; |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
238 |
|
22064 | 239 |
fun rev_apply (v, []) = v |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
240 |
| rev_apply (v, arg::args) = apply [rev_apply (v, args), arg]; |
22064 | 241 |
|
242 |
fun string_apply (v, args) = rev_apply (v, rev args); |
|
243 |
||
244 |
(*Apply an operator to the argument strings, using either the "apply" operator or |
|
245 |
direct function application.*) |
|
32994 | 246 |
fun string_of_applic t_full cma (CombConst (c, _, tvars), args) = |
22064 | 247 |
let val c = if c = "equal" then "c_fequal" else c |
30150 | 248 |
val nargs = min_arity_of cma c |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
249 |
val args1 = List.take(args, nargs) |
27187 | 250 |
handle Subscript => error ("string_of_applic: " ^ c ^ " has arity " ^ |
251 |
Int.toString nargs ^ " but is applied to " ^ |
|
252 |
space_implode ", " args) |
|
22064 | 253 |
val args2 = List.drop(args, nargs) |
35826 | 254 |
val targs = if not t_full then map SFC.string_of_fol_type tvars |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
255 |
else [] |
21513
9e9fff87dc6c
Conversion of "equal" to "=" for TSTP format; big tidy-up
paulson
parents:
21509
diff
changeset
|
256 |
in |
35826 | 257 |
string_apply (c ^ SFC.paren_pack (args1@targs), args2) |
21513
9e9fff87dc6c
Conversion of "equal" to "=" for TSTP format; big tidy-up
paulson
parents:
21509
diff
changeset
|
258 |
end |
32994 | 259 |
| string_of_applic _ _ (CombVar (v, _), args) = string_apply (v, args) |
31791 | 260 |
| string_of_applic _ _ _ = error "string_of_applic"; |
24311 | 261 |
|
31791 | 262 |
fun wrap_type_if t_full cnh (head, s, tp) = |
263 |
if head_needs_hBOOL cnh head then wrap_type t_full (s, tp) else s; |
|
264 |
||
31800 | 265 |
fun string_of_combterm (params as (t_full, cma, cnh)) t = |
22064 | 266 |
let val (head, args) = strip_comb t |
31791 | 267 |
in wrap_type_if t_full cnh (head, |
31800 | 268 |
string_of_applic t_full cma (head, map (string_of_combterm (params)) args), |
22851
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
diff
changeset
|
269 |
type_of_combterm t) |
7b7d6e1c70b6
First-order variant of the fully-typed translation
paulson
parents:
22825
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.*) |
31800 | 273 |
fun boolify params t = |
35826 | 274 |
"hBOOL" ^ SFC.paren_pack [string_of_combterm params t]; |
22064 | 275 |
|
31800 | 276 |
fun string_of_predicate (params as (_,_,cnh)) t = |
22064 | 277 |
case t of |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
278 |
(CombApp(CombApp(CombConst("equal",_,_), t1), t2)) => |
24311 | 279 |
(*DFG only: new TPTP prefers infix equality*) |
35826 | 280 |
("equal" ^ SFC.paren_pack [string_of_combterm params t1, string_of_combterm params t2]) |
24311 | 281 |
| _ => |
22064 | 282 |
case #1 (strip_comb t) of |
31800 | 283 |
CombConst(c,_,_) => if needs_hBOOL cnh c then boolify params t else string_of_combterm params t |
284 |
| _ => boolify params t; |
|
18356
443717b3a9ad
Added new type embedding methods for translating HOL clauses.
mengj
parents:
18276
diff
changeset
|
285 |
|
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
|
286 |
|
21561 | 287 |
(*** tptp format ***) |
19720 | 288 |
|
31800 | 289 |
fun tptp_of_equality params pol (t1,t2) = |
21513
9e9fff87dc6c
Conversion of "equal" to "=" for TSTP format; big tidy-up
paulson
parents:
21509
diff
changeset
|
290 |
let val eqop = if pol then " = " else " != " |
31800 | 291 |
in string_of_combterm params t1 ^ eqop ^ string_of_combterm params t2 end; |
21513
9e9fff87dc6c
Conversion of "equal" to "=" for TSTP format; big tidy-up
paulson
parents:
21509
diff
changeset
|
292 |
|
31800 | 293 |
fun tptp_literal params (Literal(pol, CombApp(CombApp(CombConst("equal",_,_), t1), t2))) = |
294 |
tptp_of_equality params pol (t1,t2) |
|
295 |
| tptp_literal params (Literal(pol,pred)) = |
|
35826 | 296 |
SFC.tptp_sign pol (string_of_predicate params pred); |
24311 | 297 |
|
22064 | 298 |
(*Given a clause, returns its literals paired with a list of literals concerning TFrees; |
299 |
the latter should only occur in conjecture clauses.*) |
|
31800 | 300 |
fun tptp_type_lits params pos (Clause{literals, ctypes_sorts, ...}) = |
301 |
(map (tptp_literal params) literals, |
|
35826 | 302 |
map (SFC.tptp_of_typeLit pos) (SFC.add_typs ctypes_sorts)); |
24311 | 303 |
|
32994 | 304 |
fun clause2tptp params (cls as Clause {axiom_name, clause_id, kind, ...}) = |
35826 | 305 |
let val (lits,tylits) = tptp_type_lits params (kind = SFC.Conjecture) cls |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
306 |
in |
35826 | 307 |
(SFC.gen_tptp_cls(clause_id,axiom_name,kind,lits,tylits), tylits) |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
308 |
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
|
309 |
|
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
|
310 |
|
21561 | 311 |
(*** dfg format ***) |
312 |
||
35826 | 313 |
fun dfg_literal params (Literal(pol,pred)) = SFC.dfg_sign pol (string_of_predicate params pred); |
19720 | 314 |
|
31800 | 315 |
fun dfg_type_lits params pos (Clause{literals, ctypes_sorts, ...}) = |
316 |
(map (dfg_literal params) literals, |
|
35826 | 317 |
map (SFC.dfg_of_typeLit pos) (SFC.add_typs ctypes_sorts)); |
19720 | 318 |
|
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
319 |
fun get_uvars (CombConst _) vars = vars |
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
320 |
| get_uvars (CombVar(v,_)) vars = (v::vars) |
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
321 |
| get_uvars (CombApp(P,Q)) vars = get_uvars P (get_uvars Q vars); |
19720 | 322 |
|
323 |
fun get_uvars_l (Literal(_,c)) = get_uvars c []; |
|
324 |
||
35826 | 325 |
fun dfg_vars (Clause {literals,...}) = SFC.union_all (map get_uvars_l literals); |
24311 | 326 |
|
31800 | 327 |
fun clause2dfg params (cls as Clause{axiom_name,clause_id,kind,ctypes_sorts,...}) = |
35826 | 328 |
let val (lits,tylits) = dfg_type_lits params (kind = SFC.Conjecture) cls |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
329 |
val vars = dfg_vars cls |
35826 | 330 |
val tvars = SFC.get_tvar_strs ctypes_sorts |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
331 |
in |
35826 | 332 |
(SFC.gen_dfg_cls(clause_id, axiom_name, kind, lits, tylits, tvars@vars), tylits) |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
333 |
end; |
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24385
diff
changeset
|
334 |
|
19720 | 335 |
|
22064 | 336 |
(** For DFG format: accumulate function and predicate declarations **) |
19720 | 337 |
|
35826 | 338 |
fun addtypes tvars tab = List.foldl SFC.add_foltype_funcs tab tvars; |
19720 | 339 |
|
32994 | 340 |
fun add_decls (t_full, cma, cnh) (CombConst (c, _, tvars), (funcs, preds)) = |
22064 | 341 |
if c = "equal" then (addtypes tvars funcs, preds) |
21561 | 342 |
else |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32955
diff
changeset
|
343 |
let val arity = min_arity_of cma c |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32955
diff
changeset
|
344 |
val ntys = if not t_full then length tvars else 0 |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32955
diff
changeset
|
345 |
val addit = Symtab.update(c, arity+ntys) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32955
diff
changeset
|
346 |
in |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32955
diff
changeset
|
347 |
if needs_hBOOL cnh c then (addtypes tvars (addit funcs), preds) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32955
diff
changeset
|
348 |
else (addtypes tvars funcs, addit preds) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32955
diff
changeset
|
349 |
end |
31800 | 350 |
| add_decls _ (CombVar(_,ctp), (funcs,preds)) = |
35826 | 351 |
(SFC.add_foltype_funcs (ctp,funcs), preds) |
31800 | 352 |
| add_decls params (CombApp(P,Q),decls) = add_decls params (P,add_decls params (Q,decls)); |
19720 | 353 |
|
31800 | 354 |
fun add_literal_decls params (Literal(_,c), decls) = add_decls params (c,decls); |
19720 | 355 |
|
31800 | 356 |
fun add_clause_decls params (Clause {literals, ...}, decls) = |
357 |
List.foldl (add_literal_decls params) decls literals |
|
27187 | 358 |
handle Symtab.DUP a => error ("function " ^ a ^ " has multiple arities") |
19720 | 359 |
|
31800 | 360 |
fun decls_of_clauses params clauses arity_clauses = |
35826 | 361 |
let val init_functab = Symtab.update (type_wrapper,2) (Symtab.update ("hAPP",2) SFC.init_functab) |
22064 | 362 |
val init_predtab = Symtab.update ("hBOOL",1) Symtab.empty |
31800 | 363 |
val (functab,predtab) = (List.foldl (add_clause_decls params) (init_functab, init_predtab) clauses) |
22064 | 364 |
in |
35826 | 365 |
(Symtab.dest (List.foldl SFC.add_arityClause_funcs functab arity_clauses), |
22064 | 366 |
Symtab.dest predtab) |
367 |
end; |
|
19720 | 368 |
|
21398
11996e938dfe
Includes type:sort constraints in its code to collect predicates in axiom clauses.
paulson
parents:
21254
diff
changeset
|
369 |
fun add_clause_preds (Clause {ctypes_sorts, ...}, preds) = |
35826 | 370 |
List.foldl SFC.add_type_sort_preds preds ctypes_sorts |
27187 | 371 |
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
|
372 |
|
11996e938dfe
Includes type:sort constraints in its code to collect predicates in axiom clauses.
paulson
parents:
21254
diff
changeset
|
373 |
(*Higher-order clauses have only the predicates hBOOL and type classes.*) |
24311 | 374 |
fun preds_of_clauses clauses clsrel_clauses arity_clauses = |
19720 | 375 |
Symtab.dest |
35826 | 376 |
(List.foldl SFC.add_classrelClause_preds |
377 |
(List.foldl SFC.add_arityClause_preds |
|
30190 | 378 |
(List.foldl add_clause_preds Symtab.empty clauses) |
24311 | 379 |
arity_clauses) |
380 |
clsrel_clauses) |
|
19720 | 381 |
|
18440
72ee07f4b56b
Literals in clauses are sorted now. Added functions for clause equivalence tests and hashing clauses to words.
mengj
parents:
18356
diff
changeset
|
382 |
|
72ee07f4b56b
Literals in clauses are sorted now. Added functions for clause equivalence tests and hashing clauses to words.
mengj
parents:
18356
diff
changeset
|
383 |
(**********************************************************************) |
19198
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
384 |
(* write clauses to files *) |
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
385 |
(**********************************************************************) |
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
386 |
|
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
387 |
val init_counters = |
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
388 |
Symtab.make [("c_COMBI", 0), ("c_COMBK", 0), |
24311 | 389 |
("c_COMBB", 0), ("c_COMBC", 0), |
24943 | 390 |
("c_COMBS", 0)]; |
24311 | 391 |
|
32994 | 392 |
fun count_combterm (CombConst (c, _, _), ct) = |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
393 |
(case Symtab.lookup ct c of NONE => ct (*no counter*) |
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
394 |
| SOME n => Symtab.update (c,n+1) ct) |
32994 | 395 |
| count_combterm (CombVar _, ct) = ct |
22078
5084f53cef39
Streamlining: removing the type argument of CombApp; abbreviating ResClause as RC
paulson
parents:
22064
diff
changeset
|
396 |
| count_combterm (CombApp(t1,t2), ct) = count_combterm(t1, count_combterm(t2, ct)); |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
397 |
|
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
398 |
fun count_literal (Literal(_,t), ct) = count_combterm(t,ct); |
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
399 |
|
30190 | 400 |
fun count_clause (Clause{literals,...}, ct) = List.foldl count_literal ct literals; |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
401 |
|
24311 | 402 |
fun count_user_clause user_lemmas (Clause{axiom_name,literals,...}, ct) = |
30190 | 403 |
if axiom_name mem_string user_lemmas then List.foldl count_literal ct literals |
21573
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
404 |
else ct; |
8a7a68c0096c
Removed the references for counting combinators. Instead they are counted in actual clauses.
paulson
parents:
21561
diff
changeset
|
405 |
|
27178
c8ddb3000743
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
25243
diff
changeset
|
406 |
fun cnf_helper_thms thy = |
35826 | 407 |
Sledgehammer_Fact_Preprocessor.cnf_rules_pairs thy |
408 |
o map Sledgehammer_Fact_Preprocessor.pairname |
|
20644 | 409 |
|
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
410 |
fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) = |
23386 | 411 |
if isFO then [] (*first-order*) |
412 |
else |
|
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
413 |
let |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
414 |
val axclauses = map #2 (make_axiom_clauses dfg thy axcls) |
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31749
diff
changeset
|
415 |
val ct0 = List.foldl count_clause init_counters conjectures |
30190 | 416 |
val ct = List.foldl (count_user_clause user_lemmas) ct0 axclauses |
33035 | 417 |
fun needed c = the (Symtab.lookup ct c) > 0 |
24311 | 418 |
val IK = if needed "c_COMBI" orelse needed "c_COMBK" |
31910 | 419 |
then cnf_helper_thms thy [comb_I,comb_K] |
24311 | 420 |
else [] |
421 |
val BC = if needed "c_COMBB" orelse needed "c_COMBC" |
|
31910 | 422 |
then cnf_helper_thms thy [comb_B,comb_C] |
21135 | 423 |
else [] |
24311 | 424 |
val S = if needed "c_COMBS" |
31910 | 425 |
then cnf_helper_thms thy [comb_S] |
24311 | 426 |
else [] |
31837 | 427 |
val other = cnf_helper_thms thy [fequal_imp_equal,equal_imp_fequal] |
20791
497e1c9d4a9f
Combinator axioms are now from Isabelle theorems, no need to use helper files.
mengj
parents:
20660
diff
changeset
|
428 |
in |
30151 | 429 |
map #2 (make_axiom_clauses dfg thy (other @ IK @ BC @ S)) |
23386 | 430 |
end; |
20791
497e1c9d4a9f
Combinator axioms are now from Isabelle theorems, no need to use helper files.
mengj
parents:
20660
diff
changeset
|
431 |
|
22064 | 432 |
(*Find the minimal arity of each function mentioned in the term. Also, note which uses |
433 |
are not at top level, to see if hBOOL is needed.*) |
|
30150 | 434 |
fun count_constants_term toplev t (const_min_arity, const_needs_hBOOL) = |
22064 | 435 |
let val (head, args) = strip_comb t |
436 |
val n = length args |
|
30150 | 437 |
val (const_min_arity, const_needs_hBOOL) = fold (count_constants_term false) args (const_min_arity, const_needs_hBOOL) |
22064 | 438 |
in |
439 |
case head of |
|
24311 | 440 |
CombConst (a,_,_) => (*predicate or function version of "equal"?*) |
441 |
let val a = if a="equal" andalso not toplev then "c_fequal" else a |
|
33029 | 442 |
val const_min_arity = Symtab.map_default (a, n) (Integer.min n) const_min_arity |
24311 | 443 |
in |
30150 | 444 |
if toplev then (const_min_arity, const_needs_hBOOL) |
445 |
else (const_min_arity, Symtab.update (a,true) (const_needs_hBOOL)) |
|
24311 | 446 |
end |
32994 | 447 |
| _ => (const_min_arity, const_needs_hBOOL) |
22064 | 448 |
end; |
449 |
||
450 |
(*A literal is a top-level term*) |
|
30150 | 451 |
fun count_constants_lit (Literal (_,t)) (const_min_arity, const_needs_hBOOL) = |
452 |
count_constants_term true t (const_min_arity, const_needs_hBOOL); |
|
22064 | 453 |
|
30150 | 454 |
fun count_constants_clause (Clause{literals,...}) (const_min_arity, const_needs_hBOOL) = |
455 |
fold count_constants_lit literals (const_min_arity, const_needs_hBOOL); |
|
22064 | 456 |
|
30150 | 457 |
fun display_arity const_needs_hBOOL (c,n) = |
35826 | 458 |
Sledgehammer_Fact_Preprocessor.trace_msg (fn () => "Constant: " ^ c ^ |
459 |
" arity:\t" ^ Int.toString n ^ |
|
30150 | 460 |
(if needs_hBOOL const_needs_hBOOL c then " needs hBOOL" else "")); |
22064 | 461 |
|
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
462 |
fun count_constants (conjectures, _, extra_clauses, helper_clauses, _, _) = |
30153
051d3825a15d
turned "read-only refs" typ_level and minimize_applies into constant values;
wenzelm
parents:
30151
diff
changeset
|
463 |
if minimize_applies then |
30150 | 464 |
let val (const_min_arity, const_needs_hBOOL) = |
465 |
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
|
466 |
|> fold count_constants_clause extra_clauses |
30149 | 467 |
|> fold count_constants_clause helper_clauses |
30150 | 468 |
val _ = List.app (display_arity const_needs_hBOOL) (Symtab.dest (const_min_arity)) |
469 |
in (const_min_arity, const_needs_hBOOL) end |
|
470 |
else (Symtab.empty, Symtab.empty); |
|
22064 | 471 |
|
31749 | 472 |
(* tptp format *) |
473 |
||
31838 | 474 |
fun tptp_write_file t_full file clauses = |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
475 |
let |
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
476 |
val (conjectures, axclauses, _, helper_clauses, |
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
477 |
classrel_clauses, arity_clauses) = clauses |
31791 | 478 |
val (cma, cnh) = count_constants clauses |
479 |
val params = (t_full, cma, cnh) |
|
480 |
val (tptp_clss,tfree_litss) = ListPair.unzip (map (clause2tptp params) conjectures) |
|
35826 | 481 |
val tfree_clss = map SFC.tptp_tfree_clause (List.foldl (uncurry (union (op =))) [] tfree_litss) |
31839
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
482 |
val _ = |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
483 |
File.write_list file ( |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
484 |
map (#1 o (clause2tptp params)) axclauses @ |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
485 |
tfree_clss @ |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
486 |
tptp_clss @ |
35826 | 487 |
map SFC.tptp_classrelClause classrel_clauses @ |
488 |
map SFC.tptp_arity_clause arity_clauses @ |
|
31839
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
489 |
map (#1 o (clause2tptp params)) helper_clauses) |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
490 |
in (length axclauses + 1, length tfree_clss + length tptp_clss) |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
491 |
end; |
19198
e6f1ff40ba99
Added tptp_write_file to write all necessary ATP input clauses to one file.
mengj
parents:
19130
diff
changeset
|
492 |
|
19720 | 493 |
|
494 |
(* dfg format *) |
|
495 |
||
31838 | 496 |
fun dfg_write_file t_full file clauses = |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
497 |
let |
31865
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
498 |
val (conjectures, axclauses, _, helper_clauses, |
5e97c4abd18e
fixed: count constants with supplementary lemmas
immler@in.tum.de
parents:
31839
diff
changeset
|
499 |
classrel_clauses, arity_clauses) = clauses |
31791 | 500 |
val (cma, cnh) = count_constants clauses |
501 |
val params = (t_full, cma, cnh) |
|
502 |
val (dfg_clss, tfree_litss) = ListPair.unzip (map (clause2dfg params) conjectures) |
|
31838 | 503 |
and probname = Path.implode (Path.base file) |
31791 | 504 |
val axstrs = map (#1 o (clause2dfg params)) axclauses |
35826 | 505 |
val tfree_clss = map SFC.dfg_tfree_clause (SFC.union_all tfree_litss) |
31791 | 506 |
val helper_clauses_strs = map (#1 o (clause2dfg params)) helper_clauses |
507 |
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
|
508 |
and ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses |
31839
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
509 |
val _ = |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
510 |
File.write_list file ( |
35826 | 511 |
SFC.string_of_start probname :: |
512 |
SFC.string_of_descrip probname :: |
|
513 |
SFC.string_of_symbols (SFC.string_of_funcs funcs) |
|
514 |
(SFC.string_of_preds (cl_preds @ ty_preds)) :: |
|
31839
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
515 |
"list_of_clauses(axioms,cnf).\n" :: |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
516 |
axstrs @ |
35826 | 517 |
map SFC.dfg_classrelClause classrel_clauses @ |
518 |
map SFC.dfg_arity_clause arity_clauses @ |
|
31839
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
519 |
helper_clauses_strs @ |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
520 |
["end_of_list.\n\nlist_of_clauses(conjectures,cnf).\n"] @ |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
521 |
tfree_clss @ |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
522 |
dfg_clss @ |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
523 |
["end_of_list.\n\n", |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
524 |
(*VarWeight=3 helps the HO problems, probably by counteracting the presence of hAPP*) |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
525 |
"list_of_settings(SPASS).\n{*\nset_flag(VarWeight,3).\n*}\nend_of_list.\n\n", |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
526 |
"end_problem.\n"]) |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
527 |
|
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
528 |
in (length axclauses + length classrel_clauses + length arity_clauses + |
26f18ec0e293
return number of first conjecture-clause and number of conjecture-clauses;
immler@in.tum.de
parents:
31838
diff
changeset
|
529 |
length helper_clauses + 1, length tfree_clss + length dfg_clss) |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
30242
diff
changeset
|
530 |
end; |
19720 | 531 |
|
33311 | 532 |
end; |
533 |