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