| author | wenzelm | 
| Sun, 21 Aug 2022 23:01:08 +0200 | |
| changeset 75954 | 7c72091abd55 | 
| parent 75341 | 72cbbb4d98f3 | 
| child 76301 | 73b120e0dbfe | 
| permissions | -rw-r--r-- | 
| 46320 | 1  | 
(* Title: HOL/Tools/ATP/atp_problem_generate.ML  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Fabian Immler, TU Muenchen  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
3  | 
Author: Makarius  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
72922
 
d78bd4432f05
tuned name generation in tptp to not depend on shadowing
 
desharna 
parents: 
72921 
diff
changeset
 | 
5  | 
Author: Martin Desharnais, UniBw Muenchen  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
|
| 
43863
 
a43d61270142
ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
 
blanchet 
parents: 
43862 
diff
changeset
 | 
7  | 
Translation of HOL to FOL for Metis and Sledgehammer.  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
*)  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
9  | 
|
| 46320 | 10  | 
signature ATP_PROBLEM_GENERATE =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
11  | 
sig  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
12  | 
  type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
13  | 
type atp_connective = ATP_Problem.atp_connective  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
14  | 
  type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula
 | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
15  | 
type atp_format = ATP_Problem.atp_format  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
16  | 
type atp_formula_role = ATP_Problem.atp_formula_role  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
17  | 
type 'a atp_problem = 'a ATP_Problem.atp_problem  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
18  | 
|
| 62718 | 19  | 
datatype mode = Metis | Sledgehammer | Sledgehammer_Completish of int | Exporter | Translator  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
20  | 
|
| 46340 | 21  | 
datatype scope = Global | Local | Assum | Chained  | 
| 55212 | 22  | 
datatype status = General | Induction | Intro | Inductive | Elim | Simp | Non_Rec_Def | Rec_Def  | 
23  | 
||
| 46340 | 24  | 
type stature = scope * status  | 
| 
42613
 
23b13b1bd565
use strings to encode type systems in ATP module, to reduce the amount of out-of-place information and also to make it easier to print the type system used
 
blanchet 
parents: 
42612 
diff
changeset
 | 
25  | 
|
| 46301 | 26  | 
datatype strictness = Strict | Non_Strict  | 
| 48183 | 27  | 
datatype uniformity = Uniform | Non_Uniform  | 
| 52026 | 28  | 
datatype ctr_optim = With_Ctr_Optim | Without_Ctr_Optim  | 
| 
42613
 
23b13b1bd565
use strings to encode type systems in ATP module, to reduce the amount of out-of-place information and also to make it easier to print the type system used
 
blanchet 
parents: 
42612 
diff
changeset
 | 
29  | 
datatype type_level =  | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
30  | 
All_Types |  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
31  | 
Undercover_Types |  | 
| 48183 | 32  | 
Nonmono_Types of strictness * uniformity |  | 
| 52026 | 33  | 
Const_Types of ctr_optim |  | 
| 43362 | 34  | 
No_Types  | 
| 44782 | 35  | 
type type_enc  | 
| 
42613
 
23b13b1bd565
use strings to encode type systems in ATP module, to reduce the amount of out-of-place information and also to make it easier to print the type system used
 
blanchet 
parents: 
42612 
diff
changeset
 | 
36  | 
|
| 
45513
 
25388cf06437
rename the lambda translation schemes, so that they are understandable out of context
 
blanchet 
parents: 
45511 
diff
changeset
 | 
37  | 
val no_lamsN : string  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
38  | 
val opaque_liftingN : string  | 
| 46365 | 39  | 
val liftingN : string  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
40  | 
val opaque_combsN : string  | 
| 46365 | 41  | 
val combsN : string  | 
42  | 
val combs_and_liftingN : string  | 
|
43  | 
val combs_or_liftingN : string  | 
|
| 
45513
 
25388cf06437
rename the lambda translation schemes, so that they are understandable out of context
 
blanchet 
parents: 
45511 
diff
changeset
 | 
44  | 
val keep_lamsN : string  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
45  | 
val schematic_var_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
46  | 
val fixed_var_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
47  | 
val tvar_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
48  | 
val tfree_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
49  | 
val const_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
50  | 
val type_const_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
51  | 
val class_prefix : string  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
52  | 
val lam_lifted_prefix : string  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
53  | 
val lam_lifted_mono_prefix : string  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
54  | 
val lam_lifted_poly_prefix : string  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
55  | 
val skolem_const_prefix : string  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
56  | 
val old_skolem_const_prefix : string  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
57  | 
val new_skolem_const_prefix : string  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
58  | 
val combinator_prefix : string  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
59  | 
val class_decl_prefix : string  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
60  | 
val type_decl_prefix : string  | 
| 
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
61  | 
val sym_decl_prefix : string  | 
| 52004 | 62  | 
val datatype_decl_prefix : string  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
63  | 
val class_memb_prefix : string  | 
| 43989 | 64  | 
val guards_sym_formula_prefix : string  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
65  | 
val tags_sym_formula_prefix : string  | 
| 
40204
 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 
blanchet 
parents: 
40145 
diff
changeset
 | 
66  | 
val fact_prefix : string  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
67  | 
val conjecture_prefix : string  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
68  | 
val helper_prefix : string  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
69  | 
val subclass_prefix : string  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
70  | 
val tcon_clause_prefix : string  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
71  | 
val tfree_clause_prefix : string  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
72  | 
val lam_fact_prefix : string  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
73  | 
val typed_helper_suffix : string  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
74  | 
val untyped_helper_suffix : string  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
75  | 
val predicator_name : string  | 
| 
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
76  | 
val app_op_name : string  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
77  | 
val type_guard_name : string  | 
| 
43104
 
