| author | wenzelm | 
| Tue, 20 Dec 2022 19:43:55 +0100 | |
| changeset 76721 | 5364cdc3ec0e | 
| parent 76301 | 73b120e0dbfe | 
| child 77427 | 4cdefee3f97f | 
| permissions | -rw-r--r-- | 
| 38047 | 1  | 
(* Title: HOL/Tools/ATP/atp_problem.ML  | 
| 38027 | 2  | 
Author: Jia Meng, Cambridge University Computer Laboratory and NICTA  | 
| 
37509
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
3  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
74760
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
4  | 
Author: Martin Desharnais, MPI-INF Saarbruecken  | 
| 
37509
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
|
| 39452 | 6  | 
Abstract representation of ATP problems and TPTP syntax.  | 
| 
37509
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
|
| 
38019
 
e207a64e1e0b
complete renaming of "Sledgehammer_TPTP_Format" to "ATP_Problem"
 
blanchet 
parents: 
38018 
diff
changeset
 | 
9  | 
signature ATP_PROBLEM =  | 
| 
37509
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
11  | 
  datatype ('a, 'b) atp_term =
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
12  | 
    ATerm of ('a * 'b list) * ('a, 'b) atp_term list |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
13  | 
    AAbs of (('a * 'b) * ('a, 'b) atp_term) * ('a, 'b) atp_term list
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
14  | 
datatype atp_quantifier = AForall | AExists  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
15  | 
datatype atp_connective = ANot | AAnd | AOr | AImplies | AIff  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
16  | 
  datatype ('a, 'b, 'c, 'd) atp_formula =
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
17  | 
    ATyQuant of atp_quantifier * ('b * 'd list) list
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
18  | 
        * ('a, 'b, 'c, 'd) atp_formula |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
19  | 
    AQuant of atp_quantifier * ('a * 'b option) list
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
20  | 
        * ('a, 'b, 'c, 'd) atp_formula |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
21  | 
    AConn of atp_connective * ('a, 'b, 'c, 'd) atp_formula list |
 | 
| 
42531
 
a462dbaa584f
added more rudimentary type support to Sledgehammer's ATP encoding
 
blanchet 
parents: 
42530 
diff
changeset
 | 
22  | 
AAtom of 'c  | 
| 
37994
 
b04307085a09
make TPTP generator accept full first-order formulas
 
blanchet 
parents: 
37993 
diff
changeset
 | 
23  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
24  | 
datatype 'a atp_type =  | 
| 54820 | 25  | 
    AType of ('a * 'a list) * 'a atp_type list |
 | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
26  | 
AFun of 'a atp_type * 'a atp_type |  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
27  | 
APi of 'a list * 'a atp_type  | 
| 42963 | 28  | 
|
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
29  | 
type term_order =  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
30  | 
    {is_lpo : bool,
 | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
31  | 
gen_weights : bool,  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
32  | 
gen_prec : bool,  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
33  | 
gen_simp : bool}  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
34  | 
|
| 74890 | 35  | 
  type syntax = {with_ite : bool, with_let : bool}
 | 
36  | 
datatype fool = Without_FOOL | With_FOOL of syntax  | 
|
| 48140 | 37  | 
datatype polymorphism = Monomorphic | Polymorphic  | 
| 72915 | 38  | 
datatype thf_flavor = THF_Lambda_Free | THF_Without_Choice | THF_With_Choice  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
39  | 
|
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
40  | 
datatype atp_format =  | 
| 
44235
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
41  | 
CNF |  | 
| 
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
42  | 
CNF_UEQ |  | 
| 
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
43  | 
FOF |  | 
| 74890 | 44  | 
TFF of polymorphism * fool |  | 
45  | 
THF of polymorphism * syntax * thf_flavor |  | 
|
| 48131 | 46  | 
DFG of polymorphism  | 
| 
44235
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
47  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
48  | 
datatype atp_formula_role =  | 
| 
50012
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
48438 
diff
changeset
 | 
49  | 
Axiom | Definition | Lemma | Hypothesis | Conjecture | Negated_Conjecture |  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
50  | 
Plain | Type_Role | Unknown  | 
| 
50012
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
48438 
diff
changeset
 | 
51  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
52  | 
datatype 'a atp_problem_line =  | 
| 48141 | 53  | 
Class_Decl of string * 'a * 'a list |  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
54  | 
Type_Decl of string * 'a * int |  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
55  | 
Sym_Decl of string * 'a * 'a atp_type |  | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
56  | 
Datatype_Decl of  | 
| 
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
57  | 
      string * ('a * 'a list) list * 'a atp_type * ('a, 'a atp_type) atp_term list * bool |
 | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
58  | 
    Class_Memb of string * ('a * 'a list) list * 'a atp_type * 'a |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
59  | 
Formula of (string * string) * atp_formula_role  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
60  | 
               * ('a, 'a atp_type, ('a, 'a atp_type) atp_term, 'a) atp_formula
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
61  | 
* (string, string atp_type) atp_term option  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
62  | 
* (string, string atp_type) atp_term list  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
63  | 
type 'a atp_problem = (string * 'a atp_problem_line list) list  | 
| 37992 | 64  | 
|
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
65  | 
val tptp_cnf : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
66  | 
val tptp_fof : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
67  | 
val tptp_tff : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
68  | 
val tptp_thf : string  | 
| 42967 | 69  | 
val tptp_has_type : string  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
70  | 
val tptp_type_of_types : string  | 
| 
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
71  | 
val tptp_bool_type : string  | 
| 
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
72  | 
val tptp_individual_type : string  | 
| 42963 | 73  | 
val tptp_fun_type : string  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
74  | 
val tptp_product_type : string  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
75  | 
val tptp_forall : string  | 
| 43678 | 76  | 
val tptp_ho_forall : string  | 
| 44650 | 77  | 
val tptp_pi_binder : string  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
78  | 
val tptp_exists : string  | 
| 43678 | 79  | 
val tptp_ho_exists : string  | 
| 44495 | 80  | 
val tptp_choice : string  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
81  | 
val tptp_ho_choice : string  | 
| 57709 | 82  | 
val tptp_hilbert_choice : string  | 
83  | 
val tptp_hilbert_the : string  | 
|
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
84  | 
val tptp_not : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
85  | 
val tptp_and : string  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
86  | 
val tptp_not_and : string  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
87  | 
val tptp_or : string  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
88  | 
val tptp_not_or : string  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
89  | 
val tptp_implies : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
90  | 
val tptp_if : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
91  | 
val tptp_iff : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
92  | 
val tptp_not_iff : string  | 
| 
73859
 
bc263f1f68cd
added support for TFX's and THF's $ite to Sledgehammer
 
desharna 
parents: 
72917 
diff
changeset
 | 
93  | 
val tptp_ite : string  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
73859 
diff
changeset
 | 
94  | 
val tptp_let : string  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
95  | 
val tptp_app : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
96  | 
val tptp_not_infix : string  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
97  | 
val tptp_equal : string  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
98  | 
val tptp_not_equal : string  | 
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
99  | 
val tptp_old_equal : string  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
100  | 
val tptp_false : string  | 
| 
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
101  | 
val tptp_true : string  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
102  | 
val tptp_lambda : string  | 
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
103  | 
val tptp_empty_list : string  | 
| 74552 | 104  | 
|
105  | 
  type tptp_builtin_desc = {arity : int, is_predicate : bool}
 | 
|
106  | 
val tptp_builtins : tptp_builtin_desc Symtab.table  | 
|
107  | 
||
| 58600 | 108  | 
val dfg_individual_type : string  | 
| 46406 | 109  | 
val isabelle_info_prefix : string  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
110  | 
val isabelle_info : bool -> string -> int -> (string, 'a) atp_term list  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
111  | 
val extract_isabelle_status : (string, 'a) atp_term list -> string option  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
112  | 
val extract_isabelle_rank : (string, 'a) atp_term list -> int  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
113  | 
val inductionN : string  | 
| 46406 | 114  | 
val introN : string  | 
| 
47148
 
7b5846065c1b
be less forceful about ":lt" to make infinite loops less likely (could still fail with mutually recursive tail rec functions)
 
blanchet 
parents: 
47146 
diff
changeset
 | 
115  | 
val inductiveN : string  | 
| 46406 | 116  | 
val elimN : string  | 
117  | 
val simpN : string  | 
|
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
118  | 
val non_rec_defN : string  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
119  | 
val rec_defN : string  | 
| 46406 | 120  | 
val rankN : string  | 
121  | 
val minimum_rank : int  | 
|
122  | 
val default_rank : int  | 
|
| 47030 | 123  | 
val default_term_order_weight : int  | 
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
124  | 
val is_tptp_equal : string -> bool  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
125  | 
val is_built_in_tptp_symbol : string -> bool  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
126  | 
val is_tptp_variable : string -> bool  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
127  | 
val is_tptp_user_symbol : string -> bool  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
128  | 
val bool_atype : (string * string) atp_type  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
129  | 
val individual_atype : (string * string) atp_type  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
130  | 
  val mk_anot : ('a, 'b, 'c, 'd) atp_formula -> ('a, 'b, 'c, 'd) atp_formula
 | 
| 75341 | 131  | 
  val mk_aconn : atp_connective -> ('a, 'b, 'c, 'd) atp_formula -> ('a, 'b, 'c, 'd) atp_formula ->
 | 
132  | 
    ('a, 'b, 'c, 'd) atp_formula
 | 
|
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
133  | 
val mk_app : (string, 'a) atp_term -> (string, 'a) atp_term -> (string, 'a) atp_term  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
134  | 
val mk_apps : (string, 'a) atp_term -> (string, 'a) atp_term list -> (string, 'a) atp_term  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
135  | 
  val mk_simple_aterm:  'a -> ('a, 'b) atp_term
 | 
| 75341 | 136  | 
val aconn_fold : bool option -> (bool option -> 'a -> 'b -> 'b) -> atp_connective * 'a list ->  | 
137  | 
'b -> 'b  | 
|
138  | 
  val aconn_map : bool option -> (bool option -> 'a -> ('b, 'c, 'd, 'e) atp_formula) ->
 | 
|
139  | 
    atp_connective * 'a list -> ('b, 'c, 'd, 'e) atp_formula
 | 
|
140  | 
val formula_fold : bool option -> (bool option -> 'c -> 'e -> 'e) ->  | 
|
141  | 
    ('a, 'b, 'c, 'd) atp_formula -> 'e -> 'e
 | 
|
142  | 
  val formula_map : ('c -> 'e) -> ('a, 'b, 'c, 'd) atp_formula -> ('a, 'b, 'e, 'd) atp_formula
 | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
143  | 
  val strip_atype : 'a atp_type -> 'a list * ('a atp_type list * 'a atp_type)
 | 
| 
45303
 
bd03b08161ac
added DFG unsorted support (like in the old days)
 
blanchet 
parents: 
45301 
diff
changeset
 | 
144  | 
val is_format_higher_order : atp_format -> bool  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
145  | 
val tptp_string_of_format : atp_format -> string  | 
| 72355 | 146  | 
val tptp_string_of_role : atp_formula_role -> string  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
147  | 
val tptp_string_of_line : atp_format -> string atp_problem_line -> string  | 
| 75341 | 148  | 
val lines_of_atp_problem : atp_format -> (unit -> (string * int) list) -> string atp_problem ->  | 
149  | 
string list  | 
|
150  | 
val ensure_cnf_problem : (string * string) atp_problem -> (string * string) atp_problem  | 
|
151  | 
val filter_cnf_ueq_problem : (string * string) atp_problem -> (string * string) atp_problem  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
152  | 
  val declared_in_atp_problem : 'a atp_problem -> ('a list * 'a list) * 'a list
 | 
| 75341 | 153  | 
val nice_atp_problem : bool -> atp_format ->  | 
154  | 
    ('a * (string * string) atp_problem_line list) list ->
 | 
|
155  | 
    ('a * string atp_problem_line list) list * (string Symtab.table * string Symtab.table) option
 | 
|
| 
37509
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
156  | 
end;  | 
| 
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
157  | 
|
| 
38019
 
e207a64e1e0b
complete renaming of "Sledgehammer_TPTP_Format" to "ATP_Problem"
 
blanchet 
parents: 
38018 
diff
changeset
 | 
158  | 
structure ATP_Problem : ATP_PROBLEM =  | 
| 
37509
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
struct  | 
| 
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
160  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43000 
diff
changeset
 | 
161  | 
open ATP_Util  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43000 
diff
changeset
 | 
162  | 
|
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
163  | 
val parens = enclose "(" ")"
 | 
| 
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
164  | 
|
| 
37643
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
165  | 
(** ATP problem **)  | 
| 
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
166  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
167  | 
datatype ('a, 'b) atp_term =
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
168  | 
  ATerm of ('a * 'b list) * ('a, 'b) atp_term list |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
169  | 
  AAbs of (('a * 'b) * ('a, 'b) atp_term) * ('a, 'b) atp_term list
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
170  | 
datatype atp_quantifier = AForall | AExists  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
171  | 
datatype atp_connective = ANot | AAnd | AOr | AImplies | AIff  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
172  | 
datatype ('a, 'b, 'c, 'd) atp_formula =
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
173  | 
  ATyQuant of atp_quantifier * ('b * 'd list) list
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
174  | 
      * ('a, 'b, 'c, 'd) atp_formula |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
175  | 
  AQuant of atp_quantifier * ('a * 'b option) list
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
176  | 
      * ('a, 'b, 'c, 'd) atp_formula |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
177  | 
  AConn of atp_connective * ('a, 'b, 'c, 'd) atp_formula list |
 | 
| 
42531
 
a462dbaa584f
added more rudimentary type support to Sledgehammer's ATP encoding
 
blanchet 
parents: 
42530 
diff
changeset
 | 
178  | 
AAtom of 'c  | 
| 
37961
 
6a48c85a211a
first step in using "fof" rather than "cnf" in TPTP problems
 
blanchet 
parents: 
37931 
diff
changeset
 | 
179  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
180  | 
datatype 'a atp_type =  | 
| 54820 | 181  | 
  AType of ('a * 'a list) * 'a atp_type list |
 | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
182  | 
AFun of 'a atp_type * 'a atp_type |  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
183  | 
APi of 'a list * 'a atp_type  | 
| 42963 | 184  | 
|
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
185  | 
type term_order =  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
186  | 
  {is_lpo : bool,
 | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
187  | 
gen_weights : bool,  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
188  | 
gen_prec : bool,  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
189  | 
gen_simp : bool}  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
190  | 
|
| 74890 | 191  | 
|
192  | 
type syntax = {with_ite : bool, with_let : bool}
 | 
|
193  | 
datatype fool = Without_FOOL | With_FOOL of syntax  | 
|
| 48140 | 194  | 
datatype polymorphism = Monomorphic | Polymorphic  | 
| 72915 | 195  | 
datatype thf_flavor = THF_Lambda_Free | THF_Without_Choice | THF_With_Choice  | 
| 44499 | 196  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
197  | 
datatype atp_format =  | 
| 
44235
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
198  | 
CNF |  | 
| 
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
199  | 
CNF_UEQ |  | 
| 
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
200  | 
FOF |  | 
| 74890 | 201  | 
TFF of polymorphism * fool |  | 
202  | 
THF of polymorphism * syntax * thf_flavor |  | 
|
| 48131 | 203  | 
DFG of polymorphism  | 
| 
44235
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
204  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
205  | 
datatype atp_formula_role =  | 
| 
50012
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
48438 
diff
changeset
 | 
206  | 
Axiom | Definition | Lemma | Hypothesis | Conjecture | Negated_Conjecture |  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
207  | 
Plain | Type_Role | Unknown  | 
| 
50012
 
01cb92151a53
track formula roles in proofs and use that to determine whether the conjecture should be negated or not
 
blanchet 
parents: 
48438 
diff
changeset
 | 
208  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
209  | 
datatype 'a atp_problem_line =  | 
| 48141 | 210  | 
Class_Decl of string * 'a * 'a list |  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
211  | 
Type_Decl of string * 'a * int |  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
212  | 
Sym_Decl of string * 'a * 'a atp_type |  | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
213  | 
Datatype_Decl of  | 
| 
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
214  | 
    string * ('a * 'a list) list * 'a atp_type * ('a, 'a atp_type) atp_term list * bool |
 | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
215  | 
  Class_Memb of string * ('a * 'a list) list * 'a atp_type * 'a |
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
216  | 
Formula of (string * string) * atp_formula_role  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
217  | 
             * ('a, 'a atp_type, ('a, 'a atp_type) atp_term, 'a) atp_formula
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
218  | 
* (string, string atp_type) atp_term option  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
219  | 
* (string, string atp_type) atp_term list  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
220  | 
type 'a atp_problem = (string * 'a atp_problem_line list) list  | 
| 
37643
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
221  | 
|
| 42722 | 222  | 
(* official TPTP syntax *)  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
223  | 
val tptp_cnf = "cnf"  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
224  | 
val tptp_fof = "fof"  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
225  | 
val tptp_tff = "tff"  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
226  | 
val tptp_thf = "thf"  | 
| 42967 | 227  | 
val tptp_has_type = ":"  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
228  | 
val tptp_type_of_types = "$tType"  | 
| 
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
229  | 
val tptp_bool_type = "$o"  | 
| 
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
230  | 
val tptp_individual_type = "$i"  | 
| 42963 | 231  | 
val tptp_fun_type = ">"  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
232  | 
val tptp_product_type = "*"  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
233  | 
val tptp_forall = "!"  | 
| 43678 | 234  | 
val tptp_ho_forall = "!!"  | 
| 44650 | 235  | 
val tptp_pi_binder = "!>"  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
236  | 
val tptp_exists = "?"  | 
| 43678 | 237  | 
val tptp_ho_exists = "??"  | 
| 44495 | 238  | 
val tptp_choice = "@+"  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
239  | 
val tptp_ho_choice = "@@+"  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
240  | 
val tptp_not = "~"  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
241  | 
val tptp_and = "&"  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
242  | 
val tptp_not_and = "~&"  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
243  | 
val tptp_or = "|"  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
244  | 
val tptp_not_or = "~|"  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
245  | 
val tptp_implies = "=>"  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
246  | 
val tptp_if = "<="  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
247  | 
val tptp_iff = "<=>"  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
248  | 
val tptp_not_iff = "<~>"  | 
| 
73859
 
bc263f1f68cd
added support for TFX's and THF's $ite to Sledgehammer
 
desharna 
parents: 
72917 
diff
changeset
 | 
249  | 
val tptp_ite = "$ite"  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
73859 
diff
changeset
 | 
250  | 
val tptp_let = "$let"  | 
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
251  | 
val tptp_app = "@"  | 
| 57709 | 252  | 
val tptp_hilbert_choice = "@+"  | 
253  | 
val tptp_hilbert_the = "@-"  | 
|
| 
42968
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
254  | 
val tptp_not_infix = "!"  | 
| 
 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 
blanchet 
parents: 
42967 
diff
changeset
 | 
255  | 
val tptp_equal = "="  | 
| 75363 | 256  | 
val tptp_not_equal = tptp_not_infix ^ tptp_equal  | 
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
257  | 
val tptp_old_equal = "equal"  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
258  | 
val tptp_false = "$false"  | 
| 
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
259  | 
val tptp_true = "$true"  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
260  | 
val tptp_lambda = "^"  | 
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
261  | 
val tptp_empty_list = "[]"  | 
| 58600 | 262  | 
|
| 74552 | 263  | 
(* A predicate has return type $o (i.e. bool) *)  | 
264  | 
type tptp_builtin_desc = {arity : int, is_predicate : bool}
 | 
|
265  | 
||
266  | 
val tptp_builtins =  | 
|
267  | 
let  | 
|
268  | 
    fun make_builtin arity is_predicate name = (name, {arity = arity, is_predicate = is_predicate})
 | 
|
269  | 
in  | 
|
270  | 
map (make_builtin 0 true) [tptp_false, tptp_true] @  | 
|
| 74898 | 271  | 
map (make_builtin 1 true) [tptp_not, tptp_ho_forall, tptp_ho_exists] @  | 
| 74552 | 272  | 
map (make_builtin 2 true) [tptp_and, tptp_or, tptp_implies, tptp_if, tptp_iff, tptp_not_iff,  | 
273  | 
tptp_equal, tptp_old_equal] @  | 
|
274  | 
map (make_builtin 2 false) [tptp_let] @  | 
|
275  | 
map (make_builtin 3 false) [tptp_ite]  | 
|
276  | 
|> Symtab.make  | 
|
277  | 
end  | 
|
| 72914 | 278  | 
|
| 58600 | 279  | 
val dfg_individual_type = "iii" (* cannot clash *)  | 
280  | 
||
| 46406 | 281  | 
val isabelle_info_prefix = "isabelle_"  | 
282  | 
||
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
283  | 
val inductionN = "induction"  | 
| 46406 | 284  | 
val introN = "intro"  | 
| 
47148
 
7b5846065c1b
be less forceful about ":lt" to make infinite loops less likely (could still fail with mutually recursive tail rec functions)
 
blanchet 
parents: 
47146 
diff
changeset
 | 
285  | 
val inductiveN = "inductive"  | 
| 46406 | 286  | 
val elimN = "elim"  | 
287  | 
val simpN = "simp"  | 
|
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
288  | 
val non_rec_defN = "non_rec_def"  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
289  | 
val rec_defN = "rec_def"  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
290  | 
|
| 46406 | 291  | 
val rankN = "rank"  | 
292  | 
||
293  | 
val minimum_rank = 0  | 
|
294  | 
val default_rank = 1000  | 
|
| 47030 | 295  | 
val default_term_order_weight = 1  | 
| 46406 | 296  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
297  | 
fun isabelle_info generate_info status rank =  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
298  | 
if generate_info then  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
299  | 
[] |> rank <> default_rank  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
300  | 
? cons (ATerm ((isabelle_info_prefix ^ rankN, []),  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
301  | 
[ATerm ((string_of_int rank, []), [])]))  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
302  | 
|> status <> "" ? cons (ATerm ((isabelle_info_prefix ^ status, []), []))  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
303  | 
else  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
304  | 
[]  | 
| 46406 | 305  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
306  | 
fun extract_isabelle_status (ATerm ((s, []), []) :: _) =  | 
| 46406 | 307  | 
try (unprefix isabelle_info_prefix) s  | 
308  | 
| extract_isabelle_status _ = NONE  | 
|
309  | 
||
310  | 
fun extract_isabelle_rank (tms as _ :: _) =  | 
|
311  | 
(case List.last tms of  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
312  | 
ATerm ((_, []), [ATerm ((rank, []), [])]) => the (Int.fromString rank)  | 
| 46406 | 313  | 
| _ => default_rank)  | 
314  | 
| extract_isabelle_rank _ = default_rank  | 
|
315  | 
||
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
316  | 
fun is_tptp_equal s = (s = tptp_equal orelse s = tptp_old_equal)  | 
| 
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
317  | 
fun is_built_in_tptp_symbol s =  | 
| 
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
318  | 
s = tptp_old_equal orelse not (Char.isAlpha (String.sub (s, 0)))  | 
| 51878 | 319  | 
fun is_tptp_variable s = s <> "" andalso Char.isUpper (String.sub (s, 0))  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
320  | 
val is_tptp_user_symbol = not o (is_tptp_variable orf is_built_in_tptp_symbol)  | 
| 42939 | 321  | 
|
| 54820 | 322  | 
val bool_atype = AType ((`I tptp_bool_type, []), [])  | 
323  | 
val individual_atype = AType ((`I tptp_individual_type, []), [])  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
324  | 
|
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
325  | 
fun raw_polarities_of_conn ANot = (SOME false, NONE)  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
326  | 
| raw_polarities_of_conn AAnd = (SOME true, SOME true)  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
327  | 
| raw_polarities_of_conn AOr = (SOME true, SOME true)  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
328  | 
| raw_polarities_of_conn AImplies = (SOME false, SOME true)  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
329  | 
| raw_polarities_of_conn AIff = (NONE, NONE)  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
330  | 
fun polarities_of_conn NONE = K (NONE, NONE)  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
331  | 
| polarities_of_conn (SOME pos) =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58600 
diff
changeset
 | 
332  | 
raw_polarities_of_conn #> not pos ? apply2 (Option.map not)  | 
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
333  | 
|
| 42942 | 334  | 
fun mk_anot (AConn (ANot, [phi])) = phi  | 
335  | 
| mk_anot phi = AConn (ANot, [phi])  | 
|
336  | 
fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])  | 
|
337  | 
||
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
338  | 
fun mk_app t u = ATerm ((tptp_app, []), [t, u])  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
339  | 
fun mk_apps f xs = fold (fn x => fn f => mk_app f x) xs f  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
340  | 
|
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
341  | 
fun mk_simple_aterm p = ATerm ((p, []), [])  | 
| 
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
342  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
343  | 
fun aconn_fold pos f (ANot, [phi]) = f (Option.map not pos) phi  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
344  | 
| aconn_fold pos f (AImplies, [phi1, phi2]) =  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
345  | 
f (Option.map not pos) phi1 #> f pos phi2  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
346  | 
| aconn_fold pos f (AAnd, phis) = fold (f pos) phis  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
347  | 
| aconn_fold pos f (AOr, phis) = fold (f pos) phis  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
348  | 
| aconn_fold _ f (_, phis) = fold (f NONE) phis  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
349  | 
|
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
350  | 
fun aconn_map pos f (ANot, [phi]) = AConn (ANot, [f (Option.map not pos) phi])  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
351  | 
| aconn_map pos f (AImplies, [phi1, phi2]) =  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
352  | 
AConn (AImplies, [f (Option.map not pos) phi1, f pos phi2])  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
353  | 
| aconn_map pos f (AAnd, phis) = AConn (AAnd, map (f pos) phis)  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
354  | 
| aconn_map pos f (AOr, phis) = AConn (AOr, map (f pos) phis)  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
355  | 
| aconn_map _ f (c, phis) = AConn (c, map (f NONE) phis)  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
356  | 
|
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
357  | 
fun formula_fold pos f =  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
358  | 
let  | 
| 44501 | 359  | 
fun fld pos (AQuant (_, _, phi)) = fld pos phi  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
360  | 
| fld pos (ATyQuant (_, _, phi)) = fld pos phi  | 
| 44501 | 361  | 
| fld pos (AConn conn) = aconn_fold pos fld conn  | 
362  | 
| fld pos (AAtom tm) = f pos tm  | 
|
363  | 
in fld pos end  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
364  | 
|
| 
42944
 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 
blanchet 
parents: 
42942 
diff
changeset
 | 
365  | 
fun formula_map f (AQuant (q, xs, phi)) = AQuant (q, xs, formula_map f phi)  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
366  | 
| formula_map f (ATyQuant (q, xs, phi)) = ATyQuant (q, xs, formula_map f phi)  | 
| 
42944
 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 
blanchet 
parents: 
42942 
diff
changeset
 | 
367  | 
| formula_map f (AConn (c, phis)) = AConn (c, map (formula_map f) phis)  | 
| 
 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 
blanchet 
parents: 
42942 
diff
changeset
 | 
368  | 
| formula_map f (AAtom tm) = AAtom (f tm)  | 
| 
 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 
blanchet 
parents: 
42942 
diff
changeset
 | 
369  | 
|
| 52025 | 370  | 
fun strip_api (APi (tys, ty)) = strip_api ty |>> append tys  | 
371  | 
| strip_api ty = ([], ty)  | 
|
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
372  | 
|
| 52025 | 373  | 
fun strip_afun (AFun (ty1, ty2)) = strip_afun ty2 |>> cons ty1  | 
374  | 
| strip_afun ty = ([], ty)  | 
|
375  | 
||
376  | 
fun strip_atype ty = ty |> strip_api ||> strip_afun  | 
|
377  | 
||
| 54820 | 378  | 
fun is_function_atype ty = snd (snd (strip_atype ty)) <> AType ((tptp_bool_type, []), [])  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
379  | 
fun is_predicate_atype ty = not (is_function_atype ty)  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
380  | 
fun is_nontrivial_predicate_atype (AType _) = false  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
381  | 
| is_nontrivial_predicate_atype ty = is_predicate_atype ty  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46414 
diff
changeset
 | 
382  | 
|
| 
45303
 
bd03b08161ac
added DFG unsorted support (like in the old days)
 
blanchet 
parents: 
45301 
diff
changeset
 | 
383  | 
fun is_format_higher_order (THF _) = true  | 
| 
 
bd03b08161ac
added DFG unsorted support (like in the old days)
 
blanchet 
parents: 
45301 
diff
changeset
 | 
384  | 
| is_format_higher_order _ = false  | 
| 74896 | 385  | 
|
386  | 
fun is_format_higher_order_with_choice (THF (_, _, THF_With_Choice)) = true  | 
|
387  | 
| is_format_higher_order_with_choice _ = false  | 
|
388  | 
||
| 44499 | 389  | 
fun is_format_typed (TFF _) = true  | 
| 44754 | 390  | 
| is_format_typed (THF _) = true  | 
| 48131 | 391  | 
| is_format_typed (DFG _) = true  | 
| 
44235
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
43987 
diff
changeset
 | 
392  | 
| is_format_typed _ = false  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
393  | 
|
| 74890 | 394  | 
fun is_format_with_fool (THF _) = true  | 
395  | 
| is_format_with_fool (TFF (_, With_FOOL _)) = true  | 
|
| 72641 | 396  | 
| is_format_with_fool _ = false  | 
397  | 
||
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
398  | 
fun tptp_string_of_role Axiom = "axiom"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
399  | 
| tptp_string_of_role Definition = "definition"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
400  | 
| tptp_string_of_role Lemma = "lemma"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
401  | 
| tptp_string_of_role Hypothesis = "hypothesis"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
402  | 
| tptp_string_of_role Conjecture = "conjecture"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
403  | 
| tptp_string_of_role Negated_Conjecture = "negated_conjecture"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
404  | 
| tptp_string_of_role Plain = "plain"  | 
| 
57255
 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 
fleury 
parents: 
56847 
diff
changeset
 | 
405  | 
| tptp_string_of_role Type_Role = "type"  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
406  | 
| tptp_string_of_role Unknown = "unknown"  | 
| 
38631
 
979a0b37f981
prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
 
blanchet 
parents: 
38613 
diff
changeset
 | 
407  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
408  | 
fun tptp_string_of_app _ func [] = func  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
409  | 
| tptp_string_of_app format func args =  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
410  | 
if is_format_higher_order format then  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
411  | 
      "(" ^ space_implode (" " ^ tptp_app ^ " ") (func :: args) ^ ")"
 | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
412  | 
else  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
413  | 
      func ^ "(" ^ commas args ^ ")"
 | 
| 44787 | 414  | 
|
| 52025 | 415  | 
fun uncurry_type (APi (tys, ty)) = APi (tys, uncurry_type ty)  | 
416  | 
| uncurry_type (ty as AFun (ty1 as AType _, ty2)) =  | 
|
417  | 
(case uncurry_type ty2 of  | 
|
| 54820 | 418  | 
AFun (ty' as AType ((s, _), tys), ty) =>  | 
419  | 
AFun (AType ((tptp_product_type, []),  | 
|
420  | 
ty1 :: (if s = tptp_product_type then tys else [ty'])), ty)  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
421  | 
| _ => ty)  | 
| 52025 | 422  | 
| uncurry_type (ty as AType _) = ty  | 
423  | 
| uncurry_type _ =  | 
|
| 42963 | 424  | 
raise Fail "unexpected higher-order type in first-order format"  | 
425  | 
||
| 
57267
 
8b87114357bd
integrated new Waldmeister code with 'sledgehammer' command
 
blanchet 
parents: 
57256 
diff
changeset
 | 
426  | 
val suffix_type_of_types = suffix (" " ^ tptp_has_type ^ " " ^ tptp_type_of_types)
 | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
427  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
428  | 
fun str_of_type format ty =  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
429  | 
let  | 
| 54829 | 430  | 
val dfg = (case format of DFG _ => true | _ => false)  | 
| 54820 | 431  | 
fun str _ (AType ((s, _), [])) =  | 
| 
46338
 
b02ff6b17599
better handling of individual type for DFG format (SPASS)
 
blanchet 
parents: 
46320 
diff
changeset
 | 
432  | 
if dfg andalso s = tptp_individual_type then dfg_individual_type else s  | 
| 
61549
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
433  | 
| str rhs (AType ((s, _), tys)) =  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
434  | 
if s = tptp_fun_type then  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
435  | 
let val [ty1, ty2] = tys in  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
436  | 
str rhs (AFun (ty1, ty2))  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
437  | 
end  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
438  | 
else  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
439  | 
let val ss = tys |> map (str false) in  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
440  | 
if s = tptp_product_type then  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
441  | 
ss |> space_implode  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
442  | 
(if dfg then ", " else " " ^ tptp_product_type ^ " ")  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
443  | 
|> (not dfg andalso length ss > 1) ? parens  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
444  | 
else  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
445  | 
tptp_string_of_app format s ss  | 
| 
 
16e2313e855c
make sure that function types are never generated as '> @ A @ B', but always as 'A > B'
 
blanchet 
parents: 
59058 
diff
changeset
 | 
446  | 
end  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
447  | 
| str rhs (AFun (ty1, ty2)) =  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
448  | 
(str false ty1 |> dfg ? parens) ^ " " ^  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
449  | 
(if dfg then "" else tptp_fun_type ^ " ") ^ str true ty2  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
450  | 
|> not rhs ? parens  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
451  | 
| str _ (APi (ss, ty)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
452  | 
if dfg then  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
453  | 
"[" ^ commas ss ^ "], " ^ str true ty  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
454  | 
else  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
455  | 
tptp_pi_binder ^ "[" ^ commas (map suffix_type_of_types ss) ^ "]: " ^  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
456  | 
str false ty  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
457  | 
in str true ty end  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
458  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
459  | 
fun string_of_type (format as THF _) ty = str_of_type format ty  | 
| 52025 | 460  | 
| string_of_type format ty = str_of_type format (uncurry_type ty)  | 
| 42963 | 461  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
462  | 
fun tptp_string_of_quantifier AForall = tptp_forall  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
463  | 
| tptp_string_of_quantifier AExists = tptp_exists  | 
| 42963 | 464  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
465  | 
fun tptp_string_of_connective ANot = tptp_not  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
466  | 
| tptp_string_of_connective AAnd = tptp_and  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
467  | 
| tptp_string_of_connective AOr = tptp_or  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
468  | 
| tptp_string_of_connective AImplies = tptp_implies  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
469  | 
| tptp_string_of_connective AIff = tptp_iff  | 
| 42963 | 470  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
471  | 
fun string_of_bound_var format (s, ty) =  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
472  | 
s ^  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
473  | 
(if is_format_typed format then  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
474  | 
" " ^ tptp_has_type ^ " " ^  | 
| 54820 | 475  | 
(ty |> the_default (AType ((tptp_individual_type, []), [])) |> string_of_type format)  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
476  | 
else  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44589 
diff
changeset
 | 
477  | 
"")  | 
| 42963 | 478  | 
|
| 74552 | 479  | 
fun tptp_string_of_term _ (ATerm ((s, []), [])) = s |> Symtab.defined tptp_builtins s ? parens  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
480  | 
| tptp_string_of_term format (ATerm ((s, tys), ts)) =  | 
| 52074 | 481  | 
let  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
482  | 
val of_type = string_of_type format  | 
| 
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
483  | 
val of_term = tptp_string_of_term format  | 
| 
74760
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
484  | 
fun app0 f types args =  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
485  | 
tptp_string_of_app format (f |> Symtab.defined tptp_builtins f ? parens)  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
486  | 
(map (of_type #> is_format_higher_order format ? parens) types @ map of_term args)  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
487  | 
fun app () = app0 s tys ts  | 
| 52074 | 488  | 
in  | 
489  | 
if s = tptp_empty_list then  | 
|
490  | 
(* used for lists in the optional "source" field of a derivation *)  | 
|
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
491  | 
"[" ^ commas (map of_term ts) ^ "]"  | 
| 52074 | 492  | 
else if is_tptp_equal s then  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
493  | 
        space_implode (" " ^ tptp_equal ^ " ") (map of_term ts)
 | 
| 72641 | 494  | 
|> (is_format_higher_order format orelse is_format_with_fool format) ? parens  | 
| 52074 | 495  | 
else if s = tptp_ho_forall orelse s = tptp_ho_exists then  | 
| 54829 | 496  | 
(case ts of  | 
| 52074 | 497  | 
[AAbs (((s', ty), tm), [])] =>  | 
| 72917 | 498  | 
(* Use syntactic sugar "!" and "?" instead of "!!" and "??" whenever possible, to work  | 
499  | 
around LEO-II, Leo-III, and Satallax parser limitation. *)  | 
|
| 52074 | 500  | 
tptp_string_of_formula format  | 
501  | 
(AQuant (if s = tptp_ho_forall then AForall else AExists,  | 
|
502  | 
[(s', SOME ty)], AAtom tm))  | 
|
| 54829 | 503  | 
| _ => app ())  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
73859 
diff
changeset
 | 
504  | 
else if s = tptp_let then  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
73859 
diff
changeset
 | 
505  | 
(case ts of  | 
| 
74760
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
506  | 
t1 :: AAbs (((s', ty), tm), []) :: ts =>  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
507  | 
let  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
508  | 
val declaration = s' ^ " : " ^ of_type ty  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
509  | 
val definition = s' ^ " := " ^ of_term t1  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
510  | 
val usage = of_term tm  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
511  | 
in  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
512  | 
if ts = [] orelse is_format_higher_order format then  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
513  | 
              app0 (s ^ "(" ^ declaration ^ ", " ^ definition ^ ", " ^ usage ^ ")") [] ts
 | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
514  | 
else  | 
| 74896 | 515  | 
error (tptp_let ^ " is special syntax and more than two arguments is only \  | 
| 
74760
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
516  | 
\supported in higher order")  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
517  | 
end  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
518  | 
| _ => error (tptp_let ^ " is special syntax and must have at least two arguments"))  | 
| 74153 | 519  | 
else if s = tptp_ite then  | 
520  | 
(case ts of  | 
|
| 
74760
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
521  | 
t1 :: t2 :: t3 :: ts =>  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
522  | 
if ts = [] orelse is_format_higher_order format then  | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
523  | 
            app0 (s ^ "(" ^ of_term t1 ^ ", " ^ of_term t2 ^ ", " ^ of_term t3 ^ ")") [] ts
 | 
| 
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
524  | 
else  | 
| 74896 | 525  | 
error (tptp_ite ^ " is special syntax and more than three arguments is only supported \  | 
| 
74760
 
ef9f95d055f6
tuned generation of TPTP with $ite/$let in higher-order logics
 
desharna 
parents: 
74552 
diff
changeset
 | 
526  | 
\in higher order")  | 
| 74896 | 527  | 
| _ => error (tptp_ite ^ " is special syntax and must have at least three arguments"))  | 
| 52074 | 528  | 
else if s = tptp_choice then  | 
| 54829 | 529  | 
(case ts of  | 
| 74896 | 530  | 
(AAbs (((s', ty), tm), args) :: ts) =>  | 
| 72917 | 531  | 
(* There is code in "ATP_Problem_Generate" to ensure that "Eps" is always applied to an  | 
532  | 
abstraction. *)  | 
|
| 74896 | 533  | 
if ts = [] orelse is_format_higher_order_with_choice format then  | 
534  | 
let val declaration = s' ^ " : " ^ of_type ty in  | 
|
535  | 
              app0 ("(" ^ tptp_choice ^ "[" ^ declaration ^ "]: " ^ of_term tm ^ ")") [] (args @ ts)
 | 
|
536  | 
end  | 
|
537  | 
else  | 
|
538  | 
error (tptp_choice ^ " is only supported in higher order")  | 
|
539  | 
| _ => error (tptp_choice ^ " must be followed by a lambda abstraction"))  | 
|
| 74552 | 540  | 
else  | 
541  | 
(case (Symtab.lookup tptp_builtins s, ts) of  | 
|
542  | 
          (SOME {arity = 1, is_predicate = true}, [t]) =>
 | 
|
543  | 
(* generate e.g. "~ t" instead of "~ @ t". *)  | 
|
544  | 
s ^ " " ^ (of_term t |> parens) |> parens  | 
|
545  | 
        | (SOME {arity = 2, is_predicate = true}, [t1, t2]) =>
 | 
|
546  | 
(* generate e.g. "t1 & t2" instead of "& @ t1 @ t2". *)  | 
|
547  | 
(of_term t1 |> parens) ^ " " ^ s ^ " " ^ (of_term t2 |> parens) |> parens  | 
|
| 54829 | 548  | 
| _ => app ())  | 
| 52074 | 549  | 
end  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
550  | 
| tptp_string_of_term (format as THF _) (AAbs (((s, ty), tm), args)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
551  | 
tptp_string_of_app format  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
552  | 
        ("(^[" ^ s ^ " : " ^ string_of_type format ty ^ "]: " ^
 | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
553  | 
tptp_string_of_term format tm ^ ")")  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
554  | 
(map (tptp_string_of_term format) args)  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
555  | 
| tptp_string_of_term _ _ =  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
556  | 
raise Fail "unexpected term in first-order format"  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
557  | 
and tptp_string_of_formula format (ATyQuant (q, xs, phi)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
558  | 
tptp_string_of_quantifier q ^  | 
| 48135 | 559  | 
"[" ^  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
560  | 
commas (map (suffix_type_of_types o string_of_type format o fst) xs) ^  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
561  | 
"]: " ^ tptp_string_of_formula format phi  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
562  | 
|> parens  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
563  | 
| tptp_string_of_formula format (AQuant (q, xs, phi)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
564  | 
tptp_string_of_quantifier q ^  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
565  | 
"[" ^ commas (map (string_of_bound_var format) xs) ^ "]: " ^  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
566  | 
tptp_string_of_formula format phi  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
567  | 
|> parens  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
568  | 
| tptp_string_of_formula format  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
569  | 
        (AConn (ANot, [AAtom (ATerm (("=" (* tptp_equal *), []), ts))])) =
 | 
| 75363 | 570  | 
    space_implode (" " ^ tptp_not_equal ^ " ") (map (tptp_string_of_term format) ts)
 | 
571  | 
|> is_format_higher_order format ? parens  | 
|
572  | 
| tptp_string_of_formula format  | 
|
573  | 
        (AConn (ANot, [AAtom (ATerm (("equal" (* tptp_old_equal *), []), ts))])) =
 | 
|
574  | 
    space_implode (" " ^ tptp_not_equal ^ " ") (map (tptp_string_of_term format) ts)
 | 
|
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
575  | 
|> is_format_higher_order format ? parens  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
576  | 
| tptp_string_of_formula format (AConn (c, [phi])) =  | 
| 75363 | 577  | 
tptp_string_of_connective c ^ " " ^ (tptp_string_of_formula format phi |> parens)  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
578  | 
|> parens  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
579  | 
| tptp_string_of_formula format (AConn (c, phis)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
580  | 
    space_implode (" " ^ tptp_string_of_connective c ^ " ")
 | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
581  | 
(map (tptp_string_of_formula format) phis)  | 
| 
52075
 
f363f54a1936
generate agsyHOL-friendly THF (to some extent -- partial application of connectives remains an issue)
 
blanchet 
parents: 
52074 
diff
changeset
 | 
582  | 
|> parens  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
583  | 
| tptp_string_of_formula format (AAtom tm) = tptp_string_of_term format tm  | 
| 
37961
 
6a48c85a211a
first step in using "fof" rather than "cnf" in TPTP problems
 
blanchet 
parents: 
37931 
diff
changeset
 | 
584  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
585  | 
fun tptp_string_of_format CNF = tptp_cnf  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
586  | 
| tptp_string_of_format CNF_UEQ = tptp_cnf  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
587  | 
| tptp_string_of_format FOF = tptp_fof  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
588  | 
| tptp_string_of_format (TFF _) = tptp_tff  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
589  | 
| tptp_string_of_format (THF _) = tptp_thf  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
590  | 
| tptp_string_of_format (DFG _) = raise Fail "non-TPTP format"  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
591  | 
|
| 54820 | 592  | 
val atype_of_types = AType ((tptp_type_of_types, []), [])  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
593  | 
|
| 48141 | 594  | 
fun nary_type_decl_type n = funpow n (curry AFun atype_of_types) atype_of_types  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
595  | 
|
| 
50522
 
19dbd7554076
generate original name as a comment in SPASS problems as well
 
blanchet 
parents: 
50521 
diff
changeset
 | 
596  | 
fun maybe_alt "" = ""  | 
| 
 
19dbd7554076
generate original name as a comment in SPASS problems as well
 
blanchet 
parents: 
50521 
diff
changeset
 | 
597  | 
| maybe_alt s = " % " ^ s  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
50012 
diff
changeset
 | 
598  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
599  | 
fun tptp_string_of_line format (Type_Decl (ident, ty, ary)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
600  | 
tptp_string_of_line format (Sym_Decl (ident, ty, nary_type_decl_type ary))  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
601  | 
| tptp_string_of_line format (Sym_Decl (ident, sym, ty)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
602  | 
    tptp_string_of_format format ^ "(" ^ ident ^ ", type,\n    " ^ sym ^
 | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
603  | 
" : " ^ string_of_type format ty ^ ").\n"  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
604  | 
| tptp_string_of_line format (Formula ((ident, alt), kind, phi, source, info)) =  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
605  | 
    tptp_string_of_format format ^ "(" ^ ident ^ ", " ^
 | 
| 
57811
 
faab5feffb42
put comments between TPTP lines to comply with TPTP BNF
 
blanchet 
parents: 
57709 
diff
changeset
 | 
606  | 
    tptp_string_of_role kind ^ "," ^ "\n    (" ^
 | 
| 
 
faab5feffb42
put comments between TPTP lines to comply with TPTP BNF
 
blanchet 
parents: 
57709 
diff
changeset
 | 
607  | 
tptp_string_of_formula format phi ^ ")" ^  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
608  | 
(case source of  | 
| 
76301
 
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
 
blanchet 
parents: 
75363 
diff
changeset
 | 
609  | 
SOME tm => ", " ^ tptp_string_of_term FOF tm  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
610  | 
| NONE => if null info then "" else ", []") ^  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
611  | 
(case info of  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61549 
diff
changeset
 | 
612  | 
[] => ""  | 
| 
76301
 
73b120e0dbfe
generate some metainformation not only for SPASS but also for Zipperposition, for experimentation
 
blanchet 
parents: 
75363 
diff
changeset
 | 
613  | 
| tms => ", [" ^ commas (map (tptp_string_of_term FOF) tms) ^ "]") ^  | 
| 
57811
 
faab5feffb42
put comments between TPTP lines to comply with TPTP BNF
 
blanchet 
parents: 
57709 
diff
changeset
 | 
614  | 
")." ^ maybe_alt alt ^ "\n"  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
615  | 
|
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
616  | 
fun tptp_lines format =  | 
| 
37643
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
617  | 
maps (fn (_, []) => []  | 
| 
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
618  | 
| (heading, lines) =>  | 
| 41491 | 619  | 
           "\n% " ^ heading ^ " (" ^ string_of_int (length lines) ^ ")\n" ::
 | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
620  | 
map (tptp_string_of_line format) lines)  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
621  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
622  | 
fun arity_of_type (APi (tys, ty)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
623  | 
arity_of_type ty |>> Integer.add (length tys)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
624  | 
| arity_of_type (AFun (_, ty)) = arity_of_type ty ||> Integer.add 1  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
625  | 
| arity_of_type _ = (0, 0)  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
626  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
627  | 
fun string_of_arity (0, n) = string_of_int n  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
628  | 
| string_of_arity (m, n) = string_of_int m ^ "+" ^ string_of_int n  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
629  | 
|
| 48143 | 630  | 
val dfg_class_inter = space_implode " & "  | 
631  | 
||
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
632  | 
fun dfg_string_of_term (ATerm ((s, tys), tms)) =  | 
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
633  | 
s ^  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
634  | 
(if null tys then ""  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
635  | 
else "<" ^ commas (map (string_of_type (DFG Polymorphic)) tys) ^ ">") ^  | 
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
636  | 
(if null tms then ""  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
637  | 
     else "(" ^ commas (map dfg_string_of_term tms) ^ ")")
 | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
638  | 
| dfg_string_of_term _ = raise Fail "unexpected atom in first-order format"  | 
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
639  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
640  | 
fun dfg_string_of_formula poly gen_simp info =  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
641  | 
let  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
642  | 
val str_of_typ = string_of_type (DFG poly)  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
643  | 
fun str_of_bound_typ (ty, []) = str_of_typ ty  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
644  | 
| str_of_bound_typ (ty, cls) = str_of_typ ty ^ " : " ^ dfg_class_inter cls  | 
| 
46379
 
de5dd84717c1
distinguish between ":lr" and ":lt" (terminating) in DFG format
 
blanchet 
parents: 
46378 
diff
changeset
 | 
645  | 
fun suffix_tag top_level s =  | 
| 
48129
 
933d43c31689
removed support for unsorted DFG, now that SPASS 3.7 is no longer supported
 
blanchet 
parents: 
48004 
diff
changeset
 | 
646  | 
if top_level then  | 
| 54829 | 647  | 
(case extract_isabelle_status info of  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48143 
diff
changeset
 | 
648  | 
SOME s' =>  | 
| 54829 | 649  | 
if s' = non_rec_defN then s ^ ":lt"  | 
650  | 
else if (s' = simpN orelse s' = rec_defN) andalso gen_simp then s ^ ":lr"  | 
|
651  | 
else s  | 
|
652  | 
| NONE => s)  | 
|
| 
46379
 
de5dd84717c1
distinguish between ":lr" and ":lt" (terminating) in DFG format
 
blanchet 
parents: 
46378 
diff
changeset
 | 
653  | 
else  | 
| 
 
de5dd84717c1
distinguish between ":lr" and ":lt" (terminating) in DFG format
 
blanchet 
parents: 
46378 
diff
changeset
 | 
654  | 
s  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
655  | 
fun str_of_atom top_level (ATerm ((s, tys), tms)) =  | 
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
656  | 
let  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
657  | 
val s' =  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
658  | 
if is_tptp_equal s then "equal" |> suffix_tag top_level  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
659  | 
else if s = tptp_true then "true"  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
660  | 
else if s = tptp_false then "false"  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
661  | 
else s  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
662  | 
in dfg_string_of_term (ATerm ((s', tys), tms)) end  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
663  | 
| str_of_atom _ _ = raise Fail "unexpected atom in first-order format"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
664  | 
fun str_of_quant AForall = "forall"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
665  | 
| str_of_quant AExists = "exists"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
666  | 
fun str_of_conn _ ANot = "not"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
667  | 
| str_of_conn _ AAnd = "and"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
668  | 
| str_of_conn _ AOr = "or"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
669  | 
| str_of_conn _ AImplies = "implies"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
670  | 
| str_of_conn top_level AIff = "equiv" |> suffix_tag top_level  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
671  | 
fun str_of_formula top_level (ATyQuant (q, xs, phi)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
672  | 
str_of_quant q ^ "_sorts([" ^ commas (map str_of_bound_typ xs) ^ "], " ^  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
673  | 
str_of_formula top_level phi ^ ")"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
674  | 
| str_of_formula top_level (AQuant (q, xs, phi)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
675  | 
str_of_quant q ^ "([" ^  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
676  | 
commas (map (string_of_bound_var (DFG poly)) xs) ^ "], " ^  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
677  | 
str_of_formula top_level phi ^ ")"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
678  | 
| str_of_formula top_level (AConn (c, phis)) =  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
679  | 
        str_of_conn top_level c ^ "(" ^
 | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
680  | 
commas (map (str_of_formula false) phis) ^ ")"  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
681  | 
| str_of_formula top_level (AAtom tm) = str_of_atom top_level tm  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
682  | 
in str_of_formula true end  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
683  | 
|
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
684  | 
fun maybe_enclose bef aft "" = "% " ^ bef ^ aft  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
685  | 
| maybe_enclose bef aft s = bef ^ s ^ aft  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
686  | 
|
| 75341 | 687  | 
fun dfg_lines poly ord_info problem =  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
688  | 
let  | 
| 75341 | 689  | 
(* hardcoded settings *)  | 
690  | 
val is_lpo = false  | 
|
691  | 
val gen_weights = true  | 
|
692  | 
val gen_prec = true  | 
|
693  | 
val gen_simp = false  | 
|
694  | 
||
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
695  | 
val typ = string_of_type (DFG poly)  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51997 
diff
changeset
 | 
696  | 
val term = dfg_string_of_term  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
697  | 
    fun spair (s, s') = "(" ^ s ^ ", " ^ s' ^ ")"
 | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
698  | 
fun tm_ary sym ty = spair (sym, string_of_arity (arity_of_type ty))  | 
| 48141 | 699  | 
fun ty_ary 0 ty = ty  | 
700  | 
      | ty_ary n ty = "(" ^ ty ^ ", " ^ string_of_int n ^ ")"
 | 
|
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
701  | 
    fun fun_typ sym ty = "function(" ^ sym ^ ", " ^ typ ty ^ ")."
 | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
702  | 
fun pred_typ sym ty =  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
703  | 
let  | 
| 52025 | 704  | 
val (ty_vars, (tys, _)) =  | 
705  | 
strip_atype ty  | 
|
706  | 
|>> (fn [] => [] | xs => ["[" ^ commas xs ^ "]"])  | 
|
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
707  | 
      in "predicate(" ^ commas (sym :: ty_vars @ map typ tys) ^ ")." end
 | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
708  | 
fun bound_tvar (ty, []) = ty  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
709  | 
| bound_tvar (ty, cls) = ty ^ " : " ^ dfg_class_inter cls  | 
| 52004 | 710  | 
fun binder_typ xs ty =  | 
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
711  | 
(if null xs then "" else "[" ^ commas (map bound_tvar xs) ^ "], ") ^  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
712  | 
typ ty  | 
| 52004 | 713  | 
    fun sort_decl xs ty cl = "sort(" ^ binder_typ xs ty ^ ", " ^ cl ^ ")."
 | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
714  | 
fun datatype_decl xs ty tms exhaust =  | 
| 
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
715  | 
      "datatype(" ^ commas (binder_typ xs ty :: map term tms @ (if exhaust then [] else ["..."])) ^
 | 
| 
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
716  | 
")."  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
717  | 
    fun subclass_of sub super = "subclass(" ^ sub ^ ", " ^ super ^ ")."
 | 
| 
50522
 
19dbd7554076
generate original name as a comment in SPASS problems as well
 
blanchet 
parents: 
50521 
diff
changeset
 | 
718  | 
fun formula pred (Formula ((ident, alt), kind, phi, _, info)) =  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
719  | 
if pred kind then  | 
| 
48129
 
933d43c31689
removed support for unsorted DFG, now that SPASS 3.7 is no longer supported
 
blanchet 
parents: 
48004 
diff
changeset
 | 
720  | 
let val rank = extract_isabelle_rank info in  | 
| 
57811
 
faab5feffb42
put comments between TPTP lines to comply with TPTP BNF
 
blanchet 
parents: 
57709 
diff
changeset
 | 
721  | 
            "formula(" ^ dfg_string_of_formula poly gen_simp info phi ^ ", " ^ ident ^
 | 
| 46406 | 722  | 
(if rank = default_rank then "" else ", " ^ string_of_int rank) ^  | 
| 
50522
 
19dbd7554076
generate original name as a comment in SPASS problems as well
 
blanchet 
parents: 
50521 
diff
changeset
 | 
723  | 
")." ^ maybe_alt alt  | 
| 
 
19dbd7554076
generate original name as a comment in SPASS problems as well
 
blanchet 
parents: 
50521 
diff
changeset
 | 
724  | 
|> SOME  | 
| 46406 | 725  | 
end  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
726  | 
else  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
727  | 
NONE  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
728  | 
| formula _ _ = NONE  | 
| 46413 | 729  | 
fun filt f = problem |> map (map_filter f o snd) |> filter_out null  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
730  | 
val func_aries =  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
731  | 
filt (fn Sym_Decl (_, sym, ty) =>  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
732  | 
if is_function_atype ty then SOME (tm_ary sym ty) else NONE  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
733  | 
| _ => NONE)  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
734  | 
|> flat |> commas |> maybe_enclose "functions [" "]."  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
735  | 
val pred_aries =  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
736  | 
filt (fn Sym_Decl (_, sym, ty) =>  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
737  | 
if is_predicate_atype ty then SOME (tm_ary sym ty) else NONE  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
738  | 
| _ => NONE)  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
739  | 
|> flat |> commas |> maybe_enclose "predicates [" "]."  | 
| 
48129
 
933d43c31689
removed support for unsorted DFG, now that SPASS 3.7 is no longer supported
 
blanchet 
parents: 
48004 
diff
changeset
 | 
740  | 
val sorts =  | 
| 52000 | 741  | 
filt (try (fn Type_Decl (_, ty, ary) => ty_ary ary ty)) @  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
742  | 
[[ty_ary 0 dfg_individual_type]]  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
743  | 
|> flat |> commas |> maybe_enclose "sorts [" "]."  | 
| 48141 | 744  | 
val classes =  | 
| 52000 | 745  | 
filt (try (fn Class_Decl (_, cl, _) => cl))  | 
| 48141 | 746  | 
|> flat |> commas |> maybe_enclose "classes [" "]."  | 
| 
48129
 
933d43c31689
removed support for unsorted DFG, now that SPASS 3.7 is no longer supported
 
blanchet 
parents: 
48004 
diff
changeset
 | 
747  | 
val ord_info = if gen_weights orelse gen_prec then ord_info () else []  | 
| 
 
933d43c31689
removed support for unsorted DFG, now that SPASS 3.7 is no longer supported
 
blanchet 
parents: 
48004 
diff
changeset
 | 
748  | 
val do_term_order_weights =  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
749  | 
(if gen_weights then ord_info else [])  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
750  | 
|> map (spair o apsnd string_of_int) |> commas  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
751  | 
|> maybe_enclose "weights [" "]."  | 
| 48141 | 752  | 
val syms = [func_aries, pred_aries, do_term_order_weights, sorts, classes]  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
753  | 
val func_decls =  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
754  | 
filt (fn Sym_Decl (_, sym, ty) =>  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
755  | 
if is_function_atype ty then SOME (fun_typ sym ty) else NONE  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
756  | 
| _ => NONE) |> flat  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
757  | 
val pred_decls =  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
758  | 
filt (fn Sym_Decl (_, sym, ty) =>  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48135 
diff
changeset
 | 
759  | 
if is_nontrivial_predicate_atype ty then SOME (pred_typ sym ty)  | 
| 
46391
 
8d8d3c1f1854
really fixed syntax bug in DFG output (cf. ef62c2fafa9e)
 
blanchet 
parents: 
46382 
diff
changeset
 | 
760  | 
else NONE  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
761  | 
| _ => NONE) |> flat  | 
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
762  | 
val datatype_decls =  | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
763  | 
filt (try (fn Datatype_Decl (_, xs, ty, tms, exhaust) => datatype_decl xs ty tms exhaust))  | 
| 
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
764  | 
|> flat  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
765  | 
val sort_decls =  | 
| 52000 | 766  | 
filt (try (fn Class_Memb (_, xs, ty, cl) => sort_decl xs ty cl)) |> flat  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
767  | 
val subclass_decls =  | 
| 52000 | 768  | 
filt (try (fn Class_Decl (_, sub, supers) =>  | 
769  | 
map (subclass_of sub) supers))  | 
|
770  | 
|> flat |> flat  | 
|
| 
51997
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
771  | 
val decls =  | 
| 
 
8dbe623a7804
added datatype declaration syntax for next-gen SPASS
 
blanchet 
parents: 
51878 
diff
changeset
 | 
772  | 
func_decls @ pred_decls @ datatype_decls @ sort_decls @ subclass_decls  | 
| 46413 | 773  | 
val axioms =  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
774  | 
filt (formula (curry (op <>) Conjecture)) |> separate [""] |> flat  | 
| 46413 | 775  | 
val conjs =  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
776  | 
filt (formula (curry (op =) Conjecture)) |> separate [""] |> flat  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
777  | 
val settings =  | 
| 47054 | 778  | 
(if is_lpo then ["set_flag(Ordering, 1)."] else []) @  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
779  | 
(if gen_prec then  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
780  | 
[ord_info |> map fst |> rev |> commas  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
781  | 
                   |> maybe_enclose "set_precedence(" ")."]
 | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
782  | 
else  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47030 
diff
changeset
 | 
783  | 
[])  | 
| 47053 | 784  | 
fun list_of _ [] = []  | 
785  | 
| list_of heading ss =  | 
|
786  | 
"list_of_" ^ heading ^ ".\n" :: map (suffix "\n") ss @  | 
|
787  | 
["end_of_list.\n\n"]  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
788  | 
in  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
789  | 
"\nbegin_problem(isabelle).\n\n" ::  | 
| 47053 | 790  | 
list_of "descriptions"  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
791  | 
            ["name({**}).", "author({**}).", "status(unknown).",
 | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
792  | 
             "description({**})."] @
 | 
| 47053 | 793  | 
list_of "symbols" syms @  | 
794  | 
list_of "declarations" decls @  | 
|
795  | 
list_of "formulae(axioms)" axioms @  | 
|
796  | 
list_of "formulae(conjectures)" conjs @  | 
|
797  | 
list_of "settings(SPASS)" settings @  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
798  | 
["end_problem.\n"]  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
799  | 
end  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
800  | 
|
| 75341 | 801  | 
fun lines_of_atp_problem format ord_info problem =  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
802  | 
"% This file was generated by Isabelle (most likely Sledgehammer)\n\  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
44787 
diff
changeset
 | 
803  | 
\% " ^ timestamp () ^ "\n" ::  | 
| 48131 | 804  | 
(case format of  | 
| 75341 | 805  | 
DFG poly => dfg_lines poly ord_info  | 
| 48131 | 806  | 
| _ => tptp_lines format) problem  | 
| 
37643
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
807  | 
|
| 42939 | 808  | 
|
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
809  | 
(** CNF (Metis) and CNF UEQ (Waldmeister) **)  | 
| 42939 | 810  | 
|
| 48140 | 811  | 
fun is_line_negated (Formula (_, _, AConn (ANot, _), _, _)) = true  | 
812  | 
| is_line_negated _ = false  | 
|
| 42939 | 813  | 
|
| 48140 | 814  | 
fun is_line_cnf_ueq (Formula (_, _, AAtom (ATerm (((s, _), _), _)), _, _)) =  | 
| 
43193
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
815  | 
is_tptp_equal s  | 
| 48140 | 816  | 
| is_line_cnf_ueq _ = false  | 
| 42939 | 817  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
818  | 
fun open_conjecture_term (ATerm (((s, s'), tys), tms)) =  | 
| 56811 | 819  | 
ATerm ((if is_tptp_variable s then (s |> Name.desymbolize (SOME false), s')  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
820  | 
else (s, s'), tys), tms |> map open_conjecture_term)  | 
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43422 
diff
changeset
 | 
821  | 
| open_conjecture_term _ = raise Fail "unexpected higher-order term"  | 
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
822  | 
fun open_formula conj =  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
823  | 
let  | 
| 
43193
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
824  | 
(* We are conveniently assuming that all bound variable names are  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
825  | 
distinct, which should be the case for the formulas we generate. *)  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
826  | 
fun opn (pos as SOME true) (AQuant (AForall, _, phi)) = opn pos phi  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
827  | 
| opn (pos as SOME false) (AQuant (AExists, _, phi)) = opn pos phi  | 
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
828  | 
| opn pos (AConn (ANot, [phi])) = mk_anot (opn (Option.map not pos) phi)  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
829  | 
| opn pos (AConn (c, [phi1, phi2])) =  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
830  | 
let val (pos1, pos2) = polarities_of_conn pos c in  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
831  | 
AConn (c, [opn pos1 phi1, opn pos2 phi2])  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
832  | 
end  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
833  | 
| opn _ (AAtom t) = AAtom (t |> conj ? open_conjecture_term)  | 
| 
43422
 
dcbedaf6f80c
added missing case in pattern matching -- solves Waldmeister "Match" exceptions that have been plaguing some users
 
blanchet 
parents: 
43295 
diff
changeset
 | 
834  | 
| opn _ phi = phi  | 
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
835  | 
in opn (SOME (not conj)) end  | 
| 
42944
 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 
blanchet 
parents: 
42942 
diff
changeset
 | 
836  | 
fun open_formula_line (Formula (ident, kind, phi, source, info)) =  | 
| 
 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 
blanchet 
parents: 
42942 
diff
changeset
 | 
837  | 
Formula (ident, kind, open_formula (kind = Conjecture) phi, source, info)  | 
| 
 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 
blanchet 
parents: 
42942 
diff
changeset
 | 
838  | 
| open_formula_line line = line  | 
| 42939 | 839  | 
|
840  | 
fun negate_conjecture_line (Formula (ident, Conjecture, phi, source, info)) =  | 
|
| 42942 | 841  | 
Formula (ident, Hypothesis, mk_anot phi, source, info)  | 
| 42939 | 842  | 
| negate_conjecture_line line = line  | 
843  | 
||
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
844  | 
exception CLAUSIFY of unit  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
845  | 
|
| 
43126
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
846  | 
(* This "clausification" only expands syntactic sugar, such as "phi => psi" to  | 
| 
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
847  | 
"~ phi | psi" and "phi <=> psi" to "~ phi | psi" and "~ psi | phi". We don't  | 
| 
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
848  | 
attempt to distribute conjunctions over disjunctions. *)  | 
| 
43193
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
849  | 
fun clausify_formula pos (phi as AAtom _) = [phi |> not pos ? mk_anot]  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
850  | 
| clausify_formula pos (AConn (ANot, [phi])) = clausify_formula (not pos) phi  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
851  | 
| clausify_formula true (AConn (AOr, [phi1, phi2])) =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58600 
diff
changeset
 | 
852  | 
(phi1, phi2) |> apply2 (clausify_formula true)  | 
| 
43193
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
853  | 
|> uncurry (map_product (mk_aconn AOr))  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
854  | 
| clausify_formula false (AConn (AAnd, [phi1, phi2])) =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58600 
diff
changeset
 | 
855  | 
(phi1, phi2) |> apply2 (clausify_formula false)  | 
| 
43193
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
856  | 
|> uncurry (map_product (mk_aconn AOr))  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
857  | 
| clausify_formula true (AConn (AImplies, [phi1, phi2])) =  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
858  | 
clausify_formula true (AConn (AOr, [mk_anot phi1, phi2]))  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
859  | 
| clausify_formula true (AConn (AIff, phis)) =  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
860  | 
clausify_formula true (AConn (AImplies, phis)) @  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
861  | 
clausify_formula true (AConn (AImplies, rev phis))  | 
| 
 
e11bd628f1a5
improved ATP clausifier so it can deal with "x => (y <=> z)"
 
blanchet 
parents: 
43163 
diff
changeset
 | 
862  | 
| clausify_formula _ _ = raise CLAUSIFY ()  | 
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
863  | 
|
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
50012 
diff
changeset
 | 
864  | 
fun clausify_formula_line (Formula ((ident, alt), kind, phi, source, info)) =  | 
| 
43126
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
865  | 
let  | 
| 
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
866  | 
val (n, phis) = phi |> try (clausify_formula true) |> these |> `length  | 
| 
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
867  | 
in  | 
| 
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
868  | 
map2 (fn phi => fn j =>  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
50012 
diff
changeset
 | 
869  | 
Formula ((ident ^ replicate_string (j - 1) "x", alt), kind, phi,  | 
| 
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
50012 
diff
changeset
 | 
870  | 
source, info))  | 
| 
43126
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
871  | 
phis (1 upto n)  | 
| 
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
872  | 
end  | 
| 
 
a7db0afd5200
clausify "<=>" (needed for some type information)
 
blanchet 
parents: 
43098 
diff
changeset
 | 
873  | 
| clausify_formula_line _ = []  | 
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
874  | 
|
| 48140 | 875  | 
fun ensure_cnf_line line =  | 
| 
43098
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
876  | 
line |> open_formula_line |> negate_conjecture_line |> clausify_formula_line  | 
| 
 
e88e974c4846
proper handling of type variable classes in new Metis
 
blanchet 
parents: 
43092 
diff
changeset
 | 
877  | 
|
| 48140 | 878  | 
fun ensure_cnf_problem problem = problem |> map (apsnd (maps ensure_cnf_line))  | 
| 
43092
 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 
blanchet 
parents: 
43085 
diff
changeset
 | 
879  | 
|
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
880  | 
fun filter_cnf_ueq_problem problem =  | 
| 
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
881  | 
problem  | 
| 48140 | 882  | 
|> map (apsnd (map open_formula_line #> filter is_line_cnf_ueq  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
883  | 
#> map negate_conjecture_line))  | 
| 
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42961 
diff
changeset
 | 
884  | 
|> (fn problem =>  | 
| 42939 | 885  | 
let  | 
| 
43824
 
0234156d3fbe
don't generate Waldmeister problems with only a conjecture, since it makes it crash sometimes
 
blanchet 
parents: 
43692 
diff
changeset
 | 
886  | 
val lines = problem |> maps snd  | 
| 48140 | 887  | 
val conjs = lines |> filter is_line_negated  | 
| 
43824
 
0234156d3fbe
don't generate Waldmeister problems with only a conjecture, since it makes it crash sometimes
 
blanchet 
parents: 
43692 
diff
changeset
 | 
888  | 
in if length conjs = 1 andalso conjs <> lines then problem else [] end)  | 
| 
38017
 
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
 
blanchet 
parents: 
38014 
diff
changeset
 | 
889  | 
|
| 
37643
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
890  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
891  | 
(** Symbol declarations **)  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
892  | 
|
| 48141 | 893  | 
fun add_declared_in_line (Class_Decl (_, cl, _)) = apfst (apfst (cons cl))  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
894  | 
| add_declared_in_line (Type_Decl (_, ty, _)) = apfst (apsnd (cons ty))  | 
| 48140 | 895  | 
| add_declared_in_line (Sym_Decl (_, sym, _)) = apsnd (cons sym)  | 
896  | 
| add_declared_in_line _ = I  | 
|
| 48141 | 897  | 
fun declared_in_atp_problem problem =  | 
898  | 
fold (fold add_declared_in_line o snd) problem (([], []), [])  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
899  | 
|
| 
37643
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
900  | 
(** Nice names **)  | 
| 
 
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
 
blanchet 
parents: 
37642 
diff
changeset
 | 
901  | 
|
| 
42227
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
902  | 
val no_qualifiers =  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
903  | 
let  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
904  | 
fun skip [] = []  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
905  | 
| skip (#"." :: cs) = skip cs  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
906  | 
| skip (c :: cs) = if Char.isAlphaNum c then skip cs else c :: keep cs  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
907  | 
and keep [] = []  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
908  | 
| keep (#"." :: cs) = skip cs  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
909  | 
| keep (c :: cs) = c :: keep cs  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
910  | 
in String.explode #> rev #> keep #> rev #> String.implode end  | 
| 
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
41769 
diff
changeset
 | 
911  | 
|
| 
42761
 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 
blanchet 
parents: 
42752 
diff
changeset
 | 
912  | 
(* Long names can slow down the ATPs. *)  | 
| 
42724
 
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
 
blanchet 
parents: 
42722 
diff
changeset
 | 
913  | 
val max_readable_name_size = 20  | 
| 
42567
 
d012947edd36
shorten readable names -- they can get really long with monomorphization, which actually slows down the ATPs
 
blanchet 
parents: 
42543 
diff
changeset
 | 
914  | 
|
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
915  | 
(* "equal" is reserved by some ATPs. "op" is also reserved, to avoid the  | 
| 
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
916  | 
unreadable "op_1", "op_2", etc., in the problem files. "eq" is reserved to  | 
| 
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
917  | 
ensure that "HOL.eq" is correctly mapped to equality (not clear whether this  | 
| 
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
918  | 
is still necessary). *)  | 
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
919  | 
val reserved_nice_names = [tptp_old_equal, "op", "eq"]  | 
| 42939 | 920  | 
|
| 
46414
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
921  | 
(* hack to get the same hashing across Mirabelle runs (see "mirabelle.pl") *)  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
922  | 
fun cleanup_mirabelle_name s =  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
923  | 
let  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
924  | 
val mirabelle_infix = "_Mirabelle_"  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
925  | 
val random_suffix_len = 10  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
926  | 
val (s1, s2) = Substring.position mirabelle_infix (Substring.full s)  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
927  | 
in  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
928  | 
if Substring.isEmpty s2 then  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
929  | 
s  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
930  | 
else  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
931  | 
Substring.string s1 ^  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
932  | 
Substring.string (Substring.triml (size mirabelle_infix + random_suffix_len) s2)  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
933  | 
end  | 
| 
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
934  | 
|
| 46378 | 935  | 
fun readable_name protect full_name s =  | 
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
936  | 
(if s = full_name then  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
937  | 
s  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
938  | 
else  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
939  | 
s |> no_qualifiers  | 
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52075 
diff
changeset
 | 
940  | 
|> unquote_tvar  | 
| 56811 | 941  | 
|> Name.desymbolize (SOME (Char.isUpper (String.sub (full_name, 0))))  | 
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
942  | 
|> (fn s =>  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
943  | 
if size s > max_readable_name_size then  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
944  | 
String.substring (s, 0, max_readable_name_size div 2 - 4) ^  | 
| 
46414
 
4ed12518fb81
improved hashing w.r.t. Mirabelle, to help debugging
 
blanchet 
parents: 
46413 
diff
changeset
 | 
945  | 
string_of_int (hash_string (cleanup_mirabelle_name full_name)) ^  | 
| 56847 | 946  | 
String.extract (s, size s - max_readable_name_size div 2 + 4, NONE)  | 
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
947  | 
else  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
948  | 
s)  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
949  | 
|> (fn s =>  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
950  | 
if member (op =) reserved_nice_names s then full_name else s))  | 
| 46378 | 951  | 
|> protect  | 
| 37624 | 952  | 
|
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
953  | 
fun nice_name _ (full_name, _) NONE = (full_name, NONE)  | 
| 46378 | 954  | 
| nice_name protect (full_name, desired_name) (SOME the_pool) =  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
955  | 
if is_built_in_tptp_symbol full_name then  | 
| 39384 | 956  | 
(full_name, SOME the_pool)  | 
| 54829 | 957  | 
else  | 
958  | 
(case Symtab.lookup (fst the_pool) full_name of  | 
|
959  | 
SOME nice_name => (nice_name, SOME the_pool)  | 
|
960  | 
| NONE =>  | 
|
961  | 
let  | 
|
962  | 
val nice_prefix = readable_name protect full_name desired_name  | 
|
963  | 
fun add j =  | 
|
964  | 
let  | 
|
965  | 
val nice_name = nice_prefix ^ (if j = 1 then "" else string_of_int j)  | 
|
966  | 
in  | 
|
967  | 
(case Symtab.lookup (snd the_pool) nice_name of  | 
|
968  | 
SOME full_name' =>  | 
|
969  | 
if full_name = full_name' then (nice_name, the_pool) else add (j + 1)  | 
|
970  | 
| NONE =>  | 
|
971  | 
(nice_name,  | 
|
972  | 
(Symtab.update_new (full_name, nice_name) (fst the_pool),  | 
|
973  | 
Symtab.update_new (nice_name, full_name) (snd the_pool))))  | 
|
974  | 
end  | 
|
975  | 
in add 1 |> apsnd SOME end)  | 
|
| 37624 | 976  | 
|
| 46378 | 977  | 
fun avoid_clash_with_dfg_keywords s =  | 
978  | 
let val n = String.size s in  | 
|
| 
46443
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
979  | 
if n < 2 orelse (n = 2 andalso String.sub (s, 0) = String.sub (s, 1)) orelse  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
980  | 
String.isSubstring "_" s then  | 
| 46378 | 981  | 
s  | 
| 47150 | 982  | 
else if is_tptp_variable s then  | 
| 
52353
 
dba3d398c322
SPASS has more Uppercase keywords than I was fearing -- better always append _
 
blanchet 
parents: 
52076 
diff
changeset
 | 
983  | 
s ^ "_"  | 
| 46378 | 984  | 
else  | 
985  | 
String.substring (s, 0, n - 1) ^  | 
|
986  | 
String.str (Char.toUpper (String.sub (s, n - 1)))  | 
|
987  | 
end  | 
|
988  | 
||
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
989  | 
fun nice_atp_problem readable_names format problem =  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
990  | 
let  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
991  | 
val empty_pool =  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
992  | 
if readable_names then SOME (Symtab.empty, Symtab.empty) else NONE  | 
| 
46643
 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 
blanchet 
parents: 
46445 
diff
changeset
 | 
993  | 
val avoid_clash =  | 
| 54829 | 994  | 
(case format of  | 
| 72588 | 995  | 
DFG _ => avoid_clash_with_dfg_keywords  | 
| 54829 | 996  | 
| _ => I)  | 
| 
46643
 
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
 
blanchet 
parents: 
46445 
diff
changeset
 | 
997  | 
val nice_name = nice_name avoid_clash  | 
| 54820 | 998  | 
|
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
52000 
diff
changeset
 | 
999  | 
fun nice_bound_tvars xs =  | 
| 52003 | 1000  | 
fold_map (nice_name o fst) xs  | 
1001  | 
##>> fold_map (fold_map nice_name o snd) xs  | 
|
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
52000 
diff
changeset
 | 
1002  | 
#>> op ~~  | 
| 54820 | 1003  | 
|
1004  | 
fun nice_type (AType ((name, clss), tys)) =  | 
|
1005  | 
nice_name name ##>> fold_map nice_name clss ##>> fold_map nice_type tys #>> AType  | 
|
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
1006  | 
| nice_type (AFun (ty1, ty2)) = nice_type ty1 ##>> nice_type ty2 #>> AFun  | 
| 54820 | 1007  | 
| nice_type (APi (names, ty)) = fold_map nice_name names ##>> nice_type ty #>> APi  | 
1008  | 
||
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
1009  | 
fun nice_term (ATerm ((name, tys), ts)) =  | 
| 54820 | 1010  | 
nice_name name ##>> fold_map nice_type tys ##>> fold_map nice_term ts #>> ATerm  | 
| 
47911
 
2168126446bb
extend ATP data structure to avoid having to perform ((non-)capture avoiding) beta reduction -- fixes a bug in the THF translation of "is_measure.simps"
 
blanchet 
parents: 
47768 
diff
changeset
 | 
1011  | 
| nice_term (AAbs (((name, ty), tm), args)) =  | 
| 54820 | 1012  | 
nice_name name ##>> nice_type ty ##>> nice_term tm ##>> fold_map nice_term args #>> AAbs  | 
1013  | 
||
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1014  | 
fun nice_formula (ATyQuant (q, xs, phi)) =  | 
| 52003 | 1015  | 
fold_map (nice_type o fst) xs  | 
1016  | 
##>> fold_map (fold_map nice_name o snd) xs  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1017  | 
##>> nice_formula phi  | 
| 48141 | 1018  | 
#>> (fn ((tys, cls), phi) => ATyQuant (q, tys ~~ cls, phi))  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1019  | 
| nice_formula (AQuant (q, xs, phi)) =  | 
| 52003 | 1020  | 
fold_map (nice_name o fst) xs  | 
1021  | 
##>> fold_map (fn (_, NONE) => pair NONE  | 
|
1022  | 
| (_, SOME ty) => nice_type ty #>> SOME) xs  | 
|
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
1023  | 
##>> nice_formula phi  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
1024  | 
#>> (fn ((ss, ts), phi) => AQuant (q, ss ~~ ts, phi))  | 
| 
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
1025  | 
| nice_formula (AConn (c, phis)) =  | 
| 52002 | 1026  | 
fold_map nice_formula phis #>> curry AConn c  | 
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
1027  | 
| nice_formula (AAtom tm) = nice_term tm #>> AAtom  | 
| 54820 | 1028  | 
|
| 48141 | 1029  | 
fun nice_line (Class_Decl (ident, cl, cls)) =  | 
| 52002 | 1030  | 
nice_name cl ##>> fold_map nice_name cls  | 
| 48141 | 1031  | 
#>> (fn (cl, cls) => Class_Decl (ident, cl, cls))  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1032  | 
| nice_line (Type_Decl (ident, ty, ary)) =  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1033  | 
nice_name ty #>> (fn ty => Type_Decl (ident, ty, ary))  | 
| 48140 | 1034  | 
| nice_line (Sym_Decl (ident, sym, ty)) =  | 
| 54820 | 1035  | 
nice_name sym ##>> nice_type ty #>> (fn (sym, ty) => Sym_Decl (ident, sym, ty))  | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
1036  | 
| nice_line (Datatype_Decl (ident, xs, ty, tms, exhaust)) =  | 
| 52002 | 1037  | 
nice_bound_tvars xs ##>> nice_type ty ##>> fold_map nice_term tms  | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
54829 
diff
changeset
 | 
1038  | 
#>> (fn ((xs, ty), tms) => Datatype_Decl (ident, xs, ty, tms, exhaust))  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1039  | 
| nice_line (Class_Memb (ident, xs, ty, cl)) =  | 
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
52000 
diff
changeset
 | 
1040  | 
nice_bound_tvars xs ##>> nice_type ty ##>> nice_name cl  | 
| 
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
52000 
diff
changeset
 | 
1041  | 
#>> (fn ((xs, ty), cl) => Class_Memb (ident, xs, ty, cl))  | 
| 48140 | 1042  | 
| nice_line (Formula (ident, kind, phi, source, info)) =  | 
| 54820 | 1043  | 
nice_formula phi #>> (fn phi => Formula (ident, kind, phi, source, info))  | 
1044  | 
||
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45938 
diff
changeset
 | 
1045  | 
fun nice_problem problem =  | 
| 54820 | 1046  | 
fold_map (fn (heading, lines) => fold_map nice_line lines #>> pair heading) problem  | 
1047  | 
in  | 
|
1048  | 
nice_problem problem empty_pool  | 
|
1049  | 
end  | 
|
| 
37509
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
1050  | 
|
| 
 
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
 
blanchet 
parents:  
diff
changeset
 | 
1051  | 
end;  |