| author | wenzelm | 
| Tue, 14 Feb 2012 21:19:39 +0100 | |
| changeset 46471 | 2289a3869c88 | 
| parent 46450 | 7560930b2e06 | 
| child 46639 | d0ef1d1562d7 | 
| 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  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
|
| 
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
 | 
6  | 
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
 | 
7  | 
*)  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
|
| 46320 | 9  | 
signature ATP_PROBLEM_GENERATE =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
11  | 
  type ('a, 'b) ho_term = ('a, 'b) ATP_Problem.ho_term
 | 
| 
43136
 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 
blanchet 
parents: 
43130 
diff
changeset
 | 
12  | 
type connective = ATP_Problem.connective  | 
| 
 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 
blanchet 
parents: 
43130 
diff
changeset
 | 
13  | 
  type ('a, 'b, 'c) formula = ('a, 'b, 'c) ATP_Problem.formula
 | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
14  | 
type atp_format = ATP_Problem.atp_format  | 
| 
42709
 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 
blanchet 
parents: 
42701 
diff
changeset
 | 
15  | 
type formula_kind = ATP_Problem.formula_kind  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
16  | 
type 'a problem = 'a ATP_Problem.problem  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
17  | 
|
| 46340 | 18  | 
datatype scope = Global | Local | Assum | Chained  | 
| 46393 | 19  | 
datatype status = General | Induct | Intro | Elim | Simp | Spec_Eq  | 
| 46340 | 20  | 
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
 | 
21  | 
|
| 
44494
 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 
blanchet 
parents: 
44493 
diff
changeset
 | 
22  | 
datatype polymorphism = Polymorphic | Raw_Monomorphic | Mangled_Monomorphic  | 
| 46301 | 23  | 
datatype strictness = Strict | Non_Strict  | 
| 44811 | 24  | 
datatype granularity = All_Vars | Positively_Naked_Vars | Ghost_Type_Arg_Vars  | 
| 
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  | 
datatype type_level =  | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
26  | 
All_Types |  | 
| 46301 | 27  | 
Noninf_Nonmono_Types of strictness * granularity |  | 
| 44811 | 28  | 
Fin_Nonmono_Types of granularity |  | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
29  | 
Const_Arg_Types |  | 
| 43362 | 30  | 
No_Types  | 
| 44782 | 31  | 
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
 | 
32  | 
|
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
33  | 
val type_tag_idempotence : bool Config.T  | 
| 
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
34  | 
val type_tag_arguments : bool Config.T  | 
| 
45513
 
25388cf06437
rename the lambda translation schemes, so that they are understandable out of context
 
blanchet 
parents: 
45511 
diff
changeset
 | 
35  | 
val no_lamsN : string  | 
| 
 
25388cf06437
rename the lambda translation schemes, so that they are understandable out of context
 
blanchet 
parents: 
45511 
diff
changeset
 | 
36  | 
val hide_lamsN : string  | 
| 46365 | 37  | 
val liftingN : string  | 
38  | 
val combsN : string  | 
|
39  | 
val combs_and_liftingN : string  | 
|
40  | 
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
 | 
41  | 
val lam_liftingN : string  | 
| 
 
25388cf06437
rename the lambda translation schemes, so that they are understandable out of context
 
blanchet 
parents: 
45511 
diff
changeset
 | 
42  | 
val keep_lamsN : string  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
43  | 
val schematic_var_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
44  | 
val fixed_var_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
45  | 
val tvar_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
46  | 
val tfree_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
47  | 
val const_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
48  | 
val type_const_prefix : string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
49  | 
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
 | 
50  | 
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
 | 
51  | 
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
 | 
52  | 
val lam_lifted_poly_prefix : string  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
53  | 
val skolem_const_prefix : string  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
54  | 
val old_skolem_const_prefix : string  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
55  | 
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
 | 
56  | 
val combinator_prefix : string  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
57  | 
val type_decl_prefix : string  | 
| 
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
58  | 
val sym_decl_prefix : string  | 
| 43989 | 59  | 
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
 | 
60  | 
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
 | 
61  | 
val fact_prefix : string  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
62  | 
val conjecture_prefix : string  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
63  | 
val helper_prefix : string  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
64  | 
val class_rel_clause_prefix : string  | 
| 
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
65  | 
val arity_clause_prefix : string  | 
| 
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
66  | 
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
 | 
67  | 
val lam_fact_prefix : string  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
68  | 
val typed_helper_suffix : string  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
69  | 
val untyped_helper_suffix : string  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
70  | 
val type_tag_idempotence_helper_name : string  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
71  | 
val predicator_name : string  | 
| 
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
72  | 
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
 | 
73  | 
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
 | 
74  | 
val type_tag_name : string  | 
| 46435 | 75  | 
val native_type_prefix : string  | 
| 43174 | 76  | 
val prefixed_predicator_name : string  | 
| 
43130
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
77  | 
val prefixed_app_op_name : string  | 
| 
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
78  | 
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
 | 
79  | 
val ascii_of : string -> string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
80  | 
val unascii_of : string -> string  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
81  | 
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
 | 
82  | 
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
 | 
83  | 
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
 | 
84  | 
val invert_const : string -> string  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
85  | 
val unproxify_const : string -> string  | 
| 43093 | 86  | 
val new_skolem_var_name_from_const : string -> string  | 
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
87  | 
val atp_irrelevant_consts : string list  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
88  | 
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
 | 
89  | 
val is_type_enc_higher_order : type_enc -> bool  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
90  | 
val polymorphism_of_type_enc : type_enc -> polymorphism  | 
| 
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
91  | 
val level_of_type_enc : type_enc -> type_level  | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
92  | 
val is_type_enc_quasi_sound : type_enc -> bool  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
93  | 
val is_type_enc_fairly_sound : type_enc -> bool  | 
| 46301 | 94  | 
val type_enc_from_string : strictness -> string -> type_enc  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
95  | 
val adjust_type_enc : atp_format -> type_enc -> type_enc  | 
| 
43136
 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 
blanchet 
parents: 
43130 
diff
changeset
 | 
96  | 
val mk_aconns :  | 
| 
 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 
blanchet 
parents: 
43130 
diff
changeset
 | 