81d1b15aa0ae
use ":" for type information (looks good in Metis's output) and handle it in new path finder
 
blanchet 
parents: 
43102 
diff
changeset
 | 
78  | 
val type_tag_name : string  | 
| 46435 | 79  | 
val native_type_prefix : string  | 
| 43174 | 80  | 
val prefixed_predicator_name : string  | 
| 
43130
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
81  | 
val prefixed_app_op_name : string  | 
| 
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
82  | 
val prefixed_type_tag_name : string  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
83  | 
val ascii_of : string -> string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
84  | 
val unascii_of : string -> string  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
85  | 
val unprefix_and_unascii : string -> string -> string option  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
86  | 
val proxy_table : (string * (string * (thm * (string * string)))) list  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
87  | 
val proxify_const : string -> (string * string) option  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
88  | 
val invert_const : string -> string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
89  | 
val unproxify_const : string -> string  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
90  | 
val new_skolem_var_name_of_const : string -> string  | 
| 48318 | 91  | 
val atp_logical_consts : string list  | 
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
92  | 
val atp_irrelevant_consts : string list  | 
| 48324 | 93  | 
val atp_widely_irrelevant_consts : string list  | 
| 
54089
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
94  | 
val is_irrelevant_const : string -> bool  | 
| 
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
95  | 
val is_widely_irrelevant_const : string -> bool  | 
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
96  | 
val atp_schematic_consts_of : term -> typ list Symtab.table  | 
| 
43828
 
e07a2c4cbad8
move lambda translation option from ATP to Sledgehammer, to avoid accidentally breaking Metis (its reconstruction code can only deal with combinators)
 
blanchet 
parents: 
43827 
diff
changeset
 | 
97  | 
val is_type_enc_higher_order : type_enc -> bool  | 
| 48131 | 98  | 
val is_type_enc_polymorphic : type_enc -> bool  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
99  | 
val level_of_type_enc : type_enc -> type_level  | 
| 48089 | 100  | 
val is_type_enc_sound : type_enc -> bool  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
101  | 
val type_enc_of_string : strictness -> string -> type_enc  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
102  | 
val adjust_type_enc : atp_format -> type_enc -> type_enc  | 
| 72921 | 103  | 
val is_first_order_lambda_free : term -> bool  | 
| 
58200
 
d95055489fce
Added translation for lambda expressions in terms.
 
steckerm 
parents: 
57547 
diff
changeset
 | 
104  | 
val do_cheaply_conceal_lambdas : typ list -> term -> term  | 
| 
43136
 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 
blanchet 
parents: 
43130 
diff
changeset
 | 
105  | 
val mk_aconns :  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
106  | 
    atp_connective -> ('a, 'b, 'c, 'd) atp_formula list
 | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
107  | 
    -> ('a, 'b, 'c, 'd) atp_formula
 | 
| 58477 | 108  | 
val unmangled_type : string -> (string, 'a) ATP_Problem.atp_term  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
109  | 
val unmangled_const : string -> string * (string, 'b) atp_term list  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
110  | 
val unmangled_const_name : string -> string list  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
111  | 
val helper_table : bool -> ((string * bool) * (status * thm) list) list  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
112  | 
val trans_lams_of_string :  | 
| 45514 | 113  | 
Proof.context -> type_enc -> string -> term list -> term list * term list  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
114  | 
val string_of_status : status -> string  | 
| 
43501
 
0e422a84d0b2
don't change the way helpers are generated for the exporter's sake
 
blanchet 
parents: 
43496 
diff
changeset
 | 
115  | 
val factsN : string  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
116  | 
val generate_atp_problem : Proof.context -> bool -> atp_format -> atp_formula_role -> type_enc ->  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
117  | 
mode -> string -> bool -> bool -> bool -> term list -> term ->  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
118  | 
((string * stature) * term) list -> string atp_problem * string Symtab.table  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
119  | 
* (string * term) list * int Symtab.table  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
120  | 
val atp_problem_term_order_info : string atp_problem -> (string * int) list  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
121  | 
end;  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
122  | 
|
| 46320 | 123  | 
structure ATP_Problem_Generate : ATP_PROBLEM_GENERATE =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
124  | 
struct  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
125  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
126  | 
open ATP_Util  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
127  | 
open ATP_Problem  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
128  | 
|
| 62718 | 129  | 
datatype mode = Metis | Sledgehammer | Sledgehammer_Completish of int | Exporter | Translator  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
130  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
131  | 
datatype scope = Global | Local | Assum | Chained  | 
| 55212 | 132  | 
datatype status = General | Induction | Intro | Inductive | Elim | Simp | Non_Rec_Def | Rec_Def  | 
133  | 
||
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
134  | 
type stature = scope * status  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
135  | 
|
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
136  | 
datatype order =  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
137  | 
First_Order |  | 
| 72657 | 138  | 
Higher_Order of thf_flavor  | 
| 48131 | 139  | 
datatype phantom_policy = Without_Phantom_Type_Vars | With_Phantom_Type_Vars  | 
140  | 
datatype polymorphism =  | 
|
141  | 
Type_Class_Polymorphic |  | 
|
142  | 
Raw_Polymorphic of phantom_policy |  | 
|
143  | 
Raw_Monomorphic |  | 
|
144  | 
Mangled_Monomorphic  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
145  | 
datatype strictness = Strict | Non_Strict  | 
| 48183 | 146  | 
datatype uniformity = Uniform | Non_Uniform  | 
| 52026 | 147  | 
datatype ctr_optim = With_Ctr_Optim | Without_Ctr_Optim  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
148  | 
datatype type_level =  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
149  | 
All_Types |  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
150  | 
Undercover_Types |  | 
| 48183 | 151  | 
Nonmono_Types of strictness * uniformity |  | 
| 52026 | 152  | 
Const_Types of ctr_optim |  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
153  | 
No_Types  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
154  | 
|
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
155  | 
datatype type_enc =  | 
| 74890 | 156  | 
Native of order * fool * polymorphism * type_level |  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
157  | 
Guards of polymorphism * type_level |  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
158  | 
Tags of polymorphism * type_level  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
159  | 
|
| 48654 | 160  | 
(* not clear whether ATPs prefer to have their negative variables tagged *)  | 
161  | 
val tag_neg_vars = false  | 
|
162  | 
||
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
163  | 
fun is_type_enc_native (Native _) = true  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
164  | 
| is_type_enc_native _ = false  | 
| 74890 | 165  | 
fun is_type_enc_full_higher_order (Native (Higher_Order THF_Lambda_Free, _, _, _)) = false  | 
166  | 
| is_type_enc_full_higher_order (Native (Higher_Order _, _, _, _)) = true  | 
|
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
167  | 
| is_type_enc_full_higher_order _ = false  | 
| 74890 | 168  | 
fun is_type_enc_fool (Native (_, With_FOOL _, _, _)) = true  | 
| 72588 | 169  | 
| is_type_enc_fool _ = false  | 
| 74890 | 170  | 
fun is_type_enc_higher_order (Native (Higher_Order _, _, _, _)) = true  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
171  | 
| is_type_enc_higher_order _ = false  | 
| 74890 | 172  | 
|
| 74896 | 173  | 
fun has_type_enc_choice (Native (Higher_Order THF_With_Choice, _, _, _)) = true  | 
174  | 
| has_type_enc_choice _ = false  | 
|
| 74890 | 175  | 
fun has_type_enc_ite (Native (_, With_FOOL {with_ite, ...}, _, _)) = with_ite
 | 
| 74758 | 176  | 
| has_type_enc_ite _ = false  | 
| 74890 | 177  | 
fun has_type_enc_let (Native (_, With_FOOL {with_let, ...}, _, _)) = with_let
 | 
178  | 
| has_type_enc_let _ = false  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
179  | 
|
| 74890 | 180  | 
fun polymorphism_of_type_enc (Native (_, _, poly, _)) = poly  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
181  | 
| polymorphism_of_type_enc (Guards (poly, _)) = poly  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
182  | 
| polymorphism_of_type_enc (Tags (poly, _)) = poly  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
183  | 
|
| 48131 | 184  | 
fun is_type_enc_polymorphic type_enc =  | 
| 54829 | 185  | 
(case polymorphism_of_type_enc type_enc of  | 
| 48131 | 186  | 
Raw_Polymorphic _ => true  | 
187  | 
| Type_Class_Polymorphic => true  | 
|
| 54829 | 188  | 
| _ => false)  | 
| 48131 | 189  | 
|
| 48201 | 190  | 
fun is_type_enc_mangling type_enc =  | 
191  | 
polymorphism_of_type_enc type_enc = Mangled_Monomorphic  | 
|
192  | 
||
| 74890 | 193  | 
fun level_of_type_enc (Native (_, _, _, level)) = level  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
194  | 
| level_of_type_enc (Guards (_, level)) = level  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
195  | 
| level_of_type_enc (Tags (_, level)) = level  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
196  | 
|
| 48183 | 197  | 
fun is_type_level_uniform (Nonmono_Types (_, Non_Uniform)) = false  | 
198  | 
| is_type_level_uniform Undercover_Types = false  | 
|
199  | 
| is_type_level_uniform _ = true  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
200  | 
|
| 48183 | 201  | 
fun is_type_level_sound (Const_Types _) = false  | 
202  | 
| is_type_level_sound No_Types = false  | 
|
203  | 
| is_type_level_sound _ = true  | 
|
| 48089 | 204  | 
val is_type_enc_sound = is_type_level_sound o level_of_type_enc  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
205  | 
|
| 48092 | 206  | 
fun is_type_level_monotonicity_based (Nonmono_Types _) = true  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
207  | 
| is_type_level_monotonicity_based _ = false  | 
| 
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
208  | 
|
| 45516 | 209  | 
val no_lamsN = "no_lams" (* used internally; undocumented *)  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
210  | 
val opaque_liftingN = "opaque_lifting"  | 
| 46365 | 211  | 
val liftingN = "lifting"  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
212  | 
val opaque_combsN = "opaque_combs"  | 
| 46365 | 213  | 
val combsN = "combs"  | 
214  | 
val combs_and_liftingN = "combs_and_lifting"  | 
|
215  | 
val combs_or_liftingN = "combs_or_lifting"  | 
|
| 
45513
 
25388cf06437
rename the lambda translation schemes, so that they are understandable out of context
 
blanchet 
parents: 
45511 
diff
changeset
 | 
216  | 
val keep_lamsN = "keep_lams"  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
217  | 
|
| 74328 | 218  | 
(* The capitalization of the TPTP output respects the capitalzation of the prefix. *)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
219  | 
val bound_var_prefix = "B_"  | 
| 74328 | 220  | 
val let_bound_var_prefix = "l_"  | 
| 47150 | 221  | 
val all_bound_var_prefix = "A_"  | 
222  | 
val exist_bound_var_prefix = "E_"  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
223  | 
val schematic_var_prefix = "V_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
224  | 
val fixed_var_prefix = "v_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
225  | 
val tvar_prefix = "T_"  | 
| 48219 | 226  | 
val tfree_prefix = "tf_"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
227  | 
val const_prefix = "c_"  | 
| 48219 | 228  | 
val type_const_prefix = "t_"  | 
| 46435 | 229  | 
val native_type_prefix = "n_"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
230  | 
val class_prefix = "cl_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
231  | 
|
| 45509 | 232  | 
(* Freshness almost guaranteed! *)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
233  | 
val atp_prefix = "ATP" ^ Long_Name.separator  | 
| 45509 | 234  | 
val atp_weak_prefix = "ATP:"  | 
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
235  | 
val atp_weak_suffix = ":ATP"  | 
| 45509 | 236  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
237  | 
val lam_lifted_prefix = atp_weak_prefix ^ "Lam"  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
238  | 
val lam_lifted_mono_prefix = lam_lifted_prefix ^ "m"  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
239  | 
val lam_lifted_poly_prefix = lam_lifted_prefix ^ "p"  | 
| 
43907
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43906 
diff
changeset
 | 
240  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
241  | 
val skolem_const_prefix = atp_prefix ^ "Sko"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
242  | 
val old_skolem_const_prefix = skolem_const_prefix ^ "o"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
243  | 
val new_skolem_const_prefix = skolem_const_prefix ^ "n"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
244  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
245  | 
val combinator_prefix = "COMB"  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
246  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
247  | 
val class_decl_prefix = "cl_"  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
248  | 
val type_decl_prefix = "ty_"  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
249  | 
val sym_decl_prefix = "sy_"  | 
| 52004 | 250  | 
val datatype_decl_prefix = "dt_"  | 
251  | 
val class_memb_prefix = "cm_"  | 
|
| 43989 | 252  | 
val guards_sym_formula_prefix = "gsy_"  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
253  | 
val tags_sym_formula_prefix = "tsy_"  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
254  | 
val uncurried_alias_eq_prefix = "unc_"  | 
| 
40204
 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 
blanchet 
parents: 
40145 
diff
changeset
 | 
255  | 
val fact_prefix = "fact_"  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
256  | 
val conjecture_prefix = "conj_"  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
257  | 
val helper_prefix = "help_"  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
258  | 
val subclass_prefix = "subcl_"  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
259  | 
val tcon_clause_prefix = "tcon_"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
260  | 
val tfree_clause_prefix = "tfree_"  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
261  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
262  | 
val lam_fact_prefix = "ATP.lambda_"  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
263  | 
val typed_helper_suffix = "_T"  | 
| 
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
264  | 
val untyped_helper_suffix = "_U"  | 
| 
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
265  | 
|
| 
44491
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
266  | 
val predicator_name = "pp"  | 
| 
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
267  | 
val app_op_name = "aa"  | 
| 
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
268  | 
val type_guard_name = "gg"  | 
| 
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
269  | 
val type_tag_name = "tt"  | 
| 
42531
 
a462dbaa584f
added more rudimentary type support to Sledgehammer's ATP encoding
 
blanchet 
parents: 
42530 
diff
changeset
 | 
270  | 
|
| 43174 | 271  | 
val prefixed_predicator_name = const_prefix ^ predicator_name  | 
| 
43130
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
272  | 
val prefixed_app_op_name = const_prefix ^ app_op_name  | 
| 
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
273  | 
val prefixed_type_tag_name = const_prefix ^ type_tag_name  | 
| 
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
274  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
275  | 
(*Escaping of special characters.  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
276  | 
Alphanumeric characters are left unchanged.  | 
| 54798 | 277  | 
The character _ goes to __.  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
278  | 
Characters in the range ASCII space to / go to _A to _P, respectively.  | 
| 54798 | 279  | 
Other characters go to _nnn where nnn is the decimal ASCII code. *)  | 
| 43093 | 280  | 
val upper_a_minus_space = Char.ord #"A" - Char.ord #" "  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
281  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
282  | 
fun ascii_of_char c =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
283  | 
if Char.isAlphaNum c then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
284  | 
String.str c  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
285  | 
else if c = #"_" then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
286  | 
"__"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
287  | 
else if #" " <= c andalso c <= #"/" then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
288  | 
"_" ^ String.str (Char.chr (Char.ord c + upper_a_minus_space))  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
289  | 
else  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
290  | 
(* fixed width, in case more digits follow *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
291  | 
"_" ^ stringN_of_int 3 (Char.ord c)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
292  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
293  | 
val ascii_of = String.translate ascii_of_char  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
294  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
295  | 
(** Remove ASCII armoring from names in proof files **)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
296  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
297  | 
(* We don't raise error exceptions because this code can run inside a worker  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
298  | 
thread. Also, the errors are impossible. *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
299  | 
val unascii_of =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
300  | 
let  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
301  | 
fun un rcs [] = String.implode (rev rcs)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
302  | 
| un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
303  | 
(* Three types of _ escapes: __, _A to _P, _nnn *)  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
304  | 
| un rcs (#"_" :: #"_" :: cs) = un (#"_" :: rcs) cs  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
305  | 
| un rcs (#"_" :: c :: cs) =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
306  | 
if #"A" <= c andalso c<= #"P" then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
307  | 
(* translation of #" " to #"/" *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
308  | 
un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
309  | 
else  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
310  | 
let val digits = List.take (c :: cs, 3) handle General.Subscript => [] in  | 
| 54829 | 311  | 
(case Int.fromString (String.implode digits) of  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
312  | 
SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))  | 
| 54829 | 313  | 
| NONE => un (c :: #"_" :: rcs) cs (* ERROR *))  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
314  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
315  | 
| un rcs (c :: cs) = un (c :: rcs) cs  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
316  | 
in un [] o String.explode end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
317  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
318  | 
(* If string s has the prefix s1, return the result of deleting it,  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
319  | 
un-ASCII'd. *)  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
320  | 
fun unprefix_and_unascii s1 s =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
321  | 
if String.isPrefix s1 s then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
322  | 
SOME (unascii_of (String.extract (s, size s1, NONE)))  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
323  | 
else  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
324  | 
NONE  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
325  | 
|
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
326  | 
val proxy_table =  | 
| 69593 | 327  | 
  [("c_False", (\<^const_name>\<open>False\<close>, (@{thm fFalse_def}, ("fFalse", \<^const_name>\<open>fFalse\<close>)))),
 | 
328  | 
   ("c_True", (\<^const_name>\<open>True\<close>, (@{thm fTrue_def}, ("fTrue", \<^const_name>\<open>fTrue\<close>)))),
 | 
|
329  | 
   ("c_Not", (\<^const_name>\<open>Not\<close>, (@{thm fNot_def}, ("fNot", \<^const_name>\<open>fNot\<close>)))),
 | 
|
330  | 
   ("c_conj", (\<^const_name>\<open>conj\<close>, (@{thm fconj_def}, ("fconj", \<^const_name>\<open>fconj\<close>)))),
 | 
|
331  | 
   ("c_disj", (\<^const_name>\<open>disj\<close>, (@{thm fdisj_def}, ("fdisj", \<^const_name>\<open>fdisj\<close>)))),
 | 
|
332  | 
   ("c_implies", (\<^const_name>\<open>implies\<close>, (@{thm fimplies_def}, ("fimplies", \<^const_name>\<open>fimplies\<close>)))),
 | 
|
333  | 
   ("equal", (\<^const_name>\<open>HOL.eq\<close>, (@{thm fequal_def}, ("fequal", \<^const_name>\<open>fequal\<close>)))),
 | 
|
334  | 
   ("c_All", (\<^const_name>\<open>All\<close>, (@{thm fAll_def}, ("fAll", \<^const_name>\<open>fAll\<close>)))),
 | 
|
| 74896 | 335  | 
   ("c_Ex", (\<^const_name>\<open>Ex\<close>, (@{thm fEx_def}, ("fEx", \<^const_name>\<open>fEx\<close>)))),
 | 
336  | 
   ("c_Choice", (\<^const_name>\<open>Hilbert_Choice.Eps\<close>, (@{thm fChoice_def}, ("fChoice", \<^const_name>\<open>fChoice\<close>))))]
 | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
337  | 
|
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
338  | 
val proxify_const = AList.lookup (op =) proxy_table #> Option.map (snd o snd)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
339  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
340  | 
(* Readable names for the more common symbolic functions. Do not mess with the  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
341  | 
table unless you know what you are doing. *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
342  | 
val const_trans_table =  | 
| 69593 | 343  | 
[(\<^const_name>\<open>False\<close>, "False"),  | 
344  | 
(\<^const_name>\<open>True\<close>, "True"),  | 
|
345  | 
(\<^const_name>\<open>Not\<close>, "Not"),  | 
|
346  | 
(\<^const_name>\<open>conj\<close>, "conj"),  | 
|
347  | 
(\<^const_name>\<open>disj\<close>, "disj"),  | 
|
348  | 
(\<^const_name>\<open>implies\<close>, "implies"),  | 
|
349  | 
(\<^const_name>\<open>HOL.eq\<close>, "equal"),  | 
|
350  | 
(\<^const_name>\<open>All\<close>, "All"),  | 
|
351  | 
(\<^const_name>\<open>Ex\<close>, "Ex"),  | 
|
352  | 
(\<^const_name>\<open>If\<close>, "If"),  | 
|
353  | 
(\<^const_name>\<open>Set.member\<close>, "member"),  | 
|
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
354  | 
(\<^const_name>\<open>HOL.Let\<close>, "Let"),  | 
| 74896 | 355  | 
(\<^const_name>\<open>Hilbert_Choice.Eps\<close>, "Choice"),  | 
| 69593 | 356  | 
(\<^const_name>\<open>Meson.COMBI\<close>, combinator_prefix ^ "I"),  | 
357  | 
(\<^const_name>\<open>Meson.COMBK\<close>, combinator_prefix ^ "K"),  | 
|
358  | 
(\<^const_name>\<open>Meson.COMBB\<close>, combinator_prefix ^ "B"),  | 
|
359  | 
(\<^const_name>\<open>Meson.COMBC\<close>, combinator_prefix ^ "C"),  | 
|
360  | 
(\<^const_name>\<open>Meson.COMBS\<close>, combinator_prefix ^ "S")]  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
361  | 
|> Symtab.make  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
362  | 
|> fold (Symtab.update o swap o snd o snd o snd) proxy_table  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
363  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
364  | 
(* Invert the table of translations between Isabelle and ATPs. *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
365  | 
val const_trans_table_inv =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
366  | 
const_trans_table |> Symtab.dest |> map swap |> Symtab.make  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
367  | 
val const_trans_table_unprox =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
368  | 
Symtab.empty  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
369  | 
|> fold (fn (_, (isa, (_, (_, atp)))) => Symtab.update (atp, isa)) proxy_table  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
370  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
371  | 
val invert_const = perhaps (Symtab.lookup const_trans_table_inv)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
372  | 
val unproxify_const = perhaps (Symtab.lookup const_trans_table_unprox)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
373  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
374  | 
fun lookup_const c =  | 
| 54829 | 375  | 
(case Symtab.lookup const_trans_table c of  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
376  | 
SOME c' => c'  | 
| 54829 | 377  | 
| NONE => ascii_of c)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
378  | 
|
| 43622 | 379  | 
fun ascii_of_indexname (v, 0) = ascii_of v  | 
380  | 
| ascii_of_indexname (v, i) = ascii_of v ^ "_" ^ string_of_int i  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
381  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
382  | 
fun make_bound_var x = bound_var_prefix ^ ascii_of x  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
383  | 
fun make_all_bound_var x = all_bound_var_prefix ^ ascii_of x  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
384  | 
fun make_exist_bound_var x = exist_bound_var_prefix ^ ascii_of x  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
385  | 
fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
386  | 
fun make_fixed_var x = fixed_var_prefix ^ ascii_of x  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
387  | 
|
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
388  | 
fun make_tvar (s, i) = tvar_prefix ^ ascii_of_indexname (unquote_tvar s, i)  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
389  | 
fun make_tfree s = tfree_prefix ^ ascii_of (unquote_tvar s)  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
390  | 
fun tvar_name ((x as (s, _)), _) = (make_tvar x, s)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
391  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
392  | 
(* "HOL.eq" and choice are mapped to the ATP's equivalents *)  | 
| 75010 | 393  | 
fun make_fixed_const \<^const_name>\<open>HOL.eq\<close> = tptp_old_equal  | 
394  | 
| make_fixed_const c = const_prefix ^ lookup_const c  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
395  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
396  | 
fun make_fixed_type_const c = type_const_prefix ^ lookup_const c  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
397  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
398  | 
fun make_class clas = class_prefix ^ ascii_of clas  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
399  | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
400  | 
fun new_skolem_var_name_of_const s =  | 
| 59877 | 401  | 
let val ss = Long_Name.explode s  | 
402  | 
in nth ss (length ss - 2) end  | 
|
| 43093 | 403  | 
|
| 48318 | 404  | 
(* These are ignored anyway by the relevance filter (unless they appear in  | 
405  | 
higher-order places) but not by the monomorphizer. *)  | 
|
406  | 
val atp_logical_consts =  | 
|
| 69593 | 407  | 
[\<^const_name>\<open>Pure.prop\<close>, \<^const_name>\<open>Pure.conjunction\<close>,  | 
408  | 
\<^const_name>\<open>Pure.all\<close>, \<^const_name>\<open>Pure.imp\<close>, \<^const_name>\<open>Pure.eq\<close>,  | 
|
409  | 
\<^const_name>\<open>Trueprop\<close>, \<^const_name>\<open>All\<close>, \<^const_name>\<open>Ex\<close>,  | 
|
410  | 
\<^const_name>\<open>Ex1\<close>, \<^const_name>\<open>Ball\<close>, \<^const_name>\<open>Bex\<close>]  | 
|
| 48318 | 411  | 
|
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
412  | 
(* These are either simplified away by "Meson.presimplify" (most of the time) or  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
413  | 
handled specially via "fFalse", "fTrue", ..., "fequal". *)  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
414  | 
val atp_irrelevant_consts =  | 
| 69593 | 415  | 
[\<^const_name>\<open>False\<close>, \<^const_name>\<open>True\<close>, \<^const_name>\<open>Not\<close>, \<^const_name>\<open>conj\<close>,  | 
416  | 
\<^const_name>\<open>disj\<close>, \<^const_name>\<open>implies\<close>, \<^const_name>\<open>HOL.eq\<close>, \<^const_name>\<open>If\<close>,  | 
|
417  | 
\<^const_name>\<open>Let\<close>]  | 
|
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
418  | 
|
| 48318 | 419  | 
val atp_widely_irrelevant_consts = atp_logical_consts @ atp_irrelevant_consts  | 
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
420  | 
|
| 
54089
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
421  | 
val atp_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_irrelevant_consts)  | 
| 
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
422  | 
val atp_widely_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_widely_irrelevant_consts)  | 
| 
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
423  | 
|
| 
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
424  | 
val is_irrelevant_const = Symtab.defined atp_irrelevant_const_tab  | 
| 
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
425  | 
val is_widely_irrelevant_const = Symtab.defined atp_widely_irrelevant_const_tab  | 
| 
 
b13f6731f873
use same relevance filter for ATP and SMT solvers -- attempting to filter out certain ground instances of polymorphic symbols like + and 0 has unexpected side-effects that lead to incompletenesses (relevant facts not being selected)
 
blanchet 
parents: 
53586 
diff
changeset
 | 
426  | 
|
| 
43258
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
427  | 
fun add_schematic_const (x as (_, T)) =  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
428  | 
Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
429  | 
val add_schematic_consts_of =  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
430  | 
Term.fold_aterms (fn Const (x as (s, _)) =>  | 
| 74949 | 431  | 
not (is_widely_irrelevant_const s) ? add_schematic_const x  | 
| 
43258
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
432  | 
| _ => I)  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
433  | 
fun atp_schematic_consts_of t = add_schematic_consts_of t Symtab.empty  | 
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
434  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
435  | 
val tvar_a_str = "'a"  | 
| 69593 | 436  | 
val tvar_a_z = ((tvar_a_str, 0), \<^sort>\<open>type\<close>)  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
437  | 
val tvar_a = TVar tvar_a_z  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
438  | 
val tvar_a_name = tvar_name tvar_a_z  | 
| 69593 | 439  | 
val itself_name = `make_fixed_type_const \<^type_name>\<open>itself\<close>  | 
| 75010 | 440  | 
val TYPE_name = `make_fixed_const \<^const_name>\<open>Pure.type\<close>  | 
| 54820 | 441  | 
val tvar_a_atype = AType ((tvar_a_name, []), [])  | 
442  | 
val a_itself_atype = AType ((itself_name, []), [tvar_a_atype])  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
443  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
444  | 
(** Definitions and functions for FOL clauses and formulas for TPTP **)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
445  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
446  | 
(** Type class membership **)  | 
| 43263 | 447  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
448  | 
(* In our data structures, [] exceptionally refers to the top class, not to  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
449  | 
the empty class. *)  | 
| 44625 | 450  | 
|
| 69593 | 451  | 
val class_of_types = the_single \<^sort>\<open>type\<close>  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
452  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
453  | 
fun normalize_classes cls = if member (op =) cls class_of_types then [] else cls  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
454  | 
|
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
455  | 
(* Arity of type constructor "s :: (arg1, ..., argN) res" *)  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
456  | 
fun make_axiom_tcon_clause (s, name, (cl, args)) =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
457  | 
let  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
458  | 
val args = args |> map normalize_classes  | 
| 74918 | 459  | 
val tvars = args |> map_index (fn (j, _) => TVar ((tvar_a_str, j + 1), \<^sort>\<open>type\<close>))  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
460  | 
in (name, args ~~ tvars, (cl, Type (s, tvars))) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
461  | 
|
| 43622 | 462  | 
(* Generate all pairs (tycon, class, sorts) such that tycon belongs to class in  | 
463  | 
theory thy provided its arguments have the corresponding sorts. *)  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
464  | 
fun class_pairs thy tycons cls =  | 
| 43093 | 465  | 
let  | 
466  | 
val alg = Sign.classes_of thy  | 
|
467  | 
fun domain_sorts tycon = Sorts.mg_domain alg tycon o single  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
468  | 
fun add_class tycon cl =  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
469  | 
cons (cl, domain_sorts tycon cl)  | 
| 43093 | 470  | 
handle Sorts.CLASS_ERROR _ => I  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
471  | 
fun try_classes tycon = (tycon, fold (add_class tycon) cls [])  | 
| 43093 | 472  | 
in map try_classes tycons end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
473  | 
|
| 48141 | 474  | 
(* Proving one (tycon, class) membership may require proving others, so  | 
475  | 
iterate. *)  | 
|
| 
57539
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
476  | 
fun all_class_pairs _ _ _ [] = ([], [])  | 
| 
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
477  | 
| all_class_pairs thy tycons old_cls cls =  | 
| 48141 | 478  | 
let  | 
| 
57539
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
479  | 
val old_cls' = cls @ old_cls  | 
| 
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
480  | 
fun maybe_insert_class s = not (member (op =) old_cls' s) ? insert (op =) s  | 
| 
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
481  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
482  | 
val pairs = class_pairs thy tycons cls  | 
| 
57539
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
483  | 
val new_cls = fold (fold (fold (fold maybe_insert_class) o snd) o snd) pairs []  | 
| 
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
484  | 
val (cls', pairs') = all_class_pairs thy tycons old_cls' new_cls  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
485  | 
in (cls' @ cls, union (op =) pairs' pairs) end  | 
| 48141 | 486  | 
|
| 51651 | 487  | 
fun tcon_clause _ _ [] = []  | 
488  | 
| tcon_clause seen n ((_, []) :: rest) = tcon_clause seen n rest  | 
|
489  | 
| tcon_clause seen n ((tcons, (ar as (cl, _)) :: ars) :: rest) =  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
490  | 
if cl = class_of_types then  | 
| 51651 | 491  | 
tcon_clause seen n ((tcons, ars) :: rest)  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
492  | 
else if member (op =) seen cl then  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
493  | 
(* multiple clauses for the same (tycon, cl) pair *)  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
494  | 
make_axiom_tcon_clause (tcons,  | 
| 54820 | 495  | 
lookup_const tcons ^ "___" ^ ascii_of cl ^ "_" ^ string_of_int n, ar) ::  | 
| 51651 | 496  | 
tcon_clause seen (n + 1) ((tcons, ars) :: rest)  | 
| 48141 | 497  | 
else  | 
| 54820 | 498  | 
make_axiom_tcon_clause (tcons, lookup_const tcons ^ "___" ^ ascii_of cl, ar) ::  | 
| 51651 | 499  | 
tcon_clause (cl :: seen) n ((tcons, ars) :: rest)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
500  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
501  | 
fun make_tcon_clauses thy tycons =  | 
| 
57539
 
353fd3326835
avoid loop in 'all_class_pairs' (caused by e.g. loading the 'Ceta' theory and calling Sledgehammer with the two facts 'fun_of_map.cases' and 'Lattices.bounded_lattice_top_class.sup_top_left' with a polymorphic type encoding)
 
blanchet 
parents: 
57533 
diff
changeset
 | 
502  | 
all_class_pairs thy tycons [] ##> tcon_clause [] 1  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
503  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
504  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
505  | 
(** Isabelle class relations **)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
506  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
507  | 
(* Generate a list ("sub", "supers") such that "sub" is a proper subclass of all
 | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
508  | 
"supers". *)  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
509  | 
fun make_subclass_pairs thy subs supers =  | 
| 48141 | 510  | 
let  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
511  | 
val class_less = curry (Sorts.class_less (Sign.classes_of thy))  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
512  | 
fun supers_of sub = (sub, filter (class_less sub) supers)  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
513  | 
in map supers_of subs |> filter_out (null o snd) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
514  | 
|
| 43859 | 515  | 
(* intermediate terms *)  | 
| 74328 | 516  | 
(* TODO: Merge IConst and IVar *)  | 
| 43859 | 517  | 
datatype iterm =  | 
| 48135 | 518  | 
IConst of (string * string) * typ * typ list |  | 
519  | 
IVar of (string * string) * typ |  | 
|
| 43859 | 520  | 
IApp of iterm * iterm |  | 
| 48135 | 521  | 
IAbs of ((string * string) * typ) * iterm  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
522  | 
|
| 74328 | 523  | 
fun alpha_rename from to =  | 
524  | 
let  | 
|
525  | 
fun traverse (tm as IConst (name, T, Ts)) =  | 
|
526  | 
if name = from then IConst (to, T, Ts) else tm  | 
|
527  | 
| traverse (tm as IVar (name, T)) =  | 
|
528  | 
if name = from then IVar (to, T) else tm  | 
|
529  | 
| traverse (tm as IApp (tm1, tm2)) =  | 
|
530  | 
let  | 
|
531  | 
val tm1' = traverse tm1  | 
|
532  | 
val tm2' = traverse tm2  | 
|
533  | 
in  | 
|
| 74988 | 534  | 
if pointer_eq (tm1, tm1') andalso pointer_eq (tm2, tm2') then tm else IApp (tm1', tm2')  | 
| 74328 | 535  | 
end  | 
536  | 
| traverse (tm as IAbs (binding as (name, _), tm1)) =  | 
|
537  | 
if name = from then  | 
|
538  | 
tm  | 
|
539  | 
else  | 
|
540  | 
let val tm1' = traverse tm1 in  | 
|
| 74988 | 541  | 
if pointer_eq (tm1, tm1') then tm else IAbs (binding, tm1')  | 
| 74328 | 542  | 
end  | 
543  | 
in  | 
|
544  | 
traverse  | 
|
545  | 
end  | 
|
546  | 
||
| 43859 | 547  | 
fun ityp_of (IConst (_, T, _)) = T  | 
548  | 
| ityp_of (IVar (_, T)) = T  | 
|
549  | 
| ityp_of (IApp (t1, _)) = snd (dest_funT (ityp_of t1))  | 
|
550  | 
| ityp_of (IAbs ((_, T), tm)) = T --> ityp_of tm  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
551  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
552  | 
(*gets the head of a combinator application, along with the list of arguments*)  | 
| 43859 | 553  | 
fun strip_iterm_comb u =  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
554  | 
let  | 
| 43859 | 555  | 
fun stripc (IApp (t, u), ts) = stripc (t, u :: ts)  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
556  | 
| stripc x = x  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
557  | 
in stripc (u, []) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
558  | 
|
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
559  | 
fun atomic_types_of T = fold_atyps (insert (op =)) T []  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
560  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
561  | 
fun new_skolem_const_name s num_T_args =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
562  | 
[new_skolem_const_prefix, s, string_of_int num_T_args]  | 
| 
46711
 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 
wenzelm 
parents: 
46643 
diff
changeset
 | 
563  | 
|> Long_Name.implode  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
564  | 
|
| 69593 | 565  | 
val alpha_to_beta = Logic.varifyT_global \<^typ>\<open>'a => 'b\<close>  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
566  | 
val alpha_to_beta_to_alpha_to_beta = alpha_to_beta --> alpha_to_beta  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
567  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
568  | 
fun robust_const_type thy s =  | 
| 45509 | 569  | 
if s = app_op_name then  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
570  | 
alpha_to_beta_to_alpha_to_beta  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
571  | 
else if String.isPrefix lam_lifted_prefix s then  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
572  | 
alpha_to_beta  | 
| 45509 | 573  | 
else  | 
574  | 
(* Old Skolems throw a "TYPE" exception here, which will be caught. *)  | 
|
575  | 
s |> Sign.the_const_type thy  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
576  | 
|
| 69593 | 577  | 
fun ary_of (Type (\<^type_name>\<open>fun\<close>, [_, T])) = 1 + ary_of T  | 
| 
52038
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
578  | 
| ary_of _ = 0  | 
| 
46642
 
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
 
blanchet 
parents: 
46639 
diff
changeset
 | 
579  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
580  | 
(* This function only makes sense if "T" is as general as possible. *)  | 
| 
51920
 
16f3b9d4e515
proper unmangling -- the bug is visible when "uncurried_aliases" is enabled with Alt-Ergo or Poly. SPASS
 
blanchet 
parents: 
51717 
diff
changeset
 | 
581  | 
fun robust_const_type_args thy (s, T) =  | 
| 45509 | 582  | 
if s = app_op_name then  | 
583  | 
let val (T1, T2) = T |> domain_type |> dest_funT in [T1, T2] end  | 
|
584  | 
else if String.isPrefix old_skolem_const_prefix s then  | 
|
585  | 
[] |> Term.add_tvarsT T |> rev |> map TVar  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
586  | 
else if String.isPrefix lam_lifted_prefix s then  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
587  | 
if String.isPrefix lam_lifted_poly_prefix s then  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
588  | 
let val (T1, T2) = T |> dest_funT in [T1, T2] end  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
589  | 
else  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
590  | 
[]  | 
| 45509 | 591  | 
else  | 
592  | 
(s, T) |> Sign.const_typargs thy  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
593  | 
|
| 54798 | 594  | 
(* Converts an Isabelle/HOL term (with combinators) into an intermediate term. Also accumulates sort  | 
595  | 
infomation. *)  | 
|
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
596  | 
fun iterm_of_term thy type_enc =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
597  | 
let  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
598  | 
fun iot true bs ((t0 as Const (\<^const_name>\<open>Let\<close>, _)) $ t1 $ (t2 as Abs (s, T, t'))) =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
599  | 
let  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
600  | 
val (t0', t0_atomics_Ts) = iot true bs t0  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
601  | 
val (t1', t1_atomics_Ts) = iot true bs t1  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
602  | 
val (t2', t2_atomics_Ts) = iot true bs t2  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
603  | 
in  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
604  | 
(IApp (IApp (t0', t1'), t2'),  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
605  | 
fold (union (op =)) [t0_atomics_Ts, t1_atomics_Ts, t2_atomics_Ts] [])  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
606  | 
end  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
607  | 
| iot true bs ((t0 as Const (\<^const_name>\<open>Let\<close>, _)) $ t1 $ t2) =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
608  | 
iot true bs (t0 $ t1 $ eta_expand (map (snd o snd) bs) t2 1)  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
609  | 
| iot fool bs (P $ Q) =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
610  | 
let  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
611  | 
val (P', P_atomics_Ts) = iot fool bs P  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
612  | 
val (Q', Q_atomics_Ts) = iot fool bs Q  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
613  | 
in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
614  | 
| iot _ _ (Const (c, T)) =  | 
| 75010 | 615  | 
(IConst (`make_fixed_const c, T, robust_const_type_args thy (c, T)),  | 
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
616  | 
atomic_types_of T)  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
617  | 
| iot _ _ (Free (s, T)) = (IConst (`make_fixed_var s, T, []), atomic_types_of T)  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
618  | 
| iot _ _ (Var (v as (s, _), T)) =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
619  | 
(if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
620  | 
let  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
621  | 
val Ts = T |> strip_type |> swap |> op ::  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
622  | 
val s' = new_skolem_const_name s (length Ts)  | 
| 75010 | 623  | 
in IConst (`make_fixed_const s', T, Ts) end  | 
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
624  | 
else  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
625  | 
IVar ((make_schematic_var v, s), T), atomic_types_of T)  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
626  | 
| iot _ bs (Bound j) =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
627  | 
nth bs j |> (fn (_, (name, T)) => (IConst (name, T, []), atomic_types_of T))  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
628  | 
| iot fool bs (Abs (s, T, t)) =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
629  | 
let  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
630  | 
fun vary s = s |> AList.defined (op =) bs s ? vary o Symbol.bump_string  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
631  | 
val s = vary s  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
632  | 
val name = `make_bound_var s  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
633  | 
val (tm, atomic_Ts) = iot fool ((s, (name, T)) :: bs) t  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
634  | 
in  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
635  | 
(IAbs ((name, T), tm), union (op =) atomic_Ts (atomic_types_of T))  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
636  | 
end  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
637  | 
in iot (is_type_enc_fool type_enc) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
638  | 
|
| 48089 | 639  | 
(* "_query" and "_at" are for the ASCII-challenged Metis and Mirabelle. *)  | 
| 44785 | 640  | 
val queries = ["?", "_query"]  | 
641  | 
val ats = ["@", "_at"]  | 
|
642  | 
||
| 
42689
 
e38590764c34
versions of ! and ? for the ASCII-challenged Mirabelle
 
blanchet 
parents: 
42688 
diff
changeset
 | 
643  | 
fun try_unsuffixes ss s =  | 
| 
 
e38590764c34
versions of ! and ? for the ASCII-challenged Mirabelle
 
blanchet 
parents: 
42688 
diff
changeset
 | 
644  | 
fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE  | 
| 
 
e38590764c34
versions of ! and ? for the ASCII-challenged Mirabelle
 
blanchet 
parents: 
42688 
diff
changeset
 | 
645  | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
646  | 
fun type_enc_of_string strictness s =  | 
| 72588 | 647  | 
let  | 
648  | 
val (poly, s) =  | 
|
649  | 
(case try (unprefix "tc_") s of  | 
|
650  | 
SOME s => (SOME Type_Class_Polymorphic, s)  | 
|
| 54829 | 651  | 
| NONE =>  | 
| 72588 | 652  | 
(case try (unprefix "poly_") s of  | 
653  | 
SOME s => (SOME (Raw_Polymorphic With_Phantom_Type_Vars), s)  | 
|
| 
57293
 
4e619ee65a61
added 'dummy_thf_ml' prover for experiments with HOLyHammer
 
blanchet 
parents: 
57292 
diff
changeset
 | 
654  | 
| NONE =>  | 
| 72588 | 655  | 
(case try (unprefix "ml_poly_") s of  | 
656  | 
SOME s => (SOME (Raw_Polymorphic Without_Phantom_Type_Vars), s)  | 
|
657  | 
| NONE =>  | 
|
658  | 
(case try (unprefix "raw_mono_") s of  | 
|
659  | 
SOME s => (SOME Raw_Monomorphic, s)  | 
|
660  | 
| NONE =>  | 
|
661  | 
(case try (unprefix "mono_") s of  | 
|
662  | 
SOME s => (SOME Mangled_Monomorphic, s)  | 
|
663  | 
| NONE => (NONE, s))))))  | 
|
664  | 
||
665  | 
val (level, s) =  | 
|
666  | 
case try_unsuffixes queries s of  | 
|
667  | 
SOME s =>  | 
|
668  | 
(case try_unsuffixes queries s of  | 
|
669  | 
SOME s => (Nonmono_Types (strictness, Non_Uniform), s)  | 
|
670  | 
| NONE => (Nonmono_Types (strictness, Uniform), s))  | 
|
671  | 
| NONE =>  | 
|
672  | 
(case try_unsuffixes ats s of  | 
|
673  | 
SOME s => (Undercover_Types, s)  | 
|
674  | 
| NONE => (All_Types, s))  | 
|
675  | 
||
676  | 
fun native_of_string s =  | 
|
677  | 
let  | 
|
| 
74891
 
db4b8dd587a5
added support for higher-order SMT proof search in Sledgehammer
 
desharna 
parents: 
74890 
diff
changeset
 | 
678  | 
val (_, s) =  | 
| 
 
db4b8dd587a5
added support for higher-order SMT proof search in Sledgehammer
 
desharna 
parents: 
74890 
diff
changeset
 | 
679  | 
(case try (unsuffix "_arith") s of  | 
| 
 
db4b8dd587a5
added support for higher-order SMT proof search in Sledgehammer
 
desharna 
parents: 
74890 
diff
changeset
 | 
680  | 
SOME s => (true, s)  | 
| 
 
db4b8dd587a5
added support for higher-order SMT proof search in Sledgehammer
 
desharna 
parents: 
74890 
diff
changeset
 | 
681  | 
| NONE => (false, s))  | 
| 
74894
 
a64a8f7d593e
refactored $ite and $let configuration and added dummy_thf_reduced prover
 
desharna 
parents: 
74891 
diff
changeset
 | 
682  | 
        val syntax = {with_ite = true, with_let = true}
 | 
| 72588 | 683  | 
val (fool, core) =  | 
684  | 
(case try (unsuffix "_fool") s of  | 
|
| 74890 | 685  | 
SOME s => (With_FOOL syntax, s)  | 
| 72588 | 686  | 
| NONE => (Without_FOOL, s))  | 
687  | 
in  | 
|
688  | 
(case (core, poly) of  | 
|
689  | 
          ("native", SOME poly) =>
 | 
|
| 54829 | 690  | 
(case (poly, level) of  | 
691  | 
(Mangled_Monomorphic, _) =>  | 
|
| 72588 | 692  | 
if is_type_level_uniform level then  | 
| 74890 | 693  | 
Native (First_Order, fool, Mangled_Monomorphic, level)  | 
| 72588 | 694  | 
else  | 
695  | 
raise Same.SAME  | 
|
| 54829 | 696  | 
| (Raw_Monomorphic, _) => raise Same.SAME  | 
| 74890 | 697  | 
| (poly, All_Types) => Native (First_Order, fool, poly, All_Types))  | 
| 72588 | 698  | 
        | ("native_higher", SOME poly) =>
 | 
| 54829 | 699  | 
(case (poly, level) of  | 
700  | 
(_, Nonmono_Types _) => raise Same.SAME  | 
|
701  | 
| (_, Undercover_Types) => raise Same.SAME  | 
|
702  | 
| (Mangled_Monomorphic, _) =>  | 
|
703  | 
if is_type_level_uniform level then  | 
|
| 74890 | 704  | 
Native (Higher_Order THF_With_Choice, With_FOOL syntax, Mangled_Monomorphic, level)  | 
| 54829 | 705  | 
else  | 
706  | 
raise Same.SAME  | 
|
707  | 
| (poly as Raw_Polymorphic _, All_Types) =>  | 
|
| 74890 | 708  | 
Native (Higher_Order THF_With_Choice, With_FOOL syntax, poly, All_Types)  | 
| 72588 | 709  | 
| _ => raise Same.SAME))  | 
710  | 
end  | 
|
711  | 
||
712  | 
fun nonnative_of_string core =  | 
|
713  | 
(case (core, poly, level) of  | 
|
714  | 
        ("guards", SOME poly, _) =>
 | 
|
715  | 
if (poly = Mangled_Monomorphic andalso level = Undercover_Types) orelse  | 
|
716  | 
poly = Type_Class_Polymorphic then  | 
|
717  | 
raise Same.SAME  | 
|
718  | 
else  | 
|
719  | 
Guards (poly, level)  | 
|
720  | 
      | ("tags", SOME poly, _) =>
 | 
|
721  | 
if (poly = Mangled_Monomorphic andalso level = Undercover_Types) orelse  | 
|
722  | 
poly = Type_Class_Polymorphic then  | 
|
723  | 
raise Same.SAME  | 
|
724  | 
else  | 
|
725  | 
Tags (poly, level)  | 
|
726  | 
      | ("args", SOME poly, All_Types (* naja *)) =>
 | 
|
727  | 
if poly = Type_Class_Polymorphic then raise Same.SAME  | 
|
728  | 
else Guards (poly, Const_Types Without_Ctr_Optim)  | 
|
729  | 
      | ("args", SOME poly, Nonmono_Types (_, Uniform) (* naja *)) =>
 | 
|
730  | 
if poly = Mangled_Monomorphic orelse poly = Type_Class_Polymorphic then  | 
|
731  | 
raise Same.SAME  | 
|
732  | 
else  | 
|
733  | 
Guards (poly, Const_Types With_Ctr_Optim)  | 
|
734  | 
      | ("erased", NONE, All_Types (* naja *)) =>
 | 
|
735  | 
Guards (Raw_Polymorphic With_Phantom_Type_Vars, No_Types)  | 
|
736  | 
| _ => raise Same.SAME)  | 
|
737  | 
in  | 
|
738  | 
if String.isPrefix "native" s then  | 
|
739  | 
native_of_string s  | 
|
740  | 
else  | 
|
741  | 
nonnative_of_string s  | 
|
742  | 
end  | 
|
| 63692 | 743  | 
  handle Same.SAME => error ("Unknown type encoding: " ^ quote s)
 | 
| 
42613
 
23b13b1bd565
use strings to encode type systems in ATP module, to reduce the amount of out-of-place information and also to make it easier to print the type system used
 
blanchet 
parents: 
42612 
diff
changeset
 | 
744  | 
|
| 72916 | 745  | 
fun min_hologic THF_Lambda_Free _ = THF_Lambda_Free  | 
746  | 
| min_hologic _ THF_Lambda_Free = THF_Lambda_Free  | 
|
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
747  | 
| min_hologic THF_Without_Choice _ = THF_Without_Choice  | 
| 
68250
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
748  | 
| min_hologic _ THF_Without_Choice = THF_Without_Choice  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
749  | 
| min_hologic _ _ = THF_With_Choice  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
750  | 
|
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
751  | 
fun adjust_hologic hologic (Higher_Order hologic') = Higher_Order (min_hologic hologic hologic')  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
752  | 
| adjust_hologic _ type_enc = type_enc  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
753  | 
|
| 74890 | 754  | 
fun adjust_syntax {with_ite = ite1, with_let = let1} {with_ite = ite2, with_let = let2} =
 | 
755  | 
  {with_ite = ite1 andalso ite2, with_let = let1 andalso let2}
 | 
|
756  | 
||
757  | 
fun adjust_fool (With_FOOL syntax) (With_FOOL syntax') = With_FOOL (adjust_syntax syntax syntax')  | 
|
758  | 
| adjust_fool _ _ = Without_FOOL  | 
|
| 72588 | 759  | 
|
| 57264 | 760  | 
fun no_type_classes Type_Class_Polymorphic = Raw_Polymorphic With_Phantom_Type_Vars  | 
| 
48134
 
fa23e699494c
robustness -- TFF1 does not support type classes
 
blanchet 
parents: 
48133 
diff
changeset
 | 
761  | 
| no_type_classes poly = poly  | 
| 
 
fa23e699494c
robustness -- TFF1 does not support type classes
 
blanchet 
parents: 
48133 
diff
changeset
 | 
762  | 
|
| 74890 | 763  | 
fun adjust_type_enc (THF (poly, syntax, hologic)) (Native (order, fool, poly', level)) =  | 
764  | 
Native (adjust_hologic hologic order, adjust_fool (With_FOOL syntax) fool,  | 
|
765  | 
(case poly of Polymorphic => no_type_classes poly' | Monomorphic => Mangled_Monomorphic),  | 
|
| 74758 | 766  | 
level)  | 
| 74890 | 767  | 
| adjust_type_enc (TFF (poly, fool)) (Native (_, fool', poly', level)) =  | 
768  | 
Native (First_Order, adjust_fool fool fool',  | 
|
769  | 
(case poly of Polymorphic => no_type_classes poly' | Monomorphic => Mangled_Monomorphic),  | 
|
770  | 
level)  | 
|
771  | 
| adjust_type_enc (DFG Polymorphic) (Native (_, _, poly, level)) =  | 
|
772  | 
Native (First_Order, Without_FOOL, poly, level)  | 
|
773  | 
| adjust_type_enc (DFG Monomorphic) (Native (_, _, _, level)) =  | 
|
774  | 
Native (First_Order, Without_FOOL, Mangled_Monomorphic, level)  | 
|
775  | 
| adjust_type_enc format (Native (_, _, poly, level)) =  | 
|
| 
48134
 
fa23e699494c
robustness -- TFF1 does not support type classes
 
blanchet 
parents: 
48133 
diff
changeset
 | 
776  | 
adjust_type_enc format (Guards (no_type_classes poly, level))  | 
| 
44416
 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 
blanchet 
parents: 
44410 
diff
changeset
 | 
777  | 
| adjust_type_enc CNF_UEQ (type_enc as Guards stuff) =  | 
| 48089 | 778  | 
(if is_type_enc_sound type_enc then Tags else Guards) stuff  | 
| 
44416
 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 
blanchet 
parents: 
44410 
diff
changeset
 | 
779  | 
| adjust_type_enc _ type_enc = type_enc  | 
| 
43101
 
1d46d85cd78b
make "prepare_atp_problem" more robust w.r.t. choice of type system
 
blanchet 
parents: 
43098 
diff
changeset
 | 
780  | 
|
| 72921 | 781  | 
fun is_first_order_lambda_free t =  | 
| 54829 | 782  | 
(case t of  | 
| 74379 | 783  | 
\<^Const_>\<open>Not for t1\<close> => is_first_order_lambda_free t1  | 
784  | 
| \<^Const_>\<open>All _ for \<open>Abs (_, _, t')\<close>\<close> => is_first_order_lambda_free t'  | 
|
785  | 
| \<^Const_>\<open>All _ for t1\<close> => is_first_order_lambda_free t1  | 
|
786  | 
| \<^Const_>\<open>Ex _ for \<open>Abs (_, _, t')\<close>\<close> => is_first_order_lambda_free t'  | 
|
787  | 
| \<^Const_>\<open>Ex _ for t1\<close> => is_first_order_lambda_free t1  | 
|
788  | 
| \<^Const_>\<open>conj for t1 t2\<close> => is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2  | 
|
789  | 
| \<^Const_>\<open>disj for t1 t2\<close> => is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2  | 
|
790  | 
| \<^Const_>\<open>implies for t1 t2\<close> =>  | 
|
791  | 
is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2  | 
|
792  | 
| \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close> for t1 t2\<close> =>  | 
|
793  | 
is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2  | 
|
| 54829 | 794  | 
| _ => not (exists_subterm (fn Abs _ => true | _ => false) t))  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
795  | 
|
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
796  | 
fun simple_translate_lambdas eta_matters do_lambdas ctxt type_enc t =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
797  | 
if not eta_matters andalso is_first_order_lambda_free t then  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
798  | 
t  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
799  | 
else  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
800  | 
let  | 
| 72921 | 801  | 
fun trans_first_order Ts t =  | 
| 54829 | 802  | 
(case t of  | 
| 74379 | 803  | 
\<^Const_>\<open>Not for t1\<close> => \<^Const>\<open>Not for \<open>trans_first_order Ts t1\<close>\<close>  | 
804  | 
| (t0 as \<^Const_>\<open>All _\<close>) $ Abs (s, T, t') =>  | 
|
| 72921 | 805  | 
t0 $ Abs (s, T, trans_first_order (T :: Ts) t')  | 
| 74379 | 806  | 
| (t0 as \<^Const_>\<open>All _\<close>) $ t1 =>  | 
| 72921 | 807  | 
trans_first_order Ts (t0 $ eta_expand Ts t1 1)  | 
| 74379 | 808  | 
| (t0 as \<^Const_>\<open>Ex _\<close>) $ Abs (s, T, t') =>  | 
| 72921 | 809  | 
t0 $ Abs (s, T, trans_first_order (T :: Ts) t')  | 
| 74379 | 810  | 
| (t0 as \<^Const_>\<open>Ex _\<close>) $ t1 =>  | 
| 72921 | 811  | 
trans_first_order Ts (t0 $ eta_expand Ts t1 1)  | 
| 74379 | 812  | 
| (t0 as \<^Const_>\<open>conj\<close>) $ t1 $ t2 =>  | 
| 72921 | 813  | 
t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2  | 
| 74379 | 814  | 
| (t0 as \<^Const_>\<open>disj\<close>) $ t1 $ t2 =>  | 
| 72921 | 815  | 
t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2  | 
| 74379 | 816  | 
| (t0 as \<^Const_>\<open>implies\<close>) $ t1 $ t2 =>  | 
| 72921 | 817  | 
t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2  | 
| 74379 | 818  | 
| (t0 as \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close>\<close>) $ t1 $ t2 =>  | 
| 72921 | 819  | 
t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
820  | 
| _ =>  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
821  | 
if not (exists_subterm (fn Abs _ => true | _ => false) t) then t  | 
| 54829 | 822  | 
else t |> Envir.eta_contract |> do_lambdas ctxt Ts)  | 
| 72921 | 823  | 
|
824  | 
fun trans_fool Ts t =  | 
|
825  | 
(case t of  | 
|
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
826  | 
(t1 as Const (\<^const_name>\<open>Let\<close>, _)) $ t2 $ t3 =>  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
827  | 
(case t3 of  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
828  | 
Abs (s3, T, t') => t1 $ trans_fool Ts t2 $ Abs (s3, T, trans_fool (T :: Ts) t')  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
829  | 
| _ => trans_fool Ts (t1 $ trans_fool Ts t2 $ eta_expand Ts t3 1))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
830  | 
| (t0 as Const (\<^const_name>\<open>All\<close>, _)) $ t1 =>  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
831  | 
(case t1 of  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
832  | 
Abs (s, T, t') => t0 $ Abs (s, T, trans_fool (T :: Ts) t')  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
833  | 
| _ => trans_fool Ts (t0 $ eta_expand Ts t1 1))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
834  | 
| (t0 as Const (\<^const_name>\<open>Ex\<close>, _)) $ t1 =>  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
835  | 
(case t1 of  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
836  | 
Abs (s, T, t') => t0 $ Abs (s, T, trans_fool (T :: Ts) t')  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
837  | 
| _ => trans_fool Ts (t0 $ eta_expand Ts t1 1))  | 
| 72921 | 838  | 
| t1 $ t2 => trans_fool Ts t1 $ trans_fool Ts t2  | 
839  | 
| Abs _ => t |> Envir.eta_contract |> do_lambdas ctxt Ts  | 
|
840  | 
| _ => t)  | 
|
841  | 
||
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
842  | 
val (t', ctxt') = yield_singleton (Variable.import_terms true) t ctxt  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
843  | 
in  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
844  | 
t' |> (if is_type_enc_fool type_enc then trans_fool else trans_first_order) []  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
845  | 
|> singleton (Variable.export_terms ctxt' ctxt)  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
846  | 
end  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
847  | 
|
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
848  | 
fun do_cheaply_conceal_lambdas Ts (t1 $ t2) =  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
849  | 
do_cheaply_conceal_lambdas Ts t1  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
850  | 
$ do_cheaply_conceal_lambdas Ts t2  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
851  | 
| do_cheaply_conceal_lambdas Ts (Abs (_, T, t)) =  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
852  | 
Const (lam_lifted_poly_prefix ^ serial_string (),  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
853  | 
T --> fastype_of1 (T :: Ts, t))  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
854  | 
| do_cheaply_conceal_lambdas _ t = t  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
855  | 
|
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
856  | 
fun concealed_bound_name j = atp_weak_prefix ^ string_of_int j  | 
| 74918 | 857  | 
fun conceal_bounds Ts t = subst_bounds (map_index (Free o apfst concealed_bound_name) Ts, t)  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
858  | 
fun reveal_bounds Ts =  | 
| 74918 | 859  | 
subst_atomic (map_index (fn (j, T) => (Free (concealed_bound_name j, T), Bound j)) Ts)  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
860  | 
|
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
861  | 
fun do_introduce_combinators ctxt Ts t =  | 
| 59632 | 862  | 
(t |> conceal_bounds Ts  | 
863  | 
|> Thm.cterm_of ctxt  | 
|
| 60328 | 864  | 
|> Meson_Clausify.introduce_combinators_in_cterm ctxt  | 
| 59632 | 865  | 
|> Thm.prop_of |> Logic.dest_equals |> snd  | 
866  | 
|> reveal_bounds Ts)  | 
|
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
867  | 
  (* A type variable of sort "{}" will make abstraction fail. *)
 | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
868  | 
handle THM _ => t |> do_cheaply_conceal_lambdas Ts  | 
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
869  | 
|
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
870  | 
val introduce_combinators = simple_translate_lambdas false do_introduce_combinators  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
871  | 
|
| 45509 | 872  | 
fun constify_lifted (t $ u) = constify_lifted t $ constify_lifted u  | 
873  | 
| constify_lifted (Abs (s, T, t)) = Abs (s, T, constify_lifted t)  | 
|
874  | 
| constify_lifted (Free (x as (s, _))) =  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
875  | 
(if String.isPrefix lam_lifted_prefix s then Const else Free) x  | 
| 45509 | 876  | 
| constify_lifted t = t  | 
877  | 
||
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
878  | 
fun is_binder true (Const (\<^const_name>\<open>Let\<close>, _) $ _) = true  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
879  | 
| is_binder _ t = Lambda_Lifting.is_quantifier t  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
880  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
881  | 
fun lift_lams_part_1 ctxt type_enc =  | 
| 49982 | 882  | 
map hol_close_form #> rpair ctxt  | 
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
883  | 
#-> Lambda_Lifting.lift_lambdas (SOME  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
884  | 
((if is_type_enc_polymorphic type_enc then lam_lifted_poly_prefix  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
885  | 
else lam_lifted_mono_prefix) ^ "_a"))  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
886  | 
(is_binder (is_type_enc_fool type_enc))  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
887  | 
#> fst  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
888  | 
|
| 72921 | 889  | 
fun lift_lams_part_2 ctxt type_enc (facts, lifted) =  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
890  | 
(facts, lifted)  | 
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
891  | 
(* Lambda-lifting sometimes leaves some lambdas around; we need some way to get rid of them *)  | 
| 72921 | 892  | 
|> apply2 (map (introduce_combinators ctxt type_enc))  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58477 
diff
changeset
 | 
893  | 
|> apply2 (map constify_lifted)  | 
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
894  | 
(* Requires bound variables not to clash with any schematic variables (as should be the case right  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
895  | 
after lambda-lifting). *)  | 
| 49982 | 896  | 
|>> map (hol_open_form (unprefix hol_close_form_prefix))  | 
897  | 
||> map (hol_open_form I)  | 
|
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
898  | 
|
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
899  | 
fun lift_lams ctxt type_enc =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
900  | 
(is_type_enc_fool type_enc ?  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
901  | 
map (simple_translate_lambdas true (fn _ => fn _ => fn t => t) ctxt type_enc))  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
902  | 
#> lift_lams_part_1 ctxt type_enc  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
903  | 
#> lift_lams_part_2 ctxt type_enc  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
904  | 
|
| 69593 | 905  | 
fun intentionalize_def (Const (\<^const_name>\<open>All\<close>, _) $ Abs (_, _, t)) =  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
906  | 
intentionalize_def t  | 
| 69593 | 907  | 
| intentionalize_def (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t $ u) =  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
908  | 
let  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
909  | 
fun lam T t = Abs (Name.uu, T, t)  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
910  | 
val (head, args) = strip_comb t ||> rev  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
911  | 
val head_T = fastype_of head  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
912  | 
val n = length args  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
913  | 
val arg_Ts = head_T |> binder_types |> take n |> rev  | 
| 74918 | 914  | 
val u = u |> subst_atomic (map_index (swap o apfst Bound) args)  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
915  | 
in HOLogic.eq_const head_T $ head $ fold lam arg_Ts u end  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
916  | 
| intentionalize_def t = t  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
917  | 
|
| 47981 | 918  | 
type ifact =  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
919  | 
  {name : string,
 | 
| 46340 | 920  | 
stature : stature,  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
921  | 
role : atp_formula_role,  | 
| 
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
922  | 
iformula : (string * string, typ, iterm, string * string) atp_formula,  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
923  | 
atomic_types : typ list}  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
924  | 
|
| 47981 | 925  | 
fun update_iformula f ({name, stature, role, iformula, atomic_types} : ifact) =
 | 
| 54798 | 926  | 
  {name = name, stature = stature, role = role, iformula = f iformula, atomic_types = atomic_types}
 | 
927  | 
: ifact  | 
|
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
928  | 
|
| 47981 | 929  | 
fun ifact_lift f ({iformula, ...} : ifact) = f iformula
 | 
| 
42558
 
3d9930cb6770
cleaned up "explicit_apply" so that it shares most of its code path with the default mode of operation
 
blanchet 
parents: 
42557 
diff
changeset
 | 
930  | 
|
| 47150 | 931  | 
fun insert_type thy get_T x xs =  | 
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43039 
diff
changeset
 | 
932  | 
let val T = get_T x in  | 
| 47150 | 933  | 
if exists (type_instance thy T o get_T) xs then xs  | 
934  | 
else x :: filter_out (type_generalization thy T o get_T) xs  | 
|
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43039 
diff
changeset
 | 
935  | 
end  | 
| 
42677
 
25496cd3c199
monotonic type inference in ATP Sledgehammer problems -- based on Claessen & al.'s CADE 2011 paper, Sect. 2.3.
 
blanchet 
parents: 
42675 
diff
changeset
 | 
936  | 
|
| 48202 | 937  | 
fun chop_fun 0 T = ([], T)  | 
| 69593 | 938  | 
| chop_fun n (Type (\<^type_name>\<open>fun\<close>, [dom_T, ran_T])) =  | 
| 48202 | 939  | 
chop_fun (n - 1) ran_T |>> cons dom_T  | 
940  | 
| chop_fun _ T = ([], T)  | 
|
| 
41136
 
30bedf58b177
implemented new type system encoding "overload_args", which is more lightweight than "const_args" (the unsound default) and hopefully almost as sound
 
blanchet 
parents: 
41134 
diff
changeset
 | 
941  | 
|
| 52032 | 942  | 
fun filter_type_args thy ctrss type_enc s ary T_args =  | 
| 45315 | 943  | 
let val poly = polymorphism_of_type_enc type_enc in  | 
| 
50968
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
944  | 
if s = type_tag_name then (* FIXME: why not "type_guard_name" as well? *)  | 
| 48202 | 945  | 
T_args  | 
| 54829 | 946  | 
else  | 
947  | 
(case type_enc of  | 
|
| 74890 | 948  | 
Native (_, _, Raw_Polymorphic _, _) => T_args  | 
949  | 
| Native (_, _, Type_Class_Polymorphic, _) => T_args  | 
|
| 54829 | 950  | 
| _ =>  | 
951  | 
let  | 
|
952  | 
fun gen_type_args _ _ [] = []  | 
|
953  | 
| gen_type_args keep strip_ty T_args =  | 
|
954  | 
let  | 
|
955  | 
val U = robust_const_type thy s  | 
|
956  | 
val (binder_Us, body_U) = strip_ty U  | 
|
957  | 
val in_U_vars = fold Term.add_tvarsT binder_Us []  | 
|
958  | 
val out_U_vars = Term.add_tvarsT body_U []  | 
|
959  | 
fun filt (U_var, T) =  | 
|
960  | 
if keep (member (op =) in_U_vars U_var,  | 
|
961  | 
member (op =) out_U_vars U_var) then  | 
|
962  | 
T  | 
|
963  | 
else  | 
|
964  | 
dummyT  | 
|
965  | 
val U_args = (s, U) |> robust_const_type_args thy  | 
|
966  | 
in map (filt o apfst dest_TVar) (U_args ~~ T_args) end  | 
|
967  | 
handle TYPE _ => T_args  | 
|
968  | 
fun is_always_ctr (s', T') =  | 
|
969  | 
s' = s andalso type_equiv thy (T', robust_const_type thy s')  | 
|
970  | 
val noninfer_type_args = gen_type_args (not o fst) (chop_fun ary)  | 
|
971  | 
val ctr_infer_type_args = gen_type_args fst strip_type  | 
|
972  | 
val level = level_of_type_enc type_enc  | 
|
973  | 
in  | 
|
| 69593 | 974  | 
if level = No_Types orelse s = \<^const_name>\<open>HOL.eq\<close> orelse  | 
| 54829 | 975  | 
(case level of Const_Types _ => s = app_op_name | _ => false) then  | 
976  | 
[]  | 
|
977  | 
else if poly = Mangled_Monomorphic then  | 
|
978  | 
T_args  | 
|
979  | 
else if level = All_Types then  | 
|
980  | 
(case type_enc of  | 
|
981  | 
Guards _ => noninfer_type_args T_args  | 
|
982  | 
| Tags _ => [])  | 
|
983  | 
else if level = Undercover_Types then  | 
|
984  | 
noninfer_type_args T_args  | 
|
985  | 
else if level <> Const_Types Without_Ctr_Optim andalso  | 
|
986  | 
exists (exists is_always_ctr) ctrss then  | 
|
987  | 
ctr_infer_type_args T_args  | 
|
988  | 
else  | 
|
989  | 
T_args  | 
|
990  | 
end)  | 
|
| 
44774
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
991  | 
end  | 
| 
42227
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
42180 
diff
changeset
 | 
992  | 
|
| 54820 | 993  | 
fun raw_atp_type_of_typ type_enc =  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
994  | 
let  | 
| 
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
995  | 
fun term (Type (s, Ts)) =  | 
| 
68250
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
996  | 
AType  | 
| 69593 | 997  | 
((if s = \<^type_name>\<open>fun\<close> andalso is_type_enc_higher_order type_enc then  | 
| 
68250
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
998  | 
`I tptp_fun_type  | 
| 72588 | 999  | 
else if s = \<^type_name>\<open>bool\<close> andalso  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1000  | 
(is_type_enc_full_higher_order type_enc orelse is_type_enc_fool type_enc) then  | 
| 
68250
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1001  | 
`I tptp_bool_type  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1002  | 
else  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1003  | 
`make_fixed_type_const s, []), map term Ts)  | 
| 54820 | 1004  | 
| term (TFree (s, _)) = AType ((`make_tfree s, []), [])  | 
| 55212 | 1005  | 
| term (TVar z) = AType ((tvar_name z, []), [])  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
1006  | 
in term end  | 
| 42562 | 1007  | 
|
| 54820 | 1008  | 
fun atp_term_of_atp_type (AType ((name, _), tys)) = ATerm ((name, []), map atp_term_of_atp_type tys)  | 
1009  | 
| atp_term_of_atp_type _ = raise Fail "unexpected type"  | 
|
| 48138 | 1010  | 
|
| 54820 | 1011  | 
fun atp_type_of_type_arg type_enc T =  | 
1012  | 
if T = dummyT then NONE else SOME (raw_atp_type_of_typ type_enc T)  | 
|
| 
43401
 
e93dfcb53535
fixed soundness bug made more visible by previous change
 
blanchet 
parents: 
43399 
diff
changeset
 | 
1013  | 
|
| 42562 | 1014  | 
(* This shouldn't clash with anything else. *)  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1015  | 
val uncurried_alias_sep = "\000"  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1016  | 
val mangled_type_sep = "\001"  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1017  | 
|
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1018  | 
val ascii_of_uncurried_alias_sep = ascii_of uncurried_alias_sep  | 
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1019  | 
|
| 54820 | 1020  | 
fun generic_mangled_type_name f (AType ((name, _), [])) = f name  | 
1021  | 
| generic_mangled_type_name f (AType ((name, _), tys)) =  | 
|
1022  | 
    f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys) ^ ")"
 | 
|
| 48138 | 1023  | 
| generic_mangled_type_name _ _ = raise Fail "unexpected type"  | 
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1024  | 
|
| 54820 | 1025  | 
fun mangled_type type_enc = generic_mangled_type_name fst o raw_atp_type_of_typ type_enc  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
1026  | 
|
| 46435 | 1027  | 
fun make_native_type s =  | 
| 54820 | 1028  | 
if s = tptp_bool_type orelse s = tptp_fun_type orelse s = tptp_individual_type then s  | 
1029  | 
else native_type_prefix ^ ascii_of s  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1030  | 
|
| 54820 | 1031  | 
fun native_atp_type_of_raw_atp_type type_enc pred_sym ary =  | 
| 42963 | 1032  | 
let  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
1033  | 
fun to_mangled_atype ty =  | 
| 54820 | 1034  | 
AType (((make_native_type (generic_mangled_type_name fst ty),  | 
1035  | 
generic_mangled_type_name snd ty), []), [])  | 
|
1036  | 
fun to_poly_atype (AType ((name, clss), tys)) = AType ((name, clss), map to_poly_atype tys)  | 
|
| 48138 | 1037  | 
| to_poly_atype _ = raise Fail "unexpected type"  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
1038  | 
val to_atype =  | 
| 48131 | 1039  | 
if is_type_enc_polymorphic type_enc then to_poly_atype  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
1040  | 
else to_mangled_atype  | 
| 42963 | 1041  | 
fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))  | 
| 
68250
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1042  | 
fun to_ho (ty as AType (((s, _), _), tys)) =  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1043  | 
if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1044  | 
| to_ho _ = raise Fail "unexpected type"  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1045  | 
fun to_lfho (ty as AType (((s, _), _), tys)) =  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1046  | 
if s = tptp_fun_type then to_afun to_ho to_lfho tys  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1047  | 
else if pred_sym then bool_atype  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1048  | 
else to_atype ty  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1049  | 
| to_lfho _ = raise Fail "unexpected type"  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
1050  | 
fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty  | 
| 48138 | 1051  | 
| to_fo ary (AType (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys  | 
1052  | 
| to_fo _ _ = raise Fail "unexpected type"  | 
|
| 
68250
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1053  | 
in  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1054  | 
if is_type_enc_full_higher_order type_enc then to_ho  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1055  | 
else if is_type_enc_higher_order type_enc then to_lfho  | 
| 
68250
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1056  | 
else to_fo ary  | 
| 
 
c45067867860
added lambda-free HO output for Ehoh (higher-order E prototype)
 
blanchet 
parents: 
67405 
diff
changeset
 | 
1057  | 
end  | 
| 42963 | 1058  | 
|
| 54820 | 1059  | 
fun native_atp_type_of_typ type_enc pred_sym ary =  | 
1060  | 
native_atp_type_of_raw_atp_type type_enc pred_sym ary o raw_atp_type_of_typ type_enc  | 
|
| 42963 | 1061  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1062  | 
(* Make atoms for sorted type variables. *)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1063  | 
fun generic_add_sorts_on_type _ [] = I  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1064  | 
| generic_add_sorts_on_type T (s :: ss) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1065  | 
generic_add_sorts_on_type T ss  | 
| 69593 | 1066  | 
#> (if s = the_single \<^sort>\<open>type\<close> then I else insert (op =) (s, T))  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1067  | 
fun add_sorts_on_tfree (T as TFree (_, S)) = generic_add_sorts_on_type T S  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1068  | 
| add_sorts_on_tfree _ = I  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1069  | 
fun add_sorts_on_tvar (T as TVar (_, S)) = generic_add_sorts_on_type T S  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1070  | 
| add_sorts_on_tvar _ = I  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1071  | 
|
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1072  | 
fun process_type_args type_enc T_args =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1073  | 
if is_type_enc_native type_enc then  | 
| 54820 | 1074  | 
(map (native_atp_type_of_typ type_enc false 0) T_args, [])  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1075  | 
else  | 
| 54820 | 1076  | 
([], map_filter (Option.map atp_term_of_atp_type o atp_type_of_type_arg type_enc) T_args)  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1077  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1078  | 
fun class_atom type_enc (cl, T) =  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1079  | 
let  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1080  | 
val cl = `make_class cl  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1081  | 
val (ty_args, tm_args) = process_type_args type_enc [T]  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1082  | 
val tm_args =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1083  | 
tm_args @  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1084  | 
(case type_enc of  | 
| 74890 | 1085  | 
Native (_, _, Raw_Polymorphic Without_Phantom_Type_Vars, _) =>  | 
| 54820 | 1086  | 
[ATerm ((TYPE_name, ty_args), [])]  | 
1087  | 
| _ => [])  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1088  | 
in AAtom (ATerm ((cl, ty_args), tm_args)) end  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1089  | 
|
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1090  | 
fun class_atoms type_enc (cls, T) =  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1091  | 
map (fn cl => class_atom type_enc (cl, T)) cls  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1092  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
1093  | 
fun class_membs_of_types type_enc add_sorts_on_typ Ts =  | 
| 54798 | 1094  | 
[] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1095  | 
|
| 48141 | 1096  | 
fun mk_aconns c = split_last #> uncurry (fold_rev (mk_aconn c))  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1097  | 
|
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1098  | 
fun mk_ahorn [] phi = phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1099  | 
| mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1100  | 
|
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1101  | 
fun mk_aquant _ [] phi = phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1102  | 
| mk_aquant q xs (phi as AQuant (q', xs', phi')) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1103  | 
if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1104  | 
| mk_aquant q xs phi = AQuant (q, xs, phi)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1105  | 
|
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1106  | 
fun mk_atyquant _ [] phi = phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1107  | 
| mk_atyquant q xs (phi as ATyQuant (q', xs', phi')) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1108  | 
if q = q' then ATyQuant (q, xs @ xs', phi') else ATyQuant (q, xs, phi)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1109  | 
| mk_atyquant q xs phi = ATyQuant (q, xs, phi)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1110  | 
|
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1111  | 
fun close_universally add_term_vars phi =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1112  | 
let  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1113  | 
fun add_formula_vars bounds (ATyQuant (_, _, phi)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1114  | 
add_formula_vars bounds phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1115  | 
| add_formula_vars bounds (AQuant (_, xs, phi)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1116  | 
add_formula_vars (map fst xs @ bounds) phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1117  | 
| add_formula_vars bounds (AConn (_, phis)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1118  | 
fold (add_formula_vars bounds) phis  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1119  | 
| add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm  | 
| 
50706
 
573d84e08f3f
close formulas in the natural order, not its reverse -- so that Skolem arguments appear in the right order in Isar proofs
 
blanchet 
parents: 
50521 
diff
changeset
 | 
1120  | 
in mk_aquant AForall (rev (add_formula_vars [] phi [])) phi end  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1121  | 
|
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1122  | 
fun add_term_vars bounds (ATerm ((name as (s, _), _), tms)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1123  | 
(if is_tptp_variable s andalso  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1124  | 
not (String.isPrefix tvar_prefix s) andalso  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1125  | 
not (member (op =) bounds name) then  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1126  | 
insert (op =) (name, NONE)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1127  | 
else  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1128  | 
I)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1129  | 
#> fold (add_term_vars bounds) tms  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1130  | 
| add_term_vars bounds (AAbs (((name, _), tm), args)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1131  | 
add_term_vars (name :: bounds) tm #> fold (add_term_vars bounds) args  | 
| 
50706
 
573d84e08f3f
close formulas in the natural order, not its reverse -- so that Skolem arguments appear in the right order in Isar proofs
 
blanchet 
parents: 
50521 
diff
changeset
 | 
1132  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1133  | 
fun close_formula_universally phi = close_universally add_term_vars phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1134  | 
|
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1135  | 
fun add_iterm_vars bounds (IApp (tm1, tm2)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1136  | 
fold (add_iterm_vars bounds) [tm1, tm2]  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1137  | 
| add_iterm_vars _ (IConst _) = I  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1138  | 
| add_iterm_vars bounds (IVar (name, T)) =  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1139  | 
not (member (op =) bounds name) ? insert (op =) (name, SOME T)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1140  | 
| add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1141  | 
|
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1142  | 
fun aliased_uncurried ary (s, s') =  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1143  | 
(s ^ ascii_of_uncurried_alias_sep ^ string_of_int ary, s' ^ string_of_int ary)  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1144  | 
fun unaliased_uncurried (s, s') =  | 
| 54829 | 1145  | 
(case space_explode uncurried_alias_sep s of  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1146  | 
[_] => (s, s')  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1147  | 
| [s1, s2] => (s1, unsuffix s2 s')  | 
| 54829 | 1148  | 
| _ => raise Fail "ill-formed explicit application alias")  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1149  | 
|
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1150  | 
fun raw_mangled_const_name type_name ty_args (s, s') =  | 
| 42963 | 1151  | 
let  | 
1152  | 
fun type_suffix f g =  | 
|
| 63694 | 1153  | 
fold_rev (prefix o g o prefix mangled_type_sep o type_name f) ty_args ""  | 
| 42963 | 1154  | 
in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1155  | 
fun mangled_const_name type_enc =  | 
| 54820 | 1156  | 
map_filter (atp_type_of_type_arg type_enc)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1157  | 
#> raw_mangled_const_name generic_mangled_type_name  | 
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1158  | 
|
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1159  | 
val parse_mangled_ident =  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
1160  | 
  Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
 | 
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1161  | 
|
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1162  | 
fun parse_mangled_type x =  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1163  | 
(parse_mangled_ident  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1164  | 
   -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
 | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
1165  | 
[] >> (ATerm o apfst (rpair []))) x  | 
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1166  | 
and parse_mangled_types x =  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1167  | 
(parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1168  | 
|
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1169  | 
fun unmangled_type s =  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1170  | 
s |> suffix ")" |> raw_explode  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1171  | 
|> Scan.finite Symbol.stopper  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1172  | 
           (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
 | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1173  | 
quote s)) parse_mangled_type))  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1174  | 
|> fst  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1175  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1176  | 
fun unmangled_const_name s =  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1177  | 
(s, s) |> unaliased_uncurried |> fst |> space_explode mangled_type_sep  | 
| 
51920
 
16f3b9d4e515
proper unmangling -- the bug is visible when "uncurried_aliases" is enabled with Alt-Ergo or Poly. SPASS
 
blanchet 
parents: 
51717 
diff
changeset
 | 
1178  | 
|
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1179  | 
fun unmangled_const s =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1180  | 
let val ss = unmangled_const_name s in  | 
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1181  | 
(hd ss, map unmangled_type (tl ss))  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1182  | 
end  | 
| 
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1183  | 
|
| 
51920
 
16f3b9d4e515
proper unmangling -- the bug is visible when "uncurried_aliases" is enabled with Alt-Ergo or Poly. SPASS
 
blanchet 
parents: 
51717 
diff
changeset
 | 
1184  | 
val unmangled_invert_const = invert_const o hd o unmangled_const_name  | 
| 
 
16f3b9d4e515
proper unmangling -- the bug is visible when "uncurried_aliases" is enabled with Alt-Ergo or Poly. SPASS
 
blanchet 
parents: 
51717 
diff
changeset
 | 
1185  | 
|
| 74901 | 1186  | 
fun vars_of_iterm (IConst ((s, _), _, _)) = [s]  | 
1187  | 
| vars_of_iterm (IVar ((s, _), _)) = [s]  | 
|
1188  | 
| vars_of_iterm (IApp (tm1, tm2)) = union (op =) (vars_of_iterm tm1) (vars_of_iterm tm2)  | 
|
1189  | 
| vars_of_iterm (IAbs (((s, _), _), tm)) = insert (op =) s (vars_of_iterm tm)  | 
|
| 72920 | 1190  | 
|
| 
72922
 
d78bd4432f05
tuned name generation in tptp to not depend on shadowing
 
desharna 
parents: 
72921 
diff
changeset
 | 
1191  | 
fun generate_unique_name gen unique n =  | 
| 72920 | 1192  | 
let val x = gen n in  | 
1193  | 
if unique x then x else generate_unique_name gen unique (n + 1)  | 
|
| 
72918
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1194  | 
end  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1195  | 
|
| 72920 | 1196  | 
fun eta_expand_quantifier_body (tm as IAbs _) = tm  | 
1197  | 
| eta_expand_quantifier_body tm =  | 
|
1198  | 
let  | 
|
| 
72922
 
d78bd4432f05
tuned name generation in tptp to not depend on shadowing
 
desharna 
parents: 
72921 
diff
changeset
 | 
1199  | 
(* We accumulate all variables because E 2.5 does not support variable shadowing. *)  | 
| 74901 | 1200  | 
val vars = vars_of_iterm tm  | 
| 72920 | 1201  | 
val x = generate_unique_name  | 
1202  | 
(fn n => "X" ^ (if n = 0 then "" else string_of_int n))  | 
|
| 
72922
 
d78bd4432f05
tuned name generation in tptp to not depend on shadowing
 
desharna 
parents: 
72921 
diff
changeset
 | 
1203  | 
(fn name => not (exists (equal name) vars)) 0  | 
| 74328 | 1204  | 
|> `(prefix bound_var_prefix)  | 
| 72920 | 1205  | 
val T = domain_type (ityp_of tm)  | 
1206  | 
in  | 
|
| 74328 | 1207  | 
IAbs ((x, T), IApp (tm, IConst (x, T, [])))  | 
| 72920 | 1208  | 
end  | 
1209  | 
||
| 
73859
 
bc263f1f68cd
added support for TFX's and THF's $ite to Sledgehammer
 
desharna 
parents: 
72922 
diff
changeset
 | 
1210  | 
fun introduce_builtins_and_proxies_in_iterm type_enc =  | 
| 
42568
 
7b9801a34836
no needless "fequal" proxies if "explicit_apply" is set + always have readable names
 
blanchet 
parents: 
42566 
diff
changeset
 | 
1211  | 
let  | 
| 
73859
 
bc263f1f68cd
added support for TFX's and THF's $ite to Sledgehammer
 
desharna 
parents: 
72922 
diff
changeset
 | 
1212  | 
val is_fool = is_type_enc_fool type_enc  | 
| 74758 | 1213  | 
val has_ite = has_type_enc_ite type_enc  | 
| 74890 | 1214  | 
val has_let = has_type_enc_let type_enc  | 
| 74896 | 1215  | 
val has_choice = has_type_enc_choice type_enc  | 
| 
72918
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1216  | 
fun tweak_ho_quant ho_quant (T as Type (_, [p_T as Type (_, [x_T, _]), _])) [] =  | 
| 72917 | 1217  | 
(* Eta-expand "!!" and "??", to work around LEO-II, Leo-III, and Satallax parser  | 
1218  | 
limitations. This works in conjunction with special code in "ATP_Problem" that uses the  | 
|
1219  | 
syntactic sugar "!" and "?" whenever possible. *)  | 
|
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1220  | 
IAbs ((`I "P", p_T),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1221  | 
IApp (IConst (`I ho_quant, T, []),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1222  | 
IAbs ((`I "X", x_T),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1223  | 
IApp (IConst (`I "P", p_T, []),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1224  | 
IConst (`I "X", x_T, [])))))  | 
| 
72918
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1225  | 
| tweak_ho_quant ho_quant T _ = IConst (`I ho_quant, T, [])  | 
| 72917 | 1226  | 
fun tweak_ho_equal T argc =  | 
1227  | 
if argc = 2 then  | 
|
1228  | 
IConst (`I tptp_equal, T, [])  | 
|
1229  | 
else  | 
|
1230  | 
(* Eta-expand partially applied THF equality, because the LEO-II and Satallax parsers  | 
|
1231  | 
complain about not being able to infer the type of "=". *)  | 
|
1232  | 
let val i_T = domain_type T in  | 
|
1233  | 
IAbs ((`I "Y", i_T),  | 
|
1234  | 
IAbs ((`I "Z", i_T),  | 
|
1235  | 
IApp (IApp (IConst (`I tptp_equal, T, []),  | 
|
1236  | 
IConst (`I "Y", i_T, [])),  | 
|
1237  | 
IConst (`I "Z", i_T, []))))  | 
|
1238  | 
end  | 
|
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1239  | 
fun intro top_level args (IApp (tm1, tm2)) =  | 
| 
72918
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1240  | 
let  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1241  | 
val tm1' = intro top_level (tm2 :: args) tm1  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1242  | 
val tm2' = intro false [] tm2  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1243  | 
val tm2'' =  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1244  | 
(case tm1' of  | 
| 74328 | 1245  | 
IApp (IConst ((s, _), _, _), _) =>  | 
1246  | 
if s = tptp_let then  | 
|
1247  | 
(case tm2' of  | 
|
1248  | 
IAbs ((name, T), tm) =>  | 
|
1249  | 
let  | 
|
1250  | 
val name' =  | 
|
1251  | 
map_prod (prefix let_bound_var_prefix o unprefix bound_var_prefix) I name  | 
|
1252  | 
in  | 
|
1253  | 
IAbs ((name', T), alpha_rename name name' tm)  | 
|
1254  | 
end  | 
|
1255  | 
| _ => error "Function abstraction expected")  | 
|
1256  | 
else  | 
|
1257  | 
tm2'  | 
|
1258  | 
| IConst ((s, _), _, _) =>  | 
|
| 74896 | 1259  | 
if s = tptp_ho_forall orelse s = tptp_ho_exists orelse s = tptp_choice then  | 
| 72920 | 1260  | 
eta_expand_quantifier_body tm2'  | 
| 
72918
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1261  | 
else  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1262  | 
tm2'  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1263  | 
| _ => tm2')  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1264  | 
in  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1265  | 
IApp (tm1', tm2'')  | 
| 
 
4bf5b4f8bd6f
proper eta-expansion to avoid lambdas in tptp fool
 
desharna 
parents: 
72917 
diff
changeset
 | 
1266  | 
end  | 
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1267  | 
| intro top_level args (IConst (name as (s, _), T, T_args)) =  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1268  | 
let val argc = length args in  | 
| 
74759
 
32e95d996acc
tuned generation of TPTP with $ite in function position
 
desharna 
parents: 
74758 
diff
changeset
 | 
1269  | 
if has_ite andalso s = "c_If" andalso argc >= 3 then  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1270  | 
IConst (`I tptp_ite, T, [])  | 
| 74890 | 1271  | 
else if has_let andalso s = "c_Let" andalso argc >= 2 then  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1272  | 
IConst (`I tptp_let, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1273  | 
else  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1274  | 
(case proxify_const s of  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1275  | 
SOME proxy_base =>  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1276  | 
let  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1277  | 
fun plain_const () = IConst (name, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1278  | 
fun proxy_const () = IConst (proxy_base |>> prefix const_prefix, T, T_args)  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1279  | 
fun handle_fool card x = if card = argc then x else proxy_const ()  | 
| 74896 | 1280  | 
fun handle_min_card card x = if argc < card then proxy_const () else x  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1281  | 
in  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1282  | 
if top_level then  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1283  | 
(case s of  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1284  | 
"c_False" => IConst (`I tptp_false, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1285  | 
| "c_True" => IConst (`I tptp_true, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1286  | 
| _ => plain_const ())  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1287  | 
else if is_type_enc_full_higher_order type_enc then  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1288  | 
(case s of  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1289  | 
"c_False" => IConst (`I tptp_false, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1290  | 
| "c_True" => IConst (`I tptp_true, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1291  | 
| "c_Not" => IConst (`I tptp_not, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1292  | 
| "c_conj" => IConst (`I tptp_and, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1293  | 
| "c_disj" => IConst (`I tptp_or, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1294  | 
| "c_implies" => IConst (`I tptp_implies, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1295  | 
| "c_All" => tweak_ho_quant tptp_ho_forall T args  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1296  | 
| "c_Ex" => tweak_ho_quant tptp_ho_exists T args  | 
| 74896 | 1297  | 
| "c_Choice" =>  | 
1298  | 
if has_choice then  | 
|
1299  | 
handle_min_card 1 (IConst (`I tptp_choice, T, []))  | 
|
1300  | 
else  | 
|
1301  | 
proxy_const ()  | 
|
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1302  | 
| s =>  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1303  | 
if is_tptp_equal s then  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1304  | 
tweak_ho_equal T argc  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1305  | 
else  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1306  | 
plain_const ())  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1307  | 
else if is_fool then  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1308  | 
(case s of  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1309  | 
"c_False" => IConst (`I tptp_false, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1310  | 
| "c_True" => IConst (`I tptp_true, T, [])  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1311  | 
| "c_Not" => handle_fool 1 (IConst (`I tptp_not, T, []))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1312  | 
| "c_conj" => handle_fool 2 (IConst (`I tptp_and, T, []))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1313  | 
| "c_disj" => handle_fool 2 (IConst (`I tptp_or, T, []))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1314  | 
| "c_implies" => handle_fool 2 (IConst (`I tptp_implies, T, []))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1315  | 
| "c_All" => handle_fool 1 (tweak_ho_quant tptp_ho_forall T args)  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1316  | 
| "c_Ex" => handle_fool 1 (tweak_ho_quant tptp_ho_exists T args)  | 
| 74899 | 1317  | 
| "c_Choice" => proxy_const ()  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1318  | 
| s =>  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1319  | 
if is_tptp_equal s then  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1320  | 
handle_fool 2 (IConst (`I tptp_equal, T, []))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1321  | 
else  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1322  | 
plain_const ())  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1323  | 
else  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1324  | 
proxy_const ()  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1325  | 
end  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1326  | 
| NONE =>  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1327  | 
if s = tptp_choice then  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1328  | 
tweak_ho_quant tptp_choice T args  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1329  | 
else  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1330  | 
IConst (name, T, T_args))  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
1331  | 
end  | 
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1332  | 
| intro _ _ (IAbs (bound, tm)) = IAbs (bound, intro false [] tm)  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1333  | 
| intro _ _ tm = tm  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1334  | 
in intro true [] end  | 
| 
42568
 
7b9801a34836
no needless "fequal" proxies if "explicit_apply" is set + always have readable names
 
blanchet 
parents: 
42566 
diff
changeset
 | 
1335  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1336  | 
fun mangle_type_args_in_const type_enc (name as (s, _)) T_args =  | 
| 48203 | 1337  | 
if String.isPrefix const_prefix s andalso is_type_enc_mangling type_enc then  | 
1338  | 
(mangled_const_name type_enc T_args name, [])  | 
|
1339  | 
else  | 
|
1340  | 
(name, T_args)  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1341  | 
fun mangle_type_args_in_iterm type_enc =  | 
| 48201 | 1342  | 
if is_type_enc_mangling type_enc then  | 
| 
44774
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1343  | 
let  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1344  | 
fun mangle (IApp (tm1, tm2)) = IApp (mangle tm1, mangle tm2)  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1345  | 
| mangle (tm as IConst (_, _, [])) = tm  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1346  | 
| mangle (IConst (name, T, T_args)) =  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1347  | 
mangle_type_args_in_const type_enc name T_args  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1348  | 
|> (fn (name, T_args) => IConst (name, T, T_args))  | 
| 
44774
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1349  | 
| mangle (IAbs (bound, tm)) = IAbs (bound, mangle tm)  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1350  | 
| mangle tm = tm  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1351  | 
in mangle end  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1352  | 
else  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1353  | 
I  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1354  | 
|
| 52032 | 1355  | 
fun filter_type_args_in_const _ _ _ _ _ [] = []  | 
1356  | 
| filter_type_args_in_const thy ctrss type_enc ary s T_args =  | 
|
| 54829 | 1357  | 
(case unprefix_and_unascii const_prefix s of  | 
1358  | 
NONE => if level_of_type_enc type_enc = No_Types orelse s = tptp_choice then [] else T_args  | 
|
1359  | 
| SOME s'' => filter_type_args thy ctrss type_enc (unmangled_invert_const s'') ary T_args)  | 
|
| 
52029
 
1eefb69cb9c1
don't recognize overloaded constants as constructors for the purpose of removing type arguments
 
blanchet 
parents: 
52028 
diff
changeset
 | 
1360  | 
|
| 52028 | 1361  | 
fun filter_type_args_in_iterm thy ctrss type_enc =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1362  | 
let  | 
| 
44774
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1363  | 
fun filt ary (IApp (tm1, tm2)) = IApp (filt (ary + 1) tm1, filt 0 tm2)  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1364  | 
| filt ary (IConst (name as (s, _), T, T_args)) =  | 
| 52032 | 1365  | 
filter_type_args_in_const thy ctrss type_enc ary s T_args  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1366  | 
|> (fn T_args => IConst (name, T, T_args))  | 
| 
44774
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1367  | 
| filt _ (IAbs (bound, tm)) = IAbs (bound, filt 0 tm)  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1368  | 
| filt _ tm = tm  | 
| 
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
1369  | 
in filt 0 end  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1370  | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
1371  | 
fun iformula_of_prop ctxt type_enc iff_for_eq =  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1372  | 
let  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1373  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1374  | 
fun do_term bs t atomic_Ts =  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
1375  | 
iterm_of_term thy type_enc bs (Envir.eta_contract t)  | 
| 
73859
 
bc263f1f68cd
added support for TFX's and THF's $ite to Sledgehammer
 
desharna 
parents: 
72922 
diff
changeset
 | 
1376  | 
|>> (introduce_builtins_and_proxies_in_iterm type_enc  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1377  | 
#> mangle_type_args_in_iterm type_enc #> AAtom)  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1378  | 
||> union (op =) atomic_Ts  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1379  | 
fun do_quant bs q pos s T t' =  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1380  | 
let  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1381  | 
val s = singleton (Name.variant_list (map fst bs)) s  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1382  | 
val universal = Option.map (q = AExists ? not) pos  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1383  | 
val name =  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1384  | 
s |> `(case universal of  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1385  | 
SOME true => make_all_bound_var  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1386  | 
| SOME false => make_exist_bound_var  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1387  | 
| NONE => make_bound_var)  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1388  | 
in  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1389  | 
do_formula ((s, (name, T)) :: bs) pos t'  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1390  | 
#>> mk_aquant q [(name, SOME T)]  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1391  | 
##> union (op =) (atomic_types_of T)  | 
| 
38518
 
54727b44e277
handle bound name conflicts gracefully in FOF translation
 
blanchet 
parents: 
38496 
diff
changeset
 | 
1392  | 
end  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1393  | 
and do_conn bs c pos1 t1 pos2 t2 =  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1394  | 
do_formula bs pos1 t1 ##>> do_formula bs pos2 t2 #>> uncurry (mk_aconn c)  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1395  | 
and do_formula bs pos t =  | 
| 54829 | 1396  | 
(case t of  | 
| 74379 | 1397  | 
\<^Const_>\<open>Trueprop for t1\<close> => do_formula bs pos t1  | 
1398  | 
| \<^Const_>\<open>Not for t1\<close> => do_formula bs (Option.map not pos) t1 #>> mk_anot  | 
|
1399  | 
| \<^Const_>\<open>All _ for \<open>Abs (s, T, t')\<close>\<close> => do_quant bs AForall pos s T t'  | 
|
1400  | 
| (t0 as \<^Const_>\<open>All _\<close>) $ t1 =>  | 
|
| 
45167
 
6bc8d260d459
gracefully handle quantifiers of the form "All $ t" where "t" is not a lambda-abstraction in higher-order translations
 
blanchet 
parents: 
44859 
diff
changeset
 | 
1401  | 
do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)  | 
| 74379 | 1402  | 
| \<^Const_>\<open>Ex _ for \<open>Abs (s, T, t')\<close>\<close> => do_quant bs AExists pos s T t'  | 
1403  | 
| (t0 as \<^Const_>\<open>Ex _\<close>) $ t1 =>  | 
|
| 
45167
 
6bc8d260d459
gracefully handle quantifiers of the form "All $ t" where "t" is not a lambda-abstraction in higher-order translations
 
blanchet 
parents: 
44859 
diff
changeset
 | 
1404  | 
do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)  | 
| 74379 | 1405  | 
| \<^Const_>\<open>conj for t1 t2\<close> => do_conn bs AAnd pos t1 pos t2  | 
1406  | 
| \<^Const_>\<open>disj for t1 t2\<close> => do_conn bs AOr pos t1 pos t2  | 
|
1407  | 
| \<^Const_>\<open>implies for t1 t2\<close> => do_conn bs AImplies (Option.map not pos) t1 pos t2  | 
|
1408  | 
| \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close> for t1 t2\<close> =>  | 
|
| 
47905
 
9b6afe0eb69c
cleaner handling of bi-implication for THF output of first-order type encodings
 
blanchet 
parents: 
47810 
diff
changeset
 | 
1409  | 
if iff_for_eq then do_conn bs AIff NONE t1 NONE t2 else do_term bs t  | 
| 54829 | 1410  | 
| _ => do_term bs t)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1411  | 
in do_formula [] end  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1412  | 
|
| 74051 | 1413  | 
fun presimplify_term simp_options ctxt t =  | 
| 47151 | 1414  | 
if exists_Const (member (op =) Meson.presimplified_consts o fst) t then  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51651 
diff
changeset
 | 
1415  | 
t |> Skip_Proof.make_thm (Proof_Context.theory_of ctxt)  | 
| 74051 | 1416  | 
|> Meson.presimplify simp_options ctxt  | 
| 59582 | 1417  | 
|> Thm.prop_of  | 
| 47151 | 1418  | 
else  | 
1419  | 
t  | 
|
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1420  | 
|
| 45514 | 1421  | 
fun preprocess_abstractions_in_terms trans_lams facts =  | 
| 43862 | 1422  | 
let  | 
| 
43863
 
a43d61270142
ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
 
blanchet 
parents: 
43862 
diff
changeset
 | 
1423  | 
val (facts, lambda_ts) =  | 
| 45514 | 1424  | 
facts |> map (snd o snd) |> trans_lams  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
1425  | 
|>> map2 (fn (name, (role, _)) => fn t => (name, (role, t))) facts  | 
| 74918 | 1426  | 
val lam_facts = lambda_ts  | 
1427  | 
|> map_index (fn (j, t) =>  | 
|
1428  | 
((lam_fact_prefix ^ Int.toString (j + 1), (Global, Non_Rec_Def)), (Axiom, t)))  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1429  | 
in (facts, lam_facts) end  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1430  | 
|
| 
62126
 
2d187ace2827
avoid generating TFF1 or polymorphic DFG constructs in Vampire or SPASS problems for goals containing schematic type variables
 
blanchet 
parents: 
61940 
diff
changeset
 | 
1431  | 
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate this in Sledgehammer to  | 
| 
 
2d187ace2827
avoid generating TFF1 or polymorphic DFG constructs in Vampire or SPASS problems for goals containing schematic type variables
 
blanchet 
parents: 
61940 
diff
changeset
 | 
1432  | 
prevent the discovery of unreplayable proofs. *)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1433  | 
fun freeze_term t =  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1434  | 
let  | 
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1435  | 
(* Freshness is desirable for completeness, but not for soundness. *)  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1436  | 
fun indexed_name (s, i) = s ^ "_" ^ string_of_int i ^ atp_weak_suffix  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1437  | 
fun freeze (t $ u) = freeze t $ freeze u  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1438  | 
| freeze (Abs (s, T, t)) = Abs (s, T, freeze t)  | 
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1439  | 
| freeze (Var (x, T)) = Free (indexed_name x, T)  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1440  | 
| freeze t = t  | 
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1441  | 
fun freeze_tvar (x, S) = TFree (indexed_name x, S)  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1442  | 
in  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1443  | 
t |> exists_subterm is_Var t ? freeze  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1444  | 
|> exists_type (exists_subtype is_TVar) t  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1445  | 
? map_types (map_type_tvar freeze_tvar)  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52071 
diff
changeset
 | 
1446  | 
end  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1447  | 
|
| 74051 | 1448  | 
fun presimp_prop simp_options ctxt type_enc t =  | 
| 
47713
 
bd0683000a0f
fix handling of atomizable conjectures without a top-level "Trueprop" (e.g. "x == (y::nat)")
 
blanchet 
parents: 
47153 
diff
changeset
 | 
1449  | 
let  | 
| 
74208
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
1450  | 
val t =  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
1451  | 
t |> Envir.beta_eta_contract  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
1452  | 
|> transform_elim_prop  | 
| 
 
1a8d8dd77513
made sure lambda-lifting works well with native let binders in Sledgehammer
 
blanchet 
parents: 
74162 
diff
changeset
 | 
1453  | 
|> Object_Logic.atomize_term ctxt  | 
| 69593 | 1454  | 
val need_trueprop = (fastype_of t = \<^typ>\<open>bool\<close>)  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1455  | 
val is_ho = is_type_enc_full_higher_order type_enc  | 
| 
47713
 
bd0683000a0f
fix handling of atomizable conjectures without a top-level "Trueprop" (e.g. "x == (y::nat)")
 
blanchet 
parents: 
47153 
diff
changeset
 | 
1456  | 
in  | 
| 
 
bd0683000a0f
fix handling of atomizable conjectures without a top-level "Trueprop" (e.g. "x == (y::nat)")
 
blanchet 
parents: 
47153 
diff
changeset
 | 
1457  | 
t |> need_trueprop ? HOLogic.mk_Trueprop  | 
| 
47954
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
1458  | 
|> (if is_ho then unextensionalize_def  | 
| 59632 | 1459  | 
else cong_extensionalize_term ctxt #> abs_extensionalize_term ctxt)  | 
| 74051 | 1460  | 
|> presimplify_term simp_options ctxt  | 
| 
47713
 
bd0683000a0f
fix handling of atomizable conjectures without a top-level "Trueprop" (e.g. "x == (y::nat)")
 
blanchet 
parents: 
47153 
diff
changeset
 | 
1461  | 
|> HOLogic.dest_Trueprop  | 
| 
 
bd0683000a0f
fix handling of atomizable conjectures without a top-level "Trueprop" (e.g. "x == (y::nat)")
 
blanchet 
parents: 
47153 
diff
changeset
 | 
1462  | 
end  | 
| 74379 | 1463  | 
handle TERM _ => \<^Const>\<open>True\<close>  | 
| 43096 | 1464  | 
|
| 55212 | 1465  | 
(* Satallax prefers "=" to "<=>" (for definitions) and Metis (CNF) requires "=" for technical  | 
1466  | 
reasons. *)  | 
|
| 
47905
 
9b6afe0eb69c
cleaner handling of bi-implication for THF output of first-order type encodings
 
blanchet 
parents: 
47810 
diff
changeset
 | 
1467  | 
fun should_use_iff_for_eq CNF _ = false  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1468  | 
| should_use_iff_for_eq (THF _) format = not (is_type_enc_full_higher_order format)  | 
| 
47905
 
9b6afe0eb69c
cleaner handling of bi-implication for THF output of first-order type encodings
 
blanchet 
parents: 
47810 
diff
changeset
 | 
1469  | 
| should_use_iff_for_eq _ _ = true  | 
| 
 
9b6afe0eb69c
cleaner handling of bi-implication for THF output of first-order type encodings
 
blanchet 
parents: 
47810 
diff
changeset
 | 
1470  | 
|
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
1471  | 
fun make_formula ctxt format type_enc iff_for_eq name stature role t =  | 
| 43096 | 1472  | 
let  | 
| 
47905
 
9b6afe0eb69c
cleaner handling of bi-implication for THF output of first-order type encodings
 
blanchet 
parents: 
47810 
diff
changeset
 | 
1473  | 
val iff_for_eq = iff_for_eq andalso should_use_iff_for_eq format type_enc  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1474  | 
val (iformula, atomic_Ts) =  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
1475  | 
iformula_of_prop ctxt type_enc iff_for_eq (SOME (role <> Conjecture)) t []  | 
| 48144 | 1476  | 
|>> close_universally add_iterm_vars  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1477  | 
in  | 
| 57263 | 1478  | 
    {name = name, stature = stature, role = role, iformula = iformula, atomic_types = atomic_Ts}
 | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1479  | 
end  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1480  | 
|
| 
57547
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1481  | 
fun make_fact ctxt format type_enc iff_for_eq ((name, stature), t) =  | 
| 
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1482  | 
(case make_formula ctxt format type_enc iff_for_eq name stature Axiom t of  | 
| 
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1483  | 
    formula as {iformula = AAtom (IConst ((s, _), _, _)), ...} =>
 | 
| 
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1484  | 
if s = tptp_true then NONE else SOME formula  | 
| 
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1485  | 
| formula => SOME formula)  | 
| 
42561
 
23ddc4e3d19c
have properly type-instantiated helper facts (combinators and If)
 
blanchet 
parents: 
42560 
diff
changeset
 | 
1486  | 
|
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1487  | 
fun make_conjecture ctxt format type_enc =  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
1488  | 
map (fn ((name, stature), (role, t)) =>  | 
| 
57547
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1489  | 
let val t' = t |> role = Conjecture ? s_not in  | 
| 
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1490  | 
make_formula ctxt format type_enc true name stature role t'  | 
| 
 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 
blanchet 
parents: 
57541 
diff
changeset
 | 
1491  | 
end)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1492  | 
|
| 
42682
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1493  | 
(** Finite and infinite type inference **)  | 
| 
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1494  | 
|
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1495  | 
fun tvar_footprint thy s ary =  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1496  | 
(case unprefix_and_unascii const_prefix s of  | 
| 54829 | 1497  | 
SOME s =>  | 
1498  | 
let fun tvars_of T = [] |> Term.add_tvarsT T |> map fst in  | 
|
1499  | 
s |> unmangled_invert_const |> robust_const_type thy |> chop_fun ary |> fst |> map tvars_of  | 
|
1500  | 
end  | 
|
1501  | 
| NONE => [])  | 
|
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1502  | 
handle TYPE _ => []  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1503  | 
|
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1504  | 
fun type_arg_cover thy pos s ary =  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1505  | 
if is_tptp_equal s then  | 
| 48186 | 1506  | 
if pos = SOME false then [] else 0 upto ary - 1  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1507  | 
else  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1508  | 
let  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1509  | 
val footprint = tvar_footprint thy s ary  | 
| 69593 | 1510  | 
val eq = (s = \<^const_name>\<open>HOL.eq\<close>)  | 
| 48080 | 1511  | 
fun cover _ [] = []  | 
1512  | 
| cover seen ((i, tvars) :: args) =  | 
|
1513  | 
cover (union (op =) seen tvars) args  | 
|
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1514  | 
|> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1515  | 
? cons i  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1516  | 
in  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1517  | 
if forall null footprint then  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1518  | 
[]  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1519  | 
else  | 
| 74918 | 1520  | 
map_index I footprint  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58477 
diff
changeset
 | 
1521  | 
|> sort (rev_order o list_ord Term_Ord.indexname_ord o apply2 snd)  | 
| 48080 | 1522  | 
|> cover []  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1523  | 
end  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1524  | 
|
| 44399 | 1525  | 
type monotonicity_info =  | 
1526  | 
  {maybe_finite_Ts : typ list,
 | 
|
1527  | 
surely_infinite_Ts : typ list,  | 
|
1528  | 
maybe_nonmono_Ts : typ list}  | 
|
1529  | 
||
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
1530  | 
(* These types witness that the type classes they belong to allow infinite  | 
| 
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
1531  | 
models and hence that any types with these type classes is monotonic. *)  | 
| 69593 | 1532  | 
val known_infinite_types = [\<^typ>\<open>nat\<close>, HOLogic.intT, HOLogic.realT, \<^typ>\<open>nat => bool\<close>]  | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
1533  | 
|
| 46301 | 1534  | 
fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T =  | 
1535  | 
strictness <> Strict andalso is_type_surely_infinite ctxt true cached_Ts T  | 
|
| 
42886
 
208ec29cc013
improved "poly_preds_{bang,query}" by picking up good witnesses for the possible infinity of common type classes and ensuring that "?'a::type" doesn't ruin everything
 
blanchet 
parents: 
42885 
diff
changeset
 | 
1536  | 
|
| 
42682
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1537  | 
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are  | 
| 
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1538  | 
dangerous because their "exhaust" properties can easily lead to unsound ATP  | 
| 
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1539  | 
proofs. On the other hand, all HOL infinite types can be given the same  | 
| 50969 | 1540  | 
models in first-order logic (via Loewenheim-Skolem). *)  | 
| 
42682
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1541  | 
|
| 57396 | 1542  | 
fun should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts, maybe_nonmono_Ts}
 | 
1543  | 
(Nonmono_Types (strictness, _)) T =  | 
|
| 47150 | 1544  | 
let val thy = Proof_Context.theory_of ctxt in  | 
1545  | 
(exists (type_intersect thy T) maybe_nonmono_Ts andalso  | 
|
1546  | 
not (exists (type_instance thy T) surely_infinite_Ts orelse  | 
|
1547  | 
(not (member (type_equiv thy) maybe_finite_Ts T) andalso  | 
|
| 57396 | 1548  | 
is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts T)))  | 
| 47150 | 1549  | 
end  | 
| 57396 | 1550  | 
| should_encode_type _ _ level _ = (level = All_Types orelse level = Undercover_Types)  | 
| 
42682
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1551  | 
|
| 44768 | 1552  | 
fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T =  | 
| 44811 | 1553  | 
should_guard_var () andalso should_encode_type ctxt mono level T  | 
| 44399 | 1554  | 
| should_guard_type _ _ _ _ _ = false  | 
| 
42682
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1555  | 
|
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1556  | 
fun is_maybe_universal_name s =  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1557  | 
String.isPrefix bound_var_prefix s orelse  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1558  | 
String.isPrefix all_bound_var_prefix s  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1559  | 
|
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
1560  | 
fun is_maybe_universal_var (IConst ((s, _), _, _)) = is_maybe_universal_name s  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1561  | 
| is_maybe_universal_var (IVar _) = true  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1562  | 
| is_maybe_universal_var _ = false  | 
| 42836 | 1563  | 
|
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1564  | 
datatype site =  | 
| 
43361
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
1565  | 
Top_Level of bool option |  | 
| 
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
1566  | 
Eq_Arg of bool option |  | 
| 48146 | 1567  | 
Arg of string * int * int |  | 
| 
43361
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
1568  | 
Elsewhere  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
1569  | 
|
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1570  | 
fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1571  | 
| should_tag_with_type ctxt mono (Tags (_, level)) site u T =  | 
| 48146 | 1572  | 
let val thy = Proof_Context.theory_of ctxt in  | 
| 54829 | 1573  | 
(case level of  | 
| 48183 | 1574  | 
Nonmono_Types (_, Non_Uniform) =>  | 
| 48146 | 1575  | 
(case (site, is_maybe_universal_var u) of  | 
| 54829 | 1576  | 
(Eq_Arg pos, true) =>  | 
1577  | 
(pos <> SOME false orelse tag_neg_vars) andalso should_encode_type ctxt mono level T  | 
|
1578  | 
| _ => false)  | 
|
| 48183 | 1579  | 
| Undercover_Types =>  | 
1580  | 
(case (site, is_maybe_universal_var u) of  | 
|
| 54829 | 1581  | 
(Eq_Arg pos, true) => pos <> SOME false  | 
1582  | 
| (Arg (s, j, ary), true) => member (op =) (type_arg_cover thy NONE s ary) j  | 
|
1583  | 
| _ => false)  | 
|
1584  | 
| _ => should_encode_type ctxt mono level T)  | 
|
| 48146 | 1585  | 
end  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1586  | 
| should_tag_with_type _ _ _ _ _ _ = false  | 
| 
42682
 
562046fd8e0c
added type homogenization, whereby all (isomorphic) infinite types are mapped to the same type (to reduce the number of different predicates/TFF-types)
 
blanchet 
parents: 
42680 
diff
changeset
 | 
1587  | 
|
| 
44450
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1588  | 
(** predicators and application operators **)  | 
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
1589  | 
|
| 72914 | 1590  | 
type sym_info = {pred_sym : bool, min_ary : int, max_ary : int, types : typ list, in_conj : bool}
 | 
| 42563 | 1591  | 
|
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1592  | 
fun default_sym_tab_entries type_enc =  | 
| 72914 | 1593  | 
let  | 
1594  | 
fun mk_sym_info pred n =  | 
|
1595  | 
      {pred_sym = pred, min_ary = n, max_ary = n, types = [], in_conj = false}
 | 
|
1596  | 
in  | 
|
| 75010 | 1597  | 
(make_fixed_const \<^const_name>\<open>undefined\<close>, mk_sym_info false 0) ::  | 
| 74552 | 1598  | 
    (map (apsnd (fn {arity, is_predicate} => mk_sym_info is_predicate arity))
 | 
1599  | 
(Symtab.dest tptp_builtins))  | 
|
| 72914 | 1600  | 
|> not (is_type_enc_fool type_enc orelse is_type_enc_full_higher_order type_enc)  | 
1601  | 
? cons (prefixed_predicator_name, mk_sym_info true 1)  | 
|
1602  | 
end  | 
|
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1603  | 
|
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1604  | 
datatype app_op_level =  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1605  | 
Min_App_Op |  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1606  | 
Sufficient_App_Op |  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1607  | 
Sufficient_App_Op_And_Predicator |  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1608  | 
Full_App_Op_And_Predicator  | 
| 46389 | 1609  | 
|
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1610  | 
fun add_iterm_syms_to_sym_table ctxt app_op_level conj_fact =  | 
| 
42558
 
3d9930cb6770
cleaned up "explicit_apply" so that it shares most of its code path with the default mode of operation
 
blanchet 
parents: 
42557 
diff
changeset
 | 
1611  | 
let  | 
| 
46642
 
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
 
blanchet 
parents: 
46639 
diff
changeset
 | 
1612  | 
val thy = Proof_Context.theory_of ctxt  | 
| 44772 | 1613  | 
fun consider_var_ary const_T var_T max_ary =  | 
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43039 
diff
changeset
 | 
1614  | 
let  | 
| 
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43039 
diff
changeset
 | 
1615  | 
fun iter ary T =  | 
| 
74900
 
21ea15129166
fixed TPTP generation of multi-arity expressions
 
desharna 
parents: 
74899 
diff
changeset
 | 
1616  | 
if ary = max_ary orelse type_instance thy var_T T orelse type_instance thy T var_T orelse  | 
| 
 
21ea15129166
fixed TPTP generation of multi-arity expressions
 
desharna 
parents: 
74899 
diff
changeset
 | 
1617  | 
not (can dest_funT T) then  | 
| 
43210
 
7384b771805d
made "explicit_apply"'s smart mode (more) complete
 
blanchet 
parents: 
43207 
diff
changeset
 | 
1618  | 
ary  | 
| 
 
7384b771805d
made "explicit_apply"'s smart mode (more) complete
 
blanchet 
parents: 
43207 
diff
changeset
 | 
1619  | 
else  | 
| 
 
7384b771805d
made "explicit_apply"'s smart mode (more) complete
 
blanchet 
parents: 
43207 
diff
changeset
 | 
1620  | 
iter (ary + 1) (range_type T)  | 
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43039 
diff
changeset
 | 
1621  | 
in iter 0 const_T end  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1622  | 
fun add_universal_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1623  | 
if (app_op_level = Sufficient_App_Op andalso can dest_funT T) orelse  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1624  | 
(app_op_level = Sufficient_App_Op_And_Predicator andalso  | 
| 69593 | 1625  | 
(can dest_funT T orelse T = \<^typ>\<open>bool\<close>)) then  | 
| 
43201
 
0c9bf1a8e0d8
make "smart" mode of "explicit_apply" smarter, by also detecting the other kind of higher-order quantification, namely "bool"s
 
blanchet 
parents: 
43198 
diff
changeset
 | 
1626  | 
let  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1627  | 
val bool_vars' =  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
1628  | 
bool_vars orelse  | 
| 69593 | 1629  | 
(app_op_level = Sufficient_App_Op_And_Predicator andalso body_type T = \<^typ>\<open>bool\<close>)  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1630  | 
          fun repair_min_ary {pred_sym, min_ary, max_ary, types, in_conj} =
 | 
| 
43201
 
0c9bf1a8e0d8
make "smart" mode of "explicit_apply" smarter, by also detecting the other kind of higher-order quantification, namely "bool"s
 
blanchet 
parents: 
43198 
diff
changeset
 | 
1631  | 
            {pred_sym = pred_sym andalso not bool_vars',
 | 
| 44772 | 1632  | 
min_ary = fold (fn T' => consider_var_ary T' T) types min_ary,  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1633  | 
max_ary = max_ary, types = types, in_conj = in_conj}  | 
| 54829 | 1634  | 
val fun_var_Ts' = fun_var_Ts |> can dest_funT T ? insert_type thy I T  | 
| 
43201
 
0c9bf1a8e0d8
make "smart" mode of "explicit_apply" smarter, by also detecting the other kind of higher-order quantification, namely "bool"s
 
blanchet 
parents: 
43198 
diff
changeset
 | 
1635  | 
in  | 
| 57533 | 1636  | 
if bool_vars' = bool_vars andalso fun_var_Ts' = fun_var_Ts then accum  | 
| 54829 | 1637  | 
else ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_ary) sym_tab)  | 
| 
43201
 
0c9bf1a8e0d8
make "smart" mode of "explicit_apply" smarter, by also detecting the other kind of higher-order quantification, namely "bool"s
 
blanchet 
parents: 
43198 
diff
changeset
 | 
1638  | 
end  | 
| 
 
0c9bf1a8e0d8
make "smart" mode of "explicit_apply" smarter, by also detecting the other kind of higher-order quantification, namely "bool"s
 
blanchet 
parents: 
43198 
diff
changeset
 | 
1639  | 
else  | 
| 
 
0c9bf1a8e0d8
make "smart" mode of "explicit_apply" smarter, by also detecting the other kind of higher-order quantification, namely "bool"s
 
blanchet 
parents: 
43198 
diff
changeset
 | 
1640  | 
accum  | 
| 54829 | 1641  | 
fun add_iterm_syms top_level tm (accum as ((bool_vars, fun_var_Ts), sym_tab)) =  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1642  | 
let val (head, args) = strip_iterm_comb tm in  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1643  | 
(case head of  | 
| 54829 | 1644  | 
IConst ((s, _), T, _) =>  | 
| 
74984
 
192876ea202d
proper treatment of $let variables in symbol table in Sledgehammer
 
desharna 
parents: 
74983 
diff
changeset
 | 
1645  | 
if is_maybe_universal_name s orelse String.isPrefix let_bound_var_prefix s then  | 
| 54829 | 1646  | 
add_universal_var T accum  | 
1647  | 
else if String.isPrefix exist_bound_var_prefix s then  | 
|
1648  | 
accum  | 
|
1649  | 
else  | 
|
1650  | 
let val ary = length args in  | 
|
1651  | 
((bool_vars, fun_var_Ts),  | 
|
1652  | 
(case Symtab.lookup sym_tab s of  | 
|
1653  | 
                 SOME {pred_sym, min_ary, max_ary, types, in_conj} =>
 | 
|
1654  | 
let  | 
|
1655  | 
val pred_sym = pred_sym andalso top_level andalso not bool_vars  | 
|
1656  | 
val types' = types |> insert_type thy I T  | 
|
1657  | 
val in_conj = in_conj orelse conj_fact  | 
|
1658  | 
val min_ary =  | 
|
1659  | 
if (app_op_level = Sufficient_App_Op orelse  | 
|
1660  | 
app_op_level = Sufficient_App_Op_And_Predicator)  | 
|
| 57533 | 1661  | 
andalso types' <> types then  | 
| 54829 | 1662  | 
fold (consider_var_ary T) fun_var_Ts min_ary  | 
1663  | 
else  | 
|
1664  | 
min_ary  | 
|
1665  | 
in  | 
|
1666  | 
                   Symtab.update (s, {pred_sym = pred_sym, min_ary = Int.min (ary, min_ary),
 | 
|
1667  | 
max_ary = Int.max (ary, max_ary), types = types', in_conj = in_conj}) sym_tab  | 
|
1668  | 
end  | 
|
1669  | 
| NONE =>  | 
|
1670  | 
let  | 
|
1671  | 
val max_ary =  | 
|
1672  | 
(case unprefix_and_unascii const_prefix s of  | 
|
1673  | 
SOME s =>  | 
|
1674  | 
(if String.isSubstring uncurried_alias_sep s then  | 
|
1675  | 
ary  | 
|
1676  | 
else  | 
|
1677  | 
(case try (ary_of o robust_const_type thy o unmangled_invert_const) s of  | 
|
1678  | 
SOME ary0 => Int.min (ary0, ary)  | 
|
1679  | 
| NONE => ary))  | 
|
1680  | 
| NONE => ary)  | 
|
1681  | 
val pred_sym = top_level andalso max_ary = ary andalso not bool_vars  | 
|
1682  | 
val min_ary =  | 
|
1683  | 
(case app_op_level of  | 
|
1684  | 
Min_App_Op => max_ary  | 
|
1685  | 
| Full_App_Op_And_Predicator => 0  | 
|
1686  | 
| _ => fold (consider_var_ary T) fun_var_Ts max_ary)  | 
|
1687  | 
in  | 
|
1688  | 
                   Symtab.update_new (s, {pred_sym = pred_sym, min_ary = min_ary, max_ary = max_ary,
 | 
|
1689  | 
types = [T], in_conj = conj_fact}) sym_tab  | 
|
1690  | 
end))  | 
|
1691  | 
end  | 
|
1692  | 
| IVar (_, T) => add_universal_var T accum  | 
|
1693  | 
| IAbs ((_, T), tm) => accum |> add_universal_var T |> add_iterm_syms false tm  | 
|
1694  | 
| _ => accum)  | 
|
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1695  | 
|> fold (add_iterm_syms false) args  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1696  | 
end  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1697  | 
in add_iterm_syms end  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1698  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
1699  | 
fun sym_table_of_facts ctxt type_enc app_op_level conjs facts =  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1700  | 
let  | 
| 54829 | 1701  | 
fun add_iterm_syms conj_fact = add_iterm_syms_to_sym_table ctxt app_op_level conj_fact true  | 
1702  | 
fun add_fact_syms conj_fact = ifact_lift (formula_fold NONE (K (add_iterm_syms conj_fact)))  | 
|
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1703  | 
in  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1704  | 
((false, []), Symtab.empty)  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1705  | 
|> fold (add_fact_syms true) conjs  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1706  | 
|> fold (add_fact_syms false) facts  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1707  | 
||> fold Symtab.update (default_sym_tab_entries type_enc)  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1708  | 
end  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1709  | 
|
| 44772 | 1710  | 
fun min_ary_of sym_tab s =  | 
| 54829 | 1711  | 
(case Symtab.lookup sym_tab s of  | 
| 42574 | 1712  | 
    SOME ({min_ary, ...} : sym_info) => min_ary
 | 
| 
42558
 
3d9930cb6770
cleaned up "explicit_apply" so that it shares most of its code path with the default mode of operation
 
blanchet 
parents: 
42557 
diff
changeset
 | 
1713  | 
| NONE =>  | 
| 54829 | 1714  | 
(case unprefix_and_unascii const_prefix s of  | 
| 
42547
 
b5eec0c99528
fixed arity of special constants if "explicit_apply" is set
 
blanchet 
parents: 
42546 
diff
changeset
 | 
1715  | 
SOME s =>  | 
| 
51920
 
16f3b9d4e515
proper unmangling -- the bug is visible when "uncurried_aliases" is enabled with Alt-Ergo or Poly. SPASS
 
blanchet 
parents: 
51717 
diff
changeset
 | 
1716  | 
let val s = s |> unmangled_invert_const in  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
1717  | 
if s = predicator_name then 1  | 
| 
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
1718  | 
else if s = app_op_name then 2  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
1719  | 
else if s = type_guard_name then 1  | 
| 
42557
 
ae0deb39a254
fixed min-arity computation when "explicit_apply" is specified
 
blanchet 
parents: 
42556 
diff
changeset
 | 
1720  | 
else 0  | 
| 
42547
 
b5eec0c99528
fixed arity of special constants if "explicit_apply" is set
 
blanchet 
parents: 
42546 
diff
changeset
 | 
1721  | 
end  | 
| 54829 | 1722  | 
| NONE => 0))  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1723  | 
|
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1724  | 
(* True if the constant ever appears outside of the top-level position in  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1725  | 
literals, or if it appears with different arities (e.g., because of different  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1726  | 
type instantiations). If false, the constant always receives all of its  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1727  | 
arguments and is used as a predicate. *)  | 
| 
42558
 
3d9930cb6770
cleaned up "explicit_apply" so that it shares most of its code path with the default mode of operation
 
blanchet 
parents: 
42557 
diff
changeset
 | 
1728  | 
fun is_pred_sym sym_tab s =  | 
| 54829 | 1729  | 
(case Symtab.lookup sym_tab s of  | 
1730  | 
    SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) => pred_sym andalso min_ary = max_ary
 | 
|
1731  | 
| NONE => false)  | 
|
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1732  | 
|
| 69593 | 1733  | 
val fTrue_iconst = IConst ((const_prefix ^ "fTrue", \<^const_name>\<open>fTrue\<close>), \<^typ>\<open>bool\<close>, [])  | 
| 75010 | 1734  | 
val predicator_iconst = IConst (`make_fixed_const predicator_name, \<^typ>\<open>bool => bool\<close>, [])  | 
| 
42542
 
024920b65ce2
perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
 
blanchet 
parents: 
42541 
diff
changeset
 | 
1735  | 
|
| 48143 | 1736  | 
fun predicatify completish tm =  | 
| 64445 | 1737  | 
if completish > 1 then  | 
| 69593 | 1738  | 
IApp (IApp (IConst (`I tptp_equal, \<^typ>\<open>bool => bool => bool\<close>, []), tm), fTrue_iconst)  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1739  | 
else  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1740  | 
IApp (predicator_iconst, tm)  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1741  | 
|
| 75010 | 1742  | 
val app_op = `make_fixed_const app_op_name  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1743  | 
|
| 43859 | 1744  | 
fun list_app head args = fold (curry (IApp o swap)) args head  | 
| 
42544
 
75cb06eee990
reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
 
blanchet 
parents: 
42543 
diff
changeset
 | 
1745  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1746  | 
fun mk_app_op type_enc head arg =  | 
| 
42544
 
75cb06eee990
reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
 
blanchet 
parents: 
42543 
diff
changeset
 | 
1747  | 
let  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1748  | 
val head_T = ityp_of head  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1749  | 
val (arg_T, res_T) = dest_funT head_T  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1750  | 
val app =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1751  | 
IConst (app_op, head_T --> head_T, [arg_T, res_T])  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1752  | 
|> mangle_type_args_in_iterm type_enc  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1753  | 
in list_app app [head, arg] end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1754  | 
|
| 
52038
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
1755  | 
fun firstorderize_fact thy ctrss type_enc uncurried_aliases completish sym_tab =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1756  | 
let  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
1757  | 
fun do_app arg head = mk_app_op type_enc head arg  | 
| 
51921
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1758  | 
fun list_app_ops (head, args) = fold do_app args head  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1759  | 
fun introduce_app_ops tm =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1760  | 
let val (head, args) = tm |> strip_iterm_comb ||> map introduce_app_ops in  | 
| 54829 | 1761  | 
(case head of  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1762  | 
IConst (name as (s, _), T, T_args) =>  | 
| 
51921
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1763  | 
let  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1764  | 
val min_ary = min_ary_of sym_tab s  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1765  | 
val ary =  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1766  | 
if uncurried_aliases andalso String.isPrefix const_prefix s then  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1767  | 
let  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1768  | 
val ary = length args  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1769  | 
(* In polymorphic native type encodings, it is impossible to  | 
| 
52038
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
1770  | 
declare a fully polymorphic symbol that takes more  | 
| 
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
1771  | 
arguments than its signature (even though such concrete  | 
| 
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
1772  | 
instances, where a type variable is instantiated by a  | 
| 
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
1773  | 
function type, are possible.) *)  | 
| 
51921
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1774  | 
val official_ary =  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1775  | 
if is_type_enc_polymorphic type_enc then  | 
| 54829 | 1776  | 
(case unprefix_and_unascii const_prefix s of  | 
| 
51921
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1777  | 
SOME s' =>  | 
| 54829 | 1778  | 
(case try (ary_of o robust_const_type thy) (invert_const s') of  | 
1779  | 
SOME ary => ary  | 
|
1780  | 
| NONE => min_ary)  | 
|
1781  | 
| NONE => min_ary)  | 
|
| 
51921
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1782  | 
else  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1783  | 
1000000000 (* irrealistically big arity *)  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1784  | 
in Int.min (ary, official_ary) end  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1785  | 
else  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1786  | 
min_ary  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1787  | 
val head =  | 
| 54829 | 1788  | 
if ary = min_ary then head else IConst (aliased_uncurried ary name, T, T_args)  | 
| 
51921
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1789  | 
in  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1790  | 
args |> chop ary |>> list_app head |> list_app_ops  | 
| 
 
bbbacaef19a6
avoid polymorphic uncurried aliases with more arguments than there are arguments in the most general type of a HOL constant, to dodge issues at declaration time
 
blanchet 
parents: 
51920 
diff
changeset
 | 
1791  | 
end  | 
| 74322 | 1792  | 
| IAbs ((name, T), tm) =>  | 
1793  | 
list_app_ops (IAbs ((name, T), introduce_app_ops tm), args)  | 
|
| 54829 | 1794  | 
| _ => list_app_ops (head, args))  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1795  | 
end  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1796  | 
fun introduce_predicators tm =  | 
| 54829 | 1797  | 
(case strip_iterm_comb tm of  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1798  | 
(IConst ((s, _), _, _), _) =>  | 
| 48143 | 1799  | 
if is_pred_sym sym_tab s then tm else predicatify completish tm  | 
| 54829 | 1800  | 
| _ => predicatify completish tm)  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1801  | 
val is_ho = is_type_enc_higher_order type_enc  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1802  | 
val is_full_ho = is_type_enc_full_higher_order type_enc  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1803  | 
val is_fool = is_type_enc_fool type_enc  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1804  | 
val do_iterm =  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1805  | 
(not is_ho ? introduce_app_ops)  | 
| 
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
1806  | 
#> (not (is_full_ho orelse is_fool) ? introduce_predicators)  | 
| 52028 | 1807  | 
#> filter_type_args_in_iterm thy ctrss type_enc  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1808  | 
in update_iformula (formula_map do_iterm) end  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1809  | 
|
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1810  | 
(** Helper facts **)  | 
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1811  | 
|
| 67091 | 1812  | 
val not_ffalse = @{lemma "\<not> fFalse" by (unfold fFalse_def) fast}
 | 
| 
44450
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1813  | 
val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast}
 | 
| 
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1814  | 
|
| 43194 | 1815  | 
(* The Boolean indicates that a fairly sound type encoding is needed. *)  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1816  | 
fun helper_table with_combs =  | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1817  | 
(if with_combs then  | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1818  | 
    [(("COMBI", false), [(Non_Rec_Def, @{thm Meson.COMBI_def})]),
 | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1819  | 
     (("COMBK", false), [(Non_Rec_Def, @{thm Meson.COMBK_def})]),
 | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1820  | 
     (("COMBB", false), [(Non_Rec_Def, @{thm Meson.COMBB_def})]),
 | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1821  | 
     (("COMBC", false), [(Non_Rec_Def, @{thm Meson.COMBC_def})]),
 | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1822  | 
     (("COMBS", false), [(Non_Rec_Def, @{thm Meson.COMBS_def})])]
 | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1823  | 
else  | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1824  | 
[]) @  | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1825  | 
[((predicator_name, false), [(General, not_ffalse), (General, ftrue)]),  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1826  | 
   (("fFalse", false), [(General, not_ffalse)]),
 | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1827  | 
   (("fFalse", true), [(General, @{thm True_or_False})]),
 | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1828  | 
   (("fTrue", false), [(General, ftrue)]),
 | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1829  | 
   (("fTrue", true), [(General, @{thm True_or_False})]),
 | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1830  | 
   (("If", true),
 | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
1831  | 
    [(Non_Rec_Def, @{thm if_True}), (Non_Rec_Def, @{thm if_False}),
 | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1832  | 
     (General, @{thm True_or_False})]),
 | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1833  | 
   (("fNot", false),
 | 
| 43194 | 1834  | 
    @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
 | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1835  | 
fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
1836  | 
|> map (pair Non_Rec_Def)),  | 
| 43194 | 1837  | 
   (("fconj", false),
 | 
| 67091 | 1838  | 
    @{lemma "\<not> P \<or> \<not> Q \<or> fconj P Q" "\<not> fconj P Q \<or> P" "\<not> fconj P Q \<or> Q" by (unfold fconj_def) fast+}
 | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1839  | 
|> map (pair General)),  | 
| 43194 | 1840  | 
   (("fdisj", false),
 | 
| 67091 | 1841  | 
    @{lemma "\<not> P \<or> fdisj P Q" "\<not> Q \<or> fdisj P Q" "\<not> fdisj P Q \<or> P \<or> Q" by (unfold fdisj_def) fast+}
 | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1842  | 
|> map (pair General)),  | 
| 43194 | 1843  | 
   (("fimplies", false),
 | 
| 67091 | 1844  | 
    @{lemma "P \<or> fimplies P Q" "\<not> Q \<or> fimplies P Q" "\<not> fimplies P Q \<or> \<not> P \<or> Q"
 | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1845  | 
by (unfold fimplies_def) fast+}  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1846  | 
|> map (pair General)),  | 
| 43678 | 1847  | 
   (("fequal", true),
 | 
1848  | 
(* This is a lie: Higher-order equality doesn't need a sound type encoding.  | 
|
1849  | 
However, this is done so for backward compatibility: Including the  | 
|
1850  | 
equality helpers by default in Metis breaks a few existing proofs. *)  | 
|
1851  | 
    @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
 | 
|
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1852  | 
fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1853  | 
|> map (pair General)),  | 
| 44003 | 1854  | 
(* Partial characterization of "fAll" and "fEx". A complete characterization  | 
1855  | 
would require the axiom of choice for replay with Metis. *)  | 
|
| 67091 | 1856  | 
   (("fAll", false), [(General, @{lemma "\<not> fAll P \<or> P x" by (auto simp: fAll_def)})]),
 | 
| 74896 | 1857  | 
   (("fEx", false), [(General, @{lemma "\<not> P x \<or> fEx P" by (auto simp: fEx_def)})]),
 | 
| 74899 | 1858  | 
   (("fChoice", true), [(General, @{thm fChoice_iff_Ex})])]
 | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1859  | 
|> map (apsnd (map (apsnd zero_var_indexes)))  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1860  | 
|
| 74899 | 1861  | 
val () =  | 
1862  | 
let  | 
|
1863  | 
fun is_skolemizable \<^Const_>\<open>Ex _ for \<open>Abs _\<close>\<close> = true  | 
|
1864  | 
| is_skolemizable _ = false  | 
|
1865  | 
||
1866  | 
fun check_no_skolemizable_thm thm =  | 
|
1867  | 
if Term.exists_subterm is_skolemizable (Thm.full_prop_of thm) then  | 
|
1868  | 
error "Theorems of the helper table cannot contain skolemizable terms because they don't \  | 
|
1869  | 
\get skolimized in metis."  | 
|
1870  | 
else  | 
|
1871  | 
()  | 
|
1872  | 
in  | 
|
1873  | 
helper_table true  | 
|
1874  | 
|> List.app (fn (_, thms) => List.app (check_no_skolemizable_thm o snd) thms)  | 
|
1875  | 
end  | 
|
1876  | 
||
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1877  | 
fun completish_helper_table with_combs =  | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1878  | 
helper_table with_combs @  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1879  | 
[((predicator_name, true),  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1880  | 
    @{thms True_or_False fTrue_ne_fFalse} |> map (pair General)),
 | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1881  | 
((app_op_name, true),  | 
| 61940 | 1882  | 
    [(General, @{lemma "\<exists>x. \<not> f x = g x \<or> f = g" by blast}),
 | 
1883  | 
     (General, @{lemma "\<exists>p. (p x \<longleftrightarrow> p y) \<longrightarrow> x = y" by blast})]),
 | 
|
| 54829 | 1884  | 
   (("fconj", false), @{thms fconj_table fconj_laws fdisj_laws} |> map (pair Non_Rec_Def)),
 | 
1885  | 
   (("fdisj", false), @{thms fdisj_table fconj_laws fdisj_laws} |> map (pair Non_Rec_Def)),
 | 
|
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1886  | 
   (("fimplies", false),
 | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1887  | 
    @{thms fimplies_table fconj_laws fdisj_laws fimplies_laws}
 | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
1888  | 
|> map (pair Non_Rec_Def)),  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1889  | 
   (("fequal", false),
 | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
1890  | 
    (@{thms fequal_table} |> map (pair Non_Rec_Def)) @
 | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1891  | 
    (@{thms fequal_laws} |> map (pair General))),
 | 
| 54829 | 1892  | 
   (("fAll", false), @{thms fAll_table fComp_law fAll_law fEx_law} |> map (pair Non_Rec_Def)),
 | 
1893  | 
   (("fEx", false), @{thms fEx_table fComp_law fAll_law fEx_law} |> map (pair Non_Rec_Def))]
 | 
|
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1894  | 
|> map (apsnd (map (apsnd zero_var_indexes)))  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1895  | 
|
| 45920 | 1896  | 
fun bound_tvars type_enc sorts Ts =  | 
| 54798 | 1897  | 
(case filter is_TVar Ts of  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1898  | 
[] => I  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
1899  | 
| Ts =>  | 
| 54798 | 1900  | 
((sorts andalso polymorphism_of_type_enc type_enc <> Type_Class_Polymorphic)  | 
1901  | 
? mk_ahorn (Ts  | 
|
1902  | 
|> class_membs_of_types type_enc add_sorts_on_tvar  | 
|
1903  | 
|> map (class_atom type_enc)))  | 
|
| 
48139
 
b755096701ec
generate type classes for polymorphic DFG format (SPASS)
 
blanchet 
parents: 
48138 
diff
changeset
 | 
1904  | 
#> (case type_enc of  | 
| 74890 | 1905  | 
Native (_, _, Type_Class_Polymorphic, _) =>  | 
| 54798 | 1906  | 
mk_atyquant AForall (map (fn TVar (z as (_, S)) =>  | 
| 
62218
 
42202671777c
avoid generating polymorphic SPASS constructs to monomorphic SPASS
 
blanchet 
parents: 
62126 
diff
changeset
 | 
1907  | 
(AType ((tvar_name z, []), []), map (`make_class) (normalize_classes S) )) Ts)  | 
| 74890 | 1908  | 
| Native (_, _, Raw_Polymorphic _, _) =>  | 
| 70931 | 1909  | 
mk_atyquant AForall (map (fn TVar (z as _) => (AType ((tvar_name z, []), []), [])) Ts)  | 
| 
62218
 
42202671777c
avoid generating polymorphic SPASS constructs to monomorphic SPASS
 
blanchet 
parents: 
62126 
diff
changeset
 | 
1910  | 
| _ => mk_aquant AForall (map (fn TVar z => (tvar_name z, NONE)) Ts)))  | 
| 
44404
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
1911  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1912  | 
fun eq_formula type_enc atomic_Ts bounds pred_sym tm1 tm2 =  | 
| 
44404
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
1913  | 
(if pred_sym then AConn (AIff, [AAtom tm1, AAtom tm2])  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
1914  | 
else AAtom (ATerm ((`I tptp_equal, []), [tm1, tm2])))  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1915  | 
|> mk_aquant AForall bounds  | 
| 45377 | 1916  | 
|> close_formula_universally  | 
| 45920 | 1917  | 
|> bound_tvars type_enc true atomic_Ts  | 
| 
44404
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
1918  | 
|
| 46406 | 1919  | 
val helper_rank = default_rank  | 
1920  | 
val min_rank = 9 * helper_rank div 10  | 
|
1921  | 
val max_rank = 4 * min_rank  | 
|
1922  | 
||
1923  | 
fun rank_of_fact_num n j = min_rank + (max_rank - min_rank) * j div n  | 
|
1924  | 
||
| 75010 | 1925  | 
val type_tag = `make_fixed_const type_tag_name  | 
| 
43130
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
1926  | 
|
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1927  | 
fun could_specialize_helpers type_enc =  | 
| 54829 | 1928  | 
not (is_type_enc_polymorphic type_enc) andalso level_of_type_enc type_enc <> No_Types  | 
| 52071 | 1929  | 
|
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1930  | 
fun should_specialize_helper type_enc t =  | 
| 54829 | 1931  | 
could_specialize_helpers type_enc andalso not (null (Term.hidden_polymorphism t))  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
1932  | 
|
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1933  | 
fun add_helper_facts_of_sym ctxt format type_enc lam_trans completish (s, {types, ...} : sym_info) =
 | 
| 54829 | 1934  | 
(case unprefix_and_unascii const_prefix s of  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1935  | 
SOME mangled_s =>  | 
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1936  | 
let  | 
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1937  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1938  | 
val unmangled_s = mangled_s |> unmangled_const_name |> hd  | 
| 48089 | 1939  | 
fun dub needs_sound j k =  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1940  | 
ascii_of unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^  | 
| 
46339
 
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
 
blanchet 
parents: 
46338 
diff
changeset
 | 
1941  | 
(if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^  | 
| 48089 | 1942  | 
(if needs_sound then typed_helper_suffix else untyped_helper_suffix)  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1943  | 
fun specialize_helper t T =  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1944  | 
if unmangled_s = app_op_name then  | 
| 54829 | 1945  | 
let val tyenv = Sign.typ_match thy (alpha_to_beta, domain_type T) Vartab.empty in  | 
| 61770 | 1946  | 
Envir.subst_term_types tyenv t  | 
| 54829 | 1947  | 
end  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1948  | 
else  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1949  | 
specialize_type thy (invert_const unmangled_s, T) t  | 
| 74758 | 1950  | 
fun dub_and_inst needs_sound (j, (status, t)) =  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1951  | 
(if should_specialize_helper type_enc t then  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1952  | 
map_filter (try (specialize_helper t)) types  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1953  | 
else  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
1954  | 
[t])  | 
| 
46339
 
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
 
blanchet 
parents: 
46338 
diff
changeset
 | 
1955  | 
|> tag_list 1  | 
| 48089 | 1956  | 
|> map (fn (k, t) => ((dub needs_sound j k, (Global, status)), t))  | 
| 74758 | 1957  | 
fun make_facts type_enc = map_filter (make_fact ctxt format type_enc false)  | 
| 48089 | 1958  | 
val sound = is_type_enc_sound type_enc  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
1959  | 
val could_specialize = could_specialize_helpers type_enc  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1960  | 
val with_combs = lam_trans <> opaque_combsN  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1961  | 
in  | 
| 48089 | 1962  | 
fold (fn ((helper_s, needs_sound), ths) =>  | 
| 74758 | 1963  | 
let  | 
| 74890 | 1964  | 
fun map_syntax f (Native (order, With_FOOL syntax, polymorphism, type_level)) =  | 
1965  | 
Native (order, With_FOOL (f syntax), polymorphism, type_level)  | 
|
1966  | 
| map_syntax _ type_enc = type_enc  | 
|
1967  | 
val remove_ite_syntax = map_syntax  | 
|
1968  | 
            (fn {with_let, ...} => {with_ite = false, with_let = with_let})
 | 
|
| 74758 | 1969  | 
in  | 
1970  | 
if (needs_sound andalso not sound) orelse  | 
|
1971  | 
(helper_s <> unmangled_s andalso  | 
|
1972  | 
(completish < 3 orelse could_specialize)) then  | 
|
1973  | 
I  | 
|
1974  | 
else  | 
|
1975  | 
ths  | 
|
1976  | 
|> map_index (apfst (curry op+ 1))  | 
|
1977  | 
|> maps (dub_and_inst needs_sound o apsnd (apsnd Thm.prop_of))  | 
|
1978  | 
|> make_facts ((helper_s = "If" ? remove_ite_syntax) type_enc)  | 
|
1979  | 
|> union (op = o apply2 #iformula)  | 
|
1980  | 
end)  | 
|
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1981  | 
((if completish >= 3 then completish_helper_table else helper_table) with_combs)  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1982  | 
end  | 
| 54829 | 1983  | 
| NONE => I)  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1984  | 
fun helper_facts_of_sym_table ctxt format type_enc lam_trans completish sym_tab =  | 
| 
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
1985  | 
Symtab.fold_rev (add_helper_facts_of_sym ctxt format type_enc lam_trans completish) sym_tab []  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
1986  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1987  | 
(***************************************************************)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1988  | 
(* Type Classes Present in the Axiom or Conjecture Clauses *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1989  | 
(***************************************************************)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1990  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1991  | 
fun set_insert (x, s) = Symtab.update (x, ()) s  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1992  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1993  | 
fun add_classes (cls, cset) = List.foldl set_insert cset (flat cls)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1994  | 
|
| 43093 | 1995  | 
fun classes_of_terms get_Ts =  | 
| 43121 | 1996  | 
map (map snd o get_Ts)  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1997  | 
#> List.foldl add_classes Symtab.empty #> Symtab.delete_safe class_of_types  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
1998  | 
#> Symtab.keys  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1999  | 
|
| 44121 | 2000  | 
val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees  | 
2001  | 
val tvar_classes_of_terms = classes_of_terms Misc_Legacy.term_tvars  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
2002  | 
|
| 52026 | 2003  | 
fun fold_type_ctrs f (Type (s, Ts)) x = fold (fold_type_ctrs f) Ts (f (s, x))  | 
2004  | 
| fold_type_ctrs _ _ x = x  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
2005  | 
|
| 
43907
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43906 
diff
changeset
 | 
2006  | 
(* Type constructors used to instantiate overloaded constants are the only ones  | 
| 
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43906 
diff
changeset
 | 
2007  | 
needed. *)  | 
| 52026 | 2008  | 
fun add_type_ctrs_in_term thy =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
2009  | 
let  | 
| 69593 | 2010  | 
fun add (Const (\<^const_name>\<open>Meson.skolem\<close>, _) $ _) = I  | 
| 43181 | 2011  | 
| add (t $ u) = add t #> add u  | 
| 
44738
 
1b333e4173a2
drop more type arguments soundly, when they can be deduced from the arg types
 
blanchet 
parents: 
44634 
diff
changeset
 | 
2012  | 
| add (Const x) =  | 
| 52026 | 2013  | 
x |> robust_const_type_args thy |> fold (fold_type_ctrs set_insert)  | 
| 43181 | 2014  | 
| add (Abs (_, _, u)) = add u  | 
2015  | 
| add _ = I  | 
|
2016  | 
in add end  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
2017  | 
|
| 52026 | 2018  | 
fun type_ctrs_of_terms thy ts =  | 
2019  | 
Symtab.keys (fold (add_type_ctrs_in_term thy) ts Symtab.empty)  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
2020  | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2021  | 
fun trans_lams_of_string ctxt type_enc lam_trans =  | 
| 45514 | 2022  | 
if lam_trans = no_lamsN then  | 
2023  | 
rpair []  | 
|
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
2024  | 
else if lam_trans = opaque_liftingN then  | 
| 45514 | 2025  | 
lift_lams ctxt type_enc ##> K []  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
2026  | 
else if lam_trans = liftingN then  | 
| 45514 | 2027  | 
lift_lams ctxt type_enc  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
2028  | 
else if lam_trans = opaque_combsN orelse lam_trans = combsN then  | 
| 72921 | 2029  | 
map (introduce_combinators ctxt type_enc) #> rpair []  | 
| 46365 | 2030  | 
else if lam_trans = combs_and_liftingN then  | 
2031  | 
lift_lams_part_1 ctxt type_enc  | 
|
| 72921 | 2032  | 
##> maps (fn t => [t, introduce_combinators ctxt type_enc (intentionalize_def t)])  | 
2033  | 
#> lift_lams_part_2 ctxt type_enc  | 
|
| 46368 | 2034  | 
else if lam_trans = combs_or_liftingN then  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
2035  | 
lift_lams_part_1 ctxt type_enc  | 
| 69593 | 2036  | 
##> map (fn t => (case head_of (strip_qnt_body \<^const_name>\<open>All\<close> t) of  | 
2037  | 
\<^term>\<open>(=) ::bool => bool => bool\<close> => t  | 
|
| 72921 | 2038  | 
| _ => introduce_combinators ctxt type_enc (intentionalize_def t)))  | 
2039  | 
#> lift_lams_part_2 ctxt type_enc  | 
|
| 45514 | 2040  | 
else if lam_trans = keep_lamsN then  | 
2041  | 
map (Envir.eta_contract) #> rpair []  | 
|
2042  | 
else  | 
|
| 63692 | 2043  | 
    error ("Unknown lambda translation scheme: " ^ quote lam_trans)
 | 
| 45514 | 2044  | 
|
| 47981 | 2045  | 
val pull_and_reorder_definitions =  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2046  | 
let  | 
| 
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2047  | 
fun add_consts (IApp (t, u)) = fold add_consts [t, u]  | 
| 
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2048  | 
| add_consts (IAbs (_, t)) = add_consts t  | 
| 
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2049  | 
| add_consts (IConst (name, _, _)) = insert (op =) name  | 
| 
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2050  | 
| add_consts (IVar _) = I  | 
| 47981 | 2051  | 
    fun consts_of_hs l_or_r ({iformula, ...} : ifact) =
 | 
| 54829 | 2052  | 
(case iformula of  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2053  | 
AAtom (IApp (IApp (IConst _, t), u)) => add_consts (l_or_r (t, u)) []  | 
| 54829 | 2054  | 
| _ => [])  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2055  | 
(* Quadratic, but usually OK. *)  | 
| 47981 | 2056  | 
fun reorder [] [] = []  | 
2057  | 
| reorder (fact :: skipped) [] =  | 
|
2058  | 
fact :: reorder [] skipped (* break cycle *)  | 
|
2059  | 
| reorder skipped (fact :: facts) =  | 
|
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2060  | 
let val rhs_consts = consts_of_hs snd fact in  | 
| 
48096
 
60a09522c65e
prevent an "Empty" exception (e.g. with Satallax, "mono_native")
 
blanchet 
parents: 
48095 
diff
changeset
 | 
2061  | 
if exists (exists (exists (member (op =) rhs_consts)  | 
| 
 
60a09522c65e
prevent an "Empty" exception (e.g. with Satallax, "mono_native")
 
blanchet 
parents: 
48095 
diff
changeset
 | 
2062  | 
o consts_of_hs fst)) [skipped, facts] then  | 
| 47981 | 2063  | 
reorder (fact :: skipped) facts  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2064  | 
else  | 
| 47981 | 2065  | 
fact :: reorder [] (facts @ skipped)  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2066  | 
end  | 
| 47981 | 2067  | 
in List.partition (curry (op =) Definition o #role) #>> reorder [] #> op @ end  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2068  | 
|
| 74379 | 2069  | 
fun s_not_prop \<^Const_>\<open>Trueprop for t\<close> = \<^Const>\<open>Trueprop for \<open>s_not t\<close>\<close>  | 
2070  | 
| s_not_prop \<^Const_>\<open>Pure.imp for t \<^prop>\<open>False\<close>\<close> = t  | 
|
2071  | 
| s_not_prop t = \<^Const>\<open>Pure.imp for t \<^prop>\<open>False\<close>\<close>  | 
|
| 
54768
 
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
 
blanchet 
parents: 
54757 
diff
changeset
 | 
2072  | 
|
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
2073  | 
fun translate_formulas simp_options ctxt prem_role format type_enc lam_trans presimp hyp_ts concl_t  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
2074  | 
facts =  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2075  | 
let  | 
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2076  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2077  | 
val trans_lams = trans_lams_of_string ctxt type_enc lam_trans  | 
| 43861 | 2078  | 
val fact_ts = facts |> map snd  | 
| 54829 | 2079  | 
(* Remove existing facts from the conjecture, as this can dramatically boost an ATP's  | 
2080  | 
performance (for some reason). *)  | 
|
| 69593 | 2081  | 
val hyp_ts = hyp_ts |> map (fn t => if member (op aconv) fact_ts t then \<^prop>\<open>True\<close> else t)  | 
| 
62126
 
2d187ace2827
avoid generating TFF1 or polymorphic DFG constructs in Vampire or SPASS problems for goals containing schematic type variables
 
blanchet 
parents: 
61940 
diff
changeset
 | 
2082  | 
|
| 
 
2d187ace2827
avoid generating TFF1 or polymorphic DFG constructs in Vampire or SPASS problems for goals containing schematic type variables
 
blanchet 
parents: 
61940 
diff
changeset
 | 
2083  | 
val hyp_ts = map freeze_term hyp_ts;  | 
| 
 
2d187ace2827
avoid generating TFF1 or polymorphic DFG constructs in Vampire or SPASS problems for goals containing schematic type variables
 
blanchet 
parents: 
61940 
diff
changeset
 | 
2084  | 
val concl_t = freeze_term concl_t;  | 
| 75006 | 2085  | 
val maybe_presimp_prop = presimp ? presimp_prop simp_options ctxt type_enc  | 
| 
62126
 
2d187ace2827
avoid generating TFF1 or polymorphic DFG constructs in Vampire or SPASS problems for goals containing schematic type variables
 
blanchet 
parents: 
61940 
diff
changeset
 | 
2086  | 
|
| 75006 | 2087  | 
val facts = facts |> map (apsnd (pair Axiom o maybe_presimp_prop))  | 
| 
43864
 
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
 
blanchet 
parents: 
43863 
diff
changeset
 | 
2088  | 
val conjs =  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2089  | 
map (pair prem_role) hyp_ts @ [(Conjecture, s_not_prop concl_t)]  | 
| 75006 | 2090  | 
|> map_index (map_prod (rpair (Local, General) o string_of_int) (apsnd maybe_presimp_prop))  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
2091  | 
val ((conjs, facts), lam_facts) =  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
2092  | 
(conjs, facts)  | 
| 45514 | 2093  | 
|> (if lam_trans = no_lamsN then  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
2094  | 
rpair []  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
2095  | 
else  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
2096  | 
op @  | 
| 45514 | 2097  | 
#> preprocess_abstractions_in_terms trans_lams  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
2098  | 
#>> chop (length conjs))  | 
| 47981 | 2099  | 
val conjs =  | 
2100  | 
conjs |> make_conjecture ctxt format type_enc  | 
|
2101  | 
|> pull_and_reorder_definitions  | 
|
2102  | 
val facts =  | 
|
| 57263 | 2103  | 
facts |> map_filter (fn (name, (_, t)) => make_fact ctxt format type_enc true (name, t))  | 
| 47981 | 2104  | 
|> pull_and_reorder_definitions  | 
| 57541 | 2105  | 
val lifted = lam_facts |> map (extract_lambda_def dest_Const o snd o snd)  | 
| 54829 | 2106  | 
val lam_facts = lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd)  | 
| 43861 | 2107  | 
val all_ts = concl_t :: hyp_ts @ fact_ts  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2108  | 
val subs = tfree_classes_of_terms all_ts  | 
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2109  | 
val supers = tvar_classes_of_terms all_ts  | 
| 52026 | 2110  | 
val tycons = type_ctrs_of_terms thy all_ts  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2111  | 
val (supers, tcon_clauses) =  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
2112  | 
if level_of_type_enc type_enc = No_Types then ([], [])  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2113  | 
else make_tcon_clauses thy tycons supers  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2114  | 
val subclass_pairs = make_subclass_pairs thy subs supers  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2115  | 
in  | 
| 57263 | 2116  | 
(union (op =) subs supers, conjs, facts @ lam_facts, subclass_pairs, tcon_clauses, lifted)  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2117  | 
end  | 
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2118  | 
|
| 75010 | 2119  | 
val type_guard = `make_fixed_const type_guard_name  | 
| 
43130
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
2120  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2121  | 
fun type_guard_iterm type_enc T tm =  | 
| 69593 | 2122  | 
IApp (IConst (type_guard, T --> \<^typ>\<open>bool\<close>, [T])  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2123  | 
|> mangle_type_args_in_iterm type_enc, tm)  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2124  | 
|
| 43421 | 2125  | 
fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2126  | 
| is_var_positively_naked_in_term name _ (ATerm (((s, _), _), tms)) accum =  | 
| 54829 | 2127  | 
accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm ((name, []), [])))  | 
| 43692 | 2128  | 
| is_var_positively_naked_in_term _ _ _ _ = true  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
2129  | 
|
| 48186 | 2130  | 
fun is_var_undercover_in_term thy name pos tm accum =  | 
2131  | 
accum orelse  | 
|
| 44811 | 2132  | 
let  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2133  | 
val var = ATerm ((name, []), [])  | 
| 48095 | 2134  | 
fun is_undercover (ATerm (_, [])) = false  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2135  | 
| is_undercover (ATerm (((s, _), _), tms)) =  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
2136  | 
let  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
2137  | 
val ary = length tms  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2138  | 
val cover = type_arg_cover thy pos s ary  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
2139  | 
in  | 
| 74918 | 2140  | 
exists (fn (j, tm) => tm = var andalso member (op =) cover j) (map_index I tms) orelse  | 
| 48095 | 2141  | 
exists is_undercover tms  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
2142  | 
end  | 
| 48095 | 2143  | 
| is_undercover _ = true  | 
2144  | 
in is_undercover tm end  | 
|
| 44811 | 2145  | 
|
| 
48087
 
94835838ed2c
removed micro-optimization whose justification I can't recall
 
blanchet 
parents: 
48081 
diff
changeset
 | 
2146  | 
fun should_guard_var_in_formula thy level pos phi (SOME true) name =  | 
| 48183 | 2147  | 
(case level of  | 
| 54829 | 2148  | 
All_Types => true  | 
2149  | 
| Undercover_Types => formula_fold pos (is_var_undercover_in_term thy name) phi false  | 
|
2150  | 
| Nonmono_Types (_, Uniform) => true  | 
|
2151  | 
| Nonmono_Types (_, Non_Uniform) =>  | 
|
2152  | 
formula_fold pos (is_var_positively_naked_in_term name) phi false  | 
|
2153  | 
| _ => false)  | 
|
| 
48087
 
94835838ed2c
removed micro-optimization whose justification I can't recall
 
blanchet 
parents: 
48081 
diff
changeset
 | 
2154  | 
| should_guard_var_in_formula _ _ _ _ _ _ = true  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
2155  | 
|
| 
48087
 
94835838ed2c
removed micro-optimization whose justification I can't recall
 
blanchet 
parents: 
48081 
diff
changeset
 | 
2156  | 
fun always_guard_var_in_formula _ _ _ _ _ _ = true  | 
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2157  | 
|
| 
44405
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
2158  | 
fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false  | 
| 44768 | 2159  | 
| should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T =  | 
| 57396 | 2160  | 
not (is_type_level_uniform level) andalso should_encode_type ctxt mono level T  | 
| 
44405
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
2161  | 
| should_generate_tag_bound_decl _ _ _ _ _ = false  | 
| 
44404
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
2162  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2163  | 
fun mk_aterm type_enc name T_args args =  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2164  | 
let val (ty_args, tm_args) = process_type_args type_enc T_args in  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2165  | 
ATerm ((name, ty_args), tm_args @ args)  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2166  | 
end  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
2167  | 
|
| 72174 | 2168  | 
fun do_bound_type type_enc =  | 
| 54820 | 2169  | 
(case type_enc of  | 
| 72174 | 2170  | 
Native _ => native_atp_type_of_typ type_enc false 0 #> SOME  | 
| 54820 | 2171  | 
| _ => K NONE)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2172  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2173  | 
fun tag_with_type ctxt mono type_enc pos T tm =  | 
| 43859 | 2174  | 
IConst (type_tag, T --> T, [T])  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2175  | 
|> mangle_type_args_in_iterm type_enc  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
2176  | 
|> atp_term_of_iterm ctxt mono type_enc pos  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2177  | 
|> (fn ATerm ((s, tys), tms) => ATerm ((s, tys), tms @ [tm])  | 
| 43692 | 2178  | 
| _ => raise Fail "unexpected lambda-abstraction")  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
2179  | 
and atp_term_of_iterm ctxt mono type_enc pos =  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2180  | 
let  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
2181  | 
fun term site u =  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42956 
diff
changeset
 | 
2182  | 
let  | 
| 43859 | 2183  | 
val (head, args) = strip_iterm_comb u  | 
| 54829 | 2184  | 
val pos = (case site of Top_Level pos => pos | Eq_Arg pos => pos | _ => NONE)  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2185  | 
val T = ityp_of u  | 
| 43677 | 2186  | 
val t =  | 
| 54829 | 2187  | 
(case head of  | 
| 43859 | 2188  | 
IConst (name as (s, _), _, T_args) =>  | 
| 
47153
 
4d4f2721b3ef
fixed eta-extension of higher-order quantifiers in THF output
 
blanchet 
parents: 
47151 
diff
changeset
 | 
2189  | 
let  | 
| 48146 | 2190  | 
val ary = length args  | 
| 54829 | 2191  | 
fun arg_site j = if is_tptp_equal s then Eq_Arg pos else Arg (s, j, ary)  | 
| 48146 | 2192  | 
in  | 
| 74918 | 2193  | 
map_index (fn (j, arg) => term (arg_site j) arg) args  | 
| 48146 | 2194  | 
|> mk_aterm type_enc name T_args  | 
2195  | 
end  | 
|
| 43859 | 2196  | 
| IVar (name, _) =>  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2197  | 
map (term Elsewhere) args |> mk_aterm type_enc name []  | 
| 43859 | 2198  | 
| IAbs ((name, T), tm) =>  | 
| 
74073
 
6c8473b4f518
fixed TFX generation when universal quantifier is used as term
 
desharna 
parents: 
74051 
diff
changeset
 | 
2199  | 
if is_type_enc_higher_order type_enc orelse is_type_enc_fool type_enc then  | 
| 
69724
 
2444c8b85aac
fixed me -- indeed this was wrong, as demonstrated by the predicate-free HO output (e.g. ehoh with keep_lams)
 
blanchet 
parents: 
69718 
diff
changeset
 | 
2200  | 
AAbs (((name, native_atp_type_of_typ type_enc false 0 T), term Elsewhere tm),  | 
| 
 
2444c8b85aac
fixed me -- indeed this was wrong, as demonstrated by the predicate-free HO output (e.g. ehoh with keep_lams)
 
blanchet 
parents: 
69718 
diff
changeset
 | 
2201  | 
map (term Elsewhere) args)  | 
| 
46818
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
2202  | 
else  | 
| 
 
2a28e66e2e4c
ensure no abstractions leak through after lambda-lifting (for formulas with higher-order occurrences of quantifiers)
 
blanchet 
parents: 
46711 
diff
changeset
 | 
2203  | 
raise Fail "unexpected lambda-abstraction"  | 
| 54829 | 2204  | 
| IApp _ => raise Fail "impossible \"IApp\"")  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2205  | 
val tag = should_tag_with_type ctxt mono type_enc site u T  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2206  | 
in t |> tag ? tag_with_type ctxt mono type_enc pos T end  | 
| 
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: 
47905 
diff
changeset
 | 
2207  | 
in term (Top_Level pos) end  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2208  | 
and formula_of_iformula ctxt mono type_enc should_guard_var =  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2209  | 
let  | 
| 44811 | 2210  | 
val thy = Proof_Context.theory_of ctxt  | 
2211  | 
val level = level_of_type_enc type_enc  | 
|
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
2212  | 
val do_term = atp_term_of_iterm ctxt mono type_enc  | 
| 
42878
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2213  | 
fun do_out_of_bound_type pos phi universal (name, T) =  | 
| 44399 | 2214  | 
if should_guard_type ctxt mono type_enc  | 
| 
48087
 
94835838ed2c
removed micro-optimization whose justification I can't recall
 
blanchet 
parents: 
48081 
diff
changeset
 | 
2215  | 
(fn () => should_guard_var thy level pos phi universal name) T then  | 
| 43859 | 2216  | 
IVar (name, T)  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2217  | 
|> type_guard_iterm type_enc T  | 
| 
43361
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
2218  | 
|> do_term pos |> AAtom |> SOME  | 
| 
44405
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
2219  | 
else if should_generate_tag_bound_decl ctxt mono type_enc universal T then  | 
| 
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
2220  | 
let  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2221  | 
val var = ATerm ((name, []), [])  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2222  | 
val tagged_var = tag_with_type ctxt mono type_enc pos T var  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2223  | 
in SOME (AAtom (ATerm ((`I tptp_equal, []), [tagged_var, var]))) end  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2224  | 
else  | 
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2225  | 
NONE  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2226  | 
fun do_formula pos (ATyQuant (q, xs, phi)) =  | 
| 54820 | 2227  | 
ATyQuant (q, map (apfst (native_atp_type_of_typ type_enc false 0)) xs, do_formula pos phi)  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2228  | 
| do_formula pos (AQuant (q, xs, phi)) =  | 
| 
42878
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2229  | 
let  | 
| 
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2230  | 
val phi = phi |> do_formula pos  | 
| 
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2231  | 
val universal = Option.map (q = AExists ? not) pos  | 
| 
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2232  | 
in  | 
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2233  | 
AQuant (q, xs |> map (apsnd (fn NONE => NONE  | 
| 72174 | 2234  | 
| SOME T => do_bound_type type_enc T)),  | 
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2235  | 
(if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2236  | 
(map_filter  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2237  | 
(fn (_, NONE) => NONE  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2238  | 
| (s, SOME T) =>  | 
| 
42878
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2239  | 
do_out_of_bound_type pos phi universal (s, T))  | 
| 
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2240  | 
xs)  | 
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2241  | 
phi)  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2242  | 
end  | 
| 
42878
 
85ac4c12a4b7
slightly fewer type predicates introduced in the lightweight encoding, based on the observation that only universal positive equalities are dangerous
 
blanchet 
parents: 
42855 
diff
changeset
 | 
2243  | 
| do_formula pos (AConn conn) = aconn_map pos do_formula conn  | 
| 
43361
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
2244  | 
| do_formula pos (AAtom tm) = AAtom (do_term pos tm)  | 
| 
43493
 
bdb11c68f142
generate type predicates for existentials/skolems, otherwise some problems might not be provable
 
blanchet 
parents: 
43423 
diff
changeset
 | 
2245  | 
in do_formula end  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2246  | 
|
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2247  | 
fun string_of_status General = ""  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2248  | 
| string_of_status Induction = inductionN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2249  | 
| string_of_status Intro = introN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2250  | 
| string_of_status Inductive = inductiveN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2251  | 
| string_of_status Elim = elimN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2252  | 
| string_of_status Simp = simpN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2253  | 
| string_of_status Non_Rec_Def = non_rec_defN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2254  | 
| string_of_status Rec_Def = rec_defN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2255  | 
|
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2256  | 
(* Each fact is given a unique fact number to avoid name clashes (e.g., because  | 
| 52032 | 2257  | 
of monomorphization). The TPTP forbids name clashes, and some of the remote  | 
2258  | 
provers might care. *)  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2259  | 
fun line_of_fact ctxt generate_info prefix encode alt freshen pos mono type_enc rank  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2260  | 
        (j, {name, stature = (_, status), role, iformula, atomic_types}) =
 | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2261  | 
Formula ((prefix ^ (if freshen then string_of_int j ^ "_" else "") ^  | 
| 50758 | 2262  | 
encode name, alt name),  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2263  | 
role,  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2264  | 
iformula  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2265  | 
|> formula_of_iformula ctxt mono type_enc  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2266  | 
should_guard_var_in_formula (if pos then SOME true else NONE)  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2267  | 
|> close_formula_universally  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
2268  | 
|> bound_tvars type_enc true atomic_types,  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2269  | 
NONE, isabelle_info generate_info (string_of_status status) (rank j))  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2270  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2271  | 
fun lines_of_subclass generate_info type_enc sub super =  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2272  | 
Formula ((subclass_prefix ^ ascii_of sub ^ "___" ^ ascii_of super, ""), Axiom,  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2273  | 
AConn (AImplies,  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2274  | 
[sub, super] |> map (fn s => class_atom type_enc (s, tvar_a)))  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2275  | 
|> bound_tvars type_enc false [tvar_a],  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2276  | 
NONE, isabelle_info generate_info inductiveN helper_rank)  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2277  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2278  | 
fun lines_of_subclass_pair generate_info type_enc (sub, supers) =  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2279  | 
if polymorphism_of_type_enc type_enc = Type_Class_Polymorphic then  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2280  | 
[Class_Decl (class_decl_prefix ^ ascii_of sub, `make_class sub, map (`make_class) supers)]  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2281  | 
else  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2282  | 
map (lines_of_subclass generate_info type_enc sub) supers  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2283  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2284  | 
fun line_of_tcon_clause generate_info type_enc (name, prems, (cl, T)) =  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2285  | 
if polymorphism_of_type_enc type_enc = Type_Class_Polymorphic then  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2286  | 
Class_Memb (class_memb_prefix ^ name,  | 
| 54820 | 2287  | 
map (fn (cls, T) => (T |> dest_TVar |> tvar_name, map (`make_class) cls)) prems,  | 
2288  | 
native_atp_type_of_typ type_enc false 0 T, `make_class cl)  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2289  | 
else  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2290  | 
Formula ((tcon_clause_prefix ^ name, ""), Axiom,  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2291  | 
mk_ahorn (maps (class_atoms type_enc) prems)  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2292  | 
(class_atom type_enc (cl, T))  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2293  | 
|> bound_tvars type_enc true (snd (dest_Type T)),  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2294  | 
NONE, isabelle_info generate_info inductiveN helper_rank)  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2295  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2296  | 
fun line_of_conjecture ctxt mono type_enc ({name, role, iformula, atomic_types, ...} : ifact) =
 | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2297  | 
Formula ((conjecture_prefix ^ name, ""), role,  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
2298  | 
iformula  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2299  | 
|> formula_of_iformula ctxt mono type_enc should_guard_var_in_formula (SOME false)  | 
| 45377 | 2300  | 
|> close_formula_universally  | 
| 46406 | 2301  | 
|> bound_tvars type_enc true atomic_types, NONE, [])  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2302  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2303  | 
fun lines_of_free_types type_enc (facts : ifact list) =  | 
| 
48185
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2304  | 
if is_type_enc_polymorphic type_enc then  | 
| 
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2305  | 
let  | 
| 54798 | 2306  | 
val type_classes = (polymorphism_of_type_enc type_enc = Type_Class_Polymorphic)  | 
| 74918 | 2307  | 
fun line (j, (cl, T)) =  | 
| 
48185
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2308  | 
if type_classes then  | 
| 
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2309  | 
Class_Memb (class_memb_prefix ^ string_of_int j, [],  | 
| 54820 | 2310  | 
native_atp_type_of_typ type_enc false 0 T, `make_class cl)  | 
| 
48185
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2311  | 
else  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2312  | 
Formula ((tfree_clause_prefix ^ string_of_int j, ""), Hypothesis,  | 
| 54820 | 2313  | 
class_atom type_enc (cl, T), NONE, [])  | 
| 
48185
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2314  | 
val membs =  | 
| 
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2315  | 
fold (union (op =)) (map #atomic_types facts) []  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2316  | 
|> class_membs_of_types type_enc add_sorts_on_tfree  | 
| 54798 | 2317  | 
in  | 
| 74918 | 2318  | 
map_index line membs  | 
| 54798 | 2319  | 
end  | 
| 
48185
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2320  | 
else  | 
| 
 
086d9bb80d46
don't generate any type class axioms for free types for monomorphic encodings
 
blanchet 
parents: 
48183 
diff
changeset
 | 
2321  | 
[]  | 
| 
42573
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2322  | 
|
| 
 
744215c3e90c
got rid of one "sym_table" in "prepare_atp_problem" now that proxies are always handled first, and tuned accordingly
 
blanchet 
parents: 
42572 
diff
changeset
 | 
2323  | 
(** Symbol declarations **)  | 
| 
42544
 
75cb06eee990
reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
 
blanchet 
parents: 
42543 
diff
changeset
 | 
2324  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2325  | 
fun decl_line_of_class phantoms s =  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2326  | 
let val name as (s, _) = `make_class s in  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2327  | 
Sym_Decl (sym_decl_prefix ^ s, name,  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2328  | 
APi ([tvar_a_name],  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2329  | 
if phantoms = Without_Phantom_Type_Vars then  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2330  | 
AFun (a_itself_atype, bool_atype)  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2331  | 
else  | 
| 
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2332  | 
bool_atype))  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2333  | 
end  | 
| 
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2334  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2335  | 
fun decl_lines_of_classes type_enc =  | 
| 54829 | 2336  | 
(case type_enc of  | 
| 74890 | 2337  | 
Native (_, _, Raw_Polymorphic phantoms, _) => map (decl_line_of_class phantoms)  | 
| 54829 | 2338  | 
| _ => K [])  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2339  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2340  | 
fun sym_decl_table_of_facts thy type_enc sym_tab (conjs, facts, extra_tms) =  | 
| 42574 | 2341  | 
let  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2342  | 
fun add_iterm_syms tm =  | 
| 43859 | 2343  | 
let val (head, args) = strip_iterm_comb tm in  | 
| 42574 | 2344  | 
(case head of  | 
| 54829 | 2345  | 
IConst ((s, s'), T, T_args) =>  | 
2346  | 
let  | 
|
2347  | 
val (pred_sym, in_conj) =  | 
|
2348  | 
(case Symtab.lookup sym_tab s of  | 
|
2349  | 
                SOME ({pred_sym, in_conj, ...} : sym_info) => (pred_sym, in_conj)
 | 
|
2350  | 
| NONE => (false, false))  | 
|
2351  | 
val decl_sym =  | 
|
| 
74662
 
44585660f39a
do not declare $let-bound variables in TPTP output
 
desharna 
parents: 
74552 
diff
changeset
 | 
2352  | 
(case type_enc of Guards _ => not pred_sym | _ => true) andalso  | 
| 
 
44585660f39a
do not declare $let-bound variables in TPTP output
 
desharna 
parents: 
74552 
diff
changeset
 | 
2353  | 
not (String.isPrefix let_bound_var_prefix s) andalso  | 
| 
 
44585660f39a
do not declare $let-bound variables in TPTP output
 
desharna 
parents: 
74552 
diff
changeset
 | 
2354  | 
is_tptp_user_symbol s  | 
| 54829 | 2355  | 
in  | 
2356  | 
if decl_sym then  | 
|
2357  | 
Symtab.map_default (s, [])  | 
|
2358  | 
(insert_type thy #3 (s', T_args, T, pred_sym, length args, in_conj))  | 
|
2359  | 
else  | 
|
2360  | 
I  | 
|
2361  | 
end  | 
|
2362  | 
| IAbs (_, tm) => add_iterm_syms tm  | 
|
2363  | 
| _ => I)  | 
|
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2364  | 
#> fold add_iterm_syms args  | 
| 42574 | 2365  | 
end  | 
| 
50968
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
2366  | 
val add_fact_syms = ifact_lift (formula_fold NONE (K add_iterm_syms))  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2367  | 
fun add_formula_var_types (ATyQuant (_, _, phi)) = add_formula_var_types phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2368  | 
| add_formula_var_types (AQuant (_, xs, phi)) =  | 
| 47150 | 2369  | 
fold (fn (_, SOME T) => insert_type thy I T | _ => I) xs  | 
| 43966 | 2370  | 
#> add_formula_var_types phi  | 
2371  | 
| add_formula_var_types (AConn (_, phis)) =  | 
|
2372  | 
fold add_formula_var_types phis  | 
|
2373  | 
| add_formula_var_types _ = I  | 
|
2374  | 
fun var_types () =  | 
|
| 48131 | 2375  | 
if is_type_enc_polymorphic type_enc then [tvar_a]  | 
| 47981 | 2376  | 
else fold (ifact_lift add_formula_var_types) (conjs @ facts) []  | 
| 
74983
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2377  | 
(* Don't add declaration for undefined_bool as bool already has fTrue and fFalse als witnesses  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2378  | 
and this declaration causes problems in FOF if undefined_bool occurs without predicator pp.  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2379  | 
*)  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2380  | 
fun add_undefined_const \<^Type>\<open>bool\<close> = I  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2381  | 
| add_undefined_const T =  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2382  | 
let  | 
| 75010 | 2383  | 
val name = `make_fixed_const \<^const_name>\<open>undefined\<close>  | 
| 
74983
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2384  | 
val ((s, s'), Ts) =  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2385  | 
if is_type_enc_mangling type_enc then  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2386  | 
(mangled_const_name type_enc [T] name, [])  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2387  | 
else  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2388  | 
(name, [T])  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2389  | 
in  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2390  | 
Symtab.map_default (s, []) (insert_type thy #3 (s', Ts, T, false, 0, false))  | 
| 
 
b87fcf474e7f
removed unconditional TPTP symbol declaration for undefined_bool in sledgehammer
 
desharna 
parents: 
74980 
diff
changeset
 | 
2391  | 
end  | 
| 44622 | 2392  | 
fun add_TYPE_const () =  | 
2393  | 
let val (s, s') = TYPE_name in  | 
|
2394  | 
Symtab.map_default (s, [])  | 
|
| 47150 | 2395  | 
(insert_type thy #3  | 
| 69593 | 2396  | 
(s', [tvar_a], \<^typ>\<open>'a itself\<close>, false, 0, false))  | 
| 44622 | 2397  | 
end  | 
| 
42698
 
ffd1ae4ff5c6
help SOS by ensuring that typing information is marked as part of the conjecture + be more precise w.r.t. typedefs in monotonicity check
 
blanchet 
parents: 
42697 
diff
changeset
 | 
2398  | 
in  | 
| 
 
ffd1ae4ff5c6
help SOS by ensuring that typing information is marked as part of the conjecture + be more precise w.r.t. typedefs in monotonicity check
 
blanchet 
parents: 
42697 
diff
changeset
 | 
2399  | 
Symtab.empty  | 
| 48089 | 2400  | 
|> is_type_enc_sound type_enc  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2401  | 
? (fold (fold add_fact_syms) [conjs, facts]  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2402  | 
#> fold add_iterm_syms extra_tms  | 
| 
43985
 
33d8b99531c2
no need for existential witnesses for sorts in TFF and THF formats
 
blanchet 
parents: 
43984 
diff
changeset
 | 
2403  | 
#> (case type_enc of  | 
| 74890 | 2404  | 
Native (_, _, Raw_Polymorphic phantoms, _) =>  | 
| 48131 | 2405  | 
phantoms = Without_Phantom_Type_Vars ? add_TYPE_const ()  | 
| 47767 | 2406  | 
| Native _ => I  | 
| 
74980
 
8dd527e97ddb
proper name mangling of "undefined" constants in Sledgehammer
 
desharna 
parents: 
74949 
diff
changeset
 | 
2407  | 
| _ =>  | 
| 
 
8dd527e97ddb
proper name mangling of "undefined" constants in Sledgehammer
 
desharna 
parents: 
74949 
diff
changeset
 | 
2408  | 
(* Add constants "undefined" as witnesses that the types are inhabited. *)  | 
| 
 
8dd527e97ddb
proper name mangling of "undefined" constants in Sledgehammer
 
desharna 
parents: 
74949 
diff
changeset
 | 
2409  | 
fold add_undefined_const (var_types ())))  | 
| 
42698
 
ffd1ae4ff5c6
help SOS by ensuring that typing information is marked as part of the conjecture + be more precise w.r.t. typedefs in monotonicity check
 
blanchet 
parents: 
42697 
diff
changeset
 | 
2410  | 
end  | 
| 42533 | 2411  | 
|
| 44399 | 2412  | 
(* We add "bool" in case the helper "True_or_False" is included later. *)  | 
| 48143 | 2413  | 
fun default_mono level completish =  | 
| 69593 | 2414  | 
  {maybe_finite_Ts = [\<^typ>\<open>bool\<close>],
 | 
| 54829 | 2415  | 
surely_infinite_Ts = (case level of Nonmono_Types (Strict, _) => [] | _ => known_infinite_types),  | 
| 69593 | 2416  | 
maybe_nonmono_Ts = [if completish >= 3 then tvar_a else \<^typ>\<open>bool\<close>]}  | 
| 44399 | 2417  | 
|
| 
50968
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
2418  | 
(* This inference is described in section 4 of Blanchette et al., "Encoding  | 
| 
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
2419  | 
monomorphic and polymorphic types", TACAS 2013. *)  | 
| 
63769
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2420  | 
fun add_iterm_mononotonicity_info ctxt level polarity tm  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2421  | 
      (mono as {maybe_finite_Ts, surely_infinite_Ts, maybe_nonmono_Ts}) =
 | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2422  | 
let  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2423  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2424  | 
|
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2425  | 
fun update_mono T mono =  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2426  | 
(case level of  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2427  | 
Nonmono_Types (strictness, _) =>  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2428  | 
if exists (type_instance thy T) surely_infinite_Ts orelse  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2429  | 
member (type_equiv thy) maybe_finite_Ts T then  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2430  | 
mono  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2431  | 
else if is_type_kind_of_surely_infinite ctxt strictness  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2432  | 
surely_infinite_Ts T then  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2433  | 
          {maybe_finite_Ts = maybe_finite_Ts,
 | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2434  | 
surely_infinite_Ts = surely_infinite_Ts |> insert_type thy I T,  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2435  | 
maybe_nonmono_Ts = maybe_nonmono_Ts}  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2436  | 
else  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2437  | 
          {maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv thy) T,
 | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2438  | 
surely_infinite_Ts = surely_infinite_Ts,  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2439  | 
maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type thy I T}  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2440  | 
| _ => mono)  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2441  | 
|
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2442  | 
fun update_mono_rec (IConst ((_, s'), Type (_, [T, _]), _)) =  | 
| 69593 | 2443  | 
if String.isPrefix \<^const_name>\<open>fequal\<close> s' then update_mono T else I  | 
| 
63769
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2444  | 
| update_mono_rec (IApp (tm1, tm2)) = fold update_mono_rec [tm1, tm2]  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2445  | 
| update_mono_rec (IAbs (_, tm)) = update_mono_rec tm  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2446  | 
| update_mono_rec _ = I  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2447  | 
in  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2448  | 
mono |>  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2449  | 
(case tm of  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2450  | 
IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2) =>  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2451  | 
((polarity <> SOME false andalso is_tptp_equal s  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2452  | 
andalso exists is_maybe_universal_var [tm1, tm2])  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2453  | 
? update_mono T)  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2454  | 
#> fold update_mono_rec [tm1, tm2]  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2455  | 
| _ => update_mono_rec tm)  | 
| 
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2456  | 
end  | 
| 47981 | 2457  | 
fun add_fact_mononotonicity_info ctxt level ({role, iformula, ...} : ifact) =
 | 
| 
63769
 
511d5ffd56ac
consider equality proxy in monotonicity analysis
 
blanchet 
parents: 
63694 
diff
changeset
 | 
2458  | 
formula_fold (SOME (role <> Conjecture)) (add_iterm_mononotonicity_info ctxt level) iformula  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2459  | 
fun mononotonicity_info_of_facts ctxt type_enc completish facts =  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
2460  | 
let val level = level_of_type_enc type_enc in  | 
| 48143 | 2461  | 
default_mono level completish  | 
| 44399 | 2462  | 
|> is_type_level_monotonicity_based level  | 
2463  | 
? fold (add_fact_mononotonicity_info ctxt level) facts  | 
|
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2464  | 
end  | 
| 
42677
 
25496cd3c199
monotonic type inference in ATP Sledgehammer problems -- based on Claessen & al.'s CADE 2011 paper, Sect. 2.3.
 
blanchet 
parents: 
42675 
diff
changeset
 | 
2465  | 
|
| 52005 | 2466  | 
fun fold_arg_types f (IApp (tm1, tm2)) =  | 
2467  | 
fold_arg_types f tm1 #> fold_term_types f tm2  | 
|
2468  | 
| fold_arg_types _ _ = I  | 
|
2469  | 
and fold_term_types f tm = f (ityp_of tm) #> fold_arg_types f tm  | 
|
2470  | 
||
| 44501 | 2471  | 
fun add_iformula_monotonic_types ctxt mono type_enc =  | 
2472  | 
let  | 
|
| 47150 | 2473  | 
val thy = Proof_Context.theory_of ctxt  | 
| 44501 | 2474  | 
val level = level_of_type_enc type_enc  | 
2475  | 
val should_encode = should_encode_type ctxt mono level  | 
|
| 47150 | 2476  | 
fun add_type T = not (should_encode T) ? insert_type thy I T  | 
| 52005 | 2477  | 
in formula_fold NONE (K (fold_term_types add_type)) end  | 
| 
50968
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
2478  | 
|
| 44501 | 2479  | 
fun add_fact_monotonic_types ctxt mono type_enc =  | 
| 
50968
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
2480  | 
ifact_lift (add_iformula_monotonic_types ctxt mono type_enc)  | 
| 
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
2481  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2482  | 
fun monotonic_types_of_facts ctxt mono type_enc facts =  | 
| 44811 | 2483  | 
let val level = level_of_type_enc type_enc in  | 
| 48131 | 2484  | 
[] |> (is_type_enc_polymorphic type_enc andalso  | 
| 48183 | 2485  | 
is_type_level_monotonicity_based level)  | 
| 44811 | 2486  | 
? fold (add_fact_monotonic_types ctxt mono type_enc) facts  | 
2487  | 
end  | 
|
| 44501 | 2488  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2489  | 
fun line_of_guards_mono_type ctxt generate_info mono type_enc T =  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2490  | 
Formula ((guards_sym_formula_prefix ^ ascii_of (mangled_type type_enc T), ""),  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2491  | 
Axiom,  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2492  | 
IConst (`make_bound_var "X", T, [])  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2493  | 
|> type_guard_iterm type_enc T  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2494  | 
|> AAtom  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2495  | 
|> formula_of_iformula ctxt mono type_enc always_guard_var_in_formula  | 
| 
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2496  | 
(SOME true)  | 
| 45377 | 2497  | 
|> close_formula_universally  | 
| 45920 | 2498  | 
|> bound_tvars type_enc true (atomic_types_of T),  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2499  | 
NONE, isabelle_info generate_info inductiveN helper_rank)  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2500  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2501  | 
fun line_of_tags_mono_type ctxt generate_info mono type_enc T =  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2502  | 
let val x_var = ATerm ((`make_bound_var "X", []), []) in  | 
| 54798 | 2503  | 
Formula ((tags_sym_formula_prefix ^ ascii_of (mangled_type type_enc T), ""), Axiom,  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2504  | 
eq_formula type_enc (atomic_types_of T) [] false  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2505  | 
(tag_with_type ctxt mono type_enc NONE T x_var) x_var,  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2506  | 
NONE, isabelle_info generate_info non_rec_defN helper_rank)  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2507  | 
end  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2508  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2509  | 
fun lines_of_mono_types ctxt generate_info mono type_enc =  | 
| 54829 | 2510  | 
(case type_enc of  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2511  | 
Native _ => K []  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2512  | 
| Guards _ => map (line_of_guards_mono_type ctxt generate_info mono type_enc)  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2513  | 
| Tags _ => map (line_of_tags_mono_type ctxt generate_info mono type_enc))  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2514  | 
|
| 72174 | 2515  | 
fun decl_line_of_sym ctxt type_enc s (s', T_args, T, pred_sym, ary, _) =  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
2516  | 
let  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2517  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2518  | 
val (T, T_args) =  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2519  | 
if null T_args then  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2520  | 
(T, [])  | 
| 54829 | 2521  | 
else  | 
2522  | 
(case unprefix_and_unascii const_prefix s of  | 
|
2523  | 
SOME s' =>  | 
|
2524  | 
let  | 
|
2525  | 
val s' = s' |> unmangled_invert_const  | 
|
2526  | 
val T = s' |> robust_const_type thy  | 
|
2527  | 
in (T, robust_const_type_args thy (s', T)) end  | 
|
2528  | 
| NONE => raise Fail "unexpected type arguments")  | 
|
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
2529  | 
in  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2530  | 
Sym_Decl (sym_decl_prefix ^ s, (s, s'),  | 
| 72174 | 2531  | 
T |> native_atp_type_of_typ type_enc pred_sym ary  | 
2532  | 
|> not (null T_args) ? curry APi (map (tvar_name o dest_TVar) T_args))  | 
|
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
2533  | 
end  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2534  | 
|
| 
50968
 
3686bc0d4df2
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
 
blanchet 
parents: 
50758 
diff
changeset
 | 
2535  | 
fun honor_conj_sym_role in_conj = (if in_conj then Hypothesis else Axiom, I)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2536  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2537  | 
fun line_of_guards_sym_decl ctxt generate_info mono type_enc n s j  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2538  | 
(s', T_args, T, _, ary, in_conj) =  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2539  | 
let  | 
| 44811 | 2540  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2541  | 
val (role, maybe_negate) = honor_conj_sym_role in_conj  | 
| 
42753
 
c9552e617acc
drop some type arguments to constants in unsound type systems + remove a few type systems that make no sense from the circulation
 
blanchet 
parents: 
42750 
diff
changeset
 | 
2542  | 
val (arg_Ts, res_T) = chop_fun ary T  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2543  | 
val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)  | 
| 74918 | 2544  | 
val bounds = map2 (fn name => fn T => IConst (name, T, [])) bound_names arg_Ts  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2545  | 
val bound_Ts =  | 
| 54829 | 2546  | 
(case level_of_type_enc type_enc of  | 
| 
48538
 
726590131ca1
bring implementation of traditional encoding in line with paper
 
blanchet 
parents: 
48438 
diff
changeset
 | 
2547  | 
All_Types => if null T_args then replicate ary NONE else map SOME arg_Ts  | 
| 
 
726590131ca1
bring implementation of traditional encoding in line with paper
 
blanchet 
parents: 
48438 
diff
changeset
 | 
2548  | 
| Undercover_Types =>  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2549  | 
let val cover = type_arg_cover thy NONE s ary in  | 
| 74918 | 2550  | 
map_index (fn (j, arg_T) => if member (op =) cover j then SOME arg_T else NONE) arg_Ts  | 
| 48080 | 2551  | 
end  | 
| 54829 | 2552  | 
| _ => replicate ary NONE)  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2553  | 
in  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2554  | 
Formula ((guards_sym_formula_prefix ^ s ^  | 
| 
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2555  | 
(if n > 1 then "_" ^ string_of_int j else ""), ""),  | 
| 
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2556  | 
role,  | 
| 43859 | 2557  | 
IConst ((s, s'), T, T_args)  | 
2558  | 
|> fold (curry (IApp o swap)) bounds  | 
|
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2559  | 
|> type_guard_iterm type_enc res_T  | 
| 42963 | 2560  | 
|> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2561  | 
|> formula_of_iformula ctxt mono type_enc  | 
| 
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52030 
diff
changeset
 | 
2562  | 
always_guard_var_in_formula (SOME true)  | 
| 45377 | 2563  | 
|> close_formula_universally  | 
| 45920 | 2564  | 
|> bound_tvars type_enc (n > 1) (atomic_types_of T)  | 
| 
42709
 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 
blanchet 
parents: 
42701 
diff
changeset
 | 
2565  | 
|> maybe_negate,  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2566  | 
NONE, isabelle_info generate_info inductiveN helper_rank)  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2567  | 
end  | 
| 
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2568  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2569  | 
fun lines_of_tags_sym_decl ctxt generate_info mono type_enc n s  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2570  | 
(j, (s', T_args, T, pred_sym, ary, in_conj)) =  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2571  | 
let  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2572  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2573  | 
val level = level_of_type_enc type_enc  | 
| 48183 | 2574  | 
val ident =  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2575  | 
tags_sym_formula_prefix ^ s ^  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
2576  | 
(if n > 1 then "_" ^ string_of_int j else "")  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2577  | 
val (role, maybe_negate) = honor_conj_sym_role in_conj  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2578  | 
val (arg_Ts, res_T) = chop_fun ary T  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2579  | 
val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2580  | 
val bounds = bound_names |> map (fn name => ATerm ((name, []), []))  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2581  | 
val cst = mk_aterm type_enc (s, s') T_args  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2582  | 
val eq = maybe_negate oo eq_formula type_enc (atomic_types_of T) [] pred_sym  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2583  | 
val tag_with = tag_with_type ctxt mono type_enc NONE  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2584  | 
fun formula c =  | 
| 
50521
 
bec828f3364e
generate comments with original names for debugging
 
blanchet 
parents: 
49982 
diff
changeset
 | 
2585  | 
[Formula ((ident, ""), role, eq (tag_with res_T c) c, NONE,  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2586  | 
isabelle_info generate_info non_rec_defN helper_rank)]  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2587  | 
in  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2588  | 
if pred_sym orelse not (should_encode_type ctxt mono level res_T) then  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2589  | 
[]  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2590  | 
else if level = Undercover_Types then  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2591  | 
let  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2592  | 
val cover = type_arg_cover thy NONE s ary  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2593  | 
fun maybe_tag (j, arg_T) = member (op =) cover j ? tag_with arg_T  | 
| 74918 | 2594  | 
val bounds = bounds |> map2 maybe_tag (map_index I arg_Ts)  | 
| 
48200
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2595  | 
in formula (cst bounds) end  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2596  | 
else  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2597  | 
formula (cst bounds)  | 
| 
 
5156cadedfa5
fiddled with "tags_at" a bit -- the only non-tuning part of the change is a slight simplification of the typing axioms
 
blanchet 
parents: 
48199 
diff
changeset
 | 
2598  | 
end  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2599  | 
|
| 42836 | 2600  | 
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd  | 
2601  | 
||
| 47150 | 2602  | 
fun rationalize_decls thy (decls as decl :: (decls' as _ :: _)) =  | 
| 45780 | 2603  | 
let  | 
| 69593 | 2604  | 
val T = result_type_of_decl decl |> map_type_tvar (fn (z, _) => TVar (z, \<^sort>\<open>type\<close>))  | 
| 45780 | 2605  | 
in  | 
| 54829 | 2606  | 
if forall (type_generalization thy T o result_type_of_decl) decls' then [decl]  | 
2607  | 
else decls  | 
|
| 45780 | 2608  | 
end  | 
2609  | 
| rationalize_decls _ decls = decls  | 
|
2610  | 
||
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2611  | 
fun lines_of_sym_decls ctxt generate_info mono type_enc (s, decls) =  | 
| 54829 | 2612  | 
(case type_enc of  | 
| 72174 | 2613  | 
Native _ => [decl_line_of_sym ctxt type_enc s (hd decls)]  | 
| 44768 | 2614  | 
| Guards (_, level) =>  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2615  | 
let  | 
| 47150 | 2616  | 
val thy = Proof_Context.theory_of ctxt  | 
2617  | 
val decls = decls |> rationalize_decls thy  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2618  | 
val n = length decls  | 
| 54829 | 2619  | 
val decls = decls |> filter (should_encode_type ctxt mono level o result_type_of_decl)  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2620  | 
in  | 
| 74918 | 2621  | 
map_index (uncurry (line_of_guards_sym_decl ctxt generate_info mono type_enc n s)) decls  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2622  | 
end  | 
| 44768 | 2623  | 
| Tags (_, level) =>  | 
| 48183 | 2624  | 
if is_type_level_uniform level then  | 
| 44768 | 2625  | 
[]  | 
2626  | 
else  | 
|
2627  | 
let val n = length decls in  | 
|
| 74918 | 2628  | 
map_index I decls  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2629  | 
|> maps (lines_of_tags_sym_decl ctxt generate_info mono type_enc n s)  | 
| 54829 | 2630  | 
end)  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2631  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2632  | 
fun lines_of_sym_decl_table ctxt generate_info mono type_enc mono_Ts sym_decl_tab =  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2633  | 
let  | 
| 
60924
 
610794dff23c
tuned signature, in accordance to sortBy in Scala;
 
wenzelm 
parents: 
60328 
diff
changeset
 | 
2634  | 
val syms = sym_decl_tab |> Symtab.dest |> sort_by fst  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2635  | 
val mono_lines = lines_of_mono_types ctxt generate_info mono type_enc mono_Ts  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2636  | 
val decl_lines = maps (lines_of_sym_decls ctxt generate_info mono type_enc) syms  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2637  | 
in mono_lines @ decl_lines end  | 
| 
42543
 
f9d402d144d4
declare TFF types so that SNARK can be used with types
 
blanchet 
parents: 
42542 
diff
changeset
 | 
2638  | 
|
| 
56636
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2639  | 
fun datatypes_of_sym_table ctxt ctrss (DFG Polymorphic) (type_enc as Native _) uncurried_aliases  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2640  | 
sym_tab =  | 
| 52032 | 2641  | 
if is_type_enc_polymorphic type_enc then  | 
2642  | 
let  | 
|
2643  | 
val thy = Proof_Context.theory_of ctxt  | 
|
| 
56636
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2644  | 
|
| 52032 | 2645  | 
fun do_ctr (s, T) =  | 
2646  | 
let  | 
|
| 75010 | 2647  | 
val s' = make_fixed_const s  | 
| 
52038
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
2648  | 
val ary = ary_of T  | 
| 
56636
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2649  | 
fun mk name = SOME (mk_aterm type_enc name (robust_const_type_args thy (s, T)) [])  | 
| 
52038
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
2650  | 
in  | 
| 
56636
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2651  | 
if T = HOLogic.boolT then  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2652  | 
(case proxify_const s' of  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2653  | 
SOME proxy_base => mk (proxy_base |>> prefix const_prefix)  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2654  | 
| NONE => NONE)  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2655  | 
else  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2656  | 
(case Symtab.lookup sym_tab s' of  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2657  | 
NONE => NONE  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2658  | 
              | SOME ({min_ary, ...} : sym_info) =>
 | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2659  | 
if ary = min_ary then mk (s', s)  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2660  | 
else if uncurried_aliases then mk (aliased_uncurried ary (s', s))  | 
| 
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2661  | 
else NONE)  | 
| 
52038
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
2662  | 
end  | 
| 
56636
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2663  | 
|
| 52032 | 2664  | 
fun datatype_of_ctrs (ctrs as (_, T1) :: _) =  | 
2665  | 
let val ctrs' = map do_ctr ctrs in  | 
|
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
56636 
diff
changeset
 | 
2666  | 
(native_atp_type_of_typ type_enc false 0 (body_type T1), map_filter I ctrs',  | 
| 
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
56636 
diff
changeset
 | 
2667  | 
forall is_some ctrs')  | 
| 52032 | 2668  | 
end  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54768 
diff
changeset
 | 
2669  | 
in  | 
| 59709 | 2670  | 
ctrss |> map datatype_of_ctrs |> filter #3  | 
| 
54788
 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 
blanchet 
parents: 
54768 
diff
changeset
 | 
2671  | 
end  | 
| 52032 | 2672  | 
else  | 
2673  | 
[]  | 
|
| 
52038
 
a354c83dee43
properly handle SPASS constructors w.r.t. partially applied functions
 
blanchet 
parents: 
52035 
diff
changeset
 | 
2674  | 
| datatypes_of_sym_table _ _ _ _ _ _ = []  | 
| 52004 | 2675  | 
|
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
56636 
diff
changeset
 | 
2676  | 
fun decl_line_of_datatype (ty as AType (((_, s'), _), ty_args), ctrs, exhaust) =  | 
| 54820 | 2677  | 
let val xs = map (fn AType ((name, _), []) => name) ty_args in  | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
56636 
diff
changeset
 | 
2678  | 
Datatype_Decl (datatype_decl_prefix ^ ascii_of s', map (rpair []) xs, ty, ctrs, exhaust)  | 
| 52025 | 2679  | 
end  | 
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
51998 
diff
changeset
 | 
2680  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2681  | 
fun pair_append (xs1, xs2) (ys1, ys2) = (xs1 @ ys1, xs2 @ ys2)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2682  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2683  | 
fun do_uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab base_s0  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2684  | 
types in_conj =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2685  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2686  | 
fun do_alias ary =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2687  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2688  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
47975
 
adc977fec17e
order LEO-II/Satallax definitions so that they build on each other (cf. Satallax's THF policy)
 
blanchet 
parents: 
47971 
diff
changeset
 | 
2689  | 
val (role, maybe_negate) = honor_conj_sym_role in_conj  | 
| 75010 | 2690  | 
val base_name = base_s0 |> `make_fixed_const  | 
| 54829 | 2691  | 
val T = (case types of [T] => T | _ => robust_const_type thy base_s0)  | 
| 
51920
 
16f3b9d4e515
proper unmangling -- the bug is visible when "uncurried_aliases" is enabled with Alt-Ergo or Poly. SPASS
 
blanchet 
parents: 
51717 
diff
changeset
 | 
2692  | 
val T_args = robust_const_type_args thy (base_s0, T)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2693  | 
val (base_name as (base_s, _), T_args) =  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2694  | 
mangle_type_args_in_const type_enc base_name T_args  | 
| 
46402
 
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
 
blanchet 
parents: 
46400 
diff
changeset
 | 
2695  | 
val base_ary = min_ary_of sym_tab0 base_s  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2696  | 
fun do_const name = IConst (name, T, T_args)  | 
| 52028 | 2697  | 
val filter_ty_args = filter_type_args_in_iterm thy ctrss type_enc  | 
| 
53586
 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 
blanchet 
parents: 
52995 
diff
changeset
 | 
2698  | 
val atp_term_of = atp_term_of_iterm ctxt mono type_enc (SOME true)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2699  | 
val name1 as (s1, _) =  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2700  | 
base_name |> ary - 1 > base_ary ? aliased_uncurried (ary - 1)  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2701  | 
val name2 as (s2, _) = base_name |> aliased_uncurried ary  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2702  | 
val (arg_Ts, _) = chop_fun ary T  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2703  | 
val bound_names =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2704  | 
1 upto ary |> map (`I o make_bound_var o string_of_int)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2705  | 
val bounds = bound_names ~~ arg_Ts  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2706  | 
val (first_bounds, last_bound) =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2707  | 
bounds |> map (fn (name, T) => IConst (name, T, [])) |> split_last  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2708  | 
val tm1 =  | 
| 
47768
 
0b2b7ff31867
don't use the native choice operator if the type encoding isn't higher-order
 
blanchet 
parents: 
47767 
diff
changeset
 | 
2709  | 
mk_app_op type_enc (list_app (do_const name1) first_bounds) last_bound  | 
| 46437 | 2710  | 
|> filter_ty_args  | 
2711  | 
val tm2 =  | 
|
2712  | 
list_app (do_const name2) (first_bounds @ [last_bound])  | 
|
2713  | 
|> filter_ty_args  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2714  | 
val eq =  | 
| 72174 | 2715  | 
eq_formula type_enc (atomic_types_of T)  | 
2716  | 
(map (apsnd (do_bound_type type_enc)) bounds) false (atp_term_of tm1) (atp_term_of tm2)  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2717  | 
in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2718  | 
([tm1, tm2],  | 
| 54829 | 2719  | 
[Formula ((uncurried_alias_eq_prefix ^ s2, ""), role, eq |> maybe_negate, NONE,  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2720  | 
isabelle_info generate_info non_rec_defN helper_rank)])  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2721  | 
|> (if ary - 1 = base_ary orelse Symtab.defined sym_tab s1 then I  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2722  | 
else pair_append (do_alias (ary - 1)))  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2723  | 
end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2724  | 
in do_alias end  | 
| 55212 | 2725  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2726  | 
fun uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab  | 
| 52026 | 2727  | 
        (s, {min_ary, types, in_conj, ...} : sym_info) =
 | 
| 54829 | 2728  | 
(case unprefix_and_unascii const_prefix s of  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2729  | 
SOME mangled_s =>  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2730  | 
if String.isSubstring uncurried_alias_sep mangled_s then  | 
| 54829 | 2731  | 
let val base_s0 = mangled_s |> unmangled_invert_const in  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2732  | 
do_uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2733  | 
base_s0 types in_conj min_ary  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2734  | 
end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2735  | 
else  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2736  | 
([], [])  | 
| 54829 | 2737  | 
| NONE => ([], []))  | 
| 55212 | 2738  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2739  | 
fun uncurried_alias_lines_of_sym_table ctxt generate_info ctrss mono type_enc uncurried_aliases  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2740  | 
sym_tab0 sym_tab =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2741  | 
([], [])  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2742  | 
|> uncurried_aliases  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2743  | 
? Symtab.fold_rev (pair_append  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2744  | 
o uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab)  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2745  | 
sym_tab  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2746  | 
|
| 
52995
 
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
 
blanchet 
parents: 
52125 
diff
changeset
 | 
2747  | 
val implicit_declsN = "Could-be-implicit typings"  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2748  | 
val explicit_declsN = "Explicit typings"  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2749  | 
val uncurried_alias_eqsN = "Uncurried aliases"  | 
| 41157 | 2750  | 
val factsN = "Relevant facts"  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2751  | 
val subclassesN = "Subclasses"  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2752  | 
val tconsN = "Type constructors"  | 
| 41157 | 2753  | 
val helpersN = "Helper facts"  | 
2754  | 
val conjsN = "Conjectures"  | 
|
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2755  | 
val free_typesN = "Free types"  | 
| 41157 | 2756  | 
|
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2757  | 
(* TFF allows implicit declarations of types, function symbols, and predicate  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2758  | 
symbols (with "$i" as the type of individuals), but some provers (e.g.,  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2759  | 
SNARK) require explicit declarations. The situation is similar for THF. *)  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2760  | 
|
| 48141 | 2761  | 
fun default_type pred_sym =  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2762  | 
let  | 
| 
48136
 
0f9939676088
removed old hack now that types and terms are cleanly distinguished in the data structure
 
blanchet 
parents: 
48135 
diff
changeset
 | 
2763  | 
fun typ 0 0 = if pred_sym then bool_atype else individual_atype  | 
| 
 
0f9939676088
removed old hack now that types and terms are cleanly distinguished in the data structure
 
blanchet 
parents: 
48135 
diff
changeset
 | 
2764  | 
| typ 0 tm_ary = AFun (individual_atype, typ 0 (tm_ary - 1))  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2765  | 
| typ ty_ary tm_ary = APi (replicate ty_ary tvar_a_name, typ 0 tm_ary)  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2766  | 
in typ end  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2767  | 
|
| 48140 | 2768  | 
fun undeclared_in_problem problem =  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2769  | 
let  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2770  | 
fun do_sym (name as (s, _)) value =  | 
| 
74662
 
44585660f39a
do not declare $let-bound variables in TPTP output
 
desharna 
parents: 
74552 
diff
changeset
 | 
2771  | 
if is_tptp_user_symbol s andalso not (String.isPrefix let_bound_var_prefix s) then  | 
| 
 
44585660f39a
do not declare $let-bound variables in TPTP output
 
desharna 
parents: 
74552 
diff
changeset
 | 
2772  | 
Symtab.default (s, (name, value))  | 
| 
 
44585660f39a
do not declare $let-bound variables in TPTP output
 
desharna 
parents: 
74552 
diff
changeset
 | 
2773  | 
else  | 
| 
 
44585660f39a
do not declare $let-bound variables in TPTP output
 
desharna 
parents: 
74552 
diff
changeset
 | 
2774  | 
I  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2775  | 
fun do_class name = apfst (apfst (do_sym name ()))  | 
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
51998 
diff
changeset
 | 
2776  | 
val do_bound_tvars = fold do_class o snd  | 
| 54820 | 2777  | 
fun do_type (AType ((name, _), tys)) =  | 
| 48141 | 2778  | 
apfst (apsnd (do_sym name (length tys))) #> fold do_type tys  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2779  | 
| do_type (AFun (ty1, ty2)) = do_type ty1 #> do_type ty2  | 
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2780  | 
| do_type (APi (_, ty)) = do_type ty  | 
| 48141 | 2781  | 
fun do_term pred_sym (ATerm ((name, tys), tms)) =  | 
| 54830 | 2782  | 
apsnd (do_sym name (fn _ => default_type pred_sym (length tys) (length tms)))  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2783  | 
#> fold do_type tys #> fold (do_term false) tms  | 
| 
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: 
47905 
diff
changeset
 | 
2784  | 
| do_term _ (AAbs (((_, ty), tm), args)) =  | 
| 
 
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: 
47905 
diff
changeset
 | 
2785  | 
do_type ty #> do_term false tm #> fold (do_term false) args  | 
| 48141 | 2786  | 
fun do_formula (ATyQuant (_, xs, phi)) =  | 
| 54791 | 2787  | 
fold (do_type o fst) xs #> fold (fold do_class o snd) xs #> do_formula phi  | 
2788  | 
| do_formula (AQuant (_, xs, phi)) = fold do_type (map_filter snd xs) #> do_formula phi  | 
|
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2789  | 
| do_formula (AConn (_, phis)) = fold do_formula phis  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2790  | 
| do_formula (AAtom tm) = do_term true tm  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2791  | 
fun do_line (Class_Decl (_, _, cls)) = fold do_class cls  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2792  | 
| do_line (Type_Decl _) = I  | 
| 48140 | 2793  | 
| do_line (Sym_Decl (_, _, ty)) = do_type ty  | 
| 
56683
 
7f4ae504e059
reintroduced '...' (nonexhaustive) syntax for SPASS-Pirate
 
blanchet 
parents: 
56636 
diff
changeset
 | 
2794  | 
| do_line (Datatype_Decl (_, xs, ty, tms, _)) =  | 
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
51998 
diff
changeset
 | 
2795  | 
fold do_bound_tvars xs #> do_type ty #> fold (do_term false) tms  | 
| 54791 | 2796  | 
| do_line (Class_Memb (_, xs, ty, cl)) = fold do_bound_tvars xs #> do_type ty #> do_class cl  | 
| 48140 | 2797  | 
| do_line (Formula (_, _, phi, _, _)) = do_formula phi  | 
| 48141 | 2798  | 
val ((cls, tys), syms) = declared_in_atp_problem problem  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2799  | 
in  | 
| 48141 | 2800  | 
((Symtab.empty, Symtab.empty), Symtab.empty)  | 
2801  | 
    |>> apfst (fold (fn (s, _) => Symtab.default (s, (("", ""), ()))) cls)
 | 
|
2802  | 
    |>> apsnd (fold (fn (s, _) => Symtab.default (s, (("", ""), 0))) tys)
 | 
|
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2803  | 
    ||> fold (fn (s, _) => Symtab.default (s, (("", ""), K tvar_a_atype))) syms
 | 
| 48140 | 2804  | 
|> fold (fold do_line o snd) problem  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2805  | 
end  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2806  | 
|
| 48141 | 2807  | 
fun declare_undeclared_in_problem heading problem =  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2808  | 
let  | 
| 48141 | 2809  | 
val ((cls, tys), syms) = undeclared_in_problem problem  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2810  | 
val decls =  | 
| 47150 | 2811  | 
      Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
 | 
| 54830 | 2812  | 
| (s, (cls, ())) => cons (Class_Decl (class_decl_prefix ^ s, cls, []))) cls [] @  | 
| 48141 | 2813  | 
      Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
 | 
| 54830 | 2814  | 
| (s, (ty, ary)) => cons (Type_Decl (type_decl_prefix ^ s, ty, ary))) tys [] @  | 
| 
48137
 
6f524f2066e3
cleanly distinguish between type declarations and symbol declarations
 
blanchet 
parents: 
48136 
diff
changeset
 | 
2815  | 
      Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
 | 
| 54830 | 2816  | 
| (s, (sym, ty)) => cons (Sym_Decl (sym_decl_prefix ^ s, sym, ty ()))) syms []  | 
| 48141 | 2817  | 
in (heading, decls) :: problem end  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2818  | 
|
| 
56636
 
bb8b37480d3f
declare 'bool' and its proxies as a datatype for SPASS-Pirate
 
blanchet 
parents: 
56254 
diff
changeset
 | 
2819  | 
val all_ctrss_of_datatypes = map (map_filter (try dest_Const) o #ctrs) o Ctr_Sugar.ctr_sugars_of  | 
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2820  | 
|
| 
48558
 
fabbed3abc1e
tweaks in preparation for type encoding evaluation
 
blanchet 
parents: 
48538 
diff
changeset
 | 
2821  | 
val app_op_and_predicator_threshold = 45  | 
| 
43259
 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 
blanchet 
parents: 
43258 
diff
changeset
 | 
2822  | 
|
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2823  | 
fun generate_atp_problem ctxt generate_info format prem_role type_enc mode lam_trans  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2824  | 
uncurried_aliases readable_names presimp hyp_ts concl_t facts =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2825  | 
let  | 
| 
44774
 
72785558a6ff
separate mangling, which can (and should) be done before the formulas are first-orderized, and type arg filtering, which must be done after once the min arities have been computed
 
blanchet 
parents: 
44773 
diff
changeset
 | 
2826  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
44416
 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 
blanchet 
parents: 
44410 
diff
changeset
 | 
2827  | 
val type_enc = type_enc |> adjust_type_enc format  | 
| 62718 | 2828  | 
val completish = (case mode of Sledgehammer_Completish k => k | _ => 0)  | 
| 46389 | 2829  | 
(* Forcing explicit applications is expensive for polymorphic encodings,  | 
2830  | 
because it takes only one existential variable ranging over "'a => 'b" to  | 
|
2831  | 
ruin everything. Hence we do it only if there are few facts (which is  | 
|
2832  | 
normally the case for "metis" and the minimizer). *)  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2833  | 
val app_op_level =  | 
| 62718 | 2834  | 
if completish > 0 then  | 
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47944 
diff
changeset
 | 
2835  | 
Full_App_Op_And_Predicator  | 
| 
75005
 
4106bc2a9cc8
optimized app_op_level selection in TPTP generation
 
desharna 
parents: 
74988 
diff
changeset
 | 
2836  | 
else if is_greater_equal  | 
| 
 
4106bc2a9cc8
optimized app_op_level selection in TPTP generation
 
desharna 
parents: 
74988 
diff
changeset
 | 
2837  | 
(compare_length_with facts (app_op_and_predicator_threshold - length hyp_ts)) then  | 
| 
54109
 
80660c529d74
addressed rare case where the same symbol would be treated alternately as a function and as a predicate -- adding "top2I top_boolI" to a problem that didn't talk about "top" was a way to trigger the issue
 
blanchet 
parents: 
54089 
diff
changeset
 | 
2838  | 
if is_type_enc_polymorphic type_enc then Min_App_Op else Sufficient_App_Op  | 
| 46370 | 2839  | 
else  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2840  | 
Sufficient_App_Op_And_Predicator  | 
| 45514 | 2841  | 
val lam_trans =  | 
| 
72659
 
f8d25850173b
reintroduced and renamed THF_Predicate_Free deleted by c7e2a9bdc585
 
desharna 
parents: 
72658 
diff
changeset
 | 
2842  | 
if lam_trans = keep_lamsN andalso not (is_type_enc_full_higher_order type_enc) then liftingN  | 
| 
54109
 
80660c529d74
addressed rare case where the same symbol would be treated alternately as a function and as a predicate -- adding "top2I top_boolI" to a problem that didn't talk about "top" was a way to trigger the issue
 
blanchet 
parents: 
54089 
diff
changeset
 | 
2843  | 
else lam_trans  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
2844  | 
val simp_options =  | 
| 74890 | 2845  | 
      {if_simps = not (has_type_enc_ite type_enc),
 | 
2846  | 
let_simps = not (has_type_enc_let type_enc)}  | 
|
| 57263 | 2847  | 
val (classes, conjs, facts, subclass_pairs, tcon_clauses, lifted) =  | 
| 
74075
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
2848  | 
translate_formulas simp_options ctxt prem_role format type_enc lam_trans presimp hyp_ts  | 
| 
 
a5bab59d580b
added support for TFX $let to Sledgehammer's TPTP output
 
desharna 
parents: 
74073 
diff
changeset
 | 
2849  | 
concl_t facts  | 
| 
54109
 
80660c529d74
addressed rare case where the same symbol would be treated alternately as a function and as a predicate -- adding "top2I top_boolI" to a problem that didn't talk about "top" was a way to trigger the issue
 
blanchet 
parents: 
54089 
diff
changeset
 | 
2850  | 
val (_, sym_tab0) = sym_table_of_facts ctxt type_enc app_op_level conjs facts  | 
| 
57481
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2851  | 
val mono = conjs @ facts |> mononotonicity_info_of_facts ctxt type_enc completish  | 
| 
54554
 
b8d0d8407c3b
eliminated Sledgehammer's dependency on old-style datatypes
 
blanchet 
parents: 
54434 
diff
changeset
 | 
2852  | 
val ctrss = all_ctrss_of_datatypes ctxt  | 
| 
46400
 
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
 
blanchet 
parents: 
46399 
diff
changeset
 | 
2853  | 
fun firstorderize in_helper =  | 
| 54830 | 2854  | 
firstorderize_fact thy ctrss type_enc (uncurried_aliases andalso not in_helper) completish  | 
2855  | 
sym_tab0  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58477 
diff
changeset
 | 
2856  | 
val (conjs, facts) = (conjs, facts) |> apply2 (map (firstorderize false))  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
2857  | 
val (ho_stuff, sym_tab) =  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2858  | 
sym_table_of_facts ctxt type_enc Min_App_Op conjs facts  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
2859  | 
val (uncurried_alias_eq_tms, uncurried_alias_eq_lines) =  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2860  | 
uncurried_alias_lines_of_sym_table ctxt generate_info ctrss mono type_enc uncurried_aliases  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2861  | 
sym_tab0 sym_tab  | 
| 
47932
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
2862  | 
val (_, sym_tab) =  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
2863  | 
(ho_stuff, sym_tab)  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
2864  | 
|> fold (add_iterm_syms_to_sym_table ctxt Min_App_Op false false)  | 
| 
 
ce4178e037a7
get ready for automatic generation of extensionality helpers
 
blanchet 
parents: 
47925 
diff
changeset
 | 
2865  | 
uncurried_alias_eq_tms  | 
| 
57481
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2866  | 
val helpers =  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73859 
diff
changeset
 | 
2867  | 
sym_tab |> helper_facts_of_sym_table ctxt format type_enc lam_trans completish  | 
| 
57481
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2868  | 
|> map (firstorderize true)  | 
| 
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2869  | 
val all_facts = helpers @ conjs @ facts  | 
| 
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2870  | 
val mono_Ts = monotonic_types_of_facts ctxt mono type_enc all_facts  | 
| 54830 | 2871  | 
val datatypes = datatypes_of_sym_table ctxt ctrss format type_enc uncurried_aliases sym_tab  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2872  | 
val class_decl_lines = decl_lines_of_classes type_enc classes  | 
| 42680 | 2873  | 
val sym_decl_lines =  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2874  | 
(conjs, helpers @ facts, uncurried_alias_eq_tms)  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2875  | 
|> sym_decl_table_of_facts thy type_enc sym_tab  | 
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2876  | 
|> lines_of_sym_decl_table ctxt generate_info mono type_enc mono_Ts  | 
| 52025 | 2877  | 
val datatype_decl_lines = map decl_line_of_datatype datatypes  | 
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
51998 
diff
changeset
 | 
2878  | 
val decl_lines = class_decl_lines @ sym_decl_lines @ datatype_decl_lines  | 
| 
57481
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2879  | 
val num_facts = length facts  | 
| 
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2880  | 
val freshen = mode <> Exporter andalso mode <> Translator  | 
| 
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2881  | 
val pos = mode <> Exporter  | 
| 
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2882  | 
val rank_of = rank_of_fact_num num_facts  | 
| 74918 | 2883  | 
val fact_lines = facts  | 
2884  | 
|> map_index (line_of_fact ctxt generate_info fact_prefix ascii_of I freshen pos mono type_enc  | 
|
2885  | 
rank_of)  | 
|
2886  | 
||
| 
61860
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2887  | 
val subclass_lines = maps (lines_of_subclass_pair generate_info type_enc) subclass_pairs  | 
| 
 
2ce3d12015b3
cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
 
blanchet 
parents: 
61770 
diff
changeset
 | 
2888  | 
val tcon_lines = map (line_of_tcon_clause generate_info type_enc) tcon_clauses  | 
| 74918 | 2889  | 
val helper_lines = helpers  | 
2890  | 
|> map_index (line_of_fact ctxt generate_info helper_prefix I (K "") false true mono type_enc  | 
|
2891  | 
(K default_rank))  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2892  | 
val free_type_lines = lines_of_free_types type_enc (facts @ conjs)  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51921 
diff
changeset
 | 
2893  | 
val conj_lines = map (line_of_conjecture ctxt mono type_enc) conjs  | 
| 
48005
 
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
 
blanchet 
parents: 
48004 
diff
changeset
 | 
2894  | 
(* Reordering these might confuse the proof reconstruction code. *)  | 
| 
57481
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2895  | 
val problem =  | 
| 
52001
 
2fb33d73c366
more work on implementing datatype output for new SPASS
 
blanchet 
parents: 
51998 
diff
changeset
 | 
2896  | 
[(explicit_declsN, decl_lines),  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2897  | 
(uncurried_alias_eqsN, uncurried_alias_eq_lines),  | 
| 46406 | 2898  | 
(factsN, fact_lines),  | 
| 
48142
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2899  | 
(subclassesN, subclass_lines),  | 
| 
 
efaff8206967
finished implementation of DFG type class output
 
blanchet 
parents: 
48141 
diff
changeset
 | 
2900  | 
(tconsN, tcon_lines),  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
2901  | 
(helpersN, helper_lines),  | 
| 48141 | 2902  | 
(free_typesN, free_type_lines),  | 
| 48146 | 2903  | 
(conjsN, conj_lines)]  | 
| 
57481
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2904  | 
val problem =  | 
| 
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2905  | 
problem  | 
| 54830 | 2906  | 
|> (case format of  | 
2907  | 
CNF => ensure_cnf_problem  | 
|
2908  | 
| CNF_UEQ => filter_cnf_ueq_problem  | 
|
2909  | 
| FOF => I  | 
|
2910  | 
| _ => declare_undeclared_in_problem implicit_declsN)  | 
|
| 
57481
 
84bbdbf1b2da
reverted 9512b867259c -- appears to break 'metis'
 
blanchet 
parents: 
57470 
diff
changeset
 | 
2911  | 
val (problem, pool) = problem |> nice_atp_problem readable_names format  | 
| 54830 | 2912  | 
    fun add_sym_ary (s, {min_ary, ...} : sym_info) = min_ary > 0 ? Symtab.insert (op =) (s, min_ary)
 | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2913  | 
in  | 
| 57263 | 2914  | 
(problem, Option.map snd pool |> the_default Symtab.empty, lifted,  | 
2915  | 
Symtab.fold add_sym_ary sym_tab Symtab.empty)  | 
|
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2916  | 
end  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2917  | 
|
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2918  | 
(* FUDGE *)  | 
| 
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2919  | 
val conj_weight = 0.0  | 
| 41770 | 2920  | 
val hyp_weight = 0.1  | 
2921  | 
val fact_min_weight = 0.2  | 
|
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2922  | 
val fact_max_weight = 1.0  | 
| 
42608
 
6ef61823b63b
make sure E type information constants are given a weight, even if they don't appear anywhere else
 
blanchet 
parents: 
42592 
diff
changeset
 | 
2923  | 
val type_info_default_weight = 0.8  | 
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2924  | 
|
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2925  | 
(* Ugly hack: may make innocent victims (collateral damage) *)  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2926  | 
fun may_be_app s args = String.isPrefix app_op_name s andalso length args = 2  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2927  | 
fun may_be_predicator s =  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2928  | 
member (op =) [predicator_name, prefixed_predicator_name] s  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2929  | 
|
| 54830 | 2930  | 
fun strip_predicator (tm as ATerm ((s, _), [tm'])) = if may_be_predicator s then tm' else tm  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2931  | 
| strip_predicator tm = tm  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2932  | 
|
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2933  | 
fun make_head_roll (ATerm ((s, _), tms)) =  | 
| 54830 | 2934  | 
if may_be_app s tms then make_head_roll (hd tms) ||> append (tl tms) else (s, tms)  | 
| 47032 | 2935  | 
  | make_head_roll _ = ("", [])
 | 
| 
46443
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2936  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2937  | 
fun strip_up_to_predicator (ATyQuant (_, _, phi)) = strip_up_to_predicator phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2938  | 
| strip_up_to_predicator (AQuant (_, _, phi)) = strip_up_to_predicator phi  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2939  | 
| strip_up_to_predicator (AConn (_, phis)) = maps strip_up_to_predicator phis  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2940  | 
| strip_up_to_predicator (AAtom tm) = [strip_predicator tm]  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2941  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2942  | 
fun strip_ahorn_etc (ATyQuant (_, _, phi)) = strip_ahorn_etc phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2943  | 
| strip_ahorn_etc (AQuant (_, _, phi)) = strip_ahorn_etc phi  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2944  | 
| strip_ahorn_etc (AConn (AImplies, [phi1, phi2])) =  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2945  | 
strip_ahorn_etc phi2 |>> append (strip_up_to_predicator phi1)  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2946  | 
| strip_ahorn_etc phi = ([], hd (strip_up_to_predicator phi))  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2947  | 
|
| 
48133
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2948  | 
fun strip_iff_etc (ATyQuant (_, _, phi)) = strip_iff_etc phi  | 
| 
 
a5ab5964065f
implement polymorphic DFG output, without type classes for now
 
blanchet 
parents: 
48132 
diff
changeset
 | 
2949  | 
| strip_iff_etc (AQuant (_, _, phi)) = strip_iff_etc phi  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58477 
diff
changeset
 | 
2950  | 
| strip_iff_etc (AConn (AIff, [phi1, phi2])) = apply2 strip_up_to_predicator (phi1, phi2)  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2951  | 
| strip_iff_etc _ = ([], [])  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2952  | 
|
| 47030 | 2953  | 
val max_term_order_weight = 2147483647  | 
| 46450 | 2954  | 
|
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47032 
diff
changeset
 | 
2955  | 
fun atp_problem_term_order_info problem =  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2956  | 
let  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47032 
diff
changeset
 | 
2957  | 
fun add_edge s s' =  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47032 
diff
changeset
 | 
2958  | 
Graph.default_node (s, ())  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47032 
diff
changeset
 | 
2959  | 
#> Graph.default_node (s', ())  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
47032 
diff
changeset
 | 
2960  | 
#> Graph.add_edge_acyclic (s, s')  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2961  | 
fun add_term_deps head (ATerm ((s, _), args)) =  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2962  | 
if is_tptp_user_symbol head then  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2963  | 
(if is_tptp_user_symbol s then perhaps (try (add_edge s head)) else I)  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2964  | 
#> fold (add_term_deps head) args  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2965  | 
else  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2966  | 
I  | 
| 
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: 
47905 
diff
changeset
 | 
2967  | 
| add_term_deps head (AAbs ((_, tm), args)) =  | 
| 
 
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: 
47905 
diff
changeset
 | 
2968  | 
add_term_deps head tm #> fold (add_term_deps head) args  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2969  | 
fun add_intro_deps pred (Formula (_, role, phi, _, info)) =  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2970  | 
if pred (role, info) then  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2971  | 
let val (hyps, concl) = strip_ahorn_etc phi in  | 
| 54829 | 2972  | 
(case make_head_roll concl of  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2973  | 
(head, args as _ :: _) => fold (add_term_deps head) (hyps @ args)  | 
| 54829 | 2974  | 
| _ => I)  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2975  | 
end  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2976  | 
else  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2977  | 
I  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2978  | 
| add_intro_deps _ _ = I  | 
| 
48132
 
9aa0fad4e864
added type arguments to "ATerm" constructor -- but don't use them yet
 
blanchet 
parents: 
48131 
diff
changeset
 | 
2979  | 
fun add_atom_eq_deps (SOME true) (ATerm ((s, _), [lhs as _, rhs])) =  | 
| 
46443
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2980  | 
if is_tptp_equal s then  | 
| 54829 | 2981  | 
(case make_head_roll lhs of  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2982  | 
(head, args as _ :: _) => fold (add_term_deps head) (rhs :: args)  | 
| 54829 | 2983  | 
| _ => I)  | 
| 
46443
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2984  | 
else  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2985  | 
I  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2986  | 
| add_atom_eq_deps _ _ = I  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2987  | 
fun add_eq_deps pred (Formula (_, role, phi, _, info)) =  | 
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
47038 
diff
changeset
 | 
2988  | 
if pred (role, info) then  | 
| 54829 | 2989  | 
(case strip_iff_etc phi of  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2990  | 
([lhs], rhs) =>  | 
| 
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2991  | 
(case make_head_roll lhs of  | 
| 54829 | 2992  | 
(head, args as _ :: _) => fold (add_term_deps head) (rhs @ args)  | 
2993  | 
| _ => I)  | 
|
2994  | 
| _ => formula_fold (SOME (role <> Conjecture)) add_atom_eq_deps phi)  | 
|
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2995  | 
else  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2996  | 
I  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
2997  | 
| add_eq_deps _ _ = I  | 
| 47971 | 2998  | 
fun has_status status (_, info) = extract_isabelle_status info = SOME status  | 
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
47038 
diff
changeset
 | 
2999  | 
fun is_conj (role, _) = (role = Conjecture orelse role = Hypothesis)  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3000  | 
val graph =  | 
| 
47039
 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 
blanchet 
parents: 
47038 
diff
changeset
 | 
3001  | 
Graph.empty  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
3002  | 
|> fold (fold (add_eq_deps (has_status non_rec_defN)) o snd) problem  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
3003  | 
|> fold (fold (add_eq_deps (has_status rec_defN orf has_status simpN  | 
| 
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48324 
diff
changeset
 | 
3004  | 
orf is_conj)) o snd) problem  | 
| 
47148
 
7b5846065c1b
be less forceful about ":lt" to make infinite loops less likely (could still fail with mutually recursive tail rec functions)
 
blanchet 
parents: 
47145 
diff
changeset
 | 
3005  | 
|> fold (fold (add_intro_deps (has_status inductiveN)) o snd) problem  | 
| 
47073
 
c73f7b0c7ebc
generate weights and precedences for predicates as well
 
blanchet 
parents: 
47046 
diff
changeset
 | 
3006  | 
|> fold (fold (add_intro_deps (has_status introN)) o snd) problem  | 
| 47030 | 3007  | 
fun next_weight w = if w + w <= max_term_order_weight then w + w else w + 1  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3008  | 
fun add_weights _ [] = I  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3009  | 
| add_weights weight syms =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3010  | 
fold (AList.update (op =) o rpair weight) syms  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3011  | 
#> add_weights (next_weight weight)  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3012  | 
(fold (union (op =) o Graph.immediate_succs graph) syms [])  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3013  | 
in  | 
| 52030 | 3014  | 
(* Sorting is not just for aesthetics: It specifies the precedence order for  | 
3015  | 
the term ordering (KBO or LPO), from smaller to larger values. *)  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58477 
diff
changeset
 | 
3016  | 
[] |> add_weights 1 (Graph.minimals graph) |> sort (int_ord o apply2 snd)  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3017  | 
end  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
3018  | 
|
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
3019  | 
end;  |