97  | 
    connective -> ('a, 'b, 'c) formula list -> ('a, 'b, 'c) formula
 | 
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
98  | 
val unmangled_const : string -> string * (string, 'b) ho_term list  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
99  | 
val unmangled_const_name : string -> string list  | 
| 43194 | 100  | 
val helper_table : ((string * bool) * thm list) list  | 
| 45514 | 101  | 
val trans_lams_from_string :  | 
102  | 
Proof.context -> type_enc -> string -> term list -> term list * term list  | 
|
| 
43501
 
0e422a84d0b2
don't change the way helpers are generated for the exporter's sake
 
blanchet 
parents: 
43496 
diff
changeset
 | 
103  | 
val factsN : string  | 
| 
40059
 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 
blanchet 
parents: 
39975 
diff
changeset
 | 
104  | 
val prepare_atp_problem :  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
105  | 
Proof.context -> atp_format -> formula_kind -> formula_kind -> type_enc  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
106  | 
-> bool -> string -> bool -> bool -> bool -> term list -> term  | 
| 46340 | 107  | 
-> ((string * stature) * term) list  | 
108  | 
-> string problem * string Symtab.table * (string * stature) list vector  | 
|
| 45551 | 109  | 
* (string * term) list * int Symtab.table  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
110  | 
val atp_problem_relevance_weights : string problem -> (string * real) list  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
111  | 
val atp_problem_kbo_weights : string problem -> (string * int) list  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
112  | 
end;  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
113  | 
|
| 46320 | 114  | 
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
 | 
115  | 
struct  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
116  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
117  | 
open ATP_Util  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
118  | 
open ATP_Problem  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
119  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
120  | 
type name = string * string  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
121  | 
|
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
122  | 
val type_tag_idempotence =  | 
| 
44540
 
968115499161
change default for generation of tag idempotence and tag argument equations
 
blanchet 
parents: 
44508 
diff
changeset
 | 
123  | 
  Attrib.setup_config_bool @{binding atp_type_tag_idempotence} (K false)
 | 
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
124  | 
val type_tag_arguments =  | 
| 
44540
 
968115499161
change default for generation of tag idempotence and tag argument equations
 
blanchet 
parents: 
44508 
diff
changeset
 | 
125  | 
  Attrib.setup_config_bool @{binding atp_type_tag_arguments} (K false)
 | 
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
126  | 
|
| 45516 | 127  | 
val no_lamsN = "no_lams" (* used internally; undocumented *)  | 
| 
45513
 
25388cf06437
rename the lambda translation schemes, so that they are understandable out of context
 
blanchet 
parents: 
45511 
diff
changeset
 | 
128  | 
val hide_lamsN = "hide_lams"  | 
| 46365 | 129  | 
val liftingN = "lifting"  | 
130  | 
val combsN = "combs"  | 
|
131  | 
val combs_and_liftingN = "combs_and_lifting"  | 
|
132  | 
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
 | 
133  | 
val keep_lamsN = "keep_lams"  | 
| 46365 | 134  | 
val lam_liftingN = "lam_lifting" (* legacy *)  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
135  | 
|
| 45779 | 136  | 
(* It's still unclear whether all TFF1 implementations will support type  | 
137  | 
signatures such as "!>[A : $tType] : $o", with ghost type variables. *)  | 
|
138  | 
val avoid_first_order_ghost_type_vars = false  | 
|
| 44622 | 139  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
140  | 
val bound_var_prefix = "B_"  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
141  | 
val all_bound_var_prefix = "BA_"  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
142  | 
val exist_bound_var_prefix = "BE_"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
143  | 
val schematic_var_prefix = "V_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
144  | 
val fixed_var_prefix = "v_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
145  | 
val tvar_prefix = "T_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
146  | 
val tfree_prefix = "t_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
147  | 
val const_prefix = "c_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
148  | 
val type_const_prefix = "tc_"  | 
| 46435 | 149  | 
val native_type_prefix = "n_"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
150  | 
val class_prefix = "cl_"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
151  | 
|
| 45509 | 152  | 
(* Freshness almost guaranteed! *)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
153  | 
val atp_prefix = "ATP" ^ Long_Name.separator  | 
| 45509 | 154  | 
val atp_weak_prefix = "ATP:"  | 
155  | 
||
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
156  | 
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
 | 
157  | 
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
 | 
158  | 
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
 | 
159  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
160  | 
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
 | 
161  | 
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
 | 
162  | 
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
 | 
163  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
164  | 
val combinator_prefix = "COMB"  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
165  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
166  | 
val type_decl_prefix = "ty_"  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
167  | 
val sym_decl_prefix = "sy_"  | 
| 43989 | 168  | 
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
 | 
169  | 
val tags_sym_formula_prefix = "tsy_"  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
170  | 
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
 | 
171  | 
val fact_prefix = "fact_"  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
172  | 
val conjecture_prefix = "conj_"  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
173  | 
val helper_prefix = "help_"  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
174  | 
val class_rel_clause_prefix = "clar_"  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
175  | 
val arity_clause_prefix = "arity_"  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
176  | 
val tfree_clause_prefix = "tfree_"  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
177  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
178  | 
val lam_fact_prefix = "ATP.lambda_"  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
179  | 
val typed_helper_suffix = "_T"  | 
| 
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
180  | 
val untyped_helper_suffix = "_U"  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
181  | 
val type_tag_idempotence_helper_name = helper_prefix ^ "ti_idem"  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
182  | 
|
| 
44491
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
183  | 
val predicator_name = "pp"  | 
| 
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
184  | 
val app_op_name = "aa"  | 
| 
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
185  | 
val type_guard_name = "gg"  | 
| 
 
ba22ed224b20
fixed bang encoding detection of which types to encode
 
blanchet 
parents: 
44463 
diff
changeset
 | 
186  | 
val type_tag_name = "tt"  | 
| 
42531
 
a462dbaa584f
added more rudimentary type support to Sledgehammer's ATP encoding
 
blanchet 
parents: 
42530 
diff
changeset
 | 
187  | 
|
| 43174 | 188  | 
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
 | 
189  | 
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
 | 
190  | 
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
 | 
191  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
192  | 
(*Escaping of special characters.  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
193  | 
Alphanumeric characters are left unchanged.  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
194  | 
The character _ goes to __  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
195  | 
Characters in the range ASCII space to / go to _A to _P, respectively.  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
196  | 
Other characters go to _nnn where nnn is the decimal ASCII code.*)  | 
| 43093 | 197  | 
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
 | 
198  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
199  | 
fun stringN_of_int 0 _ = ""  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
200  | 
| stringN_of_int k n =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
201  | 
stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
202  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
203  | 
fun ascii_of_char c =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
204  | 
if Char.isAlphaNum c then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
205  | 
String.str c  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
206  | 
else if c = #"_" then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
207  | 
"__"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
208  | 
else if #" " <= c andalso c <= #"/" then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
209  | 
"_" ^ 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
 | 
210  | 
else  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
211  | 
(* fixed width, in case more digits follow *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
212  | 
"_" ^ stringN_of_int 3 (Char.ord c)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
213  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
214  | 
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
 | 
215  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
216  | 
(** 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
 | 
217  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
218  | 
(* 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
 | 
219  | 
thread. Also, the errors are impossible. *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
220  | 
val unascii_of =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
221  | 
let  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
222  | 
fun un rcs [] = String.implode(rev rcs)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
223  | 
| un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
224  | 
(* 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
 | 
225  | 
| un rcs (#"_" :: #"_" :: cs) = un (#"_" :: rcs) cs  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
226  | 
| un rcs (#"_" :: c :: cs) =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
227  | 
if #"A" <= c andalso c<= #"P" then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
228  | 
(* translation of #" " to #"/" *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
229  | 
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
 | 
230  | 
else  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
231  | 
let val digits = List.take (c :: cs, 3) handle General.Subscript => [] in  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
232  | 
case Int.fromString (String.implode digits) of  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
233  | 
SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
234  | 
| NONE => un (c :: #"_" :: rcs) cs (* ERROR *)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
235  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
236  | 
| 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
 | 
237  | 
in un [] o String.explode end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
238  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
239  | 
(* 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
 | 
240  | 
un-ASCII'd. *)  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
241  | 
fun unprefix_and_unascii s1 s =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
242  | 
if String.isPrefix s1 s then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
243  | 
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
 | 
244  | 
else  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
245  | 
NONE  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
246  | 
|
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
247  | 
val proxy_table =  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
248  | 
  [("c_False", (@{const_name False}, (@{thm fFalse_def},
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
249  | 
       ("fFalse", @{const_name ATP.fFalse})))),
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
250  | 
   ("c_True", (@{const_name True}, (@{thm fTrue_def},
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
251  | 
       ("fTrue", @{const_name ATP.fTrue})))),
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
252  | 
   ("c_Not", (@{const_name Not}, (@{thm fNot_def},
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
253  | 
       ("fNot", @{const_name ATP.fNot})))),
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
254  | 
   ("c_conj", (@{const_name conj}, (@{thm fconj_def},
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
255  | 
       ("fconj", @{const_name ATP.fconj})))),
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
256  | 
   ("c_disj", (@{const_name disj}, (@{thm fdisj_def},
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
257  | 
       ("fdisj", @{const_name ATP.fdisj})))),
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
258  | 
   ("c_implies", (@{const_name implies}, (@{thm fimplies_def},
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
259  | 
       ("fimplies", @{const_name ATP.fimplies})))),
 | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
260  | 
   ("equal", (@{const_name HOL.eq}, (@{thm fequal_def},
 | 
| 43678 | 261  | 
       ("fequal", @{const_name ATP.fequal})))),
 | 
262  | 
   ("c_All", (@{const_name All}, (@{thm fAll_def},
 | 
|
263  | 
       ("fAll", @{const_name ATP.fAll})))),
 | 
|
264  | 
   ("c_Ex", (@{const_name Ex}, (@{thm fEx_def},
 | 
|
265  | 
       ("fEx", @{const_name ATP.fEx}))))]
 | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
266  | 
|
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
267  | 
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
 | 
268  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
269  | 
(* 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
 | 
270  | 
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
 | 
271  | 
val const_trans_table =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
272  | 
  [(@{type_name Product_Type.prod}, "prod"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
273  | 
   (@{type_name Sum_Type.sum}, "sum"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
274  | 
   (@{const_name False}, "False"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
275  | 
   (@{const_name True}, "True"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
276  | 
   (@{const_name Not}, "Not"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
277  | 
   (@{const_name conj}, "conj"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
278  | 
   (@{const_name disj}, "disj"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
279  | 
   (@{const_name implies}, "implies"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
280  | 
   (@{const_name HOL.eq}, "equal"),
 | 
| 43678 | 281  | 
   (@{const_name All}, "All"),
 | 
282  | 
   (@{const_name Ex}, "Ex"),
 | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
283  | 
   (@{const_name If}, "If"),
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
284  | 
   (@{const_name Set.member}, "member"),
 | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
285  | 
   (@{const_name Meson.COMBI}, combinator_prefix ^ "I"),
 | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
286  | 
   (@{const_name Meson.COMBK}, combinator_prefix ^ "K"),
 | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
287  | 
   (@{const_name Meson.COMBB}, combinator_prefix ^ "B"),
 | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
288  | 
   (@{const_name Meson.COMBC}, combinator_prefix ^ "C"),
 | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
289  | 
   (@{const_name Meson.COMBS}, combinator_prefix ^ "S")]
 | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
290  | 
|> Symtab.make  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
291  | 
|> 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
 | 
292  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
293  | 
(* 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
 | 
294  | 
val const_trans_table_inv =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
295  | 
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
 | 
296  | 
val const_trans_table_unprox =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
297  | 
Symtab.empty  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
298  | 
|> 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
 | 
299  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
300  | 
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
 | 
301  | 
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
 | 
302  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
303  | 
fun lookup_const c =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
304  | 
case Symtab.lookup const_trans_table c of  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
305  | 
SOME c' => c'  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
306  | 
| NONE => ascii_of c  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
307  | 
|
| 43622 | 308  | 
fun ascii_of_indexname (v, 0) = ascii_of v  | 
309  | 
| 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
 | 
310  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
311  | 
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
 | 
312  | 
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
 | 
313  | 
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
 | 
314  | 
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
 | 
315  | 
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
 | 
316  | 
|
| 43622 | 317  | 
fun make_schematic_type_var (x, i) =  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
318  | 
tvar_prefix ^ (ascii_of_indexname (unprefix "'" x, i))  | 
| 43622 | 319  | 
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (unprefix "'" x))  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
320  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
321  | 
(* "HOL.eq" and choice are mapped to the ATP's equivalents *)  | 
| 44587 | 322  | 
local  | 
323  | 
val choice_const = (fst o dest_Const o HOLogic.choice_const) Term.dummyT  | 
|
324  | 
fun default c = const_prefix ^ lookup_const c  | 
|
325  | 
in  | 
|
326  | 
  fun make_fixed_const _ @{const_name HOL.eq} = tptp_old_equal
 | 
|
| 44754 | 327  | 
| make_fixed_const (SOME (THF (_, _, THF_With_Choice))) c =  | 
328  | 
if c = choice_const then tptp_choice else default c  | 
|
| 44587 | 329  | 
| make_fixed_const _ c = default c  | 
330  | 
end  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
331  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
332  | 
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
 | 
333  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
334  | 
fun make_type_class clas = class_prefix ^ ascii_of clas  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
335  | 
|
| 43093 | 336  | 
fun new_skolem_var_name_from_const s =  | 
337  | 
let val ss = s |> space_explode Long_Name.separator in  | 
|
338  | 
nth ss (length ss - 2)  | 
|
339  | 
end  | 
|
340  | 
||
| 
43248
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
341  | 
(* 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
 | 
342  | 
handled specially via "fFalse", "fTrue", ..., "fequal". *)  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
343  | 
val atp_irrelevant_consts =  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
344  | 
  [@{const_name False}, @{const_name True}, @{const_name Not},
 | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
345  | 
   @{const_name conj}, @{const_name disj}, @{const_name implies},
 | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
346  | 
   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
 | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
347  | 
|
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
348  | 
val atp_monomorph_bad_consts =  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
349  | 
atp_irrelevant_consts @  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
350  | 
(* These are ignored anyway by the relevance filter (unless they appear in  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
351  | 
higher-order places) but not by the monomorphizer. *)  | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
352  | 
  [@{const_name all}, @{const_name "==>"}, @{const_name "=="},
 | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
353  | 
   @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
 | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
354  | 
   @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}]
 | 
| 
 
69375eaa9016
more suitable implementation of "schematic_consts_of" for monomorphizer, for ATPs
 
blanchet 
parents: 
43222 
diff
changeset
 | 
355  | 
|
| 
43258
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
356  | 
fun add_schematic_const (x as (_, T)) =  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
357  | 
Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
358  | 
val add_schematic_consts_of =  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
359  | 
Term.fold_aterms (fn Const (x as (s, _)) =>  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
360  | 
not (member (op =) atp_monomorph_bad_consts s)  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
361  | 
? add_schematic_const x  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
362  | 
| _ => I)  | 
| 
 
956895f99904
slightly faster/cleaner accumulation of polymorphic consts
 
blanchet 
parents: 
43248 
diff
changeset
 | 
363  | 
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
 | 
364  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
365  | 
(** 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
 | 
366  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
367  | 
(** Isabelle arities **)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
368  | 
|
| 44625 | 369  | 
type arity_atom = name * name * name list  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
370  | 
|
| 43263 | 371  | 
val type_class = the_single @{sort type}
 | 
372  | 
||
| 43086 | 373  | 
type arity_clause =  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
374  | 
  {name : string,
 | 
| 44625 | 375  | 
prem_atoms : arity_atom list,  | 
376  | 
concl_atom : arity_atom}  | 
|
377  | 
||
378  | 
fun add_prem_atom tvar =  | 
|
379  | 
fold (fn s => s <> type_class ? cons (`make_type_class s, `I tvar, []))  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
380  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
381  | 
(* Arity of type constructor "tcon :: (arg1, ..., argN) res" *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
382  | 
fun make_axiom_arity_clause (tcons, name, (cls, args)) =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
383  | 
let  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
384  | 
val tvars = map (prefix tvar_prefix o string_of_int) (1 upto length args)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
385  | 
val tvars_srts = ListPair.zip (tvars, args)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
386  | 
in  | 
| 43086 | 387  | 
    {name = name,
 | 
| 44625 | 388  | 
prem_atoms = [] |> fold (uncurry add_prem_atom) tvars_srts,  | 
389  | 
concl_atom = (`make_type_class cls, `make_fixed_type_const tcons,  | 
|
390  | 
tvars ~~ tvars)}  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
391  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
392  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
393  | 
fun arity_clause _ _ (_, []) = []  | 
| 43495 | 394  | 
  | arity_clause seen n (tcons, ("HOL.type", _) :: ars) =  (* ignore *)
 | 
395  | 
arity_clause seen n (tcons, ars)  | 
|
396  | 
| arity_clause seen n (tcons, (ar as (class, _)) :: ars) =  | 
|
397  | 
if member (op =) seen class then  | 
|
398  | 
(* multiple arities for the same (tycon, class) pair *)  | 
|
399  | 
make_axiom_arity_clause (tcons,  | 
|
400  | 
lookup_const tcons ^ "___" ^ ascii_of class ^ "_" ^ string_of_int n,  | 
|
401  | 
ar) ::  | 
|
402  | 
arity_clause seen (n + 1) (tcons, ars)  | 
|
403  | 
else  | 
|
404  | 
make_axiom_arity_clause (tcons, lookup_const tcons ^ "___" ^  | 
|
405  | 
ascii_of class, ar) ::  | 
|
406  | 
arity_clause (class :: seen) n (tcons, ars)  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
407  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
408  | 
fun multi_arity_clause [] = []  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
409  | 
| multi_arity_clause ((tcons, ars) :: tc_arlists) =  | 
| 44772 | 410  | 
arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
411  | 
|
| 43622 | 412  | 
(* Generate all pairs (tycon, class, sorts) such that tycon belongs to class in  | 
413  | 
theory thy provided its arguments have the corresponding sorts. *)  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
414  | 
fun type_class_pairs thy tycons classes =  | 
| 43093 | 415  | 
let  | 
416  | 
val alg = Sign.classes_of thy  | 
|
417  | 
fun domain_sorts tycon = Sorts.mg_domain alg tycon o single  | 
|
418  | 
fun add_class tycon class =  | 
|
419  | 
cons (class, domain_sorts tycon class)  | 
|
420  | 
handle Sorts.CLASS_ERROR _ => I  | 
|
421  | 
fun try_classes tycon = (tycon, fold (add_class tycon) classes [])  | 
|
422  | 
in map try_classes tycons end  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
423  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
424  | 
(*Proving one (tycon, class) membership may require proving others, so iterate.*)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
425  | 
fun iter_type_class_pairs _ _ [] = ([], [])  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
426  | 
| iter_type_class_pairs thy tycons classes =  | 
| 43263 | 427  | 
let  | 
428  | 
fun maybe_insert_class s =  | 
|
429  | 
(s <> type_class andalso not (member (op =) classes s))  | 
|
430  | 
? insert (op =) s  | 
|
431  | 
val cpairs = type_class_pairs thy tycons classes  | 
|
432  | 
val newclasses =  | 
|
433  | 
[] |> fold (fold (fold (fold maybe_insert_class) o snd) o snd) cpairs  | 
|
434  | 
val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses  | 
|
| 43266 | 435  | 
in (classes' @ classes, union (op =) cpairs' cpairs) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
436  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
437  | 
fun make_arity_clauses thy tycons =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
438  | 
iter_type_class_pairs thy tycons ##> multi_arity_clause  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
439  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
440  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
441  | 
(** Isabelle class relations **)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
442  | 
|
| 43086 | 443  | 
type class_rel_clause =  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
444  | 
  {name : string,
 | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
445  | 
subclass : name,  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
446  | 
superclass : name}  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
447  | 
|
| 43622 | 448  | 
(* Generate all pairs (sub, super) such that sub is a proper subclass of super  | 
449  | 
in theory "thy". *)  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
450  | 
fun class_pairs _ [] _ = []  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
451  | 
| class_pairs thy subs supers =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
452  | 
let  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
453  | 
val class_less = Sorts.class_less (Sign.classes_of thy)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
454  | 
fun add_super sub super = class_less (sub, super) ? cons (sub, super)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
455  | 
fun add_supers sub = fold (add_super sub) supers  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
456  | 
in fold add_supers subs [] end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
457  | 
|
| 43622 | 458  | 
fun make_class_rel_clause (sub, super) =  | 
459  | 
  {name = sub ^ "_" ^ super, subclass = `make_type_class sub,
 | 
|
| 43086 | 460  | 
superclass = `make_type_class super}  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
461  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
462  | 
fun make_class_rel_clauses thy subs supers =  | 
| 43093 | 463  | 
map make_class_rel_clause (class_pairs thy subs supers)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
464  | 
|
| 43859 | 465  | 
(* intermediate terms *)  | 
466  | 
datatype iterm =  | 
|
467  | 
IConst of name * typ * typ list |  | 
|
468  | 
IVar of name * typ |  | 
|
469  | 
IApp of iterm * iterm |  | 
|
470  | 
IAbs of (name * typ) * iterm  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
471  | 
|
| 43859 | 472  | 
fun ityp_of (IConst (_, T, _)) = T  | 
473  | 
| ityp_of (IVar (_, T)) = T  | 
|
474  | 
| ityp_of (IApp (t1, _)) = snd (dest_funT (ityp_of t1))  | 
|
475  | 
| 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
 | 
476  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
477  | 
(*gets the head of a combinator application, along with the list of arguments*)  | 
| 43859 | 478  | 
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
 | 
479  | 
let  | 
| 43859 | 480  | 
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
 | 
481  | 
| stripc x = x  | 
| 
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
482  | 
in stripc (u, []) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
483  | 
|
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
484  | 
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
 | 
485  | 
|
| 45509 | 486  | 
val tvar_a_str = "'a"  | 
487  | 
val tvar_a = TVar ((tvar_a_str, 0), HOLogic.typeS)  | 
|
488  | 
val tvar_a_name = (make_schematic_type_var (tvar_a_str, 0), tvar_a_str)  | 
|
489  | 
val itself_name = `make_fixed_type_const @{type_name itself}
 | 
|
490  | 
val TYPE_name = `(make_fixed_const NONE) @{const_name TYPE}
 | 
|
491  | 
val tvar_a_atype = AType (tvar_a_name, [])  | 
|
492  | 
val a_itself_atype = AType (itself_name, [tvar_a_atype])  | 
|
493  | 
||
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
494  | 
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
 | 
495  | 
[new_skolem_const_prefix, s, string_of_int num_T_args]  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
496  | 
|> space_implode Long_Name.separator  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
497  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
498  | 
fun robust_const_type thy s =  | 
| 45509 | 499  | 
if s = app_op_name then  | 
500  | 
    Logic.varifyT_global @{typ "('a => 'b) => 'a => 'b"}
 | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
501  | 
else if String.isPrefix lam_lifted_prefix s then  | 
| 45509 | 502  | 
    Logic.varifyT_global @{typ "'a => 'b"}
 | 
503  | 
else  | 
|
504  | 
(* Old Skolems throw a "TYPE" exception here, which will be caught. *)  | 
|
505  | 
s |> Sign.the_const_type thy  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
506  | 
|
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
507  | 
(* This function only makes sense if "T" is as general as possible. *)  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
508  | 
fun robust_const_typargs thy (s, T) =  | 
| 45509 | 509  | 
if s = app_op_name then  | 
510  | 
let val (T1, T2) = T |> domain_type |> dest_funT in [T1, T2] end  | 
|
511  | 
else if String.isPrefix old_skolem_const_prefix s then  | 
|
512  | 
[] |> 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
 | 
513  | 
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
 | 
514  | 
if String.isPrefix lam_lifted_poly_prefix s then  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
515  | 
let val (T1, T2) = T |> dest_funT in [T1, T2] end  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
516  | 
else  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
517  | 
[]  | 
| 45509 | 518  | 
else  | 
519  | 
(s, T) |> Sign.const_typargs thy  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
520  | 
|
| 43859 | 521  | 
(* Converts an Isabelle/HOL term (with combinators) into an intermediate term.  | 
522  | 
Also accumulates sort infomation. *)  | 
|
| 44495 | 523  | 
fun iterm_from_term thy format bs (P $ Q) =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
524  | 
let  | 
| 44495 | 525  | 
val (P', P_atomics_Ts) = iterm_from_term thy format bs P  | 
526  | 
val (Q', Q_atomics_Ts) = iterm_from_term thy format bs Q  | 
|
| 43859 | 527  | 
in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end  | 
| 44495 | 528  | 
| iterm_from_term thy format _ (Const (c, T)) =  | 
529  | 
(IConst (`(make_fixed_const (SOME format)) c, T,  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
530  | 
robust_const_typargs thy (c, T)),  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
531  | 
atomic_types_of T)  | 
| 44495 | 532  | 
| iterm_from_term _ _ _ (Free (s, T)) =  | 
| 45509 | 533  | 
(IConst (`make_fixed_var s, T, []), atomic_types_of T)  | 
| 44495 | 534  | 
| iterm_from_term _ format _ (Var (v as (s, _), T)) =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
535  | 
(if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
536  | 
let  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
537  | 
val Ts = T |> strip_type |> swap |> op ::  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
538  | 
val s' = new_skolem_const_name s (length Ts)  | 
| 44495 | 539  | 
in IConst (`(make_fixed_const (SOME format)) s', T, Ts) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
540  | 
else  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
541  | 
IVar ((make_schematic_var v, s), T), atomic_types_of T)  | 
| 44495 | 542  | 
| iterm_from_term _ _ bs (Bound j) =  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
543  | 
nth bs j |> (fn (_, (name, T)) => (IConst (name, T, []), atomic_types_of T))  | 
| 44495 | 544  | 
| iterm_from_term thy format bs (Abs (s, T, t)) =  | 
| 43678 | 545  | 
let  | 
546  | 
fun vary s = s |> AList.defined (op =) bs s ? vary o Symbol.bump_string  | 
|
547  | 
val s = vary s  | 
|
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
548  | 
val name = `make_bound_var s  | 
| 44495 | 549  | 
val (tm, atomic_Ts) = iterm_from_term thy format ((s, (name, T)) :: bs) t  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
550  | 
in (IAbs ((name, T), tm), union (op =) atomic_Ts (atomic_types_of T)) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
551  | 
|
| 46340 | 552  | 
datatype scope = Global | Local | Assum | Chained  | 
| 46393 | 553  | 
datatype status = General | Induct | Intro | Elim | Simp | Spec_Eq  | 
| 46340 | 554  | 
type stature = scope * status  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
555  | 
|
| 
43624
 
de026aecab9b
cleaner handling of higher-order simple types, so that it's also possible to use first-order simple types with LEO-II and company
 
blanchet 
parents: 
43623 
diff
changeset
 | 
556  | 
datatype order = First_Order | Higher_Order  | 
| 
44494
 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 
blanchet 
parents: 
44493 
diff
changeset
 | 
557  | 
datatype polymorphism = Polymorphic | Raw_Monomorphic | Mangled_Monomorphic  | 
| 46301 | 558  | 
datatype strictness = Strict | Non_Strict  | 
| 44811 | 559  | 
datatype granularity = All_Vars | Positively_Naked_Vars | Ghost_Type_Arg_Vars  | 
| 
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
 | 
560  | 
datatype type_level =  | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
561  | 
All_Types |  | 
| 46301 | 562  | 
Noninf_Nonmono_Types of strictness * granularity |  | 
| 44811 | 563  | 
Fin_Nonmono_Types of granularity |  | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
564  | 
Const_Arg_Types |  | 
| 43362 | 565  | 
No_Types  | 
| 
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
 | 
566  | 
|
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
567  | 
datatype type_enc =  | 
| 
44591
 
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
 
blanchet 
parents: 
44589 
diff
changeset
 | 
568  | 
Simple_Types of order * polymorphism * type_level |  | 
| 44768 | 569  | 
Guards of polymorphism * type_level |  | 
570  | 
Tags of polymorphism * type_level  | 
|
571  | 
||
572  | 
fun is_type_enc_higher_order (Simple_Types (Higher_Order, _, _)) = true  | 
|
573  | 
| is_type_enc_higher_order _ = false  | 
|
574  | 
||
575  | 
fun polymorphism_of_type_enc (Simple_Types (_, poly, _)) = poly  | 
|
576  | 
| polymorphism_of_type_enc (Guards (poly, _)) = poly  | 
|
577  | 
| polymorphism_of_type_enc (Tags (poly, _)) = poly  | 
|
578  | 
||
579  | 
fun level_of_type_enc (Simple_Types (_, _, level)) = level  | 
|
580  | 
| level_of_type_enc (Guards (_, level)) = level  | 
|
581  | 
| level_of_type_enc (Tags (_, level)) = level  | 
|
582  | 
||
| 44811 | 583  | 
fun granularity_of_type_level (Noninf_Nonmono_Types (_, grain)) = grain  | 
584  | 
| granularity_of_type_level (Fin_Nonmono_Types grain) = grain  | 
|
585  | 
| granularity_of_type_level _ = All_Vars  | 
|
| 44768 | 586  | 
|
587  | 
fun is_type_level_quasi_sound All_Types = true  | 
|
588  | 
| is_type_level_quasi_sound (Noninf_Nonmono_Types _) = true  | 
|
589  | 
| is_type_level_quasi_sound _ = false  | 
|
590  | 
val is_type_enc_quasi_sound = is_type_level_quasi_sound o level_of_type_enc  | 
|
591  | 
||
592  | 
fun is_type_level_fairly_sound (Fin_Nonmono_Types _) = true  | 
|
593  | 
| is_type_level_fairly_sound level = is_type_level_quasi_sound level  | 
|
594  | 
val is_type_enc_fairly_sound = is_type_level_fairly_sound o level_of_type_enc  | 
|
595  | 
||
596  | 
fun is_type_level_monotonicity_based (Noninf_Nonmono_Types _) = true  | 
|
597  | 
| is_type_level_monotonicity_based (Fin_Nonmono_Types _) = true  | 
|
598  | 
| is_type_level_monotonicity_based _ = false  | 
|
| 
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
 | 
599  | 
|
| 44785 | 600  | 
(* "_query", "_bang", and "_at" are for the ASCII-challenged Metis and  | 
601  | 
Mirabelle. *)  | 
|
602  | 
val queries = ["?", "_query"]  | 
|
603  | 
val bangs = ["!", "_bang"]  | 
|
604  | 
val ats = ["@", "_at"]  | 
|
605  | 
||
| 
42689
 
e38590764c34
versions of ! and ? for the ASCII-challenged Mirabelle
 
blanchet 
parents: 
42688 
diff
changeset
 | 
606  | 
fun try_unsuffixes ss s =  | 
| 
 
e38590764c34
versions of ! and ? for the ASCII-challenged Mirabelle
 
blanchet 
parents: 
42688 
diff
changeset
 | 
607  | 
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
 | 
608  | 
|
| 44785 | 609  | 
fun try_nonmono constr suffixes fallback s =  | 
610  | 
case try_unsuffixes suffixes s of  | 
|
611  | 
SOME s =>  | 
|
612  | 
(case try_unsuffixes suffixes s of  | 
|
| 44811 | 613  | 
SOME s => (constr Positively_Naked_Vars, s)  | 
| 44785 | 614  | 
| NONE =>  | 
615  | 
case try_unsuffixes ats s of  | 
|
| 44811 | 616  | 
SOME s => (constr Ghost_Type_Arg_Vars, s)  | 
617  | 
| NONE => (constr All_Vars, s))  | 
|
| 44785 | 618  | 
| NONE => fallback s  | 
| 44768 | 619  | 
|
| 46301 | 620  | 
fun type_enc_from_string strictness s =  | 
| 42722 | 621  | 
(case try (unprefix "poly_") s of  | 
622  | 
SOME s => (SOME Polymorphic, 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
 | 
623  | 
| NONE =>  | 
| 
44494
 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 
blanchet 
parents: 
44493 
diff
changeset
 | 
624  | 
case try (unprefix "raw_mono_") s of  | 
| 
 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 
blanchet 
parents: 
44493 
diff
changeset
 | 
625  | 
SOME s => (SOME Raw_Monomorphic, s)  | 
| 42722 | 626  | 
| NONE =>  | 
| 
44494
 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 
blanchet 
parents: 
44493 
diff
changeset
 | 
627  | 
case try (unprefix "mono_") s of  | 
| 42722 | 628  | 
SOME s => (SOME Mangled_Monomorphic, s)  | 
629  | 
| NONE => (NONE, s))  | 
|
| 44785 | 630  | 
||> (pair All_Types  | 
631  | 
|> try_nonmono Fin_Nonmono_Types bangs  | 
|
| 46301 | 632  | 
|> try_nonmono (curry Noninf_Nonmono_Types strictness) queries)  | 
| 44768 | 633  | 
|> (fn (poly, (level, core)) =>  | 
634  | 
case (core, (poly, level)) of  | 
|
| 46435 | 635  | 
           ("native", (SOME poly, _)) =>
 | 
| 44742 | 636  | 
(case (poly, level) of  | 
637  | 
(Polymorphic, All_Types) =>  | 
|
638  | 
Simple_Types (First_Order, Polymorphic, All_Types)  | 
|
639  | 
| (Mangled_Monomorphic, _) =>  | 
|
| 44811 | 640  | 
if granularity_of_type_level level = All_Vars then  | 
| 44768 | 641  | 
Simple_Types (First_Order, Mangled_Monomorphic, level)  | 
642  | 
else  | 
|
643  | 
raise Same.SAME  | 
|
| 44742 | 644  | 
| _ => raise Same.SAME)  | 
| 46435 | 645  | 
         | ("native_higher", (SOME poly, _)) =>
 | 
| 
44591
 
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
 
blanchet 
parents: 
44589 
diff
changeset
 | 
646  | 
(case (poly, level) of  | 
| 44754 | 647  | 
(Polymorphic, All_Types) =>  | 
648  | 
Simple_Types (Higher_Order, Polymorphic, All_Types)  | 
|
649  | 
| (_, Noninf_Nonmono_Types _) => raise Same.SAME  | 
|
| 44742 | 650  | 
| (Mangled_Monomorphic, _) =>  | 
| 44811 | 651  | 
if granularity_of_type_level level = All_Vars then  | 
| 44768 | 652  | 
Simple_Types (Higher_Order, Mangled_Monomorphic, level)  | 
653  | 
else  | 
|
654  | 
raise Same.SAME  | 
|
| 44742 | 655  | 
| _ => raise Same.SAME)  | 
| 44810 | 656  | 
         | ("guards", (SOME poly, _)) =>
 | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
657  | 
if poly = Mangled_Monomorphic andalso  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
658  | 
granularity_of_type_level level = Ghost_Type_Arg_Vars then  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
659  | 
raise Same.SAME  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
660  | 
else  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
661  | 
Guards (poly, level)  | 
| 44810 | 662  | 
         | ("tags", (SOME poly, _)) =>
 | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
663  | 
if granularity_of_type_level level = Ghost_Type_Arg_Vars then  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
664  | 
raise Same.SAME  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
665  | 
else  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
666  | 
Tags (poly, level)  | 
| 44768 | 667  | 
         | ("args", (SOME poly, All_Types (* naja *))) =>
 | 
668  | 
Guards (poly, Const_Arg_Types)  | 
|
669  | 
         | ("erased", (NONE, All_Types (* naja *))) =>
 | 
|
670  | 
Guards (Polymorphic, No_Types)  | 
|
| 
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
 | 
671  | 
| _ => raise Same.SAME)  | 
| 44785 | 672  | 
  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
 | 
673  | 
|
| 44754 | 674  | 
fun adjust_type_enc (THF (TPTP_Monomorphic, _, _))  | 
675  | 
(Simple_Types (order, _, level)) =  | 
|
| 
44591
 
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
 
blanchet 
parents: 
44589 
diff
changeset
 | 
676  | 
Simple_Types (order, Mangled_Monomorphic, level)  | 
| 44754 | 677  | 
| adjust_type_enc (THF _) type_enc = type_enc  | 
678  | 
| adjust_type_enc (TFF (TPTP_Monomorphic, _)) (Simple_Types (_, _, level)) =  | 
|
| 
44591
 
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
 
blanchet 
parents: 
44589 
diff
changeset
 | 
679  | 
Simple_Types (First_Order, Mangled_Monomorphic, level)  | 
| 
45303
 
bd03b08161ac
added DFG unsorted support (like in the old days)
 
blanchet 
parents: 
45301 
diff
changeset
 | 
680  | 
| adjust_type_enc (DFG DFG_Sorted) (Simple_Types (_, _, level)) =  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
681  | 
Simple_Types (First_Order, Mangled_Monomorphic, level)  | 
| 44754 | 682  | 
| adjust_type_enc (TFF _) (Simple_Types (_, poly, level)) =  | 
| 
44591
 
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
 
blanchet 
parents: 
44589 
diff
changeset
 | 
683  | 
Simple_Types (First_Order, poly, level)  | 
| 
 
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
 
blanchet 
parents: 
44589 
diff
changeset
 | 
684  | 
| adjust_type_enc format (Simple_Types (_, poly, level)) =  | 
| 44768 | 685  | 
adjust_type_enc format (Guards (poly, level))  | 
| 
44416
 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 
blanchet 
parents: 
44410 
diff
changeset
 | 
686  | 
| adjust_type_enc CNF_UEQ (type_enc as Guards stuff) =  | 
| 
 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 
blanchet 
parents: 
44410 
diff
changeset
 | 
687  | 
(if is_type_enc_fairly_sound type_enc then Tags else Guards) stuff  | 
| 
 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 
blanchet 
parents: 
44410 
diff
changeset
 | 
688  | 
| 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
 | 
689  | 
|
| 45509 | 690  | 
fun constify_lifted (t $ u) = constify_lifted t $ constify_lifted u  | 
691  | 
| constify_lifted (Abs (s, T, t)) = Abs (s, T, constify_lifted t)  | 
|
692  | 
| 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
 | 
693  | 
(if String.isPrefix lam_lifted_prefix s then Const else Free) x  | 
| 45509 | 694  | 
| constify_lifted t = t  | 
695  | 
||
| 
46375
 
d724066ff3d0
reverted e2b1a86d59fc -- broke Metis's lambda-lifting
 
blanchet 
parents: 
46371 
diff
changeset
 | 
696  | 
(* Requires bound variables not to clash with any schematic variables (as should  | 
| 
 
d724066ff3d0
reverted e2b1a86d59fc -- broke Metis's lambda-lifting
 
blanchet 
parents: 
46371 
diff
changeset
 | 
697  | 
be the case right after lambda-lifting). *)  | 
| 
46385
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
698  | 
fun open_form unprefix (t as Const (@{const_name All}, _) $ Abs (s, T, t')) =
 | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
699  | 
(case try unprefix s of  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
700  | 
SOME s =>  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
701  | 
let  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
702  | 
val names = Name.make_context (map fst (Term.add_var_names t' []))  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
703  | 
val (s, _) = Name.variant s names  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
704  | 
in open_form unprefix (subst_bound (Var ((s, 0), T), t')) end  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
705  | 
| NONE => t)  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
706  | 
| open_form _ t = t  | 
| 
46375
 
d724066ff3d0
reverted e2b1a86d59fc -- broke Metis's lambda-lifting
 
blanchet 
parents: 
46371 
diff
changeset
 | 
707  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
708  | 
fun lift_lams_part_1 ctxt type_enc =  | 
| 
45568
 
211a6e6cbc04
move eta-contraction to before translation to Metis, to ensure everything stays in sync
 
blanchet 
parents: 
45565 
diff
changeset
 | 
709  | 
map close_form #> rpair ctxt  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
710  | 
#-> Lambda_Lifting.lift_lambdas  | 
| 45564 | 711  | 
(SOME ((if polymorphism_of_type_enc type_enc = Polymorphic then  | 
712  | 
lam_lifted_poly_prefix  | 
|
713  | 
else  | 
|
714  | 
lam_lifted_mono_prefix) ^ "_a"))  | 
|
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
715  | 
Lambda_Lifting.is_quantifier  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
716  | 
#> fst  | 
| 
46385
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
717  | 
fun lift_lams_part_2 (facts, lifted) =  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
718  | 
(map (open_form (unprefix close_form_prefix) o constify_lifted) facts,  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
719  | 
map (open_form I o constify_lifted) lifted)  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
720  | 
val lift_lams = lift_lams_part_2 ooo lift_lams_part_1  | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
721  | 
|
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
722  | 
fun intentionalize_def (Const (@{const_name All}, _) $ Abs (_, _, t)) =
 | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
723  | 
intentionalize_def t  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
724  | 
  | intentionalize_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
 | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
725  | 
let  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
726  | 
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
 | 
727  | 
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
 | 
728  | 
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
 | 
729  | 
val n = length args  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
730  | 
val arg_Ts = head_T |> binder_types |> take n |> rev  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
731  | 
val u = u |> subst_atomic (args ~~ map Bound (0 upto n - 1))  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
732  | 
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
 | 
733  | 
| intentionalize_def t = t  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
734  | 
|
| 40114 | 735  | 
type translated_formula =  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
736  | 
  {name : string,
 | 
| 46340 | 737  | 
stature : stature,  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
738  | 
kind : formula_kind,  | 
| 43859 | 739  | 
iformula : (name, typ, iterm) formula,  | 
| 
43496
 
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
 
blanchet 
parents: 
43495 
diff
changeset
 | 
740  | 
atomic_types : typ list}  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
741  | 
|
| 46340 | 742  | 
fun update_iformula f ({name, stature, kind, iformula, atomic_types}
 | 
| 43859 | 743  | 
: translated_formula) =  | 
| 46340 | 744  | 
  {name = name, stature = stature, kind = kind, iformula = f iformula,
 | 
| 42562 | 745  | 
atomic_types = atomic_types} : translated_formula  | 
| 
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
 | 
746  | 
|
| 43859 | 747  | 
fun fact_lift f ({iformula, ...} : translated_formula) = 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
 | 
748  | 
|
| 
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
 | 
749  | 
fun insert_type ctxt get_T x xs =  | 
| 
 
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
 | 
750  | 
let val T = get_T x in  | 
| 44399 | 751  | 
if exists (type_instance ctxt T o get_T) xs then xs  | 
752  | 
else x :: filter_out (type_generalization ctxt 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
 | 
753  | 
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
 | 
754  | 
|
| 
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
 | 
755  | 
(* The Booleans indicate whether all type arguments should be kept. *)  | 
| 
 
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
 | 
756  | 
datatype type_arg_policy =  | 
| 45937 | 757  | 
Explicit_Type_Args of bool (* infer_from_term_args *) |  | 
| 44771 | 758  | 
Mangled_Type_Args |  | 
| 
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
 | 
759  | 
No_Type_Args  | 
| 
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
 | 
760  | 
|
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
761  | 
fun type_arg_policy monom_constrs type_enc s =  | 
| 45315 | 762  | 
let val poly = polymorphism_of_type_enc type_enc in  | 
| 
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
 | 
763  | 
if s = type_tag_name then  | 
| 45315 | 764  | 
if poly = Mangled_Monomorphic then Mangled_Type_Args  | 
765  | 
else Explicit_Type_Args false  | 
|
| 
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
 | 
766  | 
else case type_enc of  | 
| 45315 | 767  | 
Simple_Types (_, Polymorphic, _) => Explicit_Type_Args false  | 
768  | 
| Tags (_, All_Types) => No_Type_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
 | 
769  | 
| _ =>  | 
| 
 
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
 | 
770  | 
let val level = level_of_type_enc type_enc in  | 
| 
 
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
 | 
771  | 
        if level = No_Types orelse s = @{const_name HOL.eq} orelse
 | 
| 
 
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
 | 
772  | 
(s = app_op_name andalso level = Const_Arg_Types) then  | 
| 
 
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
 | 
773  | 
No_Type_Args  | 
| 45315 | 774  | 
else if poly = Mangled_Monomorphic 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
 | 
775  | 
Mangled_Type_Args  | 
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
776  | 
else if member (op =) monom_constrs s andalso  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
777  | 
granularity_of_type_level level = Positively_Naked_Vars then  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
778  | 
No_Type_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
 | 
779  | 
else  | 
| 44811 | 780  | 
Explicit_Type_Args  | 
781  | 
(level = All_Types orelse  | 
|
782  | 
granularity_of_type_level level = Ghost_Type_Arg_Vars)  | 
|
| 
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
 | 
783  | 
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
 | 
784  | 
end  | 
| 
42227
 
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
 
blanchet 
parents: 
42180 
diff
changeset
 | 
785  | 
|
| 44625 | 786  | 
(* Make atoms for sorted type variables. *)  | 
| 43263 | 787  | 
fun generic_add_sorts_on_type (_, []) = I  | 
788  | 
| generic_add_sorts_on_type ((x, i), s :: ss) =  | 
|
789  | 
generic_add_sorts_on_type ((x, i), ss)  | 
|
790  | 
    #> (if s = the_single @{sort HOL.type} then
 | 
|
| 43093 | 791  | 
I  | 
792  | 
else if i = ~1 then  | 
|
| 44623 | 793  | 
insert (op =) (`make_type_class s, `make_fixed_type_var x)  | 
| 43093 | 794  | 
else  | 
| 44623 | 795  | 
insert (op =) (`make_type_class s,  | 
796  | 
(make_schematic_type_var (x, i), x)))  | 
|
| 43263 | 797  | 
fun add_sorts_on_tfree (TFree (s, S)) = generic_add_sorts_on_type ((s, ~1), S)  | 
798  | 
| add_sorts_on_tfree _ = I  | 
|
799  | 
fun add_sorts_on_tvar (TVar z) = generic_add_sorts_on_type z  | 
|
800  | 
| add_sorts_on_tvar _ = I  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
801  | 
|
| 44625 | 802  | 
fun type_class_formula type_enc class arg =  | 
803  | 
AAtom (ATerm (class, arg ::  | 
|
804  | 
(case type_enc of  | 
|
| 44754 | 805  | 
Simple_Types (First_Order, Polymorphic, _) =>  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
806  | 
if avoid_first_order_ghost_type_vars then [ATerm (TYPE_name, [arg])]  | 
| 44754 | 807  | 
else []  | 
| 44625 | 808  | 
| _ => [])))  | 
809  | 
fun formulas_for_types type_enc add_sorts_on_typ Ts =  | 
|
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
810  | 
[] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts  | 
| 44625 | 811  | 
|> map (fn (class, name) =>  | 
812  | 
type_class_formula type_enc class (ATerm (name, [])))  | 
|
| 
41137
 
8b634031b2a5
implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
 
blanchet 
parents: 
41136 
diff
changeset
 | 
813  | 
|
| 
42534
 
46e690db16b8
fake type declarations for full-type args and mangled type encodings, so that type assumptions can be discharged
 
blanchet 
parents: 
42533 
diff
changeset
 | 
814  | 
fun mk_aconns c phis =  | 
| 
 
46e690db16b8
fake type declarations for full-type args and mangled type encodings, so that type assumptions can be discharged
 
blanchet 
parents: 
42533 
diff
changeset
 | 
815  | 
let val (phis', phi') = split_last phis in  | 
| 
 
46e690db16b8
fake type declarations for full-type args and mangled type encodings, so that type assumptions can be discharged
 
blanchet 
parents: 
42533 
diff
changeset
 | 
816  | 
fold_rev (mk_aconn c) phis' phi'  | 
| 
 
46e690db16b8
fake type declarations for full-type args and mangled type encodings, so that type assumptions can be discharged
 
blanchet 
parents: 
42533 
diff
changeset
 | 
817  | 
end  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
818  | 
fun mk_ahorn [] phi = phi  | 
| 
42534
 
46e690db16b8
fake type declarations for full-type args and mangled type encodings, so that type assumptions can be discharged
 
blanchet 
parents: 
42533 
diff
changeset
 | 
819  | 
| mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])  | 
| 
42522
 
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
 
blanchet 
parents: 
42521 
diff
changeset
 | 
820  | 
fun mk_aquant _ [] phi = phi  | 
| 
 
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
 
blanchet 
parents: 
42521 
diff
changeset
 | 
821  | 
| mk_aquant q xs (phi as AQuant (q', xs', phi')) =  | 
| 
 
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
 
blanchet 
parents: 
42521 
diff
changeset
 | 
822  | 
if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)  | 
| 
 
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
 
blanchet 
parents: 
42521 
diff
changeset
 | 
823  | 
| mk_aquant q xs phi = AQuant (q, xs, phi)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
824  | 
|
| 45315 | 825  | 
fun close_universally add_term_vars phi =  | 
| 
41145
 
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
 
blanchet 
parents: 
41140 
diff
changeset
 | 
826  | 
let  | 
| 45315 | 827  | 
fun add_formula_vars bounds (AQuant (_, xs, phi)) =  | 
828  | 
add_formula_vars (map fst xs @ bounds) phi  | 
|
829  | 
| add_formula_vars bounds (AConn (_, phis)) =  | 
|
830  | 
fold (add_formula_vars bounds) phis  | 
|
831  | 
| add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm  | 
|
832  | 
in mk_aquant AForall (add_formula_vars [] phi []) phi end  | 
|
| 
42522
 
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
 
blanchet 
parents: 
42521 
diff
changeset
 | 
833  | 
|
| 45377 | 834  | 
fun add_term_vars bounds (ATerm (name as (s, _), tms)) =  | 
835  | 
(if is_tptp_variable s andalso  | 
|
836  | 
not (String.isPrefix tvar_prefix s) andalso  | 
|
837  | 
not (member (op =) bounds name) then  | 
|
838  | 
insert (op =) (name, NONE)  | 
|
839  | 
else  | 
|
840  | 
I)  | 
|
841  | 
#> fold (add_term_vars bounds) tms  | 
|
842  | 
| add_term_vars bounds (AAbs ((name, _), tm)) =  | 
|
843  | 
add_term_vars (name :: bounds) tm  | 
|
| 45401 | 844  | 
fun close_formula_universally phi = close_universally add_term_vars phi  | 
| 45315 | 845  | 
|
846  | 
fun add_iterm_vars bounds (IApp (tm1, tm2)) =  | 
|
847  | 
fold (add_iterm_vars bounds) [tm1, tm2]  | 
|
848  | 
| add_iterm_vars _ (IConst _) = I  | 
|
849  | 
| add_iterm_vars bounds (IVar (name, T)) =  | 
|
850  | 
not (member (op =) bounds name) ? insert (op =) (name, SOME T)  | 
|
851  | 
| add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm  | 
|
852  | 
fun close_iformula_universally phi = close_universally add_iterm_vars phi  | 
|
| 
41145
 
a5ee3b8e5a90
improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
 
blanchet 
parents: 
41140 
diff
changeset
 | 
853  | 
|
| 
46338
 
b02ff6b17599
better handling of individual type for DFG format (SPASS)
 
blanchet 
parents: 
46320 
diff
changeset
 | 
854  | 
val fused_infinite_type_name = "ATP.fused_inf" (* shouldn't clash *)  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
855  | 
val fused_infinite_type = Type (fused_infinite_type_name, [])  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
856  | 
|
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
857  | 
fun tvar_name (x as (s, _)) = (make_schematic_type_var x, s)  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
858  | 
|
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
859  | 
fun ho_term_from_typ format type_enc =  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
860  | 
let  | 
| 
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
861  | 
fun term (Type (s, Ts)) =  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
862  | 
ATerm (case (is_type_enc_higher_order type_enc, s) of  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
863  | 
               (true, @{type_name bool}) => `I tptp_bool_type
 | 
| 
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
864  | 
             | (true, @{type_name fun}) => `I tptp_fun_type
 | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
865  | 
| _ => if s = fused_infinite_type_name andalso  | 
| 
44235
 
85e9dad3c187
distinguish THF syntax with and without choice (Satallax vs. LEO-II)
 
blanchet 
parents: 
44121 
diff
changeset
 | 
866  | 
is_format_typed format then  | 
| 
43178
 
b5862142d378
use "" type only in THF and TFF -- might cause strange failures if used in FOF or CNF, depending on how liberal the prover is
 
blanchet 
parents: 
43175 
diff
changeset
 | 
867  | 
`I tptp_individual_type  | 
| 
 
b5862142d378
use "" type only in THF and TFF -- might cause strange failures if used in FOF or CNF, depending on how liberal the prover is
 
blanchet 
parents: 
43175 
diff
changeset
 | 
868  | 
else  | 
| 
 
b5862142d378
use "" type only in THF and TFF -- might cause strange failures if used in FOF or CNF, depending on how liberal the prover is
 
blanchet 
parents: 
43175 
diff
changeset
 | 
869  | 
`make_fixed_type_const s,  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
870  | 
map term Ts)  | 
| 
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
871  | 
| term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
872  | 
| term (TVar (x, _)) = ATerm (tvar_name x, [])  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
873  | 
in term end  | 
| 42562 | 874  | 
|
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
875  | 
fun ho_term_for_type_arg format type_enc T =  | 
| 
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
876  | 
if T = dummyT then NONE else SOME (ho_term_from_typ format type_enc T)  | 
| 
43401
 
e93dfcb53535
fixed soundness bug made more visible by previous change
 
blanchet 
parents: 
43399 
diff
changeset
 | 
877  | 
|
| 42562 | 878  | 
(* This shouldn't clash with anything else. *)  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
879  | 
val uncurried_alias_sep = "\000"  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
880  | 
val mangled_type_sep = "\001"  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
881  | 
|
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
882  | 
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
 | 
883  | 
|
| 42562 | 884  | 
fun generic_mangled_type_name f (ATerm (name, [])) = f name  | 
885  | 
| generic_mangled_type_name f (ATerm (name, tys)) =  | 
|
| 
42761
 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 
blanchet 
parents: 
42755 
diff
changeset
 | 
886  | 
    f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
 | 
| 
 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 
blanchet 
parents: 
42755 
diff
changeset
 | 
887  | 
^ ")"  | 
| 43692 | 888  | 
| generic_mangled_type_name _ _ = raise Fail "unexpected type abstraction"  | 
| 
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
 | 
889  | 
|
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
890  | 
fun mangled_type format type_enc =  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
891  | 
generic_mangled_type_name fst o ho_term_from_typ format type_enc  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
892  | 
|
| 46435 | 893  | 
fun make_native_type s =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
894  | 
if s = tptp_bool_type orelse s = tptp_fun_type orelse  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
895  | 
s = tptp_individual_type then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
896  | 
s  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
897  | 
else  | 
| 46435 | 898  | 
native_type_prefix ^ ascii_of s  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
899  | 
|
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
900  | 
fun ho_type_from_ho_term type_enc pred_sym ary =  | 
| 42963 | 901  | 
let  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
902  | 
fun to_mangled_atype ty =  | 
| 46435 | 903  | 
AType ((make_native_type (generic_mangled_type_name fst ty),  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
904  | 
generic_mangled_type_name snd ty), [])  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
905  | 
fun to_poly_atype (ATerm (name, tys)) = AType (name, map to_poly_atype tys)  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
906  | 
| to_poly_atype _ = raise Fail "unexpected type abstraction"  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
907  | 
val to_atype =  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
908  | 
if polymorphism_of_type_enc type_enc = Polymorphic then to_poly_atype  | 
| 
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
909  | 
else to_mangled_atype  | 
| 42963 | 910  | 
fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
911  | 
fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
912  | 
| to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys  | 
| 43692 | 913  | 
| to_fo _ _ = raise Fail "unexpected type abstraction"  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
914  | 
fun to_ho (ty as ATerm ((s, _), tys)) =  | 
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
915  | 
if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty  | 
| 
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
916  | 
| to_ho _ = raise Fail "unexpected type abstraction"  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
917  | 
in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end  | 
| 42963 | 918  | 
|
| 43677 | 919  | 
fun ho_type_from_typ format type_enc pred_sym ary =  | 
| 
43676
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
920  | 
ho_type_from_ho_term type_enc pred_sym ary  | 
| 
 
3b0b448b4d69
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
 
nik 
parents: 
43628 
diff
changeset
 | 
921  | 
o ho_term_from_typ format type_enc  | 
| 42963 | 922  | 
|
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
923  | 
fun aliased_uncurried ary (s, s') =  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
924  | 
(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
 | 
925  | 
fun unaliased_uncurried (s, s') =  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
926  | 
case space_explode uncurried_alias_sep s of  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
927  | 
[_] => (s, s')  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
928  | 
| [s1, s2] => (s1, unsuffix s2 s')  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
929  | 
| _ => raise Fail "ill-formed explicit application alias"  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
930  | 
|
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
931  | 
fun raw_mangled_const_name type_name ty_args (s, s') =  | 
| 42963 | 932  | 
let  | 
933  | 
fun type_suffix f g =  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
934  | 
fold_rev (curry (op ^) o g o prefix mangled_type_sep o type_name f)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
935  | 
ty_args ""  | 
| 42963 | 936  | 
in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
937  | 
fun mangled_const_name format type_enc =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
938  | 
map_filter (ho_term_for_type_arg format type_enc)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
939  | 
#> 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
 | 
940  | 
|
| 
 
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
 | 
941  | 
val 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
 | 
942  | 
  Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
 | 
| 
 
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
 | 
943  | 
|
| 
 
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
 | 
944  | 
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
 | 
945  | 
(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
 | 
946  | 
   -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
 | 
| 
 
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
 | 
947  | 
[] >> ATerm) 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
 | 
948  | 
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
 | 
949  | 
(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
 | 
950  | 
|
| 
 
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
 | 
951  | 
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
 | 
952  | 
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
 | 
953  | 
|> 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
 | 
954  | 
           (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
 | 
955  | 
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
 | 
956  | 
|> 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
 | 
957  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
958  | 
fun unmangled_const_name s =  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
959  | 
(s, s) |> unaliased_uncurried |> fst |> space_explode mangled_type_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
 | 
960  | 
fun unmangled_const s =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
961  | 
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
 | 
962  | 
(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
 | 
963  | 
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
 | 
964  | 
|
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
965  | 
fun introduce_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
 | 
966  | 
let  | 
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
967  | 
fun tweak_ho_quant ho_quant T [IAbs _] = IConst (`I ho_quant, T, [])  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
968  | 
| tweak_ho_quant ho_quant (T as Type (_, [p_T as Type (_, [x_T, _]), _]))  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
969  | 
_ =  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
970  | 
(* Eta-expand "!!" and "??", to work around LEO-II 1.2.8 parser  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
971  | 
limitation. This works in conjuction with special code in  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
972  | 
"ATP_Problem" that uses the syntactic sugar "!" and "?" whenever  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
973  | 
possible. *)  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
974  | 
IAbs ((`I "P", p_T),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
975  | 
IApp (IConst (`I ho_quant, T, []),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
976  | 
IAbs ((`I "X", x_T),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
977  | 
IApp (IConst (`I "P", p_T, []),  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
978  | 
IConst (`I "X", x_T, [])))))  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
979  | 
| tweak_ho_quant _ _ _ = raise Fail "unexpected type for quantifier"  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
980  | 
fun intro top_level args (IApp (tm1, tm2)) =  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
981  | 
IApp (intro top_level (tm2 :: args) tm1, intro false [] tm2)  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
982  | 
| intro top_level args (IConst (name as (s, _), T, T_args)) =  | 
| 
42570
 
77f94ac04f32
cleanup proxification/unproxification and make sure that "num_atp_type_args" is called on the proxy in the reconstruction code, since "c_fequal" has one type arg but the unproxified equal has 0
 
blanchet 
parents: 
42569 
diff
changeset
 | 
983  | 
(case proxify_const s of  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
984  | 
SOME proxy_base =>  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
985  | 
if top_level orelse is_type_enc_higher_order type_enc then  | 
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
986  | 
case (top_level, s) of  | 
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
987  | 
(_, "c_False") => IConst (`I tptp_false, T, [])  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
988  | 
| (_, "c_True") => IConst (`I tptp_true, T, [])  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
989  | 
| (false, "c_Not") => IConst (`I tptp_not, T, [])  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
990  | 
| (false, "c_conj") => IConst (`I tptp_and, T, [])  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
991  | 
| (false, "c_disj") => IConst (`I tptp_or, T, [])  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
992  | 
| (false, "c_implies") => IConst (`I tptp_implies, T, [])  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
993  | 
| (false, "c_All") => tweak_ho_quant tptp_ho_forall T args  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
994  | 
| (false, "c_Ex") => tweak_ho_quant tptp_ho_exists T args  | 
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
995  | 
| (false, s) =>  | 
| 44097 | 996  | 
if is_tptp_equal s andalso length args = 2 then  | 
997  | 
IConst (`I tptp_equal, T, [])  | 
|
998  | 
else  | 
|
| 
44589
 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 
blanchet 
parents: 
44587 
diff
changeset
 | 
999  | 
(* Use a proxy even for partially applied THF0 equality,  | 
| 
 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 
blanchet 
parents: 
44587 
diff
changeset
 | 
1000  | 
because the LEO-II and Satallax parsers complain about not  | 
| 
 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 
blanchet 
parents: 
44587 
diff
changeset
 | 
1001  | 
being able to infer the type of "=". *)  | 
| 44097 | 1002  | 
IConst (proxy_base |>> prefix const_prefix, T, T_args)  | 
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1003  | 
| _ => IConst (name, T, [])  | 
| 
42569
 
5737947e4c77
make sure that fequal keeps its type arguments for mangled type systems
 
blanchet 
parents: 
42568 
diff
changeset
 | 
1004  | 
else  | 
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1005  | 
IConst (proxy_base |>> prefix const_prefix, T, T_args)  | 
| 
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
 | 
1006  | 
| NONE => if s = tptp_choice then tweak_ho_quant tptp_choice T args  | 
| 
 
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
 | 
1007  | 
else IConst (name, T, T_args))  | 
| 
43987
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1008  | 
| 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
 | 
1009  | 
| intro _ _ tm = tm  | 
| 
 
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
 
blanchet 
parents: 
43985 
diff
changeset
 | 
1010  | 
in intro true [] end  | 
| 
42568
 
7b9801a34836
no needless "fequal" proxies if "explicit_apply" is set + always have readable names
 
blanchet 
parents: 
42566 
diff
changeset
 | 
1011  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1012  | 
fun mangle_type_args_in_const format type_enc (name as (s, _)) T_args =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1013  | 
case unprefix_and_unascii const_prefix s of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1014  | 
NONE => (name, T_args)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1015  | 
| SOME s'' =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1016  | 
case type_arg_policy [] type_enc (invert_const s'') of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1017  | 
Mangled_Type_Args => (mangled_const_name format type_enc T_args name, [])  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1018  | 
| _ => (name, 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
 | 
1019  | 
fun mangle_type_args_in_iterm format type_enc =  | 
| 
 
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
 | 
1020  | 
if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then  | 
| 
 
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
 | 
1021  | 
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
 | 
1022  | 
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
 | 
1023  | 
| mangle (tm as IConst (_, _, [])) = tm  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1024  | 
| mangle (IConst (name, T, T_args)) =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1025  | 
mangle_type_args_in_const format type_enc name T_args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1026  | 
|> (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
 | 
1027  | 
| 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
 | 
1028  | 
| 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
 | 
1029  | 
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
 | 
1030  | 
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
 | 
1031  | 
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
 | 
1032  | 
|
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1033  | 
fun chop_fun 0 T = ([], T)  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1034  | 
  | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
 | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1035  | 
chop_fun (n - 1) ran_T |>> cons dom_T  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1036  | 
| chop_fun _ T = ([], T)  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1037  | 
|
| 
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
 | 
1038  | 
fun filter_const_type_args _ _ _ [] = []  | 
| 
 
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
 | 
1039  | 
| filter_const_type_args thy s ary T_args =  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1040  | 
let  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1041  | 
val U = robust_const_type thy s  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1042  | 
val arg_U_vars = fold Term.add_tvarsT (U |> chop_fun ary |> fst) []  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1043  | 
val U_args = (s, U) |> robust_const_typargs thy  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1044  | 
in  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1045  | 
U_args ~~ T_args  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1046  | 
|> map (fn (U, T) =>  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1047  | 
if member (op =) arg_U_vars (dest_TVar U) then dummyT else T)  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1048  | 
end  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1049  | 
handle TYPE _ => T_args  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1050  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1051  | 
fun filter_type_args_in_const _ _ _ _ _ [] = []  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1052  | 
| filter_type_args_in_const thy monom_constrs type_enc ary s T_args =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1053  | 
case unprefix_and_unascii const_prefix s of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1054  | 
NONE =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1055  | 
if level_of_type_enc type_enc = No_Types orelse s = tptp_choice then []  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1056  | 
else T_args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1057  | 
| SOME s'' =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1058  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1059  | 
val s'' = invert_const s''  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1060  | 
fun filter_T_args false = T_args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1061  | 
| filter_T_args true = filter_const_type_args thy s'' ary T_args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1062  | 
in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1063  | 
case type_arg_policy monom_constrs type_enc s'' of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1064  | 
Explicit_Type_Args infer_from_term_args =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1065  | 
filter_T_args infer_from_term_args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1066  | 
| No_Type_Args => []  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1067  | 
| Mangled_Type_Args => raise Fail "unexpected (un)mangled symbol"  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1068  | 
end  | 
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
1069  | 
fun filter_type_args_in_iterm thy monom_constrs type_enc =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1070  | 
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
 | 
1071  | 
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
 | 
1072  | 
| filt ary (IConst (name as (s, _), T, T_args)) =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1073  | 
filter_type_args_in_const thy monom_constrs type_enc ary s T_args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1074  | 
|> (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
 | 
1075  | 
| 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
 | 
1076  | 
| 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
 | 
1077  | 
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
 | 
1078  | 
|
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1079  | 
fun iformula_from_prop ctxt format type_enc eq_as_iff =  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1080  | 
let  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1081  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1082  | 
fun do_term bs t atomic_Ts =  | 
| 44495 | 1083  | 
iterm_from_term thy format bs (Envir.eta_contract t)  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1084  | 
|>> (introduce_proxies_in_iterm type_enc  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1085  | 
#> mangle_type_args_in_iterm format type_enc #> AAtom)  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1086  | 
||> 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
 | 
1087  | 
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
 | 
1088  | 
let  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1089  | 
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
 | 
1090  | 
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
 | 
1091  | 
val name =  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1092  | 
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
 | 
1093  | 
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
 | 
1094  | 
| 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
 | 
1095  | 
| 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
 | 
1096  | 
in  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1097  | 
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
 | 
1098  | 
#>> mk_aquant q [(name, SOME T)]  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1099  | 
##> union (op =) (atomic_types_of T)  | 
| 
38518
 
54727b44e277
handle bound name conflicts gracefully in FOF translation
 
blanchet 
parents: 
38496 
diff
changeset
 | 
1100  | 
end  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1101  | 
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
 | 
1102  | 
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
 | 
1103  | 
and do_formula bs pos t =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1104  | 
case t of  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1105  | 
        @{const Trueprop} $ t1 => do_formula bs pos t1
 | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1106  | 
      | @{const Not} $ t1 => do_formula bs (Option.map not pos) t1 #>> mk_anot
 | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1107  | 
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
 | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1108  | 
do_quant bs AForall pos s T t'  | 
| 
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
 | 
1109  | 
      | (t0 as Const (@{const_name All}, _)) $ t1 =>
 | 
| 
 
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
 | 
1110  | 
do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1111  | 
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
 | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1112  | 
do_quant bs AExists pos s T t'  | 
| 
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
 | 
1113  | 
      | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
 | 
| 
 
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
 | 
1114  | 
do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)  | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1115  | 
      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd pos t1 pos t2
 | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1116  | 
      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr pos t1 pos t2
 | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1117  | 
      | @{const HOL.implies} $ t1 $ t2 =>
 | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1118  | 
do_conn bs AImplies (Option.map not pos) t1 pos t2  | 
| 
38864
 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 
haftmann 
parents: 
38829 
diff
changeset
 | 
1119  | 
      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
 | 
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1120  | 
if eq_as_iff then do_conn bs AIff NONE t1 NONE t2 else do_term bs t  | 
| 
41140
 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 
blanchet 
parents: 
41138 
diff
changeset
 | 
1121  | 
| _ => do_term bs t  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1122  | 
in do_formula [] end  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1123  | 
|
| 46093 | 1124  | 
fun presimplify_term ctxt t =  | 
1125  | 
t |> exists_Const (member (op =) Meson.presimplified_consts o fst) t  | 
|
1126  | 
? (Skip_Proof.make_thm (Proof_Context.theory_of ctxt)  | 
|
1127  | 
#> Meson.presimplify  | 
|
1128  | 
#> prop_of)  | 
|
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1129  | 
|
| 
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
 | 
1130  | 
fun concealed_bound_name j = atp_weak_prefix ^ string_of_int j  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1131  | 
fun conceal_bounds Ts t =  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1132  | 
subst_bounds (map (Free o apfst concealed_bound_name)  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1133  | 
(0 upto length Ts - 1 ~~ Ts), t)  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1134  | 
fun reveal_bounds Ts =  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1135  | 
subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1136  | 
(0 upto length Ts - 1 ~~ Ts))  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1137  | 
|
| 43265 | 1138  | 
fun is_fun_equality (@{const_name HOL.eq},
 | 
1139  | 
                     Type (_, [Type (@{type_name fun}, _), _])) = true
 | 
|
1140  | 
| is_fun_equality _ = false  | 
|
1141  | 
||
| 
42747
 
f132d13fcf75
use the same code for extensionalization in Metis and Sledgehammer and generalize that code so that it gracefully handles negations (e.g. negated conjecture), formulas of the form (%x. t) = u, etc.
 
blanchet 
parents: 
42742 
diff
changeset
 | 
1142  | 
fun extensionalize_term ctxt t =  | 
| 43265 | 1143  | 
if exists_Const is_fun_equality t then  | 
1144  | 
let val thy = Proof_Context.theory_of ctxt in  | 
|
1145  | 
t |> cterm_of thy |> Meson.extensionalize_conv ctxt  | 
|
1146  | 
|> prop_of |> Logic.dest_equals |> snd  | 
|
1147  | 
end  | 
|
1148  | 
else  | 
|
1149  | 
t  | 
|
| 
38608
 
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
 
blanchet 
parents: 
38606 
diff
changeset
 | 
1150  | 
|
| 43862 | 1151  | 
fun simple_translate_lambdas do_lambdas ctxt t =  | 
| 
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
 | 
1152  | 
let val thy = Proof_Context.theory_of ctxt in  | 
| 
 
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
 | 
1153  | 
if Meson.is_fol_term thy t then  | 
| 
 
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
 | 
1154  | 
t  | 
| 
 
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
 | 
1155  | 
else  | 
| 
 
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
 | 
1156  | 
let  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1157  | 
fun trans Ts t =  | 
| 
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
 | 
1158  | 
case t of  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1159  | 
            @{const Not} $ t1 => @{const Not} $ trans Ts t1
 | 
| 
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
 | 
1160  | 
          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
 | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1161  | 
t0 $ Abs (s, T, trans (T :: Ts) t')  | 
| 
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
 | 
1162  | 
          | (t0 as Const (@{const_name All}, _)) $ t1 =>
 | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1163  | 
trans Ts (t0 $ eta_expand Ts t1 1)  | 
| 
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
 | 
1164  | 
          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
 | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1165  | 
t0 $ Abs (s, T, trans (T :: Ts) t')  | 
| 
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
 | 
1166  | 
          | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
 | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1167  | 
trans Ts (t0 $ eta_expand Ts t1 1)  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1168  | 
          | (t0 as @{const HOL.conj}) $ t1 $ t2 =>
 | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1169  | 
t0 $ trans Ts t1 $ trans Ts t2  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1170  | 
          | (t0 as @{const HOL.disj}) $ t1 $ t2 =>
 | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1171  | 
t0 $ trans Ts t1 $ trans Ts t2  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1172  | 
          | (t0 as @{const HOL.implies}) $ t1 $ t2 =>
 | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1173  | 
t0 $ trans Ts t1 $ trans Ts t2  | 
| 
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
 | 
1174  | 
          | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
 | 
| 
 
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
 | 
1175  | 
$ t1 $ t2 =>  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1176  | 
t0 $ trans Ts t1 $ trans Ts t2  | 
| 
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
 | 
1177  | 
| _ =>  | 
| 
 
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
 | 
1178  | 
if not (exists_subterm (fn Abs _ => true | _ => false) t) then t  | 
| 
 
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
 | 
1179  | 
else t |> Envir.eta_contract |> do_lambdas ctxt Ts  | 
| 
 
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
 | 
1180  | 
val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1181  | 
in t |> trans [] |> singleton (Variable.export_terms ctxt' ctxt) end  | 
| 
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
 | 
1182  | 
end  | 
| 
43856
 
d636b053d4ff
move more lambda-handling logic to Sledgehammer, from ATP module, for formal dependency reasons
 
blanchet 
parents: 
43830 
diff
changeset
 | 
1183  | 
|
| 43997 | 1184  | 
fun do_cheaply_conceal_lambdas Ts (t1 $ t2) =  | 
1185  | 
do_cheaply_conceal_lambdas Ts t1  | 
|
1186  | 
$ do_cheaply_conceal_lambdas Ts t2  | 
|
1187  | 
| do_cheaply_conceal_lambdas Ts (Abs (_, T, t)) =  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1188  | 
Const (lam_lifted_poly_prefix ^ serial_string (),  | 
| 45509 | 1189  | 
T --> fastype_of1 (T :: Ts, t))  | 
| 43997 | 1190  | 
| do_cheaply_conceal_lambdas _ t = t  | 
| 
43856
 
d636b053d4ff
move more lambda-handling logic to Sledgehammer, from ATP module, for formal dependency reasons
 
blanchet 
parents: 
43830 
diff
changeset
 | 
1191  | 
|
| 43862 | 1192  | 
fun do_introduce_combinators ctxt Ts t =  | 
| 42361 | 1193  | 
let val thy = Proof_Context.theory_of ctxt in  | 
| 
43905
 
1ace987e22e5
avoid calling "Term.is_first_order" (indirectly) on a term with loose de Bruijns -- this is not necessary anyway because of the Abs check in "simple_translate_lambdas"
 
blanchet 
parents: 
43864 
diff
changeset
 | 
1194  | 
t |> conceal_bounds Ts  | 
| 
 
1ace987e22e5
avoid calling "Term.is_first_order" (indirectly) on a term with loose de Bruijns -- this is not necessary anyway because of the Abs check in "simple_translate_lambdas"
 
blanchet 
parents: 
43864 
diff
changeset
 | 
1195  | 
|> cterm_of thy  | 
| 
 
1ace987e22e5
avoid calling "Term.is_first_order" (indirectly) on a term with loose de Bruijns -- this is not necessary anyway because of the Abs check in "simple_translate_lambdas"
 
blanchet 
parents: 
43864 
diff
changeset
 | 
1196  | 
|> Meson_Clausify.introduce_combinators_in_cterm  | 
| 
 
1ace987e22e5
avoid calling "Term.is_first_order" (indirectly) on a term with loose de Bruijns -- this is not necessary anyway because of the Abs check in "simple_translate_lambdas"
 
blanchet 
parents: 
43864 
diff
changeset
 | 
1197  | 
|> prop_of |> Logic.dest_equals |> snd  | 
| 
 
1ace987e22e5
avoid calling "Term.is_first_order" (indirectly) on a term with loose de Bruijns -- this is not necessary anyway because of the Abs check in "simple_translate_lambdas"
 
blanchet 
parents: 
43864 
diff
changeset
 | 
1198  | 
|> reveal_bounds Ts  | 
| 
38491
 
f7e51d981a9f
invoke Variable.export/import_term on the entire formula, to make sure that schematic variables don't get different indices in different subterms;
 
blanchet 
parents: 
38282 
diff
changeset
 | 
1199  | 
end  | 
| 43862 | 1200  | 
  (* A type variable of sort "{}" will make abstraction fail. *)
 | 
| 43997 | 1201  | 
handle THM _ => t |> do_cheaply_conceal_lambdas Ts  | 
| 43862 | 1202  | 
val introduce_combinators = simple_translate_lambdas do_introduce_combinators  | 
1203  | 
||
| 45514 | 1204  | 
fun preprocess_abstractions_in_terms trans_lams facts =  | 
| 43862 | 1205  | 
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
 | 
1206  | 
val (facts, lambda_ts) =  | 
| 45514 | 1207  | 
facts |> map (snd o snd) |> trans_lams  | 
| 
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
 | 
1208  | 
|>> map2 (fn (name, (kind, _)) => fn t => (name, (kind, t))) facts  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1209  | 
val lam_facts =  | 
| 
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
 | 
1210  | 
map2 (fn t => fn j =>  | 
| 46393 | 1211  | 
((lam_fact_prefix ^ Int.toString j, (Global, Spec_Eq)),  | 
1212  | 
(Axiom, t)))  | 
|
| 
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
 | 
1213  | 
lambda_ts (1 upto length lambda_ts)  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1214  | 
in (facts, lam_facts) end  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1215  | 
|
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1216  | 
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the  | 
| 
42353
 
7797efa897a1
correctly handle TFrees that occur in (local) facts -- Metis did the right thing here but Sledgehammer was incorrectly generating spurious preconditions such as "dense_linorder(t_a)"
 
blanchet 
parents: 
42237 
diff
changeset
 | 
1217  | 
same in Sledgehammer to prevent the discovery of unreplayable proofs. *)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1218  | 
fun freeze_term t =  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1219  | 
let  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1220  | 
fun freeze (t $ u) = freeze t $ freeze u  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1221  | 
| freeze (Abs (s, T, t)) = Abs (s, T, freeze t)  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1222  | 
| freeze (Var ((s, i), T)) =  | 
| 
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
 | 
1223  | 
Free (atp_weak_prefix ^ s ^ "_" ^ string_of_int i, T)  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1224  | 
| freeze t = t  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1225  | 
in t |> exists_subterm is_Var t ? freeze end  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1226  | 
|
| 46093 | 1227  | 
fun presimp_prop ctxt role t =  | 
| 
45202
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1228  | 
(let  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1229  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1230  | 
val t = t |> Envir.beta_eta_contract  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1231  | 
|> transform_elim_prop  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1232  | 
|> Object_Logic.atomize_term thy  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1233  | 
     val need_trueprop = (fastype_of t = @{typ bool})
 | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1234  | 
in  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1235  | 
t |> need_trueprop ? HOLogic.mk_Trueprop  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1236  | 
|> extensionalize_term ctxt  | 
| 46093 | 1237  | 
|> presimplify_term ctxt  | 
| 
45202
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1238  | 
|> HOLogic.dest_Trueprop  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1239  | 
end  | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1240  | 
   handle TERM _ => if role = Conjecture then @{term False} else @{term True})
 | 
| 
 
62b8bcf24773
avoid generating too meta theorems -- this sometimes leads to type errors, e.g. when "pp" is applied to a "prop" instead of a "bool"
 
blanchet 
parents: 
45168 
diff
changeset
 | 
1241  | 
|> pair role  | 
| 43096 | 1242  | 
|
| 46340 | 1243  | 
fun make_formula ctxt format type_enc eq_as_iff name stature kind t =  | 
| 43096 | 1244  | 
let  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1245  | 
val (iformula, atomic_Ts) =  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1246  | 
iformula_from_prop ctxt format type_enc eq_as_iff  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1247  | 
(SOME (kind <> Conjecture)) t []  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1248  | 
|>> close_iformula_universally  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1249  | 
in  | 
| 46340 | 1250  | 
    {name = name, stature = stature, kind = kind, iformula = iformula,
 | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1251  | 
atomic_types = atomic_Ts}  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1252  | 
end  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1253  | 
|
| 46340 | 1254  | 
fun make_fact ctxt format type_enc eq_as_iff ((name, stature), t) =  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1255  | 
case t |> make_formula ctxt format type_enc (eq_as_iff andalso format <> CNF)  | 
| 46340 | 1256  | 
name stature Axiom of  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1257  | 
    formula as {iformula = AAtom (IConst ((s, _), _, _)), ...} =>
 | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1258  | 
if s = tptp_true then NONE else SOME formula  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1259  | 
| formula => SOME formula  | 
| 
42561
 
23ddc4e3d19c
have properly type-instantiated helper facts (combinators and If)
 
blanchet 
parents: 
42560 
diff
changeset
 | 
1260  | 
|
| 
46385
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
1261  | 
fun s_not_trueprop (@{const Trueprop} $ t) = @{const Trueprop} $ s_not t
 | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
1262  | 
| s_not_trueprop t =  | 
| 
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
1263  | 
    if fastype_of t = @{typ bool} then s_not t
 | 
| 
46377
 
dce6c3a460a9
new try at lambda-lifting that works correctly for both Metis and Sledgehammer (cf. d724066ff3d0)
 
blanchet 
parents: 
46375 
diff
changeset
 | 
1264  | 
    else @{prop False} (* "t" is too meta *)
 | 
| 44460 | 1265  | 
|
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1266  | 
fun make_conjecture ctxt format type_enc =  | 
| 46340 | 1267  | 
map (fn ((name, stature), (kind, t)) =>  | 
| 
46385
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
1268  | 
t |> kind = Conjecture ? s_not_trueprop  | 
| 46340 | 1269  | 
|> make_formula ctxt format type_enc (format <> CNF) name stature  | 
1270  | 
kind)  | 
|
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1271  | 
|
| 
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
 | 
1272  | 
(** 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
 | 
1273  | 
|
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1274  | 
fun tvar_footprint thy s ary =  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1275  | 
(case unprefix_and_unascii const_prefix s of  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1276  | 
SOME s =>  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1277  | 
s |> invert_const |> robust_const_type thy |> chop_fun ary |> fst  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1278  | 
|> map (fn T => Term.add_tvarsT T [] |> map fst)  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1279  | 
| NONE => [])  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1280  | 
handle TYPE _ => []  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1281  | 
|
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1282  | 
fun ghost_type_args thy s ary =  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1283  | 
if is_tptp_equal s then  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1284  | 
0 upto ary - 1  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1285  | 
else  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1286  | 
let  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1287  | 
val footprint = tvar_footprint thy s ary  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1288  | 
      val eq = (s = @{const_name HOL.eq})
 | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1289  | 
fun ghosts _ [] = []  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1290  | 
| ghosts seen ((i, tvars) :: args) =  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1291  | 
ghosts (union (op =) seen tvars) args  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1292  | 
|> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1293  | 
? cons i  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1294  | 
in  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1295  | 
if forall null footprint then  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1296  | 
[]  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1297  | 
else  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1298  | 
0 upto length footprint - 1 ~~ footprint  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1299  | 
|> sort (rev_order o list_ord Term_Ord.indexname_ord o pairself snd)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1300  | 
|> ghosts []  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1301  | 
end  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1302  | 
|
| 44399 | 1303  | 
type monotonicity_info =  | 
1304  | 
  {maybe_finite_Ts : typ list,
 | 
|
1305  | 
surely_finite_Ts : typ list,  | 
|
1306  | 
maybe_infinite_Ts : typ list,  | 
|
1307  | 
surely_infinite_Ts : typ list,  | 
|
1308  | 
maybe_nonmono_Ts : typ list}  | 
|
1309  | 
||
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
1310  | 
(* These types witness that the type classes they belong to allow infinite  | 
| 
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
1311  | 
models and hence that any types with these type classes is monotonic. *)  | 
| 
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
1312  | 
val known_infinite_types =  | 
| 
44634
 
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
 
blanchet 
parents: 
44626 
diff
changeset
 | 
1313  | 
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}]
 | 
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
1314  | 
|
| 46301 | 1315  | 
fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T =  | 
1316  | 
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
 | 
1317  | 
|
| 
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
 | 
1318  | 
(* 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
 | 
1319  | 
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
 | 
1320  | 
proofs. On the other hand, all HOL infinite types can be given the same  | 
| 
 
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
 | 
1321  | 
models in first-order logic (via Löwenheim-Skolem). *)  | 
| 
 
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
 | 
1322  | 
|
| 44399 | 1323  | 
fun should_encode_type _ (_ : monotonicity_info) All_Types _ = true  | 
1324  | 
  | should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts,
 | 
|
1325  | 
maybe_nonmono_Ts, ...}  | 
|
| 46301 | 1326  | 
(Noninf_Nonmono_Types (strictness, grain)) T =  | 
| 44811 | 1327  | 
grain = Ghost_Type_Arg_Vars orelse  | 
1328  | 
(exists (type_intersect ctxt T) maybe_nonmono_Ts andalso  | 
|
1329  | 
not (exists (type_instance ctxt T) surely_infinite_Ts orelse  | 
|
| 
44859
 
237ba63d6041
fixed definition of type intersection (soundness bug)
 
blanchet 
parents: 
44853 
diff
changeset
 | 
1330  | 
(not (member (type_equiv ctxt) maybe_finite_Ts T) andalso  | 
| 46301 | 1331  | 
is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts  | 
| 44811 | 1332  | 
T)))  | 
| 44399 | 1333  | 
  | should_encode_type ctxt {surely_finite_Ts, maybe_infinite_Ts,
 | 
1334  | 
maybe_nonmono_Ts, ...}  | 
|
| 44811 | 1335  | 
(Fin_Nonmono_Types grain) T =  | 
1336  | 
grain = Ghost_Type_Arg_Vars orelse  | 
|
1337  | 
(exists (type_intersect ctxt T) maybe_nonmono_Ts andalso  | 
|
1338  | 
(exists (type_generalization ctxt T) surely_finite_Ts orelse  | 
|
| 
44859
 
237ba63d6041
fixed definition of type intersection (soundness bug)
 
blanchet 
parents: 
44853 
diff
changeset
 | 
1339  | 
(not (member (type_equiv ctxt) maybe_infinite_Ts T) andalso  | 
| 44811 | 1340  | 
is_type_surely_finite ctxt T)))  | 
| 
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
 | 
1341  | 
| should_encode_type _ _ _ _ = false  | 
| 
 
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
 | 
1342  | 
|
| 44768 | 1343  | 
fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T =  | 
| 44811 | 1344  | 
should_guard_var () andalso should_encode_type ctxt mono level T  | 
| 44399 | 1345  | 
| 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
 | 
1346  | 
|
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1347  | 
fun is_maybe_universal_var (IConst ((s, _), _, _)) =  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1348  | 
String.isPrefix bound_var_prefix s orelse  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1349  | 
String.isPrefix all_bound_var_prefix s  | 
| 
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
1350  | 
| 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
 | 
1351  | 
| is_maybe_universal_var _ = false  | 
| 42836 | 1352  | 
|
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1353  | 
datatype site =  | 
| 
43361
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
1354  | 
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
 | 
1355  | 
Eq_Arg of bool option |  | 
| 
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
1356  | 
Elsewhere  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
1357  | 
|
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1358  | 
fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1359  | 
| should_tag_with_type ctxt mono (Tags (_, level)) site u T =  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1360  | 
if granularity_of_type_level level = All_Vars then  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1361  | 
should_encode_type ctxt mono level T  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1362  | 
else  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1363  | 
(case (site, is_maybe_universal_var u) of  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1364  | 
(Eq_Arg _, true) => should_encode_type ctxt mono level T  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1365  | 
| _ => false)  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1366  | 
| 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
 | 
1367  | 
|
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
1368  | 
fun fused_type ctxt mono level =  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
1369  | 
let  | 
| 44399 | 1370  | 
val should_encode = should_encode_type ctxt mono level  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
1371  | 
fun fuse 0 T = if should_encode T then T else fused_infinite_type  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
1372  | 
      | fuse ary (Type (@{type_name fun}, [T1, T2])) =
 | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
1373  | 
fuse 0 T1 --> fuse (ary - 1) T2  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
1374  | 
| fuse _ _ = raise Fail "expected function type"  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
1375  | 
in fuse end  | 
| 
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
 | 
1376  | 
|
| 
44450
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1377  | 
(** predicators and application operators **)  | 
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
1378  | 
|
| 42574 | 1379  | 
type sym_info =  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1380  | 
  {pred_sym : bool, min_ary : int, max_ary : int, types : typ list,
 | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1381  | 
in_conj : bool}  | 
| 42563 | 1382  | 
|
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1383  | 
fun default_sym_tab_entries type_enc =  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1384  | 
  (make_fixed_const NONE @{const_name undefined},
 | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1385  | 
       {pred_sym = false, min_ary = 0, max_ary = 0, types = [],
 | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1386  | 
in_conj = false}) ::  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1387  | 
([tptp_false, tptp_true]  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1388  | 
   |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = [],
 | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1389  | 
in_conj = false})) @  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1390  | 
([tptp_equal, tptp_old_equal]  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1391  | 
   |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = [],
 | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1392  | 
in_conj = false}))  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1393  | 
|> not (is_type_enc_higher_order type_enc)  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1394  | 
? cons (prefixed_predicator_name,  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1395  | 
             {pred_sym = true, min_ary = 1, max_ary = 1, types = [],
 | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1396  | 
in_conj = false})  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1397  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1398  | 
datatype app_op_level = Incomplete_Apply | Sufficient_Apply | Full_Apply  | 
| 46389 | 1399  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1400  | 
fun sym_table_for_facts ctxt type_enc app_op_level conjs facts =  | 
| 
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
 | 
1401  | 
let  | 
| 44772 | 1402  | 
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
 | 
1403  | 
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
 | 
1404  | 
fun iter ary T =  | 
| 44399 | 1405  | 
if ary = max_ary orelse type_instance ctxt var_T T orelse  | 
1406  | 
type_instance ctxt T var_T then  | 
|
| 
43210
 
7384b771805d
made "explicit_apply"'s smart mode (more) complete
 
blanchet 
parents: 
43207 
diff
changeset
 | 
1407  | 
ary  | 
| 
 
7384b771805d
made "explicit_apply"'s smart mode (more) complete
 
blanchet 
parents: 
43207 
diff
changeset
 | 
1408  | 
else  | 
| 
 
7384b771805d
made "explicit_apply"'s smart mode (more) complete
 
blanchet 
parents: 
43207 
diff
changeset
 | 
1409  | 
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
 | 
1410  | 
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
 | 
1411  | 
fun add_universal_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1412  | 
if app_op_level = Sufficient_Apply andalso  | 
| 
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
 | 
1413  | 
         (can dest_funT T orelse T = @{typ bool}) then
 | 
| 
 
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
 | 
1414  | 
let  | 
| 
 
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
 | 
1415  | 
          val bool_vars' = bool_vars orelse body_type T = @{typ bool}
 | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1416  | 
          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
 | 
1417  | 
            {pred_sym = pred_sym andalso not bool_vars',
 | 
| 44772 | 1418  | 
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
 | 
1419  | 
max_ary = max_ary, types = types, in_conj = 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
 | 
1420  | 
val fun_var_Ts' =  | 
| 
 
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
 | 
1421  | 
fun_var_Ts |> can dest_funT T ? insert_type ctxt I T  | 
| 
 
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
 | 
1422  | 
in  | 
| 
 
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
 | 
1423  | 
if bool_vars' = bool_vars andalso  | 
| 
 
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
 | 
1424  | 
pointer_eq (fun_var_Ts', fun_var_Ts) then  | 
| 
 
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
 | 
1425  | 
accum  | 
| 
43167
 
839f599bc7ed
ensured that the logic for "explicit_apply = smart" also works on CNF (i.e. new Metis)
 
blanchet 
parents: 
43159 
diff
changeset
 | 
1426  | 
else  | 
| 44772 | 1427  | 
((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
 | 
1428  | 
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
 | 
1429  | 
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
 | 
1430  | 
accum  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1431  | 
fun add_iterm_syms conj_fact top_level tm  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1432  | 
(accum as ((bool_vars, fun_var_Ts), sym_tab)) =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1433  | 
let val (head, args) = strip_iterm_comb tm in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1434  | 
(case head of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1435  | 
IConst ((s, _), T, _) =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1436  | 
if String.isPrefix bound_var_prefix s orelse  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1437  | 
String.isPrefix all_bound_var_prefix s then  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1438  | 
add_universal_var T accum  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1439  | 
else if String.isPrefix exist_bound_var_prefix s then  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1440  | 
accum  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1441  | 
else  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1442  | 
let val ary = length args in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1443  | 
((bool_vars, fun_var_Ts),  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1444  | 
case Symtab.lookup sym_tab s of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1445  | 
                    SOME {pred_sym, min_ary, max_ary, types, in_conj} =>
 | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1446  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1447  | 
val pred_sym =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1448  | 
pred_sym andalso top_level andalso not bool_vars  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1449  | 
val types' = types |> insert_type ctxt I T  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1450  | 
val in_conj = in_conj orelse conj_fact  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1451  | 
val min_ary =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1452  | 
if app_op_level = Sufficient_Apply andalso  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1453  | 
not (pointer_eq (types', types)) then  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1454  | 
fold (consider_var_ary T) fun_var_Ts min_ary  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1455  | 
else  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1456  | 
min_ary  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1457  | 
in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1458  | 
                      Symtab.update (s, {pred_sym = pred_sym,
 | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1459  | 
min_ary = Int.min (ary, min_ary),  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1460  | 
max_ary = Int.max (ary, max_ary),  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1461  | 
types = types', in_conj = in_conj})  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1462  | 
sym_tab  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1463  | 
end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1464  | 
| NONE =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1465  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1466  | 
val pred_sym = top_level andalso not bool_vars  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1467  | 
val min_ary =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1468  | 
case app_op_level of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1469  | 
Incomplete_Apply => ary  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1470  | 
| Sufficient_Apply =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1471  | 
fold (consider_var_ary T) fun_var_Ts ary  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1472  | 
| Full_Apply => 0  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1473  | 
in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1474  | 
Symtab.update_new (s,  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1475  | 
                          {pred_sym = pred_sym, min_ary = min_ary,
 | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1476  | 
max_ary = ary, types = [T], in_conj = conj_fact})  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1477  | 
sym_tab  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1478  | 
end)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1479  | 
end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1480  | 
| IVar (_, T) => add_universal_var T accum  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1481  | 
| IAbs ((_, T), tm) =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1482  | 
accum |> add_universal_var T |> add_iterm_syms conj_fact false tm  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1483  | 
| _ => accum)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1484  | 
|> fold (add_iterm_syms conj_fact false) args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1485  | 
end  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1486  | 
fun add_fact_syms conj_fact =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1487  | 
K (add_iterm_syms conj_fact true) |> formula_fold NONE |> fact_lift  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1488  | 
in  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1489  | 
((false, []), Symtab.empty)  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1490  | 
|> fold (add_fact_syms true) conjs  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1491  | 
|> fold (add_fact_syms false) facts  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1492  | 
|> snd  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1493  | 
|> fold Symtab.update (default_sym_tab_entries type_enc)  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
1494  | 
end  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1495  | 
|
| 44772 | 1496  | 
fun min_ary_of sym_tab s =  | 
| 
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
 | 
1497  | 
case Symtab.lookup sym_tab s of  | 
| 42574 | 1498  | 
    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
 | 
1499  | 
| NONE =>  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1500  | 
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
 | 
1501  | 
SOME s =>  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1502  | 
let val s = s |> unmangled_const_name |> hd |> invert_const in  | 
| 
42966
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
1503  | 
if s = predicator_name then 1  | 
| 
 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 
blanchet 
parents: 
42963 
diff
changeset
 | 
1504  | 
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
 | 
1505  | 
else if s = type_guard_name then 1  | 
| 
42557
 
ae0deb39a254
fixed min-arity computation when "explicit_apply" is specified
 
blanchet 
parents: 
42556 
diff
changeset
 | 
1506  | 
else 0  | 
| 
42547
 
b5eec0c99528
fixed arity of special constants if "explicit_apply" is set
 
blanchet 
parents: 
42546 
diff
changeset
 | 
1507  | 
end  | 
| 
42544
 
75cb06eee990
reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
 
blanchet 
parents: 
42543 
diff
changeset
 | 
1508  | 
| NONE => 0  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1509  | 
|
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1510  | 
(* 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
 | 
1511  | 
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
 | 
1512  | 
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
 | 
1513  | 
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
 | 
1514  | 
fun is_pred_sym sym_tab s =  | 
| 
 
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
 | 
1515  | 
case Symtab.lookup sym_tab s of  | 
| 42574 | 1516  | 
    SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
 | 
1517  | 
pred_sym andalso min_ary = max_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
 | 
1518  | 
| NONE => false  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
1519  | 
|
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1520  | 
val app_op = `(make_fixed_const NONE) app_op_name  | 
| 
42568
 
7b9801a34836
no needless "fequal" proxies if "explicit_apply" is set + always have readable names
 
blanchet 
parents: 
42566 
diff
changeset
 | 
1521  | 
val predicator_combconst =  | 
| 44495 | 1522  | 
  IConst (`(make_fixed_const NONE) predicator_name, @{typ "bool => bool"}, [])
 | 
| 
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
 | 
1523  | 
|
| 43859 | 1524  | 
fun list_app head args = fold (curry (IApp o swap)) args head  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1525  | 
fun predicator tm = IApp (predicator_combconst, tm)  | 
| 
42544
 
75cb06eee990
reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
 
blanchet 
parents: 
42543 
diff
changeset
 | 
1526  | 
|
| 
46410
 
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
 
blanchet 
parents: 
46409 
diff
changeset
 | 
1527  | 
fun mk_app_op format 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
 | 
1528  | 
let  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1529  | 
val head_T = ityp_of head  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1530  | 
val (arg_T, res_T) = dest_funT head_T  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1531  | 
val app =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1532  | 
IConst (app_op, head_T --> head_T, [arg_T, res_T])  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1533  | 
|> mangle_type_args_in_iterm format type_enc  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1534  | 
in list_app app [head, arg] end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1535  | 
|
| 
46400
 
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
 
blanchet 
parents: 
46399 
diff
changeset
 | 
1536  | 
fun firstorderize_fact thy monom_constrs format type_enc sym_tab  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1537  | 
uncurried_aliases =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1538  | 
let  | 
| 
46410
 
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
 
blanchet 
parents: 
46409 
diff
changeset
 | 
1539  | 
fun do_app arg head = mk_app_op format type_enc head arg  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1540  | 
fun list_app_ops head args = fold do_app args head  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1541  | 
fun introduce_app_ops tm =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1542  | 
let val (head, args) = tm |> strip_iterm_comb ||> map introduce_app_ops in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1543  | 
case head of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1544  | 
IConst (name as (s, _), T, T_args) =>  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1545  | 
if uncurried_aliases andalso String.isPrefix const_prefix s then  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1546  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1547  | 
val ary = length args  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1548  | 
val name = name |> ary > min_ary_of sym_tab s  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
1549  | 
? aliased_uncurried ary  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1550  | 
in list_app (IConst (name, T, T_args)) args end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1551  | 
else  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1552  | 
args |> chop (min_ary_of sym_tab s)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1553  | 
|>> list_app head |-> list_app_ops  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1554  | 
| _ => list_app_ops head args  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1555  | 
end  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1556  | 
fun introduce_predicators tm =  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1557  | 
case strip_iterm_comb tm of  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1558  | 
(IConst ((s, _), _, _), _) =>  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1559  | 
if is_pred_sym sym_tab s then tm else predicator tm  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1560  | 
| _ => predicator tm  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1561  | 
val do_iterm =  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1562  | 
not (is_type_enc_higher_order type_enc)  | 
| 
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1563  | 
? (introduce_app_ops #> introduce_predicators)  | 
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
1564  | 
#> filter_type_args_in_iterm thy monom_constrs type_enc  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1565  | 
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
 | 
1566  | 
|
| 
 
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
 | 
1567  | 
(** 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
 | 
1568  | 
|
| 
44450
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1569  | 
val not_ffalse = @{lemma "~ fFalse" by (unfold fFalse_def) fast}
 | 
| 
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1570  | 
val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast}
 | 
| 
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1571  | 
|
| 43194 | 1572  | 
(* The Boolean indicates that a fairly sound type encoding is needed. *)  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1573  | 
val helper_table =  | 
| 43194 | 1574  | 
  [(("COMBI", false), @{thms Meson.COMBI_def}),
 | 
1575  | 
   (("COMBK", false), @{thms Meson.COMBK_def}),
 | 
|
1576  | 
   (("COMBB", false), @{thms Meson.COMBB_def}),
 | 
|
1577  | 
   (("COMBC", false), @{thms Meson.COMBC_def}),
 | 
|
1578  | 
   (("COMBS", false), @{thms Meson.COMBS_def}),
 | 
|
| 
44450
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1579  | 
((predicator_name, false), [not_ffalse, ftrue]),  | 
| 
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1580  | 
   (("fFalse", false), [not_ffalse]),
 | 
| 43194 | 1581  | 
   (("fFalse", true), @{thms True_or_False}),
 | 
| 
44450
 
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
 
blanchet 
parents: 
44418 
diff
changeset
 | 
1582  | 
   (("fTrue", false), [ftrue]),
 | 
| 43194 | 1583  | 
   (("fTrue", true), @{thms True_or_False}),
 | 
1584  | 
   (("fNot", false),
 | 
|
1585  | 
    @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
 | 
|
1586  | 
fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),  | 
|
1587  | 
   (("fconj", false),
 | 
|
1588  | 
    @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
 | 
|
1589  | 
by (unfold fconj_def) fast+}),  | 
|
1590  | 
   (("fdisj", false),
 | 
|
1591  | 
    @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
 | 
|
1592  | 
by (unfold fdisj_def) fast+}),  | 
|
1593  | 
   (("fimplies", false),
 | 
|
| 
43210
 
7384b771805d
made "explicit_apply"'s smart mode (more) complete
 
blanchet 
parents: 
43207 
diff
changeset
 | 
1594  | 
    @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
 | 
| 43194 | 1595  | 
by (unfold fimplies_def) fast+}),  | 
| 43678 | 1596  | 
   (("fequal", true),
 | 
1597  | 
(* This is a lie: Higher-order equality doesn't need a sound type encoding.  | 
|
1598  | 
However, this is done so for backward compatibility: Including the  | 
|
1599  | 
equality helpers by default in Metis breaks a few existing proofs. *)  | 
|
1600  | 
    @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
 | 
|
1601  | 
fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),  | 
|
| 44003 | 1602  | 
(* Partial characterization of "fAll" and "fEx". A complete characterization  | 
1603  | 
would require the axiom of choice for replay with Metis. *)  | 
|
1604  | 
   (("fAll", false), [@{lemma "~ fAll P | P x" by (auto simp: fAll_def)}]),
 | 
|
1605  | 
   (("fEx", false), [@{lemma "~ P x | fEx P" by (auto simp: fEx_def)}]),
 | 
|
| 43194 | 1606  | 
   (("If", true), @{thms if_True if_False True_or_False})]
 | 
1607  | 
|> map (apsnd (map zero_var_indexes))  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1608  | 
|
| 
45364
 
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
 
blanchet 
parents: 
45316 
diff
changeset
 | 
1609  | 
fun atype_of_type_vars (Simple_Types (_, Polymorphic, _)) = SOME atype_of_types  | 
| 
 
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
 
blanchet 
parents: 
45316 
diff
changeset
 | 
1610  | 
| atype_of_type_vars _ = NONE  | 
| 
 
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
 
blanchet 
parents: 
45316 
diff
changeset
 | 
1611  | 
|
| 45920 | 1612  | 
fun bound_tvars type_enc sorts Ts =  | 
1613  | 
(sorts ? mk_ahorn (formulas_for_types type_enc add_sorts_on_tvar Ts))  | 
|
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1614  | 
#> mk_aquant AForall  | 
| 
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1615  | 
(map_filter (fn TVar (x as (s, _), _) =>  | 
| 
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1616  | 
SOME ((make_schematic_type_var x, s),  | 
| 
45364
 
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
 
blanchet 
parents: 
45316 
diff
changeset
 | 
1617  | 
atype_of_type_vars type_enc)  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
1618  | 
| _ => NONE) Ts)  | 
| 
44404
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
1619  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1620  | 
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
 | 
1621  | 
(if pred_sym then AConn (AIff, [AAtom tm1, AAtom tm2])  | 
| 
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
1622  | 
else AAtom (ATerm (`I tptp_equal, [tm1, tm2])))  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1623  | 
|> mk_aquant AForall bounds  | 
| 45377 | 1624  | 
|> close_formula_universally  | 
| 45920 | 1625  | 
|> bound_tvars type_enc true atomic_Ts  | 
| 
44404
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
1626  | 
|
| 46406 | 1627  | 
val helper_rank = default_rank  | 
1628  | 
val min_rank = 9 * helper_rank div 10  | 
|
1629  | 
val max_rank = 4 * min_rank  | 
|
1630  | 
||
1631  | 
fun rank_of_fact_num n j = min_rank + (max_rank - min_rank) * j div n  | 
|
1632  | 
||
| 44495 | 1633  | 
val type_tag = `(make_fixed_const NONE) type_tag_name  | 
| 
43130
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
1634  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
1635  | 
fun type_tag_idempotence_fact format type_enc =  | 
| 
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
 | 
1636  | 
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
 | 
1637  | 
fun var s = ATerm (`I s, [])  | 
| 
44408
 
30ea62ab4f16
made reconstruction of type tag equalities "\?x = \?x" reliable
 
blanchet 
parents: 
44406 
diff
changeset
 | 
1638  | 
fun tag tm = ATerm (type_tag, [var "A", tm])  | 
| 
 
30ea62ab4f16
made reconstruction of type tag equalities "\?x = \?x" reliable
 
blanchet 
parents: 
44406 
diff
changeset
 | 
1639  | 
val tagged_var = tag (var "X")  | 
| 
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
 | 
1640  | 
in  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
1641  | 
Formula (type_tag_idempotence_helper_name, Axiom,  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1642  | 
eq_formula type_enc [] [] false (tag tagged_var) tagged_var,  | 
| 46406 | 1643  | 
NONE, isabelle_info format spec_eqN 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
 | 
1644  | 
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
 | 
1645  | 
|
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
1646  | 
fun should_specialize_helper type_enc t =  | 
| 
44493
 
c2602c5d4b0a
handle nonmangled monomorphich the same way as mangled monomorphic when it comes to helper -- otherwise we can end up generating too tight type guards
 
blanchet 
parents: 
44491 
diff
changeset
 | 
1647  | 
polymorphism_of_type_enc type_enc <> Polymorphic andalso  | 
| 
43628
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1648  | 
level_of_type_enc type_enc <> No_Types andalso  | 
| 
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1649  | 
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
 | 
1650  | 
|
| 
43858
 
be41d12de6fa
simplify code -- there are no lambdas in helpers anyway
 
blanchet 
parents: 
43857 
diff
changeset
 | 
1651  | 
fun helper_facts_for_sym ctxt format type_enc (s, {types, ...} : sym_info) =
 | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1652  | 
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
 | 
1653  | 
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
 | 
1654  | 
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
 | 
1655  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1656  | 
val unmangled_s = mangled_s |> unmangled_const_name |> hd  | 
| 
43628
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1657  | 
fun dub needs_fairly_sound j k =  | 
| 
46339
 
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
 
blanchet 
parents: 
46338 
diff
changeset
 | 
1658  | 
unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^  | 
| 
 
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
 
blanchet 
parents: 
46338 
diff
changeset
 | 
1659  | 
(if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^  | 
| 
 
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
 
blanchet 
parents: 
46338 
diff
changeset
 | 
1660  | 
(if needs_fairly_sound then typed_helper_suffix  | 
| 
 
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
 
blanchet 
parents: 
46338 
diff
changeset
 | 
1661  | 
else untyped_helper_suffix)  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
1662  | 
fun dub_and_inst needs_fairly_sound (th, j) =  | 
| 
43628
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1663  | 
let val t = prop_of th in  | 
| 
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1664  | 
if should_specialize_helper type_enc t then  | 
| 
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1665  | 
map (fn T => specialize_type thy (invert_const unmangled_s, T) t)  | 
| 
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1666  | 
types  | 
| 
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1667  | 
else  | 
| 
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1668  | 
[t]  | 
| 
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1669  | 
end  | 
| 
46339
 
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
 
blanchet 
parents: 
46338 
diff
changeset
 | 
1670  | 
|> tag_list 1  | 
| 46393 | 1671  | 
|> map (fn (k, t) =>  | 
1672  | 
((dub needs_fairly_sound j k, (Global, Spec_Eq)), t))  | 
|
| 
43860
 
57ef3cd4126e
more refactoring of preprocessing, so as to be able to centralize it
 
blanchet 
parents: 
43859 
diff
changeset
 | 
1673  | 
val make_facts = map_filter (make_fact ctxt format type_enc false)  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
1674  | 
val fairly_sound = is_type_enc_fairly_sound type_enc  | 
| 
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
 | 
1675  | 
in  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1676  | 
helper_table  | 
| 43194 | 1677  | 
|> maps (fn ((helper_s, needs_fairly_sound), ths) =>  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
1678  | 
if helper_s <> unmangled_s orelse  | 
| 
42894
 
ce269ee43800
further improvements to "poly_{preds,tags}_{bang,query}" -- better solution to the combinator problem + make sure type assumptions can be discharged
 
blanchet 
parents: 
42893 
diff
changeset
 | 
1679  | 
(needs_fairly_sound andalso not fairly_sound) then  | 
| 
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
 | 
1680  | 
[]  | 
| 
 
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
 | 
1681  | 
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
 | 
1682  | 
ths ~~ (1 upto length ths)  | 
| 
43628
 
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
 
blanchet 
parents: 
43626 
diff
changeset
 | 
1683  | 
|> maps (dub_and_inst needs_fairly_sound)  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
1684  | 
|> make_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
 | 
1685  | 
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
 | 
1686  | 
| NONE => []  | 
| 
43858
 
be41d12de6fa
simplify code -- there are no lambdas in helpers anyway
 
blanchet 
parents: 
43857 
diff
changeset
 | 
1687  | 
fun helper_facts_for_sym_table ctxt format type_enc sym_tab =  | 
| 
 
be41d12de6fa
simplify code -- there are no lambdas in helpers anyway
 
blanchet 
parents: 
43857 
diff
changeset
 | 
1688  | 
Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_enc) sym_tab  | 
| 
 
be41d12de6fa
simplify code -- there are no lambdas in helpers anyway
 
blanchet 
parents: 
43857 
diff
changeset
 | 
1689  | 
[]  | 
| 
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
 | 
1690  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1691  | 
(***************************************************************)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1692  | 
(* 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
 | 
1693  | 
(***************************************************************)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1694  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1695  | 
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
 | 
1696  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1697  | 
fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1698  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1699  | 
(* Remove this trivial type class (FIXME: similar code elsewhere) *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1700  | 
fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1701  | 
|
| 43093 | 1702  | 
fun classes_of_terms get_Ts =  | 
| 43121 | 1703  | 
map (map snd o get_Ts)  | 
| 43093 | 1704  | 
#> List.foldl add_classes Symtab.empty  | 
1705  | 
#> delete_type #> Symtab.keys  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1706  | 
|
| 44121 | 1707  | 
val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees  | 
1708  | 
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
 | 
1709  | 
|
| 43622 | 1710  | 
fun fold_type_constrs f (Type (s, Ts)) x =  | 
1711  | 
fold (fold_type_constrs f) Ts (f (s, x))  | 
|
| 43189 | 1712  | 
| fold_type_constrs _ _ x = x  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1713  | 
|
| 
43907
 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 
blanchet 
parents: 
43906 
diff
changeset
 | 
1714  | 
(* 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
 | 
1715  | 
needed. *)  | 
| 43189 | 1716  | 
fun add_type_constrs_in_term thy =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1717  | 
let  | 
| 
43188
 
0c36ae874fcc
fixed detection of Skolem constants in type construction detection code
 
blanchet 
parents: 
43185 
diff
changeset
 | 
1718  | 
    fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
 | 
| 43181 | 1719  | 
| 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
 | 
1720  | 
| add (Const x) =  | 
| 
 
1b333e4173a2
drop more type arguments soundly, when they can be deduced from the arg types
 
blanchet 
parents: 
44634 
diff
changeset
 | 
1721  | 
x |> robust_const_typargs thy |> fold (fold_type_constrs set_insert)  | 
| 43181 | 1722  | 
| add (Abs (_, _, u)) = add u  | 
1723  | 
| add _ = I  | 
|
1724  | 
in add end  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1725  | 
|
| 43189 | 1726  | 
fun type_constrs_of_terms thy ts =  | 
1727  | 
Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
43064 
diff
changeset
 | 
1728  | 
|
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1729  | 
fun extract_lambda_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
 | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1730  | 
let val (head, args) = strip_comb t in  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1731  | 
(head |> dest_Const |> fst,  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1732  | 
fold_rev (fn t as Var ((s, _), T) =>  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1733  | 
(fn u => Abs (s, T, abstract_over (t, u)))  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1734  | 
| _ => raise Fail "expected Var") args u)  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1735  | 
end  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1736  | 
| extract_lambda_def _ = raise Fail "malformed lifted lambda"  | 
| 45508 | 1737  | 
|
| 45514 | 1738  | 
fun trans_lams_from_string ctxt type_enc lam_trans =  | 
1739  | 
if lam_trans = no_lamsN then  | 
|
1740  | 
rpair []  | 
|
1741  | 
else if lam_trans = hide_lamsN then  | 
|
1742  | 
lift_lams ctxt type_enc ##> K []  | 
|
| 46365 | 1743  | 
else if lam_trans = liftingN orelse lam_trans = lam_liftingN then  | 
| 45514 | 1744  | 
lift_lams ctxt type_enc  | 
| 46365 | 1745  | 
else if lam_trans = combsN then  | 
| 45514 | 1746  | 
map (introduce_combinators ctxt) #> rpair []  | 
| 46365 | 1747  | 
else if lam_trans = combs_and_liftingN then  | 
1748  | 
lift_lams_part_1 ctxt type_enc  | 
|
1749  | 
##> maps (fn t => [t, introduce_combinators ctxt (intentionalize_def t)])  | 
|
1750  | 
#> lift_lams_part_2  | 
|
| 46368 | 1751  | 
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
 | 
1752  | 
lift_lams_part_1 ctxt type_enc  | 
| 46368 | 1753  | 
    ##> map (fn t => case head_of (strip_qnt_body @{const_name All} t) of
 | 
1754  | 
                       @{term "op =::bool => bool => bool"} => t
 | 
|
1755  | 
| _ => introduce_combinators ctxt (intentionalize_def t))  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1756  | 
#> lift_lams_part_2  | 
| 45514 | 1757  | 
else if lam_trans = keep_lamsN then  | 
1758  | 
map (Envir.eta_contract) #> rpair []  | 
|
1759  | 
else  | 
|
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45516 
diff
changeset
 | 
1760  | 
    error ("Unknown lambda translation scheme: " ^ quote lam_trans ^ ".")
 | 
| 45514 | 1761  | 
|
1762  | 
fun translate_formulas ctxt format prem_kind type_enc lam_trans presimp hyp_ts  | 
|
1763  | 
concl_t 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
 | 
1764  | 
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
 | 
1765  | 
val thy = Proof_Context.theory_of ctxt  | 
| 45514 | 1766  | 
val trans_lams = trans_lams_from_string ctxt type_enc lam_trans  | 
| 43861 | 1767  | 
val fact_ts = facts |> map snd  | 
| 
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
 | 
1768  | 
(* Remove existing facts from the conjecture, as this can dramatically  | 
| 
 
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
 | 
1769  | 
boost an ATP's performance (for some reason). *)  | 
| 43192 | 1770  | 
val hyp_ts =  | 
1771  | 
hyp_ts  | 
|
1772  | 
      |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
 | 
|
| 
43864
 
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
 
blanchet 
parents: 
43863 
diff
changeset
 | 
1773  | 
val facts = facts |> map (apsnd (pair Axiom))  | 
| 
 
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
 
blanchet 
parents: 
43863 
diff
changeset
 | 
1774  | 
val conjs =  | 
| 
46385
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
46384 
diff
changeset
 | 
1775  | 
map (pair prem_kind) hyp_ts @ [(Conjecture, s_not_trueprop concl_t)]  | 
| 
45168
 
0e8e662013f9
freeze conjecture schematics before applying lambda-translation, which sometimes calls "close_form" and ruins it for freezing
 
blanchet 
parents: 
45167 
diff
changeset
 | 
1776  | 
|> map (apsnd freeze_term)  | 
| 46340 | 1777  | 
|> map2 (pair o rpair (Local, General) o string_of_int)  | 
1778  | 
(0 upto length hyp_ts)  | 
|
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1779  | 
val ((conjs, facts), lam_facts) =  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1780  | 
(conjs, facts)  | 
| 46093 | 1781  | 
|> presimp ? pairself (map (apsnd (uncurry (presimp_prop ctxt))))  | 
| 45514 | 1782  | 
|> (if lam_trans = no_lamsN then  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1783  | 
rpair []  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1784  | 
else  | 
| 
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1785  | 
op @  | 
| 45514 | 1786  | 
#> preprocess_abstractions_in_terms trans_lams  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
1787  | 
#>> chop (length conjs))  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
1788  | 
val conjs = conjs |> make_conjecture ctxt format type_enc  | 
| 
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
 | 
1789  | 
val (fact_names, facts) =  | 
| 
43864
 
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
 
blanchet 
parents: 
43863 
diff
changeset
 | 
1790  | 
facts  | 
| 
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
 | 
1791  | 
|> map_filter (fn (name, (_, t)) =>  | 
| 
 
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
 | 
1792  | 
make_fact ctxt format type_enc true (name, t)  | 
| 
 
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
 | 
1793  | 
|> Option.map (pair name))  | 
| 43861 | 1794  | 
|> ListPair.unzip  | 
| 
46375
 
d724066ff3d0
reverted e2b1a86d59fc -- broke Metis's lambda-lifting
 
blanchet 
parents: 
46371 
diff
changeset
 | 
1795  | 
val lifted = lam_facts |> map (extract_lambda_def o snd o snd)  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1796  | 
val lam_facts =  | 
| 
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1797  | 
lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd)  | 
| 43861 | 1798  | 
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
 | 
1799  | 
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
 | 
1800  | 
val supers = tvar_classes_of_terms all_ts  | 
| 43189 | 1801  | 
val tycons = type_constrs_of_terms thy all_ts  | 
| 43861 | 1802  | 
val (supers, arity_clauses) =  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
1803  | 
if level_of_type_enc type_enc = No_Types then ([], [])  | 
| 
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
 | 
1804  | 
else make_arity_clauses thy tycons supers  | 
| 43861 | 1805  | 
val class_rel_clauses = make_class_rel_clauses 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
 | 
1806  | 
in  | 
| 45508 | 1807  | 
(fact_names |> map single, union (op =) subs supers, conjs,  | 
| 
45554
 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 
blanchet 
parents: 
45551 
diff
changeset
 | 
1808  | 
facts @ lam_facts, class_rel_clauses, arity_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
 | 
1809  | 
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
 | 
1810  | 
|
| 44495 | 1811  | 
val type_guard = `(make_fixed_const NONE) type_guard_name  | 
| 
43130
 
d73fc2e55308
implemented missing hAPP and ti cases of new path finder
 
blanchet 
parents: 
43129 
diff
changeset
 | 
1812  | 
|
| 
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
 | 
1813  | 
fun type_guard_iterm format type_enc T tm =  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
1814  | 
  IApp (IConst (type_guard, T --> @{typ bool}, [T])
 | 
| 
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
 | 
1815  | 
|> mangle_type_args_in_iterm format 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
 | 
1816  | 
|
| 43421 | 1817  | 
fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum  | 
1818  | 
| is_var_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =  | 
|
| 
43000
 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 
blanchet 
parents: 
42998 
diff
changeset
 | 
1819  | 
accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))  | 
| 43692 | 1820  | 
| 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
 | 
1821  | 
|
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1822  | 
fun is_var_ghost_type_arg_in_term thy polym_constrs name pos tm accum =  | 
| 44811 | 1823  | 
is_var_positively_naked_in_term name pos tm accum orelse  | 
1824  | 
let  | 
|
1825  | 
val var = ATerm (name, [])  | 
|
1826  | 
fun is_nasty_in_term (ATerm (_, [])) = false  | 
|
1827  | 
| is_nasty_in_term (ATerm ((s, _), tms)) =  | 
|
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1828  | 
let  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1829  | 
val ary = length tms  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1830  | 
val polym_constr = member (op =) polym_constrs s  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1831  | 
val ghosts = ghost_type_args thy s ary  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1832  | 
in  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1833  | 
exists (fn (j, tm) =>  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1834  | 
if polym_constr then  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1835  | 
member (op =) ghosts j andalso  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1836  | 
(tm = var orelse is_nasty_in_term tm)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1837  | 
else  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1838  | 
tm = var andalso member (op =) ghosts j)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1839  | 
(0 upto ary - 1 ~~ tms)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1840  | 
orelse (not polym_constr andalso exists is_nasty_in_term tms)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1841  | 
end  | 
| 44811 | 1842  | 
| is_nasty_in_term _ = true  | 
1843  | 
in is_nasty_in_term tm end  | 
|
1844  | 
||
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1845  | 
fun should_guard_var_in_formula thy polym_constrs level pos phi (SOME true)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1846  | 
name =  | 
| 44811 | 1847  | 
(case granularity_of_type_level level of  | 
1848  | 
All_Vars => true  | 
|
1849  | 
| Positively_Naked_Vars =>  | 
|
1850  | 
formula_fold pos (is_var_positively_naked_in_term name) phi false  | 
|
1851  | 
| Ghost_Type_Arg_Vars =>  | 
|
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1852  | 
formula_fold pos (is_var_ghost_type_arg_in_term thy polym_constrs name)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1853  | 
phi false)  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1854  | 
| should_guard_var_in_formula _ _ _ _ _ _ _ = true  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1855  | 
|
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1856  | 
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
 | 
1857  | 
|
| 
44405
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
1858  | 
fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false  | 
| 44768 | 1859  | 
| should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T =  | 
| 44811 | 1860  | 
granularity_of_type_level level <> All_Vars andalso  | 
| 44782 | 1861  | 
should_encode_type ctxt mono level T  | 
| 
44405
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
1862  | 
| should_generate_tag_bound_decl _ _ _ _ _ = false  | 
| 
44404
 
3111af540141
tuning, plus started implementing tag equation generation for existential variables
 
blanchet 
parents: 
44403 
diff
changeset
 | 
1863  | 
|
| 43677 | 1864  | 
fun mk_aterm format type_enc name T_args args =  | 
1865  | 
ATerm (name, map_filter (ho_term_for_type_arg format type_enc) T_args @ args)  | 
|
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
1866  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1867  | 
fun do_bound_type ctxt format mono type_enc =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1868  | 
case type_enc of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1869  | 
Simple_Types (_, _, level) =>  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1870  | 
fused_type ctxt mono level 0  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1871  | 
#> ho_type_from_typ format type_enc false 0 #> SOME  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1872  | 
| _ => K NONE  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1873  | 
|
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1874  | 
fun tag_with_type ctxt format mono type_enc pos T tm =  | 
| 43859 | 1875  | 
IConst (type_tag, T --> T, [T])  | 
| 
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
 | 
1876  | 
|> mangle_type_args_in_iterm format type_enc  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1877  | 
|> ho_term_from_iterm ctxt format mono type_enc pos  | 
| 43692 | 1878  | 
|> (fn ATerm (s, tms) => ATerm (s, tms @ [tm])  | 
1879  | 
| _ => raise Fail "unexpected lambda-abstraction")  | 
|
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1880  | 
and ho_term_from_iterm ctxt format mono type_enc =  | 
| 
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
 | 
1881  | 
let  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1882  | 
fun term site u =  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42956 
diff
changeset
 | 
1883  | 
let  | 
| 43859 | 1884  | 
val (head, args) = strip_iterm_comb u  | 
| 43677 | 1885  | 
val pos =  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1886  | 
case site of  | 
| 43677 | 1887  | 
Top_Level pos => pos  | 
1888  | 
| Eq_Arg pos => pos  | 
|
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
1889  | 
| _ => NONE  | 
| 43677 | 1890  | 
val t =  | 
1891  | 
case head of  | 
|
| 43859 | 1892  | 
IConst (name as (s, _), _, T_args) =>  | 
| 43677 | 1893  | 
let  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1894  | 
val arg_site = if is_tptp_equal s then Eq_Arg pos else Elsewhere  | 
| 43677 | 1895  | 
in  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1896  | 
map (term arg_site) args |> mk_aterm format type_enc name T_args  | 
| 43677 | 1897  | 
end  | 
| 43859 | 1898  | 
| IVar (name, _) =>  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1899  | 
map (term Elsewhere) args |> mk_aterm format type_enc name []  | 
| 43859 | 1900  | 
| IAbs ((name, T), tm) =>  | 
1901  | 
AAbs ((name, ho_type_from_typ format type_enc true 0 T),  | 
|
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1902  | 
term Elsewhere tm)  | 
| 43859 | 1903  | 
| IApp _ => raise Fail "impossible \"IApp\""  | 
1904  | 
val T = ityp_of u  | 
|
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42956 
diff
changeset
 | 
1905  | 
in  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1906  | 
if should_tag_with_type ctxt mono type_enc site u T then  | 
| 
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1907  | 
tag_with_type ctxt format mono type_enc pos T t  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1908  | 
else  | 
| 
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1909  | 
t  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42956 
diff
changeset
 | 
1910  | 
end  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1911  | 
in term o Top_Level end  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1912  | 
and formula_from_iformula ctxt polym_constrs format mono type_enc  | 
| 
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1913  | 
should_guard_var =  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
1914  | 
let  | 
| 44811 | 1915  | 
val thy = Proof_Context.theory_of ctxt  | 
1916  | 
val level = level_of_type_enc type_enc  | 
|
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1917  | 
val do_term = ho_term_from_iterm ctxt format 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
 | 
1918  | 
fun do_out_of_bound_type pos phi universal (name, T) =  | 
| 44399 | 1919  | 
if should_guard_type ctxt mono type_enc  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1920  | 
(fn () => should_guard_var thy polym_constrs level pos phi  | 
| 
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
1921  | 
universal name) T then  | 
| 43859 | 1922  | 
IVar (name, T)  | 
| 
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
 | 
1923  | 
|> type_guard_iterm format type_enc T  | 
| 
43361
 
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
 
blanchet 
parents: 
43324 
diff
changeset
 | 
1924  | 
|> do_term pos |> AAtom |> SOME  | 
| 
44405
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
1925  | 
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
 | 
1926  | 
let  | 
| 
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
1927  | 
val var = ATerm (name, [])  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
1928  | 
val tagged_var = tag_with_type ctxt format mono type_enc pos T var  | 
| 
44405
 
6fe1a89bb69a
generate tag equations for existential variables
 
blanchet 
parents: 
44404 
diff
changeset
 | 
1929  | 
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
 | 
1930  | 
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
 | 
1931  | 
NONE  | 
| 
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
 | 
1932  | 
fun do_formula pos (AQuant (q, xs, phi)) =  | 
| 
 
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
 | 
1933  | 
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
 | 
1934  | 
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
 | 
1935  | 
val universal = Option.map (q = AExists ? not) pos  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
1936  | 
val do_bound_type = do_bound_type ctxt format 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
 | 
1937  | 
in  | 
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
1938  | 
AQuant (q, xs |> map (apsnd (fn NONE => NONE  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
1939  | 
| SOME T => do_bound_type T)),  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
1940  | 
(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
 | 
1941  | 
(map_filter  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
1942  | 
(fn (_, NONE) => NONE  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
1943  | 
| (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
 | 
1944  | 
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
 | 
1945  | 
xs)  | 
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
1946  | 
phi)  | 
| 
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
1947  | 
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
 | 
1948  | 
| 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
 | 
1949  | 
| 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
 | 
1950  | 
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
 | 
1951  | 
|
| 
 
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
 | 
1952  | 
(* Each fact is given a unique fact number to avoid name clashes (e.g., because  | 
| 
 
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
 | 
1953  | 
of monomorphization). The TPTP explicitly forbids name clashes, and some of  | 
| 
 
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
 | 
1954  | 
the remote provers might care. *)  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1955  | 
fun formula_line_for_fact ctxt polym_constrs format prefix encode freshen pos  | 
| 46406 | 1956  | 
        mono type_enc rank (j, {name, stature, kind, iformula, atomic_types}) =
 | 
| 
43864
 
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
 
blanchet 
parents: 
43863 
diff
changeset
 | 
1957  | 
(prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name, kind,  | 
| 43859 | 1958  | 
iformula  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1959  | 
|> formula_from_iformula ctxt polym_constrs format mono type_enc  | 
| 
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1960  | 
should_guard_var_in_formula (if pos then SOME true else NONE)  | 
| 45377 | 1961  | 
|> close_formula_universally  | 
| 45920 | 1962  | 
|> bound_tvars type_enc true atomic_types,  | 
| 
43493
 
bdb11c68f142
generate type predicates for existentials/skolems, otherwise some problems might not be provable
 
blanchet 
parents: 
43423 
diff
changeset
 | 
1963  | 
NONE,  | 
| 46406 | 1964  | 
let val rank = rank j in  | 
1965  | 
case snd stature of  | 
|
1966  | 
Intro => isabelle_info format introN rank  | 
|
1967  | 
| Elim => isabelle_info format elimN rank  | 
|
1968  | 
| Simp => isabelle_info format simpN rank  | 
|
1969  | 
| Spec_Eq => isabelle_info format spec_eqN rank  | 
|
1970  | 
| _ => isabelle_info format "" rank  | 
|
1971  | 
end)  | 
|
| 
43493
 
bdb11c68f142
generate type predicates for existentials/skolems, otherwise some problems might not be provable
 
blanchet 
parents: 
43423 
diff
changeset
 | 
1972  | 
|> Formula  | 
| 
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
 | 
1973  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
1974  | 
fun formula_line_for_class_rel_clause format type_enc  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
1975  | 
        ({name, subclass, superclass, ...} : class_rel_clause) =
 | 
| 44622 | 1976  | 
let val ty_arg = ATerm (tvar_a_name, []) in  | 
| 
42577
 
78414ec6fa4e
made the format (TFF or FOF) of the TPTP problem a global argument of the problem again and have the ATPs report which formats they support
 
blanchet 
parents: 
42576 
diff
changeset
 | 
1977  | 
Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,  | 
| 44622 | 1978  | 
AConn (AImplies,  | 
| 44625 | 1979  | 
[type_class_formula type_enc subclass ty_arg,  | 
1980  | 
type_class_formula type_enc superclass ty_arg])  | 
|
| 45830 | 1981  | 
|> mk_aquant AForall  | 
1982  | 
[(tvar_a_name, atype_of_type_vars type_enc)],  | 
|
| 46406 | 1983  | 
NONE, isabelle_info format introN 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
 | 
1984  | 
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
 | 
1985  | 
|
| 44625 | 1986  | 
fun formula_from_arity_atom type_enc (class, t, args) =  | 
1987  | 
ATerm (t, map (fn arg => ATerm (arg, [])) args)  | 
|
1988  | 
|> type_class_formula type_enc class  | 
|
| 
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
 | 
1989  | 
|
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
1990  | 
fun formula_line_for_arity_clause format type_enc  | 
| 
45364
 
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
 
blanchet 
parents: 
45316 
diff
changeset
 | 
1991  | 
        ({name, prem_atoms, concl_atom} : arity_clause) =
 | 
| 43495 | 1992  | 
Formula (arity_clause_prefix ^ name, Axiom,  | 
| 44625 | 1993  | 
mk_ahorn (map (formula_from_arity_atom type_enc) prem_atoms)  | 
1994  | 
(formula_from_arity_atom type_enc concl_atom)  | 
|
| 
45364
 
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
 
blanchet 
parents: 
45316 
diff
changeset
 | 
1995  | 
|> mk_aquant AForall  | 
| 
 
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
 
blanchet 
parents: 
45316 
diff
changeset
 | 
1996  | 
(map (rpair (atype_of_type_vars type_enc)) (#3 concl_atom)),  | 
| 46406 | 1997  | 
NONE, isabelle_info format introN 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
 | 
1998  | 
|
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
1999  | 
fun formula_line_for_conjecture ctxt polym_constrs format mono type_enc  | 
| 43859 | 2000  | 
        ({name, kind, iformula, atomic_types, ...} : translated_formula) =
 | 
| 
42577
 
78414ec6fa4e
made the format (TFF or FOF) of the TPTP problem a global argument of the problem again and have the ATPs report which formats they support
 
blanchet 
parents: 
42576 
diff
changeset
 | 
2001  | 
Formula (conjecture_prefix ^ name, kind,  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
2002  | 
iformula  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2003  | 
|> formula_from_iformula ctxt polym_constrs format mono type_enc  | 
| 
45316
 
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
 
blanchet 
parents: 
45315 
diff
changeset
 | 
2004  | 
should_guard_var_in_formula (SOME false)  | 
| 45377 | 2005  | 
|> close_formula_universally  | 
| 46406 | 2006  | 
|> 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
 | 
2007  | 
|
| 44625 | 2008  | 
fun formula_line_for_free_type j phi =  | 
| 46406 | 2009  | 
Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis, phi, NONE, [])  | 
| 44626 | 2010  | 
fun formula_lines_for_free_types type_enc (facts : translated_formula list) =  | 
| 
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
 | 
2011  | 
let  | 
| 44625 | 2012  | 
val phis =  | 
2013  | 
fold (union (op =)) (map #atomic_types facts) []  | 
|
2014  | 
|> formulas_for_types type_enc add_sorts_on_tfree  | 
|
2015  | 
in map2 formula_line_for_free_type (0 upto length phis - 1) phis 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
 | 
2016  | 
|
| 
 
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
 | 
2017  | 
(** 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
 | 
2018  | 
|
| 44754 | 2019  | 
fun decl_line_for_class order s =  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2020  | 
let val name as (s, _) = `make_type_class s in  | 
| 44622 | 2021  | 
Decl (sym_decl_prefix ^ s, name,  | 
| 45779 | 2022  | 
if order = First_Order then  | 
2023  | 
ATyAbs ([tvar_a_name],  | 
|
2024  | 
if avoid_first_order_ghost_type_vars then  | 
|
2025  | 
AFun (a_itself_atype, bool_atype)  | 
|
2026  | 
else  | 
|
2027  | 
bool_atype)  | 
|
| 44622 | 2028  | 
else  | 
| 44754 | 2029  | 
AFun (atype_of_types, bool_atype))  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2030  | 
end  | 
| 
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2031  | 
|
| 
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2032  | 
fun decl_lines_for_classes type_enc classes =  | 
| 
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2033  | 
case type_enc of  | 
| 44754 | 2034  | 
Simple_Types (order, Polymorphic, _) =>  | 
2035  | 
map (decl_line_for_class order) classes  | 
|
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2036  | 
| _ => []  | 
| 
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2037  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2038  | 
fun sym_decl_table_for_facts ctxt format type_enc sym_tab  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2039  | 
(conjs, facts, extra_tms) =  | 
| 42574 | 2040  | 
let  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2041  | 
fun add_iterm_syms tm =  | 
| 43859 | 2042  | 
let val (head, args) = strip_iterm_comb tm in  | 
| 42574 | 2043  | 
(case head of  | 
| 43859 | 2044  | 
IConst ((s, s'), T, T_args) =>  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2045  | 
let  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2046  | 
val (pred_sym, in_conj) =  | 
| 
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2047  | 
case Symtab.lookup sym_tab s of  | 
| 44853 | 2048  | 
                 SOME ({pred_sym, in_conj, ...} : sym_info) =>
 | 
2049  | 
(pred_sym, in_conj)  | 
|
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2050  | 
| NONE => (false, false)  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2051  | 
val decl_sym =  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2052  | 
(case type_enc of  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2053  | 
Guards _ => not pred_sym  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2054  | 
| _ => true) andalso  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2055  | 
is_tptp_user_symbol s  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2056  | 
in  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2057  | 
if decl_sym then  | 
| 
42576
 
a8a80a2a34be
merge symbol declarations that are type-instances of each other -- useful for type system "Args true" with monomorphization turned off
 
blanchet 
parents: 
42575 
diff
changeset
 | 
2058  | 
Symtab.map_default (s, [])  | 
| 
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
 | 
2059  | 
(insert_type ctxt #3 (s', T_args, T, pred_sym, length args,  | 
| 
 
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
 | 
2060  | 
in_conj))  | 
| 42574 | 2061  | 
else  | 
2062  | 
I  | 
|
2063  | 
end  | 
|
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2064  | 
| IAbs (_, tm) => add_iterm_syms tm  | 
| 42574 | 2065  | 
| _ => I)  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2066  | 
#> fold add_iterm_syms args  | 
| 42574 | 2067  | 
end  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2068  | 
val add_fact_syms = K add_iterm_syms |> formula_fold NONE |> fact_lift  | 
| 43966 | 2069  | 
fun add_formula_var_types (AQuant (_, xs, phi)) =  | 
2070  | 
fold (fn (_, SOME T) => insert_type ctxt I T | _ => I) xs  | 
|
2071  | 
#> add_formula_var_types phi  | 
|
2072  | 
| add_formula_var_types (AConn (_, phis)) =  | 
|
2073  | 
fold add_formula_var_types phis  | 
|
2074  | 
| add_formula_var_types _ = I  | 
|
2075  | 
fun var_types () =  | 
|
2076  | 
if polymorphism_of_type_enc type_enc = Polymorphic then [tvar_a]  | 
|
2077  | 
else fold (fact_lift add_formula_var_types) (conjs @ facts) []  | 
|
2078  | 
fun add_undefined_const T =  | 
|
| 43984 | 2079  | 
let  | 
2080  | 
val (s, s') =  | 
|
| 44622 | 2081  | 
          `(make_fixed_const NONE) @{const_name undefined}
 | 
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2082  | 
          |> (case type_arg_policy [] type_enc @{const_name undefined} of
 | 
| 44771 | 2083  | 
Mangled_Type_Args => mangled_const_name format type_enc [T]  | 
| 44001 | 2084  | 
| _ => I)  | 
| 43984 | 2085  | 
in  | 
2086  | 
Symtab.map_default (s, [])  | 
|
2087  | 
(insert_type ctxt #3 (s', [T], T, false, 0, false))  | 
|
2088  | 
end  | 
|
| 44622 | 2089  | 
fun add_TYPE_const () =  | 
2090  | 
let val (s, s') = TYPE_name in  | 
|
2091  | 
Symtab.map_default (s, [])  | 
|
2092  | 
(insert_type ctxt #3  | 
|
2093  | 
                         (s', [tvar_a], @{typ "'a itself"}, false, 0, false))
 | 
|
2094  | 
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
 | 
2095  | 
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
 | 
2096  | 
Symtab.empty  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
2097  | 
|> is_type_enc_fairly_sound type_enc  | 
| 
44829
 
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
 
blanchet 
parents: 
44814 
diff
changeset
 | 
2098  | 
? (fold (fold add_fact_syms) [conjs, facts]  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2099  | 
#> 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
 | 
2100  | 
#> (case type_enc of  | 
| 45829 | 2101  | 
Simple_Types (First_Order, Polymorphic, _) =>  | 
2102  | 
if avoid_first_order_ghost_type_vars then add_TYPE_const ()  | 
|
2103  | 
else I  | 
|
| 44786 | 2104  | 
| Simple_Types _ => I  | 
| 
43985
 
33d8b99531c2
no need for existential witnesses for sorts in TFF and THF formats
 
blanchet 
parents: 
43984 
diff
changeset
 | 
2105  | 
| _ => 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
 | 
2106  | 
end  | 
| 42533 | 2107  | 
|
| 44399 | 2108  | 
(* We add "bool" in case the helper "True_or_False" is included later. *)  | 
| 
44634
 
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
 
blanchet 
parents: 
44626 
diff
changeset
 | 
2109  | 
fun default_mono level =  | 
| 44399 | 2110  | 
  {maybe_finite_Ts = [@{typ bool}],
 | 
2111  | 
   surely_finite_Ts = [@{typ bool}],
 | 
|
2112  | 
maybe_infinite_Ts = known_infinite_types,  | 
|
| 
44634
 
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
 
blanchet 
parents: 
44626 
diff
changeset
 | 
2113  | 
surely_infinite_Ts =  | 
| 
 
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
 
blanchet 
parents: 
44626 
diff
changeset
 | 
2114  | 
case level of  | 
| 46301 | 2115  | 
Noninf_Nonmono_Types (Strict, _) => []  | 
| 
44634
 
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
 
blanchet 
parents: 
44626 
diff
changeset
 | 
2116  | 
| _ => known_infinite_types,  | 
| 44399 | 2117  | 
   maybe_nonmono_Ts = [@{typ bool}]}
 | 
2118  | 
||
| 42685 | 2119  | 
(* This inference is described in section 2.3 of Claessen et al.'s "Sorting it  | 
2120  | 
out with monotonicity" paper presented at CADE 2011. *)  | 
|
| 44399 | 2121  | 
fun add_iterm_mononotonicity_info _ _ (SOME false) _ mono = mono  | 
2122  | 
| add_iterm_mononotonicity_info ctxt level _  | 
|
2123  | 
(IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2))  | 
|
2124  | 
        (mono as {maybe_finite_Ts, surely_finite_Ts, maybe_infinite_Ts,
 | 
|
2125  | 
surely_infinite_Ts, maybe_nonmono_Ts}) =  | 
|
| 
44403
 
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
 
blanchet 
parents: 
44402 
diff
changeset
 | 
2126  | 
if is_tptp_equal s andalso exists is_maybe_universal_var [tm1, tm2] then  | 
| 44399 | 2127  | 
case level of  | 
| 46301 | 2128  | 
Noninf_Nonmono_Types (strictness, _) =>  | 
| 44399 | 2129  | 
if exists (type_instance ctxt T) surely_infinite_Ts orelse  | 
| 
44859
 
237ba63d6041
fixed definition of type intersection (soundness bug)
 
blanchet 
parents: 
44853 
diff
changeset
 | 
2130  | 
member (type_equiv ctxt) maybe_finite_Ts T then  | 
| 44399 | 2131  | 
mono  | 
| 46301 | 2132  | 
else if is_type_kind_of_surely_infinite ctxt strictness  | 
| 44500 | 2133  | 
surely_infinite_Ts T then  | 
| 44399 | 2134  | 
          {maybe_finite_Ts = maybe_finite_Ts,
 | 
2135  | 
surely_finite_Ts = surely_finite_Ts,  | 
|
2136  | 
maybe_infinite_Ts = maybe_infinite_Ts,  | 
|
2137  | 
surely_infinite_Ts = surely_infinite_Ts |> insert_type ctxt I T,  | 
|
2138  | 
maybe_nonmono_Ts = maybe_nonmono_Ts}  | 
|
2139  | 
else  | 
|
| 
44859
 
237ba63d6041
fixed definition of type intersection (soundness bug)
 
blanchet 
parents: 
44853 
diff
changeset
 | 
2140  | 
          {maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv ctxt) T,
 | 
| 44399 | 2141  | 
surely_finite_Ts = surely_finite_Ts,  | 
2142  | 
maybe_infinite_Ts = maybe_infinite_Ts,  | 
|
2143  | 
surely_infinite_Ts = surely_infinite_Ts,  | 
|
2144  | 
maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type ctxt I T}  | 
|
| 44768 | 2145  | 
| Fin_Nonmono_Types _ =>  | 
| 44399 | 2146  | 
if exists (type_instance ctxt T) surely_finite_Ts orelse  | 
| 
44859
 
237ba63d6041
fixed definition of type intersection (soundness bug)
 
blanchet 
parents: 
44853 
diff
changeset
 | 
2147  | 
member (type_equiv ctxt) maybe_infinite_Ts T then  | 
| 44399 | 2148  | 
mono  | 
2149  | 
else if is_type_surely_finite ctxt T then  | 
|
2150  | 
          {maybe_finite_Ts = maybe_finite_Ts,
 | 
|
2151  | 
surely_finite_Ts = surely_finite_Ts |> insert_type ctxt I T,  | 
|
2152  | 
maybe_infinite_Ts = maybe_infinite_Ts,  | 
|
2153  | 
surely_infinite_Ts = surely_infinite_Ts,  | 
|
2154  | 
maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type ctxt I T}  | 
|
2155  | 
else  | 
|
2156  | 
          {maybe_finite_Ts = maybe_finite_Ts,
 | 
|
2157  | 
surely_finite_Ts = surely_finite_Ts,  | 
|
| 
44859
 
237ba63d6041
fixed definition of type intersection (soundness bug)
 
blanchet 
parents: 
44853 
diff
changeset
 | 
2158  | 
maybe_infinite_Ts = maybe_infinite_Ts |> insert (type_equiv ctxt) T,  | 
| 44399 | 2159  | 
surely_infinite_Ts = surely_infinite_Ts,  | 
2160  | 
maybe_nonmono_Ts = maybe_nonmono_Ts}  | 
|
2161  | 
| _ => mono  | 
|
2162  | 
else  | 
|
2163  | 
mono  | 
|
2164  | 
| add_iterm_mononotonicity_info _ _ _ _ mono = mono  | 
|
2165  | 
fun add_fact_mononotonicity_info ctxt level  | 
|
2166  | 
        ({kind, iformula, ...} : translated_formula) =
 | 
|
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2167  | 
formula_fold (SOME (kind <> Conjecture))  | 
| 44399 | 2168  | 
(add_iterm_mononotonicity_info ctxt level) iformula  | 
2169  | 
fun mononotonicity_info_for_facts ctxt type_enc facts =  | 
|
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
2170  | 
let val level = level_of_type_enc type_enc in  | 
| 
44634
 
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
 
blanchet 
parents: 
44626 
diff
changeset
 | 
2171  | 
default_mono level  | 
| 44399 | 2172  | 
|> is_type_level_monotonicity_based level  | 
2173  | 
? 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
 | 
2174  | 
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
 | 
2175  | 
|
| 44501 | 2176  | 
fun add_iformula_monotonic_types ctxt mono type_enc =  | 
2177  | 
let  | 
|
2178  | 
val level = level_of_type_enc type_enc  | 
|
2179  | 
val should_encode = should_encode_type ctxt mono level  | 
|
| 
44504
 
6f29df8d2007
fixed inverted logic and improve precision when handling monotonic types in polymorphic encodings
 
blanchet 
parents: 
44502 
diff
changeset
 | 
2180  | 
fun add_type T = not (should_encode T) ? insert_type ctxt I T  | 
| 44506 | 2181  | 
fun add_args (IApp (tm1, tm2)) = add_args tm1 #> add_term tm2  | 
2182  | 
| add_args _ = I  | 
|
2183  | 
and add_term tm = add_type (ityp_of tm) #> add_args tm  | 
|
| 44501 | 2184  | 
in formula_fold NONE (K add_term) end  | 
2185  | 
fun add_fact_monotonic_types ctxt mono type_enc =  | 
|
2186  | 
add_iformula_monotonic_types ctxt mono type_enc |> fact_lift  | 
|
2187  | 
fun monotonic_types_for_facts ctxt mono type_enc facts =  | 
|
| 44811 | 2188  | 
let val level = level_of_type_enc type_enc in  | 
2189  | 
[] |> (polymorphism_of_type_enc type_enc = Polymorphic andalso  | 
|
2190  | 
is_type_level_monotonicity_based level andalso  | 
|
2191  | 
granularity_of_type_level level <> Ghost_Type_Arg_Vars)  | 
|
2192  | 
? fold (add_fact_monotonic_types ctxt mono type_enc) facts  | 
|
2193  | 
end  | 
|
| 44501 | 2194  | 
|
| 44399 | 2195  | 
fun formula_line_for_guards_mono_type ctxt format mono type_enc T =  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2196  | 
Formula (guards_sym_formula_prefix ^  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2197  | 
ascii_of (mangled_type format type_enc T),  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2198  | 
Axiom,  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2199  | 
IConst (`make_bound_var "X", T, [])  | 
| 
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
 | 
2200  | 
|> type_guard_iterm format type_enc T  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2201  | 
|> AAtom  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2202  | 
|> formula_from_iformula ctxt [] format mono type_enc  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
2203  | 
always_guard_var_in_formula (SOME true)  | 
| 45377 | 2204  | 
|> close_formula_universally  | 
| 45920 | 2205  | 
|> bound_tvars type_enc true (atomic_types_of T),  | 
| 46406 | 2206  | 
NONE, isabelle_info format introN helper_rank)  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2207  | 
|
| 44399 | 2208  | 
fun formula_line_for_tags_mono_type ctxt format mono type_enc T =  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2209  | 
let val x_var = ATerm (`make_bound_var "X", []) in  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2210  | 
Formula (tags_sym_formula_prefix ^  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2211  | 
ascii_of (mangled_type format type_enc T),  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2212  | 
Axiom,  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2213  | 
eq_formula type_enc (atomic_types_of T) [] false  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
2214  | 
(tag_with_type ctxt format mono type_enc NONE T x_var) x_var,  | 
| 46406 | 2215  | 
NONE, isabelle_info format spec_eqN helper_rank)  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2216  | 
end  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2217  | 
|
| 44399 | 2218  | 
fun problem_lines_for_mono_types ctxt format mono type_enc Ts =  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2219  | 
case type_enc of  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2220  | 
Simple_Types _ => []  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2221  | 
| Guards _ =>  | 
| 44399 | 2222  | 
map (formula_line_for_guards_mono_type ctxt format mono type_enc) Ts  | 
2223  | 
| Tags _ => map (formula_line_for_tags_mono_type ctxt format mono type_enc) Ts  | 
|
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2224  | 
|
| 44399 | 2225  | 
fun decl_line_for_sym ctxt format mono type_enc s  | 
| 
42994
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
2226  | 
(s', T_args, T, pred_sym, ary, _) =  | 
| 
 
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
 
blanchet 
parents: 
42966 
diff
changeset
 | 
2227  | 
let  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2228  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2229  | 
val (T, T_args) =  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2230  | 
if null T_args then  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2231  | 
(T, [])  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45509 
diff
changeset
 | 
2232  | 
else case unprefix_and_unascii const_prefix s of  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2233  | 
SOME s' =>  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2234  | 
let  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2235  | 
val s' = s' |> invert_const  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2236  | 
val T = s' |> robust_const_type thy  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2237  | 
in (T, robust_const_typargs thy (s', T)) end  | 
| 45509 | 2238  | 
| 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
 | 
2239  | 
in  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2240  | 
Decl (sym_decl_prefix ^ s, (s, s'),  | 
| 
44594
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2241  | 
T |> fused_type ctxt mono (level_of_type_enc type_enc) ary  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2242  | 
|> ho_type_from_typ format type_enc pred_sym ary  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2243  | 
|> not (null T_args)  | 
| 
 
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
 
blanchet 
parents: 
44593 
diff
changeset
 | 
2244  | 
? curry ATyAbs (map (tvar_name o fst 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
 | 
2245  | 
end  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2246  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2247  | 
fun honor_conj_sym_kind in_conj conj_sym_kind =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2248  | 
if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2249  | 
else (Axiom, I)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2250  | 
|
| 44399 | 2251  | 
fun formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono type_enc n s  | 
2252  | 
j (s', T_args, T, _, ary, in_conj) =  | 
|
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2253  | 
let  | 
| 44811 | 2254  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2255  | 
val (kind, maybe_negate) = honor_conj_sym_kind in_conj conj_sym_kind  | 
| 
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
 | 
2256  | 
val (arg_Ts, res_T) = chop_fun ary T  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2257  | 
val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2258  | 
val bounds =  | 
| 43859 | 2259  | 
bound_names ~~ arg_Ts |> map (fn (name, T) => IConst (name, T, []))  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2260  | 
val bound_Ts =  | 
| 44811 | 2261  | 
if exists (curry (op =) dummyT) T_args then  | 
2262  | 
case level_of_type_enc type_enc of  | 
|
2263  | 
All_Types => map SOME arg_Ts  | 
|
2264  | 
| level =>  | 
|
2265  | 
if granularity_of_type_level level = Ghost_Type_Arg_Vars then  | 
|
2266  | 
let val ghosts = ghost_type_args thy s ary in  | 
|
2267  | 
map2 (fn j => if member (op =) ghosts j then SOME else K NONE)  | 
|
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2268  | 
(0 upto ary - 1) arg_Ts  | 
| 44811 | 2269  | 
end  | 
2270  | 
else  | 
|
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2271  | 
replicate ary NONE  | 
| 44811 | 2272  | 
else  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2273  | 
replicate ary NONE  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2274  | 
in  | 
| 43989 | 2275  | 
Formula (guards_sym_formula_prefix ^ s ^  | 
| 
42709
 
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
 
blanchet 
parents: 
42701 
diff
changeset
 | 
2276  | 
(if n > 1 then "_" ^ string_of_int j else ""), kind,  | 
| 43859 | 2277  | 
IConst ((s, s'), T, T_args)  | 
2278  | 
|> fold (curry (IApp o swap)) bounds  | 
|
| 
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
 | 
2279  | 
|> type_guard_iterm format type_enc res_T  | 
| 42963 | 2280  | 
|> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2281  | 
|> formula_from_iformula ctxt [] format mono type_enc  | 
| 
45948
 
f88f502d635f
extend previous optimizations to guard-based encodings
 
blanchet 
parents: 
45947 
diff
changeset
 | 
2282  | 
always_guard_var_in_formula (SOME true)  | 
| 45377 | 2283  | 
|> close_formula_universally  | 
| 45920 | 2284  | 
|> 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
 | 
2285  | 
|> maybe_negate,  | 
| 46406 | 2286  | 
NONE, isabelle_info format introN helper_rank)  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2287  | 
end  | 
| 
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2288  | 
|
| 44811 | 2289  | 
fun formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono type_enc n s  | 
2290  | 
(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
 | 
2291  | 
let  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2292  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2293  | 
val level = level_of_type_enc type_enc  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2294  | 
val grain = granularity_of_type_level level  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2295  | 
val ident_base =  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2296  | 
tags_sym_formula_prefix ^ s ^  | 
| 
43125
 
ddf63baabdec
distinguish different kinds of typing informations in the fact name
 
blanchet 
parents: 
43121 
diff
changeset
 | 
2297  | 
(if n > 1 then "_" ^ string_of_int j else "")  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2298  | 
val (kind, maybe_negate) = honor_conj_sym_kind in_conj conj_sym_kind  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2299  | 
val (arg_Ts, res_T) = chop_fun ary T  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2300  | 
val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2301  | 
val bounds = bound_names |> map (fn name => ATerm (name, []))  | 
| 43677 | 2302  | 
val cst = mk_aterm format type_enc (s, s') T_args  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2303  | 
val eq = maybe_negate oo eq_formula type_enc (atomic_types_of T) [] pred_sym  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2304  | 
val should_encode = should_encode_type ctxt mono level  | 
| 
45949
 
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
 
blanchet 
parents: 
45948 
diff
changeset
 | 
2305  | 
val tag_with = tag_with_type ctxt format mono type_enc NONE  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2306  | 
val add_formula_for_res =  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2307  | 
if should_encode res_T then  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2308  | 
let  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2309  | 
val tagged_bounds =  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2310  | 
if grain = Ghost_Type_Arg_Vars then  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2311  | 
let val ghosts = ghost_type_args thy s ary in  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2312  | 
map2 (fn (j, arg_T) => member (op =) ghosts j ? tag_with arg_T)  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2313  | 
(0 upto ary - 1 ~~ arg_Ts) bounds  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2314  | 
end  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2315  | 
else  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2316  | 
bounds  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2317  | 
in  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2318  | 
cons (Formula (ident_base ^ "_res", kind,  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2319  | 
eq (tag_with res_T (cst bounds)) (cst tagged_bounds),  | 
| 46406 | 2320  | 
NONE, isabelle_info format spec_eqN helper_rank))  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2321  | 
end  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2322  | 
else  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2323  | 
I  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2324  | 
fun add_formula_for_arg k =  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2325  | 
let val arg_T = nth arg_Ts k in  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2326  | 
if should_encode arg_T then  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2327  | 
case chop k bounds of  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2328  | 
(bounds1, bound :: bounds2) =>  | 
| 
42852
 
40649ab0cead
honor "conj_sym_kind" also for tag symbol declarations
 
blanchet 
parents: 
42851 
diff
changeset
 | 
2329  | 
cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), kind,  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2330  | 
eq (cst (bounds1 @ tag_with arg_T bound :: bounds2))  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2331  | 
(cst bounds),  | 
| 46406 | 2332  | 
NONE, isabelle_info format spec_eqN helper_rank))  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2333  | 
| _ => raise Fail "expected nonempty tail"  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2334  | 
else  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2335  | 
I  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2336  | 
end  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2337  | 
in  | 
| 
42834
 
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
 
blanchet 
parents: 
42832 
diff
changeset
 | 
2338  | 
[] |> not pred_sym ? add_formula_for_res  | 
| 
44814
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2339  | 
|> (Config.get ctxt type_tag_arguments andalso  | 
| 
 
52318464c73b
also implemented ghost version of the tagged encodings
 
blanchet 
parents: 
44812 
diff
changeset
 | 
2340  | 
grain = Positively_Naked_Vars)  | 
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
2341  | 
? fold add_formula_for_arg (ary - 1 downto 0)  | 
| 
42829
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2342  | 
end  | 
| 
 
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
 
blanchet 
parents: 
42828 
diff
changeset
 | 
2343  | 
|
| 42836 | 2344  | 
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd  | 
2345  | 
||
| 45780 | 2346  | 
fun rationalize_decls ctxt (decls as decl :: (decls' as _ :: _)) =  | 
2347  | 
let  | 
|
2348  | 
val T = result_type_of_decl decl  | 
|
2349  | 
|> map_type_tvar (fn (z, _) => TVar (z, HOLogic.typeS))  | 
|
2350  | 
in  | 
|
2351  | 
if forall (type_generalization ctxt T o result_type_of_decl) decls' then  | 
|
2352  | 
[decl]  | 
|
2353  | 
else  | 
|
2354  | 
decls  | 
|
2355  | 
end  | 
|
2356  | 
| rationalize_decls _ decls = decls  | 
|
2357  | 
||
| 44399 | 2358  | 
fun problem_lines_for_sym_decls ctxt format conj_sym_kind mono type_enc  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2359  | 
(s, decls) =  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
2360  | 
case type_enc of  | 
| 
45831
 
8f8fce7bd32b
avoid multiple type decls in TFF (improves on cef82dc1462d)
 
blanchet 
parents: 
45830 
diff
changeset
 | 
2361  | 
Simple_Types _ => [decl_line_for_sym ctxt format mono type_enc s (hd decls)]  | 
| 44768 | 2362  | 
| Guards (_, level) =>  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2363  | 
let  | 
| 45780 | 2364  | 
val decls = decls |> rationalize_decls ctxt  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2365  | 
val n = length decls  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2366  | 
val decls =  | 
| 44399 | 2367  | 
decls |> filter (should_encode_type ctxt mono level  | 
| 
43401
 
e93dfcb53535
fixed soundness bug made more visible by previous change
 
blanchet 
parents: 
43399 
diff
changeset
 | 
2368  | 
o result_type_of_decl)  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2369  | 
in  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2370  | 
(0 upto length decls - 1, decls)  | 
| 44399 | 2371  | 
|-> map2 (formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono  | 
2372  | 
type_enc n s)  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2373  | 
end  | 
| 44768 | 2374  | 
| Tags (_, level) =>  | 
| 44811 | 2375  | 
if granularity_of_type_level level = All_Vars then  | 
| 44768 | 2376  | 
[]  | 
2377  | 
else  | 
|
2378  | 
let val n = length decls in  | 
|
2379  | 
(0 upto n - 1 ~~ decls)  | 
|
| 44811 | 2380  | 
|> maps (formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono  | 
2381  | 
type_enc n s)  | 
|
| 44768 | 2382  | 
end  | 
| 
42579
 
2552c09b1a72
implement the new ATP type system in Sledgehammer
 
blanchet 
parents: 
42577 
diff
changeset
 | 
2383  | 
|
| 44399 | 2384  | 
fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono type_enc  | 
| 44501 | 2385  | 
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
 | 
2386  | 
let  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2387  | 
val syms = sym_decl_tab |> Symtab.dest |> sort_wrt fst  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2388  | 
val mono_lines =  | 
| 44399 | 2389  | 
problem_lines_for_mono_types ctxt format mono type_enc mono_Ts  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2390  | 
val decl_lines =  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2391  | 
fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2392  | 
mono type_enc)  | 
| 
44396
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2393  | 
syms []  | 
| 
 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 
blanchet 
parents: 
44394 
diff
changeset
 | 
2394  | 
in mono_lines @ decl_lines end  | 
| 
42543
 
f9d402d144d4
declare TFF types so that SNARK can be used with types
 
blanchet 
parents: 
42542 
diff
changeset
 | 
2395  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2396  | 
fun pair_append (xs1, xs2) (ys1, ys2) = (xs1 @ ys1, xs2 @ ys2)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2397  | 
|
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2398  | 
fun do_uncurried_alias_lines_for_sym ctxt monom_constrs format conj_sym_kind  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2399  | 
mono type_enc sym_tab0 sym_tab base_s0 types in_conj =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2400  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2401  | 
fun do_alias ary =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2402  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2403  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2404  | 
val (kind, maybe_negate) = honor_conj_sym_kind in_conj conj_sym_kind  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2405  | 
val base_name = base_s0 |> `(make_fixed_const (SOME format))  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2406  | 
val T = case types of [T] => T | _ => robust_const_type thy base_s0  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2407  | 
val T_args = robust_const_typargs thy (base_s0, T)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2408  | 
val (base_name as (base_s, _), T_args) =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2409  | 
mangle_type_args_in_const format type_enc base_name T_args  | 
| 
46402
 
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
 
blanchet 
parents: 
46400 
diff
changeset
 | 
2410  | 
val base_ary = min_ary_of sym_tab0 base_s  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2411  | 
fun do_const name = IConst (name, T, T_args)  | 
| 46437 | 2412  | 
val filter_ty_args =  | 
| 
46410
 
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
 
blanchet 
parents: 
46409 
diff
changeset
 | 
2413  | 
filter_type_args_in_iterm thy monom_constrs type_enc  | 
| 46437 | 2414  | 
val ho_term_of =  | 
2415  | 
ho_term_from_iterm ctxt format mono type_enc (SOME true)  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2416  | 
val name1 as (s1, _) =  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2417  | 
base_name |> ary - 1 > base_ary ? aliased_uncurried (ary - 1)  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2418  | 
val name2 as (s2, _) = base_name |> aliased_uncurried ary  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2419  | 
val (arg_Ts, _) = chop_fun ary T  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2420  | 
val bound_names =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2421  | 
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
 | 
2422  | 
val bounds = bound_names ~~ arg_Ts  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2423  | 
val (first_bounds, last_bound) =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2424  | 
bounds |> map (fn (name, T) => IConst (name, T, [])) |> split_last  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2425  | 
val tm1 =  | 
| 
46410
 
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
 
blanchet 
parents: 
46409 
diff
changeset
 | 
2426  | 
mk_app_op format type_enc (list_app (do_const name1) first_bounds)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2427  | 
last_bound  | 
| 46437 | 2428  | 
|> filter_ty_args  | 
2429  | 
val tm2 =  | 
|
2430  | 
list_app (do_const name2) (first_bounds @ [last_bound])  | 
|
2431  | 
|> filter_ty_args  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2432  | 
val do_bound_type = do_bound_type ctxt format mono type_enc  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2433  | 
val eq =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2434  | 
eq_formula type_enc (atomic_types_of T)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2435  | 
(map (apsnd do_bound_type) bounds) false  | 
| 46437 | 2436  | 
(ho_term_of tm1) (ho_term_of tm2)  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2437  | 
in  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2438  | 
([tm1, tm2],  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2439  | 
[Formula (uncurried_alias_eq_prefix ^ s2, kind, eq |> maybe_negate,  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2440  | 
NONE, isabelle_info format spec_eqN helper_rank)])  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2441  | 
|> (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
 | 
2442  | 
else pair_append (do_alias (ary - 1)))  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2443  | 
end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2444  | 
in do_alias end  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2445  | 
fun uncurried_alias_lines_for_sym ctxt monom_constrs format conj_sym_kind mono  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2446  | 
type_enc sym_tab0 sym_tab  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2447  | 
        (s, {min_ary, types, in_conj, ...} : sym_info) =
 | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2448  | 
case unprefix_and_unascii const_prefix s of  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2449  | 
SOME mangled_s =>  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2450  | 
if String.isSubstring uncurried_alias_sep mangled_s then  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2451  | 
let  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2452  | 
val base_s0 = mangled_s |> unmangled_const_name |> hd |> invert_const  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2453  | 
in  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2454  | 
do_uncurried_alias_lines_for_sym ctxt monom_constrs format conj_sym_kind  | 
| 
46402
 
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
 
blanchet 
parents: 
46400 
diff
changeset
 | 
2455  | 
mono type_enc sym_tab0 sym_tab base_s0 types in_conj min_ary  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2456  | 
end  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2457  | 
else  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2458  | 
([], [])  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2459  | 
| NONE => ([], [])  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2460  | 
fun uncurried_alias_lines_for_sym_table ctxt monom_constrs format conj_sym_kind  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2461  | 
mono type_enc uncurried_aliases sym_tab0 sym_tab =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2462  | 
([], [])  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2463  | 
|> uncurried_aliases  | 
| 
46402
 
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
 
blanchet 
parents: 
46400 
diff
changeset
 | 
2464  | 
? Symtab.fold_rev  | 
| 
 
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
 
blanchet 
parents: 
46400 
diff
changeset
 | 
2465  | 
(pair_append  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2466  | 
o uncurried_alias_lines_for_sym ctxt monom_constrs format  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2467  | 
conj_sym_kind mono type_enc sym_tab0 sym_tab) sym_tab  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2468  | 
|
| 44768 | 2469  | 
fun needs_type_tag_idempotence ctxt (Tags (poly, level)) =  | 
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
2470  | 
Config.get ctxt type_tag_idempotence andalso  | 
| 44768 | 2471  | 
is_type_level_monotonicity_based level andalso  | 
2472  | 
poly <> Mangled_Monomorphic  | 
|
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
2473  | 
| needs_type_tag_idempotence _ _ = false  | 
| 
42831
 
c9b0968484fb
more work on "shallow" encoding + adjustments to other encodings
 
blanchet 
parents: 
42830 
diff
changeset
 | 
2474  | 
|
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2475  | 
val implicit_declsN = "Should-be-implicit typings"  | 
| 
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2476  | 
val explicit_declsN = "Explicit typings"  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2477  | 
val uncurried_alias_eqsN = "Uncurried aliases"  | 
| 41157 | 2478  | 
val factsN = "Relevant facts"  | 
2479  | 
val class_relsN = "Class relationships"  | 
|
| 
42543
 
f9d402d144d4
declare TFF types so that SNARK can be used with types
 
blanchet 
parents: 
42542 
diff
changeset
 | 
2480  | 
val aritiesN = "Arities"  | 
| 41157 | 2481  | 
val helpersN = "Helper facts"  | 
2482  | 
val conjsN = "Conjectures"  | 
|
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2483  | 
val free_typesN = "Type variables"  | 
| 41157 | 2484  | 
|
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2485  | 
(* 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
 | 
2486  | 
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
 | 
2487  | 
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
 | 
2488  | 
|
| 
45875
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2489  | 
fun default_type type_enc pred_sym s =  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2490  | 
let  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2491  | 
val ind =  | 
| 
45875
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2492  | 
case type_enc of  | 
| 
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2493  | 
Simple_Types _ =>  | 
| 
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2494  | 
if String.isPrefix type_const_prefix s then atype_of_types  | 
| 
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2495  | 
else individual_atype  | 
| 
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2496  | 
| _ => individual_atype  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2497  | 
fun typ 0 = if pred_sym then bool_atype else ind  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2498  | 
| typ ary = AFun (ind, typ (ary - 1))  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2499  | 
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
 | 
2500  | 
|
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2501  | 
fun nary_type_constr_type n =  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2502  | 
funpow n (curry AFun atype_of_types) atype_of_types  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2503  | 
|
| 
45875
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2504  | 
fun undeclared_syms_in_problem type_enc problem =  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2505  | 
let  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2506  | 
val declared = declared_syms_in_problem problem  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2507  | 
fun do_sym (name as (s, _)) ty =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2508  | 
if is_tptp_user_symbol s andalso not (member (op =) declared name) then  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2509  | 
AList.default (op =) (name, ty)  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2510  | 
else  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2511  | 
I  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2512  | 
fun do_type (AType (name, tys)) =  | 
| 
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2513  | 
do_sym name (fn () => nary_type_constr_type (length tys))  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2514  | 
#> fold do_type tys  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2515  | 
| do_type (AFun (ty1, ty2)) = do_type ty1 #> do_type ty2  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2516  | 
| do_type (ATyAbs (_, ty)) = do_type ty  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2517  | 
fun do_term pred_sym (ATerm (name as (s, _), tms)) =  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2518  | 
do_sym name (fn _ => default_type type_enc pred_sym s (length tms))  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2519  | 
#> fold (do_term false) tms  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2520  | 
| do_term _ (AAbs ((_, ty), tm)) = do_type ty #> do_term false tm  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2521  | 
fun do_formula (AQuant (_, xs, phi)) =  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2522  | 
fold do_type (map_filter snd xs) #> do_formula phi  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2523  | 
| 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
 | 
2524  | 
| do_formula (AAtom tm) = do_term true tm  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2525  | 
fun do_problem_line (Decl (_, _, ty)) = do_type ty  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2526  | 
| do_problem_line (Formula (_, _, phi, _, _)) = do_formula phi  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2527  | 
in  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2528  | 
fold (fold do_problem_line o snd) problem []  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2529  | 
|> filter_out (is_built_in_tptp_symbol o fst o fst)  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2530  | 
end  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2531  | 
|
| 
45875
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2532  | 
fun declare_undeclared_syms_in_atp_problem type_enc problem =  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2533  | 
let  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2534  | 
val decls =  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2535  | 
problem  | 
| 
45875
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2536  | 
|> undeclared_syms_in_problem type_enc  | 
| 
45828
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2537  | 
|> sort_wrt (fst o fst)  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2538  | 
|> map (fn (x as (s, _), ty) => Decl (type_decl_prefix ^ s, x, ty ()))  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2539  | 
in (implicit_declsN, decls) :: problem end  | 
| 
 
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
 
blanchet 
parents: 
45780 
diff
changeset
 | 
2540  | 
|
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2541  | 
fun exists_subdtype P =  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2542  | 
let  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2543  | 
fun ex U = P U orelse  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2544  | 
(case U of Datatype.DtType (_, Us) => exists ex Us | _ => false)  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2545  | 
in ex end  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2546  | 
|
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2547  | 
fun is_poly_constr (_, Us) =  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2548  | 
exists (exists_subdtype (fn Datatype.DtTFree _ => true | _ => false)) Us  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2549  | 
|
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2550  | 
fun all_constrs_of_polymorphic_datatypes thy =  | 
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2551  | 
Symtab.fold (snd  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2552  | 
#> #descr  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2553  | 
#> maps (snd #> #3)  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2554  | 
#> (fn cs => exists is_poly_constr cs ? append cs))  | 
| 
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2555  | 
(Datatype.get_all thy) []  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2556  | 
|> List.partition is_poly_constr  | 
| 
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2557  | 
|> pairself (map fst)  | 
| 
45945
 
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
 
blanchet 
parents: 
45939 
diff
changeset
 | 
2558  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2559  | 
val app_op_threshold = 50  | 
| 
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
 | 
2560  | 
|
| 
44397
 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 
blanchet 
parents: 
44396 
diff
changeset
 | 
2561  | 
fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_enc exporter  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2562  | 
lam_trans uncurried_aliases readable_names preproc  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2563  | 
hyp_ts concl_t facts =  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2564  | 
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
 | 
2565  | 
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
 | 
2566  | 
val type_enc = type_enc |> adjust_type_enc format  | 
| 46389 | 2567  | 
(* Forcing explicit applications is expensive for polymorphic encodings,  | 
2568  | 
because it takes only one existential variable ranging over "'a => 'b" to  | 
|
2569  | 
ruin everything. Hence we do it only if there are few facts (which is  | 
|
2570  | 
normally the case for "metis" and the minimizer). *)  | 
|
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2571  | 
val app_op_level =  | 
| 46370 | 2572  | 
if polymorphism_of_type_enc type_enc = Polymorphic andalso  | 
| 
46392
 
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
 
blanchet 
parents: 
46389 
diff
changeset
 | 
2573  | 
length facts >= app_op_threshold then  | 
| 46389 | 2574  | 
Incomplete_Apply  | 
| 46370 | 2575  | 
else  | 
| 46389 | 2576  | 
Sufficient_Apply  | 
| 45514 | 2577  | 
val lam_trans =  | 
| 45520 | 2578  | 
if lam_trans = keep_lamsN andalso  | 
2579  | 
not (is_type_enc_higher_order type_enc) then  | 
|
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45516 
diff
changeset
 | 
2580  | 
        error ("Lambda translation scheme incompatible with first-order \
 | 
| 
44088
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
2581  | 
\encoding.")  | 
| 
 
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
 
blanchet 
parents: 
44003 
diff
changeset
 | 
2582  | 
else  | 
| 45514 | 2583  | 
lam_trans  | 
| 45508 | 2584  | 
val (fact_names, classes, conjs, facts, class_rel_clauses, arity_clauses,  | 
2585  | 
lifted) =  | 
|
| 45514 | 2586  | 
translate_formulas ctxt format prem_kind type_enc lam_trans preproc hyp_ts  | 
2587  | 
concl_t facts  | 
|
| 
46402
 
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
 
blanchet 
parents: 
46400 
diff
changeset
 | 
2588  | 
val sym_tab0 = sym_table_for_facts ctxt type_enc app_op_level conjs facts  | 
| 44399 | 2589  | 
val mono = conjs @ facts |> mononotonicity_info_for_facts ctxt type_enc  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2590  | 
val (polym_constrs, monom_constrs) =  | 
| 
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2591  | 
all_constrs_of_polymorphic_datatypes thy  | 
| 
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2592  | 
|>> map (make_fixed_const (SOME format))  | 
| 
46400
 
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
 
blanchet 
parents: 
46399 
diff
changeset
 | 
2593  | 
fun firstorderize in_helper =  | 
| 
46402
 
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
 
blanchet 
parents: 
46400 
diff
changeset
 | 
2594  | 
firstorderize_fact thy monom_constrs format type_enc sym_tab0  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2595  | 
(uncurried_aliases andalso not in_helper)  | 
| 
46400
 
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
 
blanchet 
parents: 
46399 
diff
changeset
 | 
2596  | 
val (conjs, facts) = (conjs, facts) |> pairself (map (firstorderize false))  | 
| 46389 | 2597  | 
val sym_tab = sym_table_for_facts ctxt type_enc Incomplete_Apply conjs 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
 | 
2598  | 
val helpers =  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
2599  | 
sym_tab |> helper_facts_for_sym_table ctxt format type_enc  | 
| 
46400
 
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
 
blanchet 
parents: 
46399 
diff
changeset
 | 
2600  | 
|> map (firstorderize true)  | 
| 44501 | 2601  | 
val mono_Ts =  | 
| 45937 | 2602  | 
helpers @ conjs @ facts |> monotonic_types_for_facts ctxt mono type_enc  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2603  | 
val class_decl_lines = decl_lines_for_classes type_enc classes  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2604  | 
val (uncurried_alias_eq_tms, uncurried_alias_eq_lines) =  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2605  | 
uncurried_alias_lines_for_sym_table ctxt monom_constrs format  | 
| 
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2606  | 
conj_sym_kind mono type_enc uncurried_aliases sym_tab0 sym_tab  | 
| 42680 | 2607  | 
val sym_decl_lines =  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2608  | 
(conjs, helpers @ facts, uncurried_alias_eq_tms)  | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
2609  | 
|> sym_decl_table_for_facts ctxt format type_enc sym_tab  | 
| 44399 | 2610  | 
|> problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono  | 
| 44501 | 2611  | 
type_enc mono_Ts  | 
| 46406 | 2612  | 
val num_facts = length facts  | 
2613  | 
val fact_lines =  | 
|
2614  | 
map (formula_line_for_fact ctxt polym_constrs format fact_prefix  | 
|
2615  | 
ascii_of (not exporter) (not exporter) mono type_enc  | 
|
2616  | 
(rank_of_fact_num num_facts))  | 
|
2617  | 
(0 upto num_facts - 1 ~~ facts)  | 
|
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
2618  | 
val helper_lines =  | 
| 
42956
 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 
blanchet 
parents: 
42951 
diff
changeset
 | 
2619  | 
0 upto length helpers - 1 ~~ helpers  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2620  | 
|> map (formula_line_for_fact ctxt polym_constrs format helper_prefix I  | 
| 46406 | 2621  | 
false true mono type_enc (K default_rank))  | 
| 
44496
 
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
 
blanchet 
parents: 
44495 
diff
changeset
 | 
2622  | 
|> (if needs_type_tag_idempotence ctxt type_enc then  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
2623  | 
cons (type_tag_idempotence_fact format type_enc)  | 
| 
43159
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
2624  | 
else  | 
| 
 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 
blanchet 
parents: 
43139 
diff
changeset
 | 
2625  | 
I)  | 
| 
42522
 
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
 
blanchet 
parents: 
42521 
diff
changeset
 | 
2626  | 
(* Reordering these might confuse the proof reconstruction code or the SPASS  | 
| 43039 | 2627  | 
FLOTTER hack. *)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2628  | 
val problem =  | 
| 
44595
 
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
 
blanchet 
parents: 
44594 
diff
changeset
 | 
2629  | 
[(explicit_declsN, class_decl_lines @ sym_decl_lines),  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46406 
diff
changeset
 | 
2630  | 
(uncurried_alias_eqsN, uncurried_alias_eq_lines),  | 
| 46406 | 2631  | 
(factsN, fact_lines),  | 
| 
44593
 
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
 
blanchet 
parents: 
44591 
diff
changeset
 | 
2632  | 
(class_relsN,  | 
| 
45301
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
2633  | 
map (formula_line_for_class_rel_clause format type_enc)  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
2634  | 
class_rel_clauses),  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
2635  | 
(aritiesN,  | 
| 
 
866b075aa99b
added sorted DFG output for coming version of SPASS
 
blanchet 
parents: 
45299 
diff
changeset
 | 
2636  | 
map (formula_line_for_arity_clause format type_enc) arity_clauses),  | 
| 
42881
 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 
blanchet 
parents: 
42879 
diff
changeset
 | 
2637  | 
(helpersN, helper_lines),  | 
| 
42962
 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 
blanchet 
parents: 
42956 
diff
changeset
 | 
2638  | 
(conjsN,  | 
| 
45947
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2639  | 
map (formula_line_for_conjecture ctxt polym_constrs format mono  | 
| 
 
7eccf8147f57
treat polymorphic constructors specially in @? encodings
 
blanchet 
parents: 
45946 
diff
changeset
 | 
2640  | 
type_enc) conjs),  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43624 
diff
changeset
 | 
2641  | 
(free_typesN, formula_lines_for_free_types type_enc (facts @ conjs))]  | 
| 
42543
 
f9d402d144d4
declare TFF types so that SNARK can be used with types
 
blanchet 
parents: 
42542 
diff
changeset
 | 
2642  | 
val problem =  | 
| 
42561
 
23ddc4e3d19c
have properly type-instantiated helper facts (combinators and If)
 
blanchet 
parents: 
42560 
diff
changeset
 | 
2643  | 
problem  | 
| 
43092
 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 
blanchet 
parents: 
43086 
diff
changeset
 | 
2644  | 
|> (case format of  | 
| 
 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 
blanchet 
parents: 
43086 
diff
changeset
 | 
2645  | 
CNF => ensure_cnf_problem  | 
| 
 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 
blanchet 
parents: 
43086 
diff
changeset
 | 
2646  | 
| CNF_UEQ => filter_cnf_ueq_problem  | 
| 
44589
 
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
 
blanchet 
parents: 
44587 
diff
changeset
 | 
2647  | 
| FOF => I  | 
| 44754 | 2648  | 
| TFF (_, TPTP_Implicit) => I  | 
2649  | 
| THF (_, TPTP_Implicit, _) => I  | 
|
| 
45875
 
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
 
blanchet 
parents: 
45831 
diff
changeset
 | 
2650  | 
| _ => declare_undeclared_syms_in_atp_problem type_enc)  | 
| 
45939
 
711fec5b4f61
don't try to avoid SPASS keywords; instead, just suffix an underscore to all generated identifiers
 
blanchet 
parents: 
45937 
diff
changeset
 | 
2651  | 
val (problem, pool) = problem |> nice_atp_problem readable_names format  | 
| 44772 | 2652  | 
    fun add_sym_ary (s, {min_ary, ...} : sym_info) =
 | 
| 
44773
 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 
blanchet 
parents: 
44772 
diff
changeset
 | 
2653  | 
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
 | 
2654  | 
in  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2655  | 
(problem,  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2656  | 
case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,  | 
| 
42755
 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 
blanchet 
parents: 
42754 
diff
changeset
 | 
2657  | 
fact_names |> Vector.fromList,  | 
| 45508 | 2658  | 
lifted,  | 
| 44772 | 2659  | 
Symtab.empty |> Symtab.fold add_sym_ary sym_tab)  | 
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2660  | 
end  | 
| 
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2661  | 
|
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2662  | 
(* FUDGE *)  | 
| 
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2663  | 
val conj_weight = 0.0  | 
| 41770 | 2664  | 
val hyp_weight = 0.1  | 
2665  | 
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
 | 
2666  | 
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
 | 
2667  | 
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
 | 
2668  | 
|
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2669  | 
(* Weights are from 0.0 (most important) to 1.0 (least important). *)  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2670  | 
fun atp_problem_relevance_weights problem =  | 
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2671  | 
let  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2672  | 
fun add_term_weights weight (ATerm (s, tms)) =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2673  | 
is_tptp_user_symbol s ? Symtab.default (s, weight)  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2674  | 
#> fold (add_term_weights weight) tms  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2675  | 
| add_term_weights weight (AAbs (_, tm)) = add_term_weights weight tm  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2676  | 
fun add_line_weights weight (Formula (_, _, phi, _, _)) =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2677  | 
formula_fold NONE (K (add_term_weights weight)) phi  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2678  | 
| add_line_weights _ _ = I  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2679  | 
fun add_conjectures_weights [] = I  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2680  | 
| add_conjectures_weights conjs =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2681  | 
let val (hyps, conj) = split_last conjs in  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2682  | 
add_line_weights conj_weight conj  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2683  | 
#> fold (add_line_weights hyp_weight) hyps  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2684  | 
end  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2685  | 
fun add_facts_weights facts =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2686  | 
let  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2687  | 
val num_facts = length facts  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2688  | 
fun weight_of j =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2689  | 
fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2690  | 
/ Real.fromInt num_facts  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2691  | 
in  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2692  | 
map weight_of (0 upto num_facts - 1) ~~ facts  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2693  | 
|> fold (uncurry add_line_weights)  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2694  | 
end  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2695  | 
val get = these o AList.lookup (op =) problem  | 
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2696  | 
in  | 
| 
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
 | 
2697  | 
Symtab.empty  | 
| 
 
6ef61823b63b
make sure E type information constants are given a weight, even if they don't appear anywhere else
 
blanchet 
parents: 
42592 
diff
changeset
 | 
2698  | 
|> add_conjectures_weights (get free_typesN @ get conjsN)  | 
| 
 
6ef61823b63b
make sure E type information constants are given a weight, even if they don't appear anywhere else
 
blanchet 
parents: 
42592 
diff
changeset
 | 
2699  | 
|> add_facts_weights (get factsN)  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2700  | 
|> fold (fold (add_line_weights type_info_default_weight) o get)  | 
| 
42998
 
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
 
blanchet 
parents: 
42994 
diff
changeset
 | 
2701  | 
[explicit_declsN, class_relsN, aritiesN]  | 
| 
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
 | 
2702  | 
|> Symtab.dest  | 
| 
 
6ef61823b63b
make sure E type information constants are given a weight, even if they don't appear anywhere else
 
blanchet 
parents: 
42592 
diff
changeset
 | 
2703  | 
|> sort (prod_ord Real.compare string_ord o pairself swap)  | 
| 
 
6ef61823b63b
make sure E type information constants are given a weight, even if they don't appear anywhere else
 
blanchet 
parents: 
42592 
diff
changeset
 | 
2704  | 
end  | 
| 
41313
 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 
blanchet 
parents: 
41211 
diff
changeset
 | 
2705  | 
|
| 
46443
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2706  | 
fun have_head_rolling (ATerm (s, args)) =  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2707  | 
(* ugly hack: may make innocent victims *)  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2708  | 
if String.isPrefix app_op_name s andalso length args = 2 then  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2709  | 
have_head_rolling (hd args) ||> append (tl args)  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2710  | 
else  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2711  | 
(s, args)  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2712  | 
  | have_head_rolling _ = ("", [])
 | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2713  | 
|
| 46450 | 2714  | 
val max_kbo_weight = 2147483647  | 
2715  | 
||
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2716  | 
fun atp_problem_kbo_weights problem =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2717  | 
let  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2718  | 
fun add_term_deps head (ATerm (s, args)) =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2719  | 
is_tptp_user_symbol s ? perhaps (try (Graph.add_edge_acyclic (s, head)))  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2720  | 
#> fold (add_term_deps head) args  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2721  | 
| add_term_deps head (AAbs (_, tm)) = add_term_deps head tm  | 
| 
46443
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2722  | 
fun add_eq_deps (ATerm (s, [lhs as _, rhs])) =  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2723  | 
if is_tptp_equal s then  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2724  | 
let val (head, rest) = have_head_rolling lhs in  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2725  | 
fold (add_term_deps head) (rhs :: rest)  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2726  | 
end  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2727  | 
else  | 
| 
 
c86276014571
improved KBO weights -- beware of explicit applications
 
blanchet 
parents: 
46442 
diff
changeset
 | 
2728  | 
I  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2729  | 
| add_eq_deps _ = I  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2730  | 
fun add_line_deps _ (Decl (_, s, ty)) =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2731  | 
is_function_type ty ? Graph.default_node (s, ())  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2732  | 
| add_line_deps status (Formula (_, _, phi, _, info)) =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2733  | 
if extract_isabelle_status info = SOME status then  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2734  | 
formula_fold NONE (K add_eq_deps) phi  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2735  | 
else  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2736  | 
I  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2737  | 
val graph =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2738  | 
Graph.empty |> fold (fold (add_line_deps spec_eqN) o snd) problem  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2739  | 
|> fold (fold (add_line_deps simpN) o snd) problem  | 
| 46450 | 2740  | 
fun next_weight w = if w + w <= max_kbo_weight then w + w else w + 1  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2741  | 
fun add_weights _ [] = I  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2742  | 
| add_weights weight syms =  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2743  | 
fold (AList.update (op =) o rpair weight) syms  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2744  | 
#> add_weights (next_weight weight)  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2745  | 
(fold (union (op =) o Graph.immediate_succs graph) syms [])  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2746  | 
in  | 
| 46446 | 2747  | 
[] |> add_weights 1 (Graph.minimals graph) |> sort (int_ord o pairself snd)  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2748  | 
end  | 
| 
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46437 
diff
changeset
 | 
2749  | 
|
| 
38282
 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 
blanchet 
parents:  
diff
changeset
 | 
2750  | 
end;  |