author | wenzelm |
Mon, 27 Feb 2012 16:56:25 +0100 | |
changeset 46711 | f745bcc4a1e5 |
parent 46643 | a88bccd2b567 |
child 46818 | 2a28e66e2e4c |
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] |
46711
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents:
46643
diff
changeset
|
496 |
|> Long_Name.implode |
43085
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 |
|
46642
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
507 |
val robust_const_ary = |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
508 |
let |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
509 |
fun ary (Type (@{type_name fun}, [_, T])) = 1 + ary T |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
510 |
| ary _ = 0 |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
511 |
in ary oo robust_const_type end |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
512 |
|
44594
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
513 |
(* 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
|
514 |
fun robust_const_typargs thy (s, T) = |
45509 | 515 |
if s = app_op_name then |
516 |
let val (T1, T2) = T |> domain_type |> dest_funT in [T1, T2] end |
|
517 |
else if String.isPrefix old_skolem_const_prefix s then |
|
518 |
[] |> 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
|
519 |
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
|
520 |
if String.isPrefix lam_lifted_poly_prefix s then |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
521 |
let val (T1, T2) = T |> dest_funT in [T1, T2] end |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
522 |
else |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
523 |
[] |
45509 | 524 |
else |
525 |
(s, T) |> Sign.const_typargs thy |
|
44594
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
526 |
|
43859 | 527 |
(* Converts an Isabelle/HOL term (with combinators) into an intermediate term. |
528 |
Also accumulates sort infomation. *) |
|
44495 | 529 |
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
|
530 |
let |
44495 | 531 |
val (P', P_atomics_Ts) = iterm_from_term thy format bs P |
532 |
val (Q', Q_atomics_Ts) = iterm_from_term thy format bs Q |
|
43859 | 533 |
in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end |
44495 | 534 |
| iterm_from_term thy format _ (Const (c, T)) = |
535 |
(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
|
536 |
robust_const_typargs thy (c, T)), |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
537 |
atomic_types_of T) |
44495 | 538 |
| iterm_from_term _ _ _ (Free (s, T)) = |
45509 | 539 |
(IConst (`make_fixed_var s, T, []), atomic_types_of T) |
44495 | 540 |
| 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
|
541 |
(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
|
542 |
let |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
543 |
val Ts = T |> strip_type |> swap |> op :: |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
544 |
val s' = new_skolem_const_name s (length Ts) |
44495 | 545 |
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
|
546 |
else |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
547 |
IVar ((make_schematic_var v, s), T), atomic_types_of T) |
44495 | 548 |
| iterm_from_term _ _ bs (Bound j) = |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
549 |
nth bs j |> (fn (_, (name, T)) => (IConst (name, T, []), atomic_types_of T)) |
44495 | 550 |
| iterm_from_term thy format bs (Abs (s, T, t)) = |
43678 | 551 |
let |
552 |
fun vary s = s |> AList.defined (op =) bs s ? vary o Symbol.bump_string |
|
553 |
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
|
554 |
val name = `make_bound_var s |
44495 | 555 |
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
|
556 |
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
|
557 |
|
46340 | 558 |
datatype scope = Global | Local | Assum | Chained |
46393 | 559 |
datatype status = General | Induct | Intro | Elim | Simp | Spec_Eq |
46340 | 560 |
type stature = scope * status |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
561 |
|
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
|
562 |
datatype order = First_Order | Higher_Order |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44493
diff
changeset
|
563 |
datatype polymorphism = Polymorphic | Raw_Monomorphic | Mangled_Monomorphic |
46301 | 564 |
datatype strictness = Strict | Non_Strict |
44811 | 565 |
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
|
566 |
datatype type_level = |
44397
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44396
diff
changeset
|
567 |
All_Types | |
46301 | 568 |
Noninf_Nonmono_Types of strictness * granularity | |
44811 | 569 |
Fin_Nonmono_Types of granularity | |
44397
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44396
diff
changeset
|
570 |
Const_Arg_Types | |
43362 | 571 |
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
|
572 |
|
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
573 |
datatype type_enc = |
44591
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
blanchet
parents:
44589
diff
changeset
|
574 |
Simple_Types of order * polymorphism * type_level | |
44768 | 575 |
Guards of polymorphism * type_level | |
576 |
Tags of polymorphism * type_level |
|
577 |
||
578 |
fun is_type_enc_higher_order (Simple_Types (Higher_Order, _, _)) = true |
|
579 |
| is_type_enc_higher_order _ = false |
|
580 |
||
581 |
fun polymorphism_of_type_enc (Simple_Types (_, poly, _)) = poly |
|
582 |
| polymorphism_of_type_enc (Guards (poly, _)) = poly |
|
583 |
| polymorphism_of_type_enc (Tags (poly, _)) = poly |
|
584 |
||
585 |
fun level_of_type_enc (Simple_Types (_, _, level)) = level |
|
586 |
| level_of_type_enc (Guards (_, level)) = level |
|
587 |
| level_of_type_enc (Tags (_, level)) = level |
|
588 |
||
44811 | 589 |
fun granularity_of_type_level (Noninf_Nonmono_Types (_, grain)) = grain |
590 |
| granularity_of_type_level (Fin_Nonmono_Types grain) = grain |
|
591 |
| granularity_of_type_level _ = All_Vars |
|
44768 | 592 |
|
593 |
fun is_type_level_quasi_sound All_Types = true |
|
594 |
| is_type_level_quasi_sound (Noninf_Nonmono_Types _) = true |
|
595 |
| is_type_level_quasi_sound _ = false |
|
596 |
val is_type_enc_quasi_sound = is_type_level_quasi_sound o level_of_type_enc |
|
597 |
||
598 |
fun is_type_level_fairly_sound (Fin_Nonmono_Types _) = true |
|
599 |
| is_type_level_fairly_sound level = is_type_level_quasi_sound level |
|
600 |
val is_type_enc_fairly_sound = is_type_level_fairly_sound o level_of_type_enc |
|
601 |
||
602 |
fun is_type_level_monotonicity_based (Noninf_Nonmono_Types _) = true |
|
603 |
| is_type_level_monotonicity_based (Fin_Nonmono_Types _) = true |
|
604 |
| 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
|
605 |
|
44785 | 606 |
(* "_query", "_bang", and "_at" are for the ASCII-challenged Metis and |
607 |
Mirabelle. *) |
|
608 |
val queries = ["?", "_query"] |
|
609 |
val bangs = ["!", "_bang"] |
|
610 |
val ats = ["@", "_at"] |
|
611 |
||
42689
e38590764c34
versions of ! and ? for the ASCII-challenged Mirabelle
blanchet
parents:
42688
diff
changeset
|
612 |
fun try_unsuffixes ss s = |
e38590764c34
versions of ! and ? for the ASCII-challenged Mirabelle
blanchet
parents:
42688
diff
changeset
|
613 |
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
|
614 |
|
44785 | 615 |
fun try_nonmono constr suffixes fallback s = |
616 |
case try_unsuffixes suffixes s of |
|
617 |
SOME s => |
|
618 |
(case try_unsuffixes suffixes s of |
|
44811 | 619 |
SOME s => (constr Positively_Naked_Vars, s) |
44785 | 620 |
| NONE => |
621 |
case try_unsuffixes ats s of |
|
44811 | 622 |
SOME s => (constr Ghost_Type_Arg_Vars, s) |
623 |
| NONE => (constr All_Vars, s)) |
|
44785 | 624 |
| NONE => fallback s |
44768 | 625 |
|
46301 | 626 |
fun type_enc_from_string strictness s = |
42722 | 627 |
(case try (unprefix "poly_") s of |
628 |
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
|
629 |
| NONE => |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44493
diff
changeset
|
630 |
case try (unprefix "raw_mono_") s of |
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44493
diff
changeset
|
631 |
SOME s => (SOME Raw_Monomorphic, s) |
42722 | 632 |
| NONE => |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44493
diff
changeset
|
633 |
case try (unprefix "mono_") s of |
42722 | 634 |
SOME s => (SOME Mangled_Monomorphic, s) |
635 |
| NONE => (NONE, s)) |
|
44785 | 636 |
||> (pair All_Types |
637 |
|> try_nonmono Fin_Nonmono_Types bangs |
|
46301 | 638 |
|> try_nonmono (curry Noninf_Nonmono_Types strictness) queries) |
44768 | 639 |
|> (fn (poly, (level, core)) => |
640 |
case (core, (poly, level)) of |
|
46435 | 641 |
("native", (SOME poly, _)) => |
44742 | 642 |
(case (poly, level) of |
643 |
(Polymorphic, All_Types) => |
|
644 |
Simple_Types (First_Order, Polymorphic, All_Types) |
|
645 |
| (Mangled_Monomorphic, _) => |
|
44811 | 646 |
if granularity_of_type_level level = All_Vars then |
44768 | 647 |
Simple_Types (First_Order, Mangled_Monomorphic, level) |
648 |
else |
|
649 |
raise Same.SAME |
|
44742 | 650 |
| _ => raise Same.SAME) |
46435 | 651 |
| ("native_higher", (SOME poly, _)) => |
44591
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
blanchet
parents:
44589
diff
changeset
|
652 |
(case (poly, level) of |
44754 | 653 |
(Polymorphic, All_Types) => |
654 |
Simple_Types (Higher_Order, Polymorphic, All_Types) |
|
655 |
| (_, Noninf_Nonmono_Types _) => raise Same.SAME |
|
44742 | 656 |
| (Mangled_Monomorphic, _) => |
44811 | 657 |
if granularity_of_type_level level = All_Vars then |
44768 | 658 |
Simple_Types (Higher_Order, Mangled_Monomorphic, level) |
659 |
else |
|
660 |
raise Same.SAME |
|
44742 | 661 |
| _ => raise Same.SAME) |
44810 | 662 |
| ("guards", (SOME poly, _)) => |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
663 |
if poly = Mangled_Monomorphic andalso |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
664 |
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
|
665 |
raise Same.SAME |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
666 |
else |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
667 |
Guards (poly, level) |
44810 | 668 |
| ("tags", (SOME poly, _)) => |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
669 |
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
|
670 |
raise Same.SAME |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
671 |
else |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
672 |
Tags (poly, level) |
44768 | 673 |
| ("args", (SOME poly, All_Types (* naja *))) => |
674 |
Guards (poly, Const_Arg_Types) |
|
675 |
| ("erased", (NONE, All_Types (* naja *))) => |
|
676 |
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
|
677 |
| _ => raise Same.SAME) |
44785 | 678 |
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
|
679 |
|
44754 | 680 |
fun adjust_type_enc (THF (TPTP_Monomorphic, _, _)) |
681 |
(Simple_Types (order, _, level)) = |
|
44591
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
blanchet
parents:
44589
diff
changeset
|
682 |
Simple_Types (order, Mangled_Monomorphic, level) |
44754 | 683 |
| adjust_type_enc (THF _) type_enc = type_enc |
684 |
| 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
|
685 |
Simple_Types (First_Order, Mangled_Monomorphic, level) |
45303
bd03b08161ac
added DFG unsorted support (like in the old days)
blanchet
parents:
45301
diff
changeset
|
686 |
| 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
|
687 |
Simple_Types (First_Order, Mangled_Monomorphic, level) |
44754 | 688 |
| 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
|
689 |
Simple_Types (First_Order, poly, level) |
0b107d11f634
extended simple types with polymorphism -- the implementation still needs some work though
blanchet
parents:
44589
diff
changeset
|
690 |
| adjust_type_enc format (Simple_Types (_, poly, level)) = |
44768 | 691 |
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
|
692 |
| 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
|
693 |
(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
|
694 |
| 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
|
695 |
|
45509 | 696 |
fun constify_lifted (t $ u) = constify_lifted t $ constify_lifted u |
697 |
| constify_lifted (Abs (s, T, t)) = Abs (s, T, constify_lifted t) |
|
698 |
| 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
|
699 |
(if String.isPrefix lam_lifted_prefix s then Const else Free) x |
45509 | 700 |
| constify_lifted t = t |
701 |
||
46375
d724066ff3d0
reverted e2b1a86d59fc -- broke Metis's lambda-lifting
blanchet
parents:
46371
diff
changeset
|
702 |
(* 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
|
703 |
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
|
704 |
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
|
705 |
(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
|
706 |
SOME s => |
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents:
46384
diff
changeset
|
707 |
let |
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents:
46384
diff
changeset
|
708 |
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
|
709 |
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
|
710 |
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
|
711 |
| NONE => t) |
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents:
46384
diff
changeset
|
712 |
| open_form _ t = t |
46375
d724066ff3d0
reverted e2b1a86d59fc -- broke Metis's lambda-lifting
blanchet
parents:
46371
diff
changeset
|
713 |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45551
diff
changeset
|
714 |
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
|
715 |
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
|
716 |
#-> Lambda_Lifting.lift_lambdas |
45564 | 717 |
(SOME ((if polymorphism_of_type_enc type_enc = Polymorphic then |
718 |
lam_lifted_poly_prefix |
|
719 |
else |
|
720 |
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
|
721 |
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
|
722 |
#> fst |
46385
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents:
46384
diff
changeset
|
723 |
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
|
724 |
(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
|
725 |
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
|
726 |
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
|
727 |
|
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
blanchet
parents:
44003
diff
changeset
|
728 |
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
|
729 |
intentionalize_def t |
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
blanchet
parents:
44003
diff
changeset
|
730 |
| 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
|
731 |
let |
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
blanchet
parents:
44003
diff
changeset
|
732 |
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
|
733 |
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
|
734 |
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
|
735 |
val n = length args |
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
blanchet
parents:
44003
diff
changeset
|
736 |
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
|
737 |
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
|
738 |
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
|
739 |
| intentionalize_def t = t |
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
blanchet
parents:
44003
diff
changeset
|
740 |
|
40114 | 741 |
type translated_formula = |
43496
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
blanchet
parents:
43495
diff
changeset
|
742 |
{name : string, |
46340 | 743 |
stature : stature, |
43496
92f5a4c78b37
remove historical bloat -- another benefit of merging Metis's and Sledgehammer's translations
blanchet
parents:
43495
diff
changeset
|
744 |
kind : formula_kind, |
43859 | 745 |
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
|
746 |
atomic_types : typ list} |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
747 |
|
46340 | 748 |
fun update_iformula f ({name, stature, kind, iformula, atomic_types} |
43859 | 749 |
: translated_formula) = |
46340 | 750 |
{name = name, stature = stature, kind = kind, iformula = f iformula, |
42562 | 751 |
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
|
752 |
|
43859 | 753 |
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
|
754 |
|
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
|
755 |
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
|
756 |
let val T = get_T x in |
44399 | 757 |
if exists (type_instance ctxt T o get_T) xs then xs |
758 |
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
|
759 |
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
|
760 |
|
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
|
761 |
(* 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
|
762 |
datatype type_arg_policy = |
45937 | 763 |
Explicit_Type_Args of bool (* infer_from_term_args *) | |
44771 | 764 |
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
|
765 |
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
|
766 |
|
45945
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
767 |
fun type_arg_policy monom_constrs type_enc s = |
45315 | 768 |
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
|
769 |
if s = type_tag_name then |
45315 | 770 |
if poly = Mangled_Monomorphic then Mangled_Type_Args |
771 |
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
|
772 |
else case type_enc of |
45315 | 773 |
Simple_Types (_, Polymorphic, _) => Explicit_Type_Args false |
774 |
| 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
|
775 |
| _ => |
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
|
776 |
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
|
777 |
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
|
778 |
(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
|
779 |
No_Type_Args |
45315 | 780 |
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
|
781 |
Mangled_Type_Args |
45945
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
782 |
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
|
783 |
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
|
784 |
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
|
785 |
else |
44811 | 786 |
Explicit_Type_Args |
787 |
(level = All_Types orelse |
|
788 |
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
|
789 |
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
|
790 |
end |
42227
662b50b7126f
if "monomorphize" is enabled, mangle the type information in the names by default
blanchet
parents:
42180
diff
changeset
|
791 |
|
44625 | 792 |
(* Make atoms for sorted type variables. *) |
43263 | 793 |
fun generic_add_sorts_on_type (_, []) = I |
794 |
| generic_add_sorts_on_type ((x, i), s :: ss) = |
|
795 |
generic_add_sorts_on_type ((x, i), ss) |
|
796 |
#> (if s = the_single @{sort HOL.type} then |
|
43093 | 797 |
I |
798 |
else if i = ~1 then |
|
44623 | 799 |
insert (op =) (`make_type_class s, `make_fixed_type_var x) |
43093 | 800 |
else |
44623 | 801 |
insert (op =) (`make_type_class s, |
802 |
(make_schematic_type_var (x, i), x))) |
|
43263 | 803 |
fun add_sorts_on_tfree (TFree (s, S)) = generic_add_sorts_on_type ((s, ~1), S) |
804 |
| add_sorts_on_tfree _ = I |
|
805 |
fun add_sorts_on_tvar (TVar z) = generic_add_sorts_on_type z |
|
806 |
| add_sorts_on_tvar _ = I |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
807 |
|
44625 | 808 |
fun type_class_formula type_enc class arg = |
809 |
AAtom (ATerm (class, arg :: |
|
810 |
(case type_enc of |
|
44754 | 811 |
Simple_Types (First_Order, Polymorphic, _) => |
45301
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45299
diff
changeset
|
812 |
if avoid_first_order_ghost_type_vars then [ATerm (TYPE_name, [arg])] |
44754 | 813 |
else [] |
44625 | 814 |
| _ => []))) |
815 |
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
|
816 |
[] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts |
44625 | 817 |
|> map (fn (class, name) => |
818 |
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
|
819 |
|
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
|
820 |
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
|
821 |
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
|
822 |
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
|
823 |
end |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
824 |
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
|
825 |
| 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
|
826 |
fun mk_aquant _ [] phi = phi |
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
blanchet
parents:
42521
diff
changeset
|
827 |
| 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
|
828 |
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
|
829 |
| 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
|
830 |
|
45315 | 831 |
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
|
832 |
let |
45315 | 833 |
fun add_formula_vars bounds (AQuant (_, xs, phi)) = |
834 |
add_formula_vars (map fst xs @ bounds) phi |
|
835 |
| add_formula_vars bounds (AConn (_, phis)) = |
|
836 |
fold (add_formula_vars bounds) phis |
|
837 |
| add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm |
|
838 |
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
|
839 |
|
45377 | 840 |
fun add_term_vars bounds (ATerm (name as (s, _), tms)) = |
841 |
(if is_tptp_variable s andalso |
|
842 |
not (String.isPrefix tvar_prefix s) andalso |
|
843 |
not (member (op =) bounds name) then |
|
844 |
insert (op =) (name, NONE) |
|
845 |
else |
|
846 |
I) |
|
847 |
#> fold (add_term_vars bounds) tms |
|
848 |
| add_term_vars bounds (AAbs ((name, _), tm)) = |
|
849 |
add_term_vars (name :: bounds) tm |
|
45401 | 850 |
fun close_formula_universally phi = close_universally add_term_vars phi |
45315 | 851 |
|
852 |
fun add_iterm_vars bounds (IApp (tm1, tm2)) = |
|
853 |
fold (add_iterm_vars bounds) [tm1, tm2] |
|
854 |
| add_iterm_vars _ (IConst _) = I |
|
855 |
| add_iterm_vars bounds (IVar (name, T)) = |
|
856 |
not (member (op =) bounds name) ? insert (op =) (name, SOME T) |
|
857 |
| add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm |
|
858 |
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
|
859 |
|
46338
b02ff6b17599
better handling of individual type for DFG format (SPASS)
blanchet
parents:
46320
diff
changeset
|
860 |
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
|
861 |
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
|
862 |
|
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
863 |
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
|
864 |
|
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
|
865 |
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
|
866 |
let |
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
blanchet
parents:
42966
diff
changeset
|
867 |
fun term (Type (s, Ts)) = |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
868 |
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
|
869 |
(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
|
870 |
| (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
|
871 |
| _ => 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
|
872 |
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
|
873 |
`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
|
874 |
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
|
875 |
`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
|
876 |
map term Ts) |
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
blanchet
parents:
42966
diff
changeset
|
877 |
| 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
|
878 |
| 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
|
879 |
in term end |
42562 | 880 |
|
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
|
881 |
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
|
882 |
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
|
883 |
|
42562 | 884 |
(* This shouldn't clash with anything else. *) |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
885 |
val uncurried_alias_sep = "\000" |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
886 |
val mangled_type_sep = "\001" |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
887 |
|
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
888 |
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
|
889 |
|
42562 | 890 |
fun generic_mangled_type_name f (ATerm (name, [])) = f name |
891 |
| 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
|
892 |
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
|
893 |
^ ")" |
43692 | 894 |
| 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
|
895 |
|
44396
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
896 |
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
|
897 |
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
|
898 |
|
46435 | 899 |
fun make_native_type s = |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
900 |
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
|
901 |
s = tptp_individual_type then |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
902 |
s |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
903 |
else |
46435 | 904 |
native_type_prefix ^ ascii_of s |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
905 |
|
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
|
906 |
fun ho_type_from_ho_term type_enc pred_sym ary = |
42963 | 907 |
let |
44593
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
blanchet
parents:
44591
diff
changeset
|
908 |
fun to_mangled_atype ty = |
46435 | 909 |
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
|
910 |
generic_mangled_type_name snd ty), []) |
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
blanchet
parents:
44591
diff
changeset
|
911 |
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
|
912 |
| to_poly_atype _ = raise Fail "unexpected type abstraction" |
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
blanchet
parents:
44591
diff
changeset
|
913 |
val to_atype = |
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
blanchet
parents:
44591
diff
changeset
|
914 |
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
|
915 |
else to_mangled_atype |
42963 | 916 |
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
|
917 |
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
|
918 |
| to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys |
43692 | 919 |
| 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
|
920 |
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
|
921 |
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
|
922 |
| to_ho _ = raise Fail "unexpected type abstraction" |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
923 |
in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end |
42963 | 924 |
|
43677 | 925 |
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
|
926 |
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
|
927 |
o ho_term_from_typ format type_enc |
42963 | 928 |
|
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
929 |
fun aliased_uncurried ary (s, s') = |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
930 |
(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
|
931 |
fun unaliased_uncurried (s, s') = |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
932 |
case space_explode uncurried_alias_sep s of |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
933 |
[_] => (s, s') |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
934 |
| [s1, s2] => (s1, unsuffix s2 s') |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
935 |
| _ => raise Fail "ill-formed explicit application alias" |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
936 |
|
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
937 |
fun raw_mangled_const_name type_name ty_args (s, s') = |
42963 | 938 |
let |
939 |
fun type_suffix f g = |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
940 |
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
|
941 |
ty_args "" |
42963 | 942 |
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
|
943 |
fun mangled_const_name format type_enc = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
944 |
map_filter (ho_term_for_type_arg format type_enc) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
945 |
#> 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
|
946 |
|
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 |
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
|
948 |
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
|
949 |
|
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 |
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
|
951 |
(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
|
952 |
-- 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
|
953 |
[] >> 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
|
954 |
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
|
955 |
(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
|
956 |
|
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 |
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
|
958 |
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
|
959 |
|> 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
|
960 |
(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
|
961 |
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
|
962 |
|> 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
|
963 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
964 |
fun unmangled_const_name s = |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
965 |
(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
|
966 |
fun unmangled_const s = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
967 |
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
|
968 |
(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
|
969 |
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
|
970 |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
971 |
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
|
972 |
let |
43987
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
973 |
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
|
974 |
| 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
|
975 |
_ = |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
976 |
(* 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
|
977 |
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
|
978 |
"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
|
979 |
possible. *) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
980 |
IAbs ((`I "P", p_T), |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
981 |
IApp (IConst (`I ho_quant, T, []), |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
982 |
IAbs ((`I "X", x_T), |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
983 |
IApp (IConst (`I "P", p_T, []), |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
984 |
IConst (`I "X", x_T, []))))) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
985 |
| 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
|
986 |
fun intro top_level args (IApp (tm1, tm2)) = |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
987 |
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
|
988 |
| 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
|
989 |
(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
|
990 |
SOME proxy_base => |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
991 |
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
|
992 |
case (top_level, s) of |
43987
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
993 |
(_, "c_False") => IConst (`I tptp_false, T, []) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
994 |
| (_, "c_True") => IConst (`I tptp_true, T, []) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
995 |
| (false, "c_Not") => IConst (`I tptp_not, T, []) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
996 |
| (false, "c_conj") => IConst (`I tptp_and, T, []) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
997 |
| (false, "c_disj") => IConst (`I tptp_or, T, []) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
998 |
| (false, "c_implies") => IConst (`I tptp_implies, T, []) |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
999 |
| (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
|
1000 |
| (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
|
1001 |
| (false, s) => |
44097 | 1002 |
if is_tptp_equal s andalso length args = 2 then |
1003 |
IConst (`I tptp_equal, T, []) |
|
1004 |
else |
|
44589
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents:
44587
diff
changeset
|
1005 |
(* Use a proxy even for partially applied THF0 equality, |
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents:
44587
diff
changeset
|
1006 |
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
|
1007 |
being able to infer the type of "=". *) |
44097 | 1008 |
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
|
1009 |
| _ => IConst (name, T, []) |
42569
5737947e4c77
make sure that fequal keeps its type arguments for mangled type systems
blanchet
parents:
42568
diff
changeset
|
1010 |
else |
43987
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
1011 |
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
|
1012 |
| 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
|
1013 |
else IConst (name, T, T_args)) |
43987
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
1014 |
| 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
|
1015 |
| intro _ _ tm = tm |
2850b7dc27a4
further worked around LEO-II parser limitation, with eta-expansion
blanchet
parents:
43985
diff
changeset
|
1016 |
in intro true [] end |
42568
7b9801a34836
no needless "fequal" proxies if "explicit_apply" is set + always have readable names
blanchet
parents:
42566
diff
changeset
|
1017 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1018 |
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
|
1019 |
case unprefix_and_unascii const_prefix s of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1020 |
NONE => (name, T_args) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1021 |
| SOME s'' => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1022 |
case type_arg_policy [] type_enc (invert_const s'') of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1023 |
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
|
1024 |
| _ => (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
|
1025 |
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
|
1026 |
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
|
1027 |
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
|
1028 |
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
|
1029 |
| mangle (tm as IConst (_, _, [])) = tm |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1030 |
| mangle (IConst (name, T, T_args)) = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1031 |
mangle_type_args_in_const format type_enc name T_args |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1032 |
|> (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
|
1033 |
| 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
|
1034 |
| 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
|
1035 |
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
|
1036 |
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
|
1037 |
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
|
1038 |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1039 |
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
|
1040 |
| 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
|
1041 |
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
|
1042 |
| chop_fun _ T = ([], T) |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1043 |
|
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
|
1044 |
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
|
1045 |
| 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
|
1046 |
let |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1047 |
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
|
1048 |
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
|
1049 |
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
|
1050 |
in |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1051 |
U_args ~~ T_args |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1052 |
|> map (fn (U, T) => |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1053 |
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
|
1054 |
end |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1055 |
handle TYPE _ => T_args |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1056 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1057 |
fun filter_type_args_in_const _ _ _ _ _ [] = [] |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1058 |
| 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
|
1059 |
case unprefix_and_unascii const_prefix s of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1060 |
NONE => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1061 |
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
|
1062 |
else T_args |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1063 |
| SOME s'' => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1064 |
let |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1065 |
val s'' = invert_const s'' |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1066 |
fun filter_T_args false = T_args |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1067 |
| 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
|
1068 |
in |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1069 |
case type_arg_policy monom_constrs type_enc s'' of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1070 |
Explicit_Type_Args infer_from_term_args => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1071 |
filter_T_args infer_from_term_args |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1072 |
| No_Type_Args => [] |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1073 |
| Mangled_Type_Args => raise Fail "unexpected (un)mangled symbol" |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1074 |
end |
45945
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
1075 |
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
|
1076 |
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
|
1077 |
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
|
1078 |
| filt ary (IConst (name as (s, _), T, T_args)) = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1079 |
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
|
1080 |
|> (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
|
1081 |
| 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
|
1082 |
| 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
|
1083 |
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
|
1084 |
|
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1085 |
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
|
1086 |
let |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1087 |
val thy = Proof_Context.theory_of ctxt |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1088 |
fun do_term bs t atomic_Ts = |
44495 | 1089 |
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
|
1090 |
|>> (introduce_proxies_in_iterm type_enc |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1091 |
#> 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
|
1092 |
||> 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
|
1093 |
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
|
1094 |
let |
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
blanchet
parents:
44402
diff
changeset
|
1095 |
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
|
1096 |
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
|
1097 |
val name = |
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
blanchet
parents:
44402
diff
changeset
|
1098 |
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
|
1099 |
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
|
1100 |
| 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
|
1101 |
| 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
|
1102 |
in |
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
blanchet
parents:
44402
diff
changeset
|
1103 |
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
|
1104 |
#>> mk_aquant q [(name, SOME T)] |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1105 |
##> union (op =) (atomic_types_of T) |
38518
54727b44e277
handle bound name conflicts gracefully in FOF translation
blanchet
parents:
38496
diff
changeset
|
1106 |
end |
44403
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
blanchet
parents:
44402
diff
changeset
|
1107 |
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
|
1108 |
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
|
1109 |
and do_formula bs pos t = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1110 |
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
|
1111 |
@{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
|
1112 |
| @{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
|
1113 |
| 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
|
1114 |
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
|
1115 |
| (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
|
1116 |
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
|
1117 |
| 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
|
1118 |
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
|
1119 |
| (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
|
1120 |
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
|
1121 |
| @{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
|
1122 |
| @{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
|
1123 |
| @{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
|
1124 |
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
|
1125 |
| 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
|
1126 |
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
|
1127 |
| _ => do_term bs t |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1128 |
in do_formula [] end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1129 |
|
46093 | 1130 |
fun presimplify_term ctxt t = |
1131 |
t |> exists_Const (member (op =) Meson.presimplified_consts o fst) t |
|
1132 |
? (Skip_Proof.make_thm (Proof_Context.theory_of ctxt) |
|
1133 |
#> Meson.presimplify |
|
1134 |
#> prop_of) |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1135 |
|
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
|
1136 |
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
|
1137 |
fun conceal_bounds Ts t = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1138 |
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
|
1139 |
(0 upto length Ts - 1 ~~ Ts), t) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1140 |
fun reveal_bounds Ts = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1141 |
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
|
1142 |
(0 upto length Ts - 1 ~~ Ts)) |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1143 |
|
43265 | 1144 |
fun is_fun_equality (@{const_name HOL.eq}, |
1145 |
Type (_, [Type (@{type_name fun}, _), _])) = true |
|
1146 |
| is_fun_equality _ = false |
|
1147 |
||
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
|
1148 |
fun extensionalize_term ctxt t = |
43265 | 1149 |
if exists_Const is_fun_equality t then |
1150 |
let val thy = Proof_Context.theory_of ctxt in |
|
1151 |
t |> cterm_of thy |> Meson.extensionalize_conv ctxt |
|
1152 |
|> prop_of |> Logic.dest_equals |> snd |
|
1153 |
end |
|
1154 |
else |
|
1155 |
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
|
1156 |
|
43862 | 1157 |
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
|
1158 |
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
|
1159 |
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
|
1160 |
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
|
1161 |
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
|
1162 |
let |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1163 |
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
|
1164 |
case t of |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1165 |
@{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
|
1166 |
| (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
|
1167 |
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
|
1168 |
| (t0 as Const (@{const_name All}, _)) $ t1 => |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1169 |
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
|
1170 |
| (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
|
1171 |
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
|
1172 |
| (t0 as Const (@{const_name Ex}, _)) $ t1 => |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1173 |
trans Ts (t0 $ eta_expand Ts t1 1) |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1174 |
| (t0 as @{const HOL.conj}) $ t1 $ t2 => |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1175 |
t0 $ trans Ts t1 $ trans Ts t2 |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1176 |
| (t0 as @{const HOL.disj}) $ t1 $ t2 => |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1177 |
t0 $ trans Ts t1 $ trans Ts t2 |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1178 |
| (t0 as @{const HOL.implies}) $ t1 $ t2 => |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1179 |
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
|
1180 |
| (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
|
1181 |
$ t1 $ t2 => |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1182 |
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
|
1183 |
| _ => |
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
|
1184 |
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
|
1185 |
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
|
1186 |
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
|
1187 |
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
|
1188 |
end |
43856
d636b053d4ff
move more lambda-handling logic to Sledgehammer, from ATP module, for formal dependency reasons
blanchet
parents:
43830
diff
changeset
|
1189 |
|
43997 | 1190 |
fun do_cheaply_conceal_lambdas Ts (t1 $ t2) = |
1191 |
do_cheaply_conceal_lambdas Ts t1 |
|
1192 |
$ do_cheaply_conceal_lambdas Ts t2 |
|
1193 |
| 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
|
1194 |
Const (lam_lifted_poly_prefix ^ serial_string (), |
45509 | 1195 |
T --> fastype_of1 (T :: Ts, t)) |
43997 | 1196 |
| 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
|
1197 |
|
43862 | 1198 |
fun do_introduce_combinators ctxt Ts t = |
42361 | 1199 |
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
|
1200 |
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
|
1201 |
|> 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
|
1202 |
|> 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
|
1203 |
|> 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
|
1204 |
|> 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
|
1205 |
end |
43862 | 1206 |
(* A type variable of sort "{}" will make abstraction fail. *) |
43997 | 1207 |
handle THM _ => t |> do_cheaply_conceal_lambdas Ts |
43862 | 1208 |
val introduce_combinators = simple_translate_lambdas do_introduce_combinators |
1209 |
||
45514 | 1210 |
fun preprocess_abstractions_in_terms trans_lams facts = |
43862 | 1211 |
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
|
1212 |
val (facts, lambda_ts) = |
45514 | 1213 |
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
|
1214 |
|>> 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
|
1215 |
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
|
1216 |
map2 (fn t => fn j => |
46393 | 1217 |
((lam_fact_prefix ^ Int.toString j, (Global, Spec_Eq)), |
1218 |
(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
|
1219 |
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
|
1220 |
in (facts, lam_facts) end |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1221 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1222 |
(* 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
|
1223 |
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
|
1224 |
fun freeze_term t = |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1225 |
let |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1226 |
fun freeze (t $ u) = freeze t $ freeze u |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1227 |
| freeze (Abs (s, T, t)) = Abs (s, T, freeze t) |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1228 |
| 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
|
1229 |
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
|
1230 |
| freeze t = t |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1231 |
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
|
1232 |
|
46093 | 1233 |
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
|
1234 |
(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
|
1235 |
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
|
1236 |
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
|
1237 |
|> 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
|
1238 |
|> 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
|
1239 |
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
|
1240 |
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
|
1241 |
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
|
1242 |
|> extensionalize_term ctxt |
46093 | 1243 |
|> 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
|
1244 |
|> 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
|
1245 |
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
|
1246 |
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
|
1247 |
|> pair role |
43096 | 1248 |
|
46340 | 1249 |
fun make_formula ctxt format type_enc eq_as_iff name stature kind t = |
43096 | 1250 |
let |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1251 |
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
|
1252 |
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
|
1253 |
(SOME (kind <> Conjecture)) t [] |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1254 |
|>> close_iformula_universally |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1255 |
in |
46340 | 1256 |
{name = name, stature = stature, kind = kind, iformula = iformula, |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1257 |
atomic_types = atomic_Ts} |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1258 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1259 |
|
46340 | 1260 |
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
|
1261 |
case t |> make_formula ctxt format type_enc (eq_as_iff andalso format <> CNF) |
46340 | 1262 |
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
|
1263 |
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
|
1264 |
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
|
1265 |
| formula => SOME formula |
42561
23ddc4e3d19c
have properly type-instantiated helper facts (combinators and If)
blanchet
parents:
42560
diff
changeset
|
1266 |
|
46385
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents:
46384
diff
changeset
|
1267 |
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
|
1268 |
| s_not_trueprop t = |
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents:
46384
diff
changeset
|
1269 |
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
|
1270 |
else @{prop False} (* "t" is too meta *) |
44460 | 1271 |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1272 |
fun make_conjecture ctxt format type_enc = |
46340 | 1273 |
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
|
1274 |
t |> kind = Conjecture ? s_not_trueprop |
46340 | 1275 |
|> make_formula ctxt format type_enc (format <> CNF) name stature |
1276 |
kind) |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1277 |
|
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
|
1278 |
(** 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
|
1279 |
|
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1280 |
fun tvar_footprint thy s ary = |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1281 |
(case unprefix_and_unascii const_prefix s of |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1282 |
SOME s => |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1283 |
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
|
1284 |
|> map (fn T => Term.add_tvarsT T [] |> map fst) |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1285 |
| NONE => []) |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1286 |
handle TYPE _ => [] |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1287 |
|
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1288 |
fun ghost_type_args thy s ary = |
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1289 |
if is_tptp_equal s then |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1290 |
0 upto ary - 1 |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1291 |
else |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1292 |
let |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1293 |
val footprint = tvar_footprint thy s ary |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1294 |
val eq = (s = @{const_name HOL.eq}) |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1295 |
fun ghosts _ [] = [] |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1296 |
| ghosts seen ((i, tvars) :: args) = |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1297 |
ghosts (union (op =) seen tvars) args |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1298 |
|> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars) |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1299 |
? cons i |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1300 |
in |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1301 |
if forall null footprint then |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1302 |
[] |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1303 |
else |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1304 |
0 upto length footprint - 1 ~~ footprint |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1305 |
|> 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
|
1306 |
|> ghosts [] |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1307 |
end |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1308 |
|
44399 | 1309 |
type monotonicity_info = |
1310 |
{maybe_finite_Ts : typ list, |
|
1311 |
surely_finite_Ts : typ list, |
|
1312 |
maybe_infinite_Ts : typ list, |
|
1313 |
surely_infinite_Ts : typ list, |
|
1314 |
maybe_nonmono_Ts : typ list} |
|
1315 |
||
44397
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44396
diff
changeset
|
1316 |
(* These types witness that the type classes they belong to allow infinite |
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44396
diff
changeset
|
1317 |
models and hence that any types with these type classes is monotonic. *) |
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44396
diff
changeset
|
1318 |
val known_infinite_types = |
44634
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
blanchet
parents:
44626
diff
changeset
|
1319 |
[@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}] |
44397
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44396
diff
changeset
|
1320 |
|
46301 | 1321 |
fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T = |
1322 |
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
|
1323 |
|
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
|
1324 |
(* 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
|
1325 |
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
|
1326 |
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
|
1327 |
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
|
1328 |
|
44399 | 1329 |
fun should_encode_type _ (_ : monotonicity_info) All_Types _ = true |
1330 |
| should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts, |
|
1331 |
maybe_nonmono_Ts, ...} |
|
46301 | 1332 |
(Noninf_Nonmono_Types (strictness, grain)) T = |
44811 | 1333 |
grain = Ghost_Type_Arg_Vars orelse |
1334 |
(exists (type_intersect ctxt T) maybe_nonmono_Ts andalso |
|
1335 |
not (exists (type_instance ctxt T) surely_infinite_Ts orelse |
|
44859
237ba63d6041
fixed definition of type intersection (soundness bug)
blanchet
parents:
44853
diff
changeset
|
1336 |
(not (member (type_equiv ctxt) maybe_finite_Ts T) andalso |
46301 | 1337 |
is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts |
44811 | 1338 |
T))) |
44399 | 1339 |
| should_encode_type ctxt {surely_finite_Ts, maybe_infinite_Ts, |
1340 |
maybe_nonmono_Ts, ...} |
|
44811 | 1341 |
(Fin_Nonmono_Types grain) T = |
1342 |
grain = Ghost_Type_Arg_Vars orelse |
|
1343 |
(exists (type_intersect ctxt T) maybe_nonmono_Ts andalso |
|
1344 |
(exists (type_generalization ctxt T) surely_finite_Ts orelse |
|
44859
237ba63d6041
fixed definition of type intersection (soundness bug)
blanchet
parents:
44853
diff
changeset
|
1345 |
(not (member (type_equiv ctxt) maybe_infinite_Ts T) andalso |
44811 | 1346 |
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
|
1347 |
| 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
|
1348 |
|
44768 | 1349 |
fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T = |
44811 | 1350 |
should_guard_var () andalso should_encode_type ctxt mono level T |
44399 | 1351 |
| 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
|
1352 |
|
44403
15160cdc4688
precisely distinguish between universal and existential quantifiers, instead of assuming the worst (universal), for monotonicity analysis
blanchet
parents:
44402
diff
changeset
|
1353 |
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
|
1354 |
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
|
1355 |
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
|
1356 |
| 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
|
1357 |
| is_maybe_universal_var _ = false |
42836 | 1358 |
|
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1359 |
datatype site = |
43361
e37b54d429f5
revived the lightweight "poly_tags_{query,bang}" type encodings by fixing their soundness bug
blanchet
parents:
43324
diff
changeset
|
1360 |
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
|
1361 |
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
|
1362 |
Elsewhere |
42829
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
1363 |
|
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1364 |
fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1365 |
| 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
|
1366 |
if granularity_of_type_level level = All_Vars then |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1367 |
should_encode_type ctxt mono level T |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1368 |
else |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1369 |
(case (site, is_maybe_universal_var u) of |
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1370 |
(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
|
1371 |
| _ => false) |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1372 |
| 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
|
1373 |
|
44594
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
1374 |
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
|
1375 |
let |
44399 | 1376 |
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
|
1377 |
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
|
1378 |
| fuse ary (Type (@{type_name fun}, [T1, T2])) = |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
1379 |
fuse 0 T1 --> fuse (ary - 1) T2 |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
1380 |
| fuse _ _ = raise Fail "expected function type" |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
1381 |
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
|
1382 |
|
44450
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
blanchet
parents:
44418
diff
changeset
|
1383 |
(** predicators and application operators **) |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
1384 |
|
42574 | 1385 |
type sym_info = |
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1386 |
{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
|
1387 |
in_conj : bool} |
42563 | 1388 |
|
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1389 |
fun default_sym_tab_entries type_enc = |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1390 |
(make_fixed_const NONE @{const_name undefined}, |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1391 |
{pred_sym = false, min_ary = 0, max_ary = 0, 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 |
([tptp_false, tptp_true] |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1394 |
|> 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
|
1395 |
in_conj = false})) @ |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1396 |
([tptp_equal, tptp_old_equal] |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1397 |
|> 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
|
1398 |
in_conj = false})) |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1399 |
|> not (is_type_enc_higher_order type_enc) |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1400 |
? cons (prefixed_predicator_name, |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1401 |
{pred_sym = true, min_ary = 1, max_ary = 1, types = [], |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1402 |
in_conj = false}) |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1403 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1404 |
datatype app_op_level = Incomplete_Apply | Sufficient_Apply | Full_Apply |
46389 | 1405 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1406 |
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
|
1407 |
let |
46642
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1408 |
val thy = Proof_Context.theory_of ctxt |
44772 | 1409 |
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
|
1410 |
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
|
1411 |
fun iter ary T = |
44399 | 1412 |
if ary = max_ary orelse type_instance ctxt var_T T orelse |
1413 |
type_instance ctxt T var_T then |
|
43210
7384b771805d
made "explicit_apply"'s smart mode (more) complete
blanchet
parents:
43207
diff
changeset
|
1414 |
ary |
7384b771805d
made "explicit_apply"'s smart mode (more) complete
blanchet
parents:
43207
diff
changeset
|
1415 |
else |
7384b771805d
made "explicit_apply"'s smart mode (more) complete
blanchet
parents:
43207
diff
changeset
|
1416 |
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
|
1417 |
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
|
1418 |
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
|
1419 |
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
|
1420 |
(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
|
1421 |
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
|
1422 |
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
|
1423 |
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
|
1424 |
{pred_sym = pred_sym andalso not bool_vars', |
44772 | 1425 |
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
|
1426 |
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
|
1427 |
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
|
1428 |
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
|
1429 |
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
|
1430 |
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
|
1431 |
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
|
1432 |
accum |
43167
839f599bc7ed
ensured that the logic for "explicit_apply = smart" also works on CNF (i.e. new Metis)
blanchet
parents:
43159
diff
changeset
|
1433 |
else |
44772 | 1434 |
((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
|
1435 |
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
|
1436 |
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
|
1437 |
accum |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1438 |
fun add_iterm_syms conj_fact top_level tm |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1439 |
(accum as ((bool_vars, fun_var_Ts), sym_tab)) = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1440 |
let val (head, args) = strip_iterm_comb tm in |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1441 |
(case head of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1442 |
IConst ((s, _), T, _) => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1443 |
if String.isPrefix bound_var_prefix s orelse |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1444 |
String.isPrefix all_bound_var_prefix s then |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1445 |
add_universal_var T accum |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1446 |
else if String.isPrefix exist_bound_var_prefix s then |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1447 |
accum |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1448 |
else |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1449 |
let val ary = length args in |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1450 |
((bool_vars, fun_var_Ts), |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1451 |
case Symtab.lookup sym_tab s of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1452 |
SOME {pred_sym, min_ary, max_ary, types, in_conj} => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1453 |
let |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1454 |
val pred_sym = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1455 |
pred_sym andalso top_level andalso not bool_vars |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1456 |
val types' = types |> insert_type ctxt I T |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1457 |
val in_conj = in_conj orelse conj_fact |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1458 |
val min_ary = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1459 |
if app_op_level = Sufficient_Apply andalso |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1460 |
not (pointer_eq (types', types)) then |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1461 |
fold (consider_var_ary T) fun_var_Ts min_ary |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1462 |
else |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1463 |
min_ary |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1464 |
in |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1465 |
Symtab.update (s, {pred_sym = pred_sym, |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1466 |
min_ary = Int.min (ary, min_ary), |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1467 |
max_ary = Int.max (ary, max_ary), |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1468 |
types = types', in_conj = in_conj}) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1469 |
sym_tab |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1470 |
end |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1471 |
| NONE => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1472 |
let |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1473 |
val pred_sym = top_level andalso not bool_vars |
46639
d0ef1d1562d7
fixed arity bug with "If" helpers for "If" that returns a function
blanchet
parents:
46450
diff
changeset
|
1474 |
val ary = |
d0ef1d1562d7
fixed arity bug with "If" helpers for "If" that returns a function
blanchet
parents:
46450
diff
changeset
|
1475 |
case unprefix_and_unascii const_prefix s of |
d0ef1d1562d7
fixed arity bug with "If" helpers for "If" that returns a function
blanchet
parents:
46450
diff
changeset
|
1476 |
SOME s => |
46642
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1477 |
(if String.isSubstring uncurried_alias_sep s then |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1478 |
ary |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1479 |
else case try (robust_const_ary thy |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1480 |
o invert_const o hd |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1481 |
o unmangled_const_name) s of |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1482 |
SOME ary0 => Int.min (ary0, ary) |
37a055f37224
general solution to the arity bug that occasionally plagues Sledgehammer -- short story, lots of things go kaputt when a polymorphic symbol's arity in the translation is higher than the arity of the fully polymorphic HOL constant
blanchet
parents:
46639
diff
changeset
|
1483 |
| NONE => ary) |
46639
d0ef1d1562d7
fixed arity bug with "If" helpers for "If" that returns a function
blanchet
parents:
46450
diff
changeset
|
1484 |
| NONE => ary |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1485 |
val min_ary = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1486 |
case app_op_level of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1487 |
Incomplete_Apply => ary |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1488 |
| Sufficient_Apply => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1489 |
fold (consider_var_ary T) fun_var_Ts ary |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1490 |
| Full_Apply => 0 |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1491 |
in |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1492 |
Symtab.update_new (s, |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1493 |
{pred_sym = pred_sym, min_ary = min_ary, |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1494 |
max_ary = ary, types = [T], in_conj = conj_fact}) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1495 |
sym_tab |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1496 |
end) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1497 |
end |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1498 |
| IVar (_, T) => add_universal_var T accum |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1499 |
| IAbs ((_, T), tm) => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1500 |
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
|
1501 |
| _ => accum) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1502 |
|> fold (add_iterm_syms conj_fact false) args |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1503 |
end |
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1504 |
fun add_fact_syms conj_fact = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1505 |
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
|
1506 |
in |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1507 |
((false, []), Symtab.empty) |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1508 |
|> fold (add_fact_syms true) conjs |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1509 |
|> fold (add_fact_syms false) facts |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1510 |
|> snd |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1511 |
|> fold Symtab.update (default_sym_tab_entries type_enc) |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
1512 |
end |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1513 |
|
44772 | 1514 |
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
|
1515 |
case Symtab.lookup sym_tab s of |
42574 | 1516 |
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
|
1517 |
| NONE => |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1518 |
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
|
1519 |
SOME s => |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1520 |
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
|
1521 |
if s = predicator_name then 1 |
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
blanchet
parents:
42963
diff
changeset
|
1522 |
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
|
1523 |
else if s = type_guard_name then 1 |
42557
ae0deb39a254
fixed min-arity computation when "explicit_apply" is specified
blanchet
parents:
42556
diff
changeset
|
1524 |
else 0 |
42547
b5eec0c99528
fixed arity of special constants if "explicit_apply" is set
blanchet
parents:
42546
diff
changeset
|
1525 |
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
|
1526 |
| NONE => 0 |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1527 |
|
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1528 |
(* 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
|
1529 |
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
|
1530 |
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
|
1531 |
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
|
1532 |
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
|
1533 |
case Symtab.lookup sym_tab s of |
42574 | 1534 |
SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) => |
1535 |
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
|
1536 |
| NONE => false |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
1537 |
|
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1538 |
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
|
1539 |
val predicator_combconst = |
44495 | 1540 |
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
|
1541 |
|
43859 | 1542 |
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
|
1543 |
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
|
1544 |
|
46410
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
blanchet
parents:
46409
diff
changeset
|
1545 |
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
|
1546 |
let |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1547 |
val head_T = ityp_of head |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1548 |
val (arg_T, res_T) = dest_funT head_T |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1549 |
val app = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1550 |
IConst (app_op, head_T --> head_T, [arg_T, res_T]) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1551 |
|> mangle_type_args_in_iterm format type_enc |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1552 |
in list_app app [head, arg] end |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1553 |
|
46400
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
blanchet
parents:
46399
diff
changeset
|
1554 |
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
|
1555 |
uncurried_aliases = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1556 |
let |
46410
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
blanchet
parents:
46409
diff
changeset
|
1557 |
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
|
1558 |
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
|
1559 |
fun introduce_app_ops tm = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1560 |
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
|
1561 |
case head of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1562 |
IConst (name as (s, _), T, T_args) => |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
1563 |
if uncurried_aliases andalso String.isPrefix const_prefix s then |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1564 |
let |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1565 |
val ary = length args |
46639
d0ef1d1562d7
fixed arity bug with "If" helpers for "If" that returns a function
blanchet
parents:
46450
diff
changeset
|
1566 |
val name = |
d0ef1d1562d7
fixed arity bug with "If" helpers for "If" that returns a function
blanchet
parents:
46450
diff
changeset
|
1567 |
name |> ary > min_ary_of sym_tab s ? aliased_uncurried ary |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1568 |
in list_app (IConst (name, T, T_args)) args end |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1569 |
else |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1570 |
args |> chop (min_ary_of sym_tab s) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1571 |
|>> list_app head |-> list_app_ops |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1572 |
| _ => list_app_ops head args |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1573 |
end |
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1574 |
fun introduce_predicators tm = |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1575 |
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
|
1576 |
(IConst ((s, _), _, _), _) => |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1577 |
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
|
1578 |
| _ => predicator tm |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1579 |
val do_iterm = |
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1580 |
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
|
1581 |
? (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
|
1582 |
#> 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
|
1583 |
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
|
1584 |
|
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
|
1585 |
(** 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
|
1586 |
|
44450
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
blanchet
parents:
44418
diff
changeset
|
1587 |
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
|
1588 |
val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast} |
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
blanchet
parents:
44418
diff
changeset
|
1589 |
|
43194 | 1590 |
(* 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
|
1591 |
val helper_table = |
43194 | 1592 |
[(("COMBI", false), @{thms Meson.COMBI_def}), |
1593 |
(("COMBK", false), @{thms Meson.COMBK_def}), |
|
1594 |
(("COMBB", false), @{thms Meson.COMBB_def}), |
|
1595 |
(("COMBC", false), @{thms Meson.COMBC_def}), |
|
1596 |
(("COMBS", false), @{thms Meson.COMBS_def}), |
|
44450
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
blanchet
parents:
44418
diff
changeset
|
1597 |
((predicator_name, false), [not_ffalse, ftrue]), |
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
blanchet
parents:
44418
diff
changeset
|
1598 |
(("fFalse", false), [not_ffalse]), |
43194 | 1599 |
(("fFalse", true), @{thms True_or_False}), |
44450
d848dd7b21f4
fixed "hBOOL" of existential variables, and generate more helpers
blanchet
parents:
44418
diff
changeset
|
1600 |
(("fTrue", false), [ftrue]), |
43194 | 1601 |
(("fTrue", true), @{thms True_or_False}), |
1602 |
(("fNot", false), |
|
1603 |
@{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1] |
|
1604 |
fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}), |
|
1605 |
(("fconj", false), |
|
1606 |
@{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q" |
|
1607 |
by (unfold fconj_def) fast+}), |
|
1608 |
(("fdisj", false), |
|
1609 |
@{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q" |
|
1610 |
by (unfold fdisj_def) fast+}), |
|
1611 |
(("fimplies", false), |
|
43210
7384b771805d
made "explicit_apply"'s smart mode (more) complete
blanchet
parents:
43207
diff
changeset
|
1612 |
@{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q" |
43194 | 1613 |
by (unfold fimplies_def) fast+}), |
43678 | 1614 |
(("fequal", true), |
1615 |
(* This is a lie: Higher-order equality doesn't need a sound type encoding. |
|
1616 |
However, this is done so for backward compatibility: Including the |
|
1617 |
equality helpers by default in Metis breaks a few existing proofs. *) |
|
1618 |
@{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1] |
|
1619 |
fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}), |
|
44003 | 1620 |
(* Partial characterization of "fAll" and "fEx". A complete characterization |
1621 |
would require the axiom of choice for replay with Metis. *) |
|
1622 |
(("fAll", false), [@{lemma "~ fAll P | P x" by (auto simp: fAll_def)}]), |
|
1623 |
(("fEx", false), [@{lemma "~ P x | fEx P" by (auto simp: fEx_def)}]), |
|
43194 | 1624 |
(("If", true), @{thms if_True if_False True_or_False})] |
1625 |
|> map (apsnd (map zero_var_indexes)) |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1626 |
|
45364
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
blanchet
parents:
45316
diff
changeset
|
1627 |
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
|
1628 |
| atype_of_type_vars _ = NONE |
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
blanchet
parents:
45316
diff
changeset
|
1629 |
|
45920 | 1630 |
fun bound_tvars type_enc sorts Ts = |
1631 |
(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
|
1632 |
#> mk_aquant AForall |
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1633 |
(map_filter (fn TVar (x as (s, _), _) => |
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1634 |
SOME ((make_schematic_type_var x, s), |
45364
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
blanchet
parents:
45316
diff
changeset
|
1635 |
atype_of_type_vars type_enc) |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
1636 |
| _ => NONE) Ts) |
44404
3111af540141
tuning, plus started implementing tag equation generation for existential variables
blanchet
parents:
44403
diff
changeset
|
1637 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1638 |
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
|
1639 |
(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
|
1640 |
else AAtom (ATerm (`I tptp_equal, [tm1, tm2]))) |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1641 |
|> mk_aquant AForall bounds |
45377 | 1642 |
|> close_formula_universally |
45920 | 1643 |
|> bound_tvars type_enc true atomic_Ts |
44404
3111af540141
tuning, plus started implementing tag equation generation for existential variables
blanchet
parents:
44403
diff
changeset
|
1644 |
|
46406 | 1645 |
val helper_rank = default_rank |
1646 |
val min_rank = 9 * helper_rank div 10 |
|
1647 |
val max_rank = 4 * min_rank |
|
1648 |
||
1649 |
fun rank_of_fact_num n j = min_rank + (max_rank - min_rank) * j div n |
|
1650 |
||
44495 | 1651 |
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
|
1652 |
|
45301
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45299
diff
changeset
|
1653 |
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
|
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 |
fun var s = ATerm (`I s, []) |
44408
30ea62ab4f16
made reconstruction of type tag equalities "\?x = \?x" reliable
blanchet
parents:
44406
diff
changeset
|
1656 |
fun tag tm = ATerm (type_tag, [var "A", tm]) |
30ea62ab4f16
made reconstruction of type tag equalities "\?x = \?x" reliable
blanchet
parents:
44406
diff
changeset
|
1657 |
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
|
1658 |
in |
43159
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents:
43139
diff
changeset
|
1659 |
Formula (type_tag_idempotence_helper_name, Axiom, |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1660 |
eq_formula type_enc [] [] false (tag tagged_var) tagged_var, |
46406 | 1661 |
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
|
1662 |
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
|
1663 |
|
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
1664 |
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
|
1665 |
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
|
1666 |
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
|
1667 |
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
|
1668 |
|
43858
be41d12de6fa
simplify code -- there are no lambdas in helpers anyway
blanchet
parents:
43857
diff
changeset
|
1669 |
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
|
1670 |
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
|
1671 |
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
|
1672 |
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
|
1673 |
val thy = Proof_Context.theory_of ctxt |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1674 |
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
|
1675 |
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
|
1676 |
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
|
1677 |
(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
|
1678 |
(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
|
1679 |
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
|
1680 |
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
|
1681 |
let val t = prop_of th in |
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
blanchet
parents:
43626
diff
changeset
|
1682 |
if should_specialize_helper type_enc t then |
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
blanchet
parents:
43626
diff
changeset
|
1683 |
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
|
1684 |
types |
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
blanchet
parents:
43626
diff
changeset
|
1685 |
else |
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
blanchet
parents:
43626
diff
changeset
|
1686 |
[t] |
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
blanchet
parents:
43626
diff
changeset
|
1687 |
end |
46339
6268c5b3efdc
generate left-to-right rewrite tag for combinator helpers for SPASS 3.8
blanchet
parents:
46338
diff
changeset
|
1688 |
|> tag_list 1 |
46393 | 1689 |
|> map (fn (k, t) => |
1690 |
((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
|
1691 |
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
|
1692 |
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
|
1693 |
in |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1694 |
helper_table |
43194 | 1695 |
|> 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
|
1696 |
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
|
1697 |
(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
|
1698 |
[] |
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
|
1699 |
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
|
1700 |
ths ~~ (1 upto length ths) |
43628
996b2022ff78
further repair "mangled_tags", now that tags are also mangled
blanchet
parents:
43626
diff
changeset
|
1701 |
|> 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
|
1702 |
|> 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
|
1703 |
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
|
1704 |
| NONE => [] |
43858
be41d12de6fa
simplify code -- there are no lambdas in helpers anyway
blanchet
parents:
43857
diff
changeset
|
1705 |
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
|
1706 |
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
|
1707 |
[] |
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
|
1708 |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1709 |
(***************************************************************) |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1710 |
(* 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
|
1711 |
(***************************************************************) |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1712 |
|
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1713 |
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
|
1714 |
|
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1715 |
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
|
1716 |
|
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1717 |
(* 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
|
1718 |
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
|
1719 |
|
43093 | 1720 |
fun classes_of_terms get_Ts = |
43121 | 1721 |
map (map snd o get_Ts) |
43093 | 1722 |
#> List.foldl add_classes Symtab.empty |
1723 |
#> delete_type #> Symtab.keys |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1724 |
|
44121 | 1725 |
val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees |
1726 |
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
|
1727 |
|
43622 | 1728 |
fun fold_type_constrs f (Type (s, Ts)) x = |
1729 |
fold (fold_type_constrs f) Ts (f (s, x)) |
|
43189 | 1730 |
| fold_type_constrs _ _ x = x |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1731 |
|
43907
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
blanchet
parents:
43906
diff
changeset
|
1732 |
(* 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
|
1733 |
needed. *) |
43189 | 1734 |
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
|
1735 |
let |
43188
0c36ae874fcc
fixed detection of Skolem constants in type construction detection code
blanchet
parents:
43185
diff
changeset
|
1736 |
fun add (Const (@{const_name Meson.skolem}, _) $ _) = I |
43181 | 1737 |
| 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
|
1738 |
| add (Const x) = |
1b333e4173a2
drop more type arguments soundly, when they can be deduced from the arg types
blanchet
parents:
44634
diff
changeset
|
1739 |
x |> robust_const_typargs thy |> fold (fold_type_constrs set_insert) |
43181 | 1740 |
| add (Abs (_, _, u)) = add u |
1741 |
| add _ = I |
|
1742 |
in add end |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
1743 |
|
43189 | 1744 |
fun type_constrs_of_terms thy ts = |
1745 |
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
|
1746 |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1747 |
fun extract_lambda_def (Const (@{const_name HOL.eq}, _) $ t $ u) = |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1748 |
let val (head, args) = strip_comb t in |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1749 |
(head |> dest_Const |> fst, |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1750 |
fold_rev (fn t as Var ((s, _), T) => |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1751 |
(fn u => Abs (s, T, abstract_over (t, u))) |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1752 |
| _ => raise Fail "expected Var") args u) |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1753 |
end |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1754 |
| extract_lambda_def _ = raise Fail "malformed lifted lambda" |
45508 | 1755 |
|
45514 | 1756 |
fun trans_lams_from_string ctxt type_enc lam_trans = |
1757 |
if lam_trans = no_lamsN then |
|
1758 |
rpair [] |
|
1759 |
else if lam_trans = hide_lamsN then |
|
1760 |
lift_lams ctxt type_enc ##> K [] |
|
46365 | 1761 |
else if lam_trans = liftingN orelse lam_trans = lam_liftingN then |
45514 | 1762 |
lift_lams ctxt type_enc |
46365 | 1763 |
else if lam_trans = combsN then |
45514 | 1764 |
map (introduce_combinators ctxt) #> rpair [] |
46365 | 1765 |
else if lam_trans = combs_and_liftingN then |
1766 |
lift_lams_part_1 ctxt type_enc |
|
1767 |
##> maps (fn t => [t, introduce_combinators ctxt (intentionalize_def t)]) |
|
1768 |
#> lift_lams_part_2 |
|
46368 | 1769 |
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
|
1770 |
lift_lams_part_1 ctxt type_enc |
46368 | 1771 |
##> map (fn t => case head_of (strip_qnt_body @{const_name All} t) of |
1772 |
@{term "op =::bool => bool => bool"} => t |
|
1773 |
| _ => 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
|
1774 |
#> lift_lams_part_2 |
45514 | 1775 |
else if lam_trans = keep_lamsN then |
1776 |
map (Envir.eta_contract) #> rpair [] |
|
1777 |
else |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45516
diff
changeset
|
1778 |
error ("Unknown lambda translation scheme: " ^ quote lam_trans ^ ".") |
45514 | 1779 |
|
1780 |
fun translate_formulas ctxt format prem_kind type_enc lam_trans presimp hyp_ts |
|
1781 |
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
|
1782 |
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
|
1783 |
val thy = Proof_Context.theory_of ctxt |
45514 | 1784 |
val trans_lams = trans_lams_from_string ctxt type_enc lam_trans |
43861 | 1785 |
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
|
1786 |
(* 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
|
1787 |
boost an ATP's performance (for some reason). *) |
43192 | 1788 |
val hyp_ts = |
1789 |
hyp_ts |
|
1790 |
|> 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
|
1791 |
val facts = facts |> map (apsnd (pair Axiom)) |
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents:
43863
diff
changeset
|
1792 |
val conjs = |
46385
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents:
46384
diff
changeset
|
1793 |
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
|
1794 |
|> map (apsnd freeze_term) |
46340 | 1795 |
|> map2 (pair o rpair (Local, General) o string_of_int) |
1796 |
(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
|
1797 |
val ((conjs, facts), lam_facts) = |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1798 |
(conjs, facts) |
46093 | 1799 |
|> presimp ? pairself (map (apsnd (uncurry (presimp_prop ctxt)))) |
45514 | 1800 |
|> (if lam_trans = no_lamsN then |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1801 |
rpair [] |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1802 |
else |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1803 |
op @ |
45514 | 1804 |
#> preprocess_abstractions_in_terms trans_lams |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
1805 |
#>> chop (length conjs)) |
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
1806 |
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
|
1807 |
val (fact_names, facts) = |
43864
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents:
43863
diff
changeset
|
1808 |
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
|
1809 |
|> 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
|
1810 |
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
|
1811 |
|> Option.map (pair name)) |
43861 | 1812 |
|> ListPair.unzip |
46375
d724066ff3d0
reverted e2b1a86d59fc -- broke Metis's lambda-lifting
blanchet
parents:
46371
diff
changeset
|
1813 |
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
|
1814 |
val lam_facts = |
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45551
diff
changeset
|
1815 |
lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd) |
43861 | 1816 |
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
|
1817 |
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
|
1818 |
val supers = tvar_classes_of_terms all_ts |
43189 | 1819 |
val tycons = type_constrs_of_terms thy all_ts |
43861 | 1820 |
val (supers, arity_clauses) = |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
1821 |
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
|
1822 |
else make_arity_clauses thy tycons supers |
43861 | 1823 |
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
|
1824 |
in |
45508 | 1825 |
(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
|
1826 |
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
|
1827 |
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
|
1828 |
|
44495 | 1829 |
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
|
1830 |
|
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
|
1831 |
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
|
1832 |
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
|
1833 |
|> 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
|
1834 |
|
43421 | 1835 |
fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum |
1836 |
| 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
|
1837 |
accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, []))) |
43692 | 1838 |
| 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
|
1839 |
|
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1840 |
fun is_var_ghost_type_arg_in_term thy polym_constrs name pos tm accum = |
44811 | 1841 |
is_var_positively_naked_in_term name pos tm accum orelse |
1842 |
let |
|
1843 |
val var = ATerm (name, []) |
|
1844 |
fun is_nasty_in_term (ATerm (_, [])) = false |
|
1845 |
| is_nasty_in_term (ATerm ((s, _), tms)) = |
|
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1846 |
let |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1847 |
val ary = length tms |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1848 |
val polym_constr = member (op =) polym_constrs s |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1849 |
val ghosts = ghost_type_args thy s ary |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1850 |
in |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1851 |
exists (fn (j, tm) => |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1852 |
if polym_constr then |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1853 |
member (op =) ghosts j andalso |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1854 |
(tm = var orelse is_nasty_in_term tm) |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1855 |
else |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1856 |
tm = var andalso member (op =) ghosts j) |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1857 |
(0 upto ary - 1 ~~ tms) |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1858 |
orelse (not polym_constr andalso exists is_nasty_in_term tms) |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1859 |
end |
44811 | 1860 |
| is_nasty_in_term _ = true |
1861 |
in is_nasty_in_term tm end |
|
1862 |
||
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1863 |
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
|
1864 |
name = |
44811 | 1865 |
(case granularity_of_type_level level of |
1866 |
All_Vars => true |
|
1867 |
| Positively_Naked_Vars => |
|
1868 |
formula_fold pos (is_var_positively_naked_in_term name) phi false |
|
1869 |
| Ghost_Type_Arg_Vars => |
|
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1870 |
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
|
1871 |
phi false) |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1872 |
| should_guard_var_in_formula _ _ _ _ _ _ _ = true |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1873 |
|
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1874 |
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
|
1875 |
|
44405
6fe1a89bb69a
generate tag equations for existential variables
blanchet
parents:
44404
diff
changeset
|
1876 |
fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false |
44768 | 1877 |
| should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T = |
44811 | 1878 |
granularity_of_type_level level <> All_Vars andalso |
44782 | 1879 |
should_encode_type ctxt mono level T |
44405
6fe1a89bb69a
generate tag equations for existential variables
blanchet
parents:
44404
diff
changeset
|
1880 |
| should_generate_tag_bound_decl _ _ _ _ _ = false |
44404
3111af540141
tuning, plus started implementing tag equation generation for existential variables
blanchet
parents:
44403
diff
changeset
|
1881 |
|
43677 | 1882 |
fun mk_aterm format type_enc name T_args args = |
1883 |
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
|
1884 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1885 |
fun do_bound_type ctxt format mono type_enc = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1886 |
case type_enc of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1887 |
Simple_Types (_, _, level) => |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1888 |
fused_type ctxt mono level 0 |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1889 |
#> ho_type_from_typ format type_enc false 0 #> SOME |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1890 |
| _ => K NONE |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1891 |
|
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1892 |
fun tag_with_type ctxt format mono type_enc pos T tm = |
43859 | 1893 |
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
|
1894 |
|> mangle_type_args_in_iterm format type_enc |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1895 |
|> ho_term_from_iterm ctxt format mono type_enc pos |
43692 | 1896 |
|> (fn ATerm (s, tms) => ATerm (s, tms @ [tm]) |
1897 |
| _ => raise Fail "unexpected lambda-abstraction") |
|
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1898 |
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
|
1899 |
let |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1900 |
fun term site u = |
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42956
diff
changeset
|
1901 |
let |
43859 | 1902 |
val (head, args) = strip_iterm_comb u |
43677 | 1903 |
val pos = |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1904 |
case site of |
43677 | 1905 |
Top_Level pos => pos |
1906 |
| Eq_Arg pos => pos |
|
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
1907 |
| _ => NONE |
43677 | 1908 |
val t = |
1909 |
case head of |
|
43859 | 1910 |
IConst (name as (s, _), _, T_args) => |
43677 | 1911 |
let |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1912 |
val arg_site = if is_tptp_equal s then Eq_Arg pos else Elsewhere |
43677 | 1913 |
in |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1914 |
map (term arg_site) args |> mk_aterm format type_enc name T_args |
43677 | 1915 |
end |
43859 | 1916 |
| IVar (name, _) => |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1917 |
map (term Elsewhere) args |> mk_aterm format type_enc name [] |
43859 | 1918 |
| IAbs ((name, T), tm) => |
1919 |
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
|
1920 |
term Elsewhere tm) |
43859 | 1921 |
| IApp _ => raise Fail "impossible \"IApp\"" |
1922 |
val T = ityp_of u |
|
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42956
diff
changeset
|
1923 |
in |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1924 |
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
|
1925 |
tag_with_type ctxt format mono type_enc pos T t |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1926 |
else |
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1927 |
t |
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42956
diff
changeset
|
1928 |
end |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1929 |
in term o Top_Level end |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1930 |
and formula_from_iformula ctxt polym_constrs format mono type_enc |
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1931 |
should_guard_var = |
42829
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
1932 |
let |
44811 | 1933 |
val thy = Proof_Context.theory_of ctxt |
1934 |
val level = level_of_type_enc type_enc |
|
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1935 |
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
|
1936 |
fun do_out_of_bound_type pos phi universal (name, T) = |
44399 | 1937 |
if should_guard_type ctxt mono type_enc |
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1938 |
(fn () => should_guard_var thy polym_constrs level pos phi |
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
1939 |
universal name) T then |
43859 | 1940 |
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
|
1941 |
|> 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
|
1942 |
|> do_term pos |> AAtom |> SOME |
44405
6fe1a89bb69a
generate tag equations for existential variables
blanchet
parents:
44404
diff
changeset
|
1943 |
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
|
1944 |
let |
6fe1a89bb69a
generate tag equations for existential variables
blanchet
parents:
44404
diff
changeset
|
1945 |
val var = ATerm (name, []) |
45949
70b9d1e9fddc
killed "guard@?" encodings -- they were found to be unsound
blanchet
parents:
45948
diff
changeset
|
1946 |
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
|
1947 |
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
|
1948 |
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
|
1949 |
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
|
1950 |
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
|
1951 |
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
|
1952 |
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
|
1953 |
val universal = Option.map (q = AExists ? not) pos |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
1954 |
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
|
1955 |
in |
42834
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
blanchet
parents:
42832
diff
changeset
|
1956 |
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
|
1957 |
| 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
|
1958 |
(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
|
1959 |
(map_filter |
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
blanchet
parents:
42832
diff
changeset
|
1960 |
(fn (_, NONE) => NONE |
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
blanchet
parents:
42832
diff
changeset
|
1961 |
| (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
|
1962 |
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
|
1963 |
xs) |
42834
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
blanchet
parents:
42832
diff
changeset
|
1964 |
phi) |
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
blanchet
parents:
42832
diff
changeset
|
1965 |
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
|
1966 |
| 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
|
1967 |
| 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
|
1968 |
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
|
1969 |
|
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
|
1970 |
(* 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
|
1971 |
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
|
1972 |
the remote provers might care. *) |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1973 |
fun formula_line_for_fact ctxt polym_constrs format prefix encode freshen pos |
46406 | 1974 |
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
|
1975 |
(prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name, kind, |
43859 | 1976 |
iformula |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1977 |
|> formula_from_iformula ctxt polym_constrs format mono type_enc |
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
1978 |
should_guard_var_in_formula (if pos then SOME true else NONE) |
45377 | 1979 |
|> close_formula_universally |
45920 | 1980 |
|> 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
|
1981 |
NONE, |
46406 | 1982 |
let val rank = rank j in |
1983 |
case snd stature of |
|
1984 |
Intro => isabelle_info format introN rank |
|
1985 |
| Elim => isabelle_info format elimN rank |
|
1986 |
| Simp => isabelle_info format simpN rank |
|
1987 |
| Spec_Eq => isabelle_info format spec_eqN rank |
|
1988 |
| _ => isabelle_info format "" rank |
|
1989 |
end) |
|
43493
bdb11c68f142
generate type predicates for existentials/skolems, otherwise some problems might not be provable
blanchet
parents:
43423
diff
changeset
|
1990 |
|> 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
|
1991 |
|
45301
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45299
diff
changeset
|
1992 |
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
|
1993 |
({name, subclass, superclass, ...} : class_rel_clause) = |
44622 | 1994 |
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
|
1995 |
Formula (class_rel_clause_prefix ^ ascii_of name, Axiom, |
44622 | 1996 |
AConn (AImplies, |
44625 | 1997 |
[type_class_formula type_enc subclass ty_arg, |
1998 |
type_class_formula type_enc superclass ty_arg]) |
|
45830 | 1999 |
|> mk_aquant AForall |
2000 |
[(tvar_a_name, atype_of_type_vars type_enc)], |
|
46406 | 2001 |
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
|
2002 |
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
|
2003 |
|
44625 | 2004 |
fun formula_from_arity_atom type_enc (class, t, args) = |
2005 |
ATerm (t, map (fn arg => ATerm (arg, [])) args) |
|
2006 |
|> 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
|
2007 |
|
45301
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45299
diff
changeset
|
2008 |
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
|
2009 |
({name, prem_atoms, concl_atom} : arity_clause) = |
43495 | 2010 |
Formula (arity_clause_prefix ^ name, Axiom, |
44625 | 2011 |
mk_ahorn (map (formula_from_arity_atom type_enc) prem_atoms) |
2012 |
(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
|
2013 |
|> mk_aquant AForall |
d00339ae7fff
repaired quantification over type variables for non-TFF1/THF encodings
blanchet
parents:
45316
diff
changeset
|
2014 |
(map (rpair (atype_of_type_vars type_enc)) (#3 concl_atom)), |
46406 | 2015 |
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
|
2016 |
|
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2017 |
fun formula_line_for_conjecture ctxt polym_constrs format mono type_enc |
43859 | 2018 |
({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
|
2019 |
Formula (conjecture_prefix ^ name, kind, |
45316
08d84bdd5b37
improve handling of bound type variables (esp. for TFF1)
blanchet
parents:
45315
diff
changeset
|
2020 |
iformula |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2021 |
|> 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
|
2022 |
should_guard_var_in_formula (SOME false) |
45377 | 2023 |
|> close_formula_universally |
46406 | 2024 |
|> 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
|
2025 |
|
44625 | 2026 |
fun formula_line_for_free_type j phi = |
46406 | 2027 |
Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis, phi, NONE, []) |
44626 | 2028 |
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
|
2029 |
let |
44625 | 2030 |
val phis = |
2031 |
fold (union (op =)) (map #atomic_types facts) [] |
|
2032 |
|> formulas_for_types type_enc add_sorts_on_tfree |
|
2033 |
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
|
2034 |
|
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
|
2035 |
(** 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
|
2036 |
|
44754 | 2037 |
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
|
2038 |
let val name as (s, _) = `make_type_class s in |
44622 | 2039 |
Decl (sym_decl_prefix ^ s, name, |
45779 | 2040 |
if order = First_Order then |
2041 |
ATyAbs ([tvar_a_name], |
|
2042 |
if avoid_first_order_ghost_type_vars then |
|
2043 |
AFun (a_itself_atype, bool_atype) |
|
2044 |
else |
|
2045 |
bool_atype) |
|
44622 | 2046 |
else |
44754 | 2047 |
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
|
2048 |
end |
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
blanchet
parents:
44594
diff
changeset
|
2049 |
|
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
blanchet
parents:
44594
diff
changeset
|
2050 |
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
|
2051 |
case type_enc of |
44754 | 2052 |
Simple_Types (order, Polymorphic, _) => |
2053 |
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
|
2054 |
| _ => [] |
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
blanchet
parents:
44594
diff
changeset
|
2055 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2056 |
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
|
2057 |
(conjs, facts, extra_tms) = |
42574 | 2058 |
let |
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2059 |
fun add_iterm_syms tm = |
43859 | 2060 |
let val (head, args) = strip_iterm_comb tm in |
42574 | 2061 |
(case head of |
43859 | 2062 |
IConst ((s, s'), T, T_args) => |
44594
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2063 |
let |
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2064 |
val (pred_sym, in_conj) = |
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2065 |
case Symtab.lookup sym_tab s of |
44853 | 2066 |
SOME ({pred_sym, in_conj, ...} : sym_info) => |
2067 |
(pred_sym, in_conj) |
|
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2068 |
| NONE => (false, false) |
44594
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2069 |
val decl_sym = |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2070 |
(case type_enc of |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2071 |
Guards _ => not pred_sym |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2072 |
| _ => true) andalso |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2073 |
is_tptp_user_symbol s |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2074 |
in |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2075 |
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
|
2076 |
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
|
2077 |
(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
|
2078 |
in_conj)) |
42574 | 2079 |
else |
2080 |
I |
|
2081 |
end |
|
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2082 |
| IAbs (_, tm) => add_iterm_syms tm |
42574 | 2083 |
| _ => I) |
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2084 |
#> fold add_iterm_syms args |
42574 | 2085 |
end |
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2086 |
val add_fact_syms = K add_iterm_syms |> formula_fold NONE |> fact_lift |
43966 | 2087 |
fun add_formula_var_types (AQuant (_, xs, phi)) = |
2088 |
fold (fn (_, SOME T) => insert_type ctxt I T | _ => I) xs |
|
2089 |
#> add_formula_var_types phi |
|
2090 |
| add_formula_var_types (AConn (_, phis)) = |
|
2091 |
fold add_formula_var_types phis |
|
2092 |
| add_formula_var_types _ = I |
|
2093 |
fun var_types () = |
|
2094 |
if polymorphism_of_type_enc type_enc = Polymorphic then [tvar_a] |
|
2095 |
else fold (fact_lift add_formula_var_types) (conjs @ facts) [] |
|
2096 |
fun add_undefined_const T = |
|
43984 | 2097 |
let |
2098 |
val (s, s') = |
|
44622 | 2099 |
`(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
|
2100 |
|> (case type_arg_policy [] type_enc @{const_name undefined} of |
44771 | 2101 |
Mangled_Type_Args => mangled_const_name format type_enc [T] |
44001 | 2102 |
| _ => I) |
43984 | 2103 |
in |
2104 |
Symtab.map_default (s, []) |
|
2105 |
(insert_type ctxt #3 (s', [T], T, false, 0, false)) |
|
2106 |
end |
|
44622 | 2107 |
fun add_TYPE_const () = |
2108 |
let val (s, s') = TYPE_name in |
|
2109 |
Symtab.map_default (s, []) |
|
2110 |
(insert_type ctxt #3 |
|
2111 |
(s', [tvar_a], @{typ "'a itself"}, false, 0, false)) |
|
2112 |
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
|
2113 |
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
|
2114 |
Symtab.empty |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
2115 |
|> is_type_enc_fairly_sound type_enc |
44829
5a2cd5db0a11
fixed computation of "in_conj" for polymorphic encodings
blanchet
parents:
44814
diff
changeset
|
2116 |
? (fold (fold add_fact_syms) [conjs, facts] |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2117 |
#> 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
|
2118 |
#> (case type_enc of |
45829 | 2119 |
Simple_Types (First_Order, Polymorphic, _) => |
2120 |
if avoid_first_order_ghost_type_vars then add_TYPE_const () |
|
2121 |
else I |
|
44786 | 2122 |
| Simple_Types _ => I |
43985
33d8b99531c2
no need for existential witnesses for sorts in TFF and THF formats
blanchet
parents:
43984
diff
changeset
|
2123 |
| _ => 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
|
2124 |
end |
42533 | 2125 |
|
44399 | 2126 |
(* 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
|
2127 |
fun default_mono level = |
44399 | 2128 |
{maybe_finite_Ts = [@{typ bool}], |
2129 |
surely_finite_Ts = [@{typ bool}], |
|
2130 |
maybe_infinite_Ts = known_infinite_types, |
|
44634
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
blanchet
parents:
44626
diff
changeset
|
2131 |
surely_infinite_Ts = |
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
blanchet
parents:
44626
diff
changeset
|
2132 |
case level of |
46301 | 2133 |
Noninf_Nonmono_Types (Strict, _) => [] |
44634
2ac4ff398bc3
make "sound" sound and "unsound" more sound, based on evaluation
blanchet
parents:
44626
diff
changeset
|
2134 |
| _ => known_infinite_types, |
44399 | 2135 |
maybe_nonmono_Ts = [@{typ bool}]} |
2136 |
||
42685 | 2137 |
(* This inference is described in section 2.3 of Claessen et al.'s "Sorting it |
2138 |
out with monotonicity" paper presented at CADE 2011. *) |
|
44399 | 2139 |
fun add_iterm_mononotonicity_info _ _ (SOME false) _ mono = mono |
2140 |
| add_iterm_mononotonicity_info ctxt level _ |
|
2141 |
(IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2)) |
|
2142 |
(mono as {maybe_finite_Ts, surely_finite_Ts, maybe_infinite_Ts, |
|
2143 |
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
|
2144 |
if is_tptp_equal s andalso exists is_maybe_universal_var [tm1, tm2] then |
44399 | 2145 |
case level of |
46301 | 2146 |
Noninf_Nonmono_Types (strictness, _) => |
44399 | 2147 |
if exists (type_instance ctxt T) surely_infinite_Ts orelse |
44859
237ba63d6041
fixed definition of type intersection (soundness bug)
blanchet
parents:
44853
diff
changeset
|
2148 |
member (type_equiv ctxt) maybe_finite_Ts T then |
44399 | 2149 |
mono |
46301 | 2150 |
else if is_type_kind_of_surely_infinite ctxt strictness |
44500 | 2151 |
surely_infinite_Ts T then |
44399 | 2152 |
{maybe_finite_Ts = maybe_finite_Ts, |
2153 |
surely_finite_Ts = surely_finite_Ts, |
|
2154 |
maybe_infinite_Ts = maybe_infinite_Ts, |
|
2155 |
surely_infinite_Ts = surely_infinite_Ts |> insert_type ctxt I T, |
|
2156 |
maybe_nonmono_Ts = maybe_nonmono_Ts} |
|
2157 |
else |
|
44859
237ba63d6041
fixed definition of type intersection (soundness bug)
blanchet
parents:
44853
diff
changeset
|
2158 |
{maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv ctxt) T, |
44399 | 2159 |
surely_finite_Ts = surely_finite_Ts, |
2160 |
maybe_infinite_Ts = maybe_infinite_Ts, |
|
2161 |
surely_infinite_Ts = surely_infinite_Ts, |
|
2162 |
maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type ctxt I T} |
|
44768 | 2163 |
| Fin_Nonmono_Types _ => |
44399 | 2164 |
if exists (type_instance ctxt T) surely_finite_Ts orelse |
44859
237ba63d6041
fixed definition of type intersection (soundness bug)
blanchet
parents:
44853
diff
changeset
|
2165 |
member (type_equiv ctxt) maybe_infinite_Ts T then |
44399 | 2166 |
mono |
2167 |
else if is_type_surely_finite ctxt T then |
|
2168 |
{maybe_finite_Ts = maybe_finite_Ts, |
|
2169 |
surely_finite_Ts = surely_finite_Ts |> insert_type ctxt I T, |
|
2170 |
maybe_infinite_Ts = maybe_infinite_Ts, |
|
2171 |
surely_infinite_Ts = surely_infinite_Ts, |
|
2172 |
maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type ctxt I T} |
|
2173 |
else |
|
2174 |
{maybe_finite_Ts = maybe_finite_Ts, |
|
2175 |
surely_finite_Ts = surely_finite_Ts, |
|
44859
237ba63d6041
fixed definition of type intersection (soundness bug)
blanchet
parents:
44853
diff
changeset
|
2176 |
maybe_infinite_Ts = maybe_infinite_Ts |> insert (type_equiv ctxt) T, |
44399 | 2177 |
surely_infinite_Ts = surely_infinite_Ts, |
2178 |
maybe_nonmono_Ts = maybe_nonmono_Ts} |
|
2179 |
| _ => mono |
|
2180 |
else |
|
2181 |
mono |
|
2182 |
| add_iterm_mononotonicity_info _ _ _ _ mono = mono |
|
2183 |
fun add_fact_mononotonicity_info ctxt level |
|
2184 |
({kind, iformula, ...} : translated_formula) = |
|
42834
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
blanchet
parents:
42832
diff
changeset
|
2185 |
formula_fold (SOME (kind <> Conjecture)) |
44399 | 2186 |
(add_iterm_mononotonicity_info ctxt level) iformula |
2187 |
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
|
2188 |
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
|
2189 |
default_mono level |
44399 | 2190 |
|> is_type_level_monotonicity_based level |
2191 |
? 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
|
2192 |
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
|
2193 |
|
44501 | 2194 |
fun add_iformula_monotonic_types ctxt mono type_enc = |
2195 |
let |
|
2196 |
val level = level_of_type_enc type_enc |
|
2197 |
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
|
2198 |
fun add_type T = not (should_encode T) ? insert_type ctxt I T |
44506 | 2199 |
fun add_args (IApp (tm1, tm2)) = add_args tm1 #> add_term tm2 |
2200 |
| add_args _ = I |
|
2201 |
and add_term tm = add_type (ityp_of tm) #> add_args tm |
|
44501 | 2202 |
in formula_fold NONE (K add_term) end |
2203 |
fun add_fact_monotonic_types ctxt mono type_enc = |
|
2204 |
add_iformula_monotonic_types ctxt mono type_enc |> fact_lift |
|
2205 |
fun monotonic_types_for_facts ctxt mono type_enc facts = |
|
44811 | 2206 |
let val level = level_of_type_enc type_enc in |
2207 |
[] |> (polymorphism_of_type_enc type_enc = Polymorphic andalso |
|
2208 |
is_type_level_monotonicity_based level andalso |
|
2209 |
granularity_of_type_level level <> Ghost_Type_Arg_Vars) |
|
2210 |
? fold (add_fact_monotonic_types ctxt mono type_enc) facts |
|
2211 |
end |
|
44501 | 2212 |
|
44399 | 2213 |
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
|
2214 |
Formula (guards_sym_formula_prefix ^ |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2215 |
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
|
2216 |
Axiom, |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2217 |
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
|
2218 |
|> 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
|
2219 |
|> AAtom |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2220 |
|> formula_from_iformula ctxt [] format mono type_enc |
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
2221 |
always_guard_var_in_formula (SOME true) |
45377 | 2222 |
|> close_formula_universally |
45920 | 2223 |
|> bound_tvars type_enc true (atomic_types_of T), |
46406 | 2224 |
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
|
2225 |
|
44399 | 2226 |
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
|
2227 |
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
|
2228 |
Formula (tags_sym_formula_prefix ^ |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2229 |
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
|
2230 |
Axiom, |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2231 |
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
|
2232 |
(tag_with_type ctxt format mono type_enc NONE T x_var) x_var, |
46406 | 2233 |
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
|
2234 |
end |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2235 |
|
44399 | 2236 |
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
|
2237 |
case type_enc of |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2238 |
Simple_Types _ => [] |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2239 |
| Guards _ => |
44399 | 2240 |
map (formula_line_for_guards_mono_type ctxt format mono type_enc) Ts |
2241 |
| 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
|
2242 |
|
44399 | 2243 |
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
|
2244 |
(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
|
2245 |
let |
44594
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2246 |
val thy = Proof_Context.theory_of ctxt |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2247 |
val (T, T_args) = |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2248 |
if null T_args then |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2249 |
(T, []) |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45509
diff
changeset
|
2250 |
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
|
2251 |
SOME s' => |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2252 |
let |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2253 |
val s' = s' |> invert_const |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2254 |
val T = s' |> robust_const_type thy |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2255 |
in (T, robust_const_typargs thy (s', T)) end |
45509 | 2256 |
| 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
|
2257 |
in |
42998
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2258 |
Decl (sym_decl_prefix ^ s, (s, s'), |
44594
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2259 |
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
|
2260 |
|> 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
|
2261 |
|> not (null T_args) |
ae82943481e9
added type abstractions (for declaring polymorphic constants) to TFF syntax
blanchet
parents:
44593
diff
changeset
|
2262 |
? 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
|
2263 |
end |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42577
diff
changeset
|
2264 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2265 |
fun honor_conj_sym_kind in_conj conj_sym_kind = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2266 |
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
|
2267 |
else (Axiom, I) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2268 |
|
44399 | 2269 |
fun formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono type_enc n s |
2270 |
j (s', T_args, T, _, ary, in_conj) = |
|
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42577
diff
changeset
|
2271 |
let |
44811 | 2272 |
val thy = Proof_Context.theory_of ctxt |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2273 |
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
|
2274 |
val (arg_Ts, res_T) = chop_fun ary T |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2275 |
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
|
2276 |
val bounds = |
43859 | 2277 |
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
|
2278 |
val bound_Ts = |
44811 | 2279 |
if exists (curry (op =) dummyT) T_args then |
2280 |
case level_of_type_enc type_enc of |
|
2281 |
All_Types => map SOME arg_Ts |
|
2282 |
| level => |
|
2283 |
if granularity_of_type_level level = Ghost_Type_Arg_Vars then |
|
2284 |
let val ghosts = ghost_type_args thy s ary in |
|
2285 |
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
|
2286 |
(0 upto ary - 1) arg_Ts |
44811 | 2287 |
end |
2288 |
else |
|
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2289 |
replicate ary NONE |
44811 | 2290 |
else |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2291 |
replicate ary NONE |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42577
diff
changeset
|
2292 |
in |
43989 | 2293 |
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
|
2294 |
(if n > 1 then "_" ^ string_of_int j else ""), kind, |
43859 | 2295 |
IConst ((s, s'), T, T_args) |
2296 |
|> 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
|
2297 |
|> type_guard_iterm format type_enc res_T |
42963 | 2298 |
|> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts) |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2299 |
|> formula_from_iformula ctxt [] format mono type_enc |
45948
f88f502d635f
extend previous optimizations to guard-based encodings
blanchet
parents:
45947
diff
changeset
|
2300 |
always_guard_var_in_formula (SOME true) |
45377 | 2301 |
|> close_formula_universally |
45920 | 2302 |
|> 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
|
2303 |
|> maybe_negate, |
46406 | 2304 |
NONE, isabelle_info format introN helper_rank) |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42577
diff
changeset
|
2305 |
end |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42577
diff
changeset
|
2306 |
|
44811 | 2307 |
fun formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono type_enc n s |
2308 |
(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
|
2309 |
let |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2310 |
val thy = Proof_Context.theory_of ctxt |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2311 |
val level = level_of_type_enc type_enc |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2312 |
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
|
2313 |
val ident_base = |
44396
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2314 |
tags_sym_formula_prefix ^ s ^ |
43125
ddf63baabdec
distinguish different kinds of typing informations in the fact name
blanchet
parents:
43121
diff
changeset
|
2315 |
(if n > 1 then "_" ^ string_of_int j else "") |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2316 |
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
|
2317 |
val (arg_Ts, res_T) = chop_fun ary T |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2318 |
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
|
2319 |
val bounds = bound_names |> map (fn name => ATerm (name, [])) |
43677 | 2320 |
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
|
2321 |
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
|
2322 |
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
|
2323 |
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
|
2324 |
val add_formula_for_res = |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2325 |
if should_encode res_T then |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2326 |
let |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2327 |
val tagged_bounds = |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2328 |
if grain = Ghost_Type_Arg_Vars then |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2329 |
let val ghosts = ghost_type_args thy s ary in |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2330 |
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
|
2331 |
(0 upto ary - 1 ~~ arg_Ts) bounds |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2332 |
end |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2333 |
else |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2334 |
bounds |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2335 |
in |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2336 |
cons (Formula (ident_base ^ "_res", kind, |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2337 |
eq (tag_with res_T (cst bounds)) (cst tagged_bounds), |
46406 | 2338 |
NONE, isabelle_info format spec_eqN helper_rank)) |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2339 |
end |
42829
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2340 |
else |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2341 |
I |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2342 |
fun add_formula_for_arg k = |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2343 |
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
|
2344 |
if should_encode arg_T then |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2345 |
case chop k bounds of |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2346 |
(bounds1, bound :: bounds2) => |
42852
40649ab0cead
honor "conj_sym_kind" also for tag symbol declarations
blanchet
parents:
42851
diff
changeset
|
2347 |
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
|
2348 |
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
|
2349 |
(cst bounds), |
46406 | 2350 |
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
|
2351 |
| _ => raise Fail "expected nonempty tail" |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2352 |
else |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2353 |
I |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2354 |
end |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2355 |
in |
42834
40f7691d0539
implemented thin versions of "preds" type systems + fixed various issues with type args
blanchet
parents:
42832
diff
changeset
|
2356 |
[] |> not pred_sym ? add_formula_for_res |
44814
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2357 |
|> (Config.get ctxt type_tag_arguments andalso |
52318464c73b
also implemented ghost version of the tagged encodings
blanchet
parents:
44812
diff
changeset
|
2358 |
grain = Positively_Naked_Vars) |
44496
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
blanchet
parents:
44495
diff
changeset
|
2359 |
? 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
|
2360 |
end |
1558741f8a72
started implementing "shallow" type systems, based on ideas by Claessen et al.
blanchet
parents:
42828
diff
changeset
|
2361 |
|
42836 | 2362 |
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd |
2363 |
||
45780 | 2364 |
fun rationalize_decls ctxt (decls as decl :: (decls' as _ :: _)) = |
2365 |
let |
|
2366 |
val T = result_type_of_decl decl |
|
2367 |
|> map_type_tvar (fn (z, _) => TVar (z, HOLogic.typeS)) |
|
2368 |
in |
|
2369 |
if forall (type_generalization ctxt T o result_type_of_decl) decls' then |
|
2370 |
[decl] |
|
2371 |
else |
|
2372 |
decls |
|
2373 |
end |
|
2374 |
| rationalize_decls _ decls = decls |
|
2375 |
||
44399 | 2376 |
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
|
2377 |
(s, decls) = |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43624
diff
changeset
|
2378 |
case type_enc of |
45831
8f8fce7bd32b
avoid multiple type decls in TFF (improves on cef82dc1462d)
blanchet
parents:
45830
diff
changeset
|
2379 |
Simple_Types _ => [decl_line_for_sym ctxt format mono type_enc s (hd decls)] |
44768 | 2380 |
| Guards (_, level) => |
42998
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2381 |
let |
45780 | 2382 |
val decls = decls |> rationalize_decls ctxt |
42998
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2383 |
val n = length decls |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2384 |
val decls = |
44399 | 2385 |
decls |> filter (should_encode_type ctxt mono level |
43401
e93dfcb53535
fixed soundness bug made more visible by previous change
blanchet
parents:
43399
diff
changeset
|
2386 |
o result_type_of_decl) |
42998
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2387 |
in |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2388 |
(0 upto length decls - 1, decls) |
44399 | 2389 |
|-> map2 (formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono |
2390 |
type_enc n s) |
|
42998
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2391 |
end |
44768 | 2392 |
| Tags (_, level) => |
44811 | 2393 |
if granularity_of_type_level level = All_Vars then |
44768 | 2394 |
[] |
2395 |
else |
|
2396 |
let val n = length decls in |
|
2397 |
(0 upto n - 1 ~~ decls) |
|
44811 | 2398 |
|> maps (formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono |
2399 |
type_enc n s) |
|
44768 | 2400 |
end |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42577
diff
changeset
|
2401 |
|
44399 | 2402 |
fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono type_enc |
44501 | 2403 |
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
|
2404 |
let |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2405 |
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
|
2406 |
val mono_lines = |
44399 | 2407 |
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
|
2408 |
val decl_lines = |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2409 |
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
|
2410 |
mono type_enc) |
44396
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2411 |
syms [] |
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
blanchet
parents:
44394
diff
changeset
|
2412 |
in mono_lines @ decl_lines end |
42543
f9d402d144d4
declare TFF types so that SNARK can be used with types
blanchet
parents:
42542
diff
changeset
|
2413 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2414 |
fun pair_append (xs1, xs2) (ys1, ys2) = (xs1 @ ys1, xs2 @ ys2) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2415 |
|
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2416 |
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
|
2417 |
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
|
2418 |
let |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2419 |
fun do_alias ary = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2420 |
let |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2421 |
val thy = Proof_Context.theory_of ctxt |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2422 |
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
|
2423 |
val base_name = base_s0 |> `(make_fixed_const (SOME format)) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2424 |
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
|
2425 |
val T_args = robust_const_typargs thy (base_s0, T) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2426 |
val (base_name as (base_s, _), T_args) = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2427 |
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
|
2428 |
val base_ary = min_ary_of sym_tab0 base_s |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2429 |
fun do_const name = IConst (name, T, T_args) |
46437 | 2430 |
val filter_ty_args = |
46410
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
blanchet
parents:
46409
diff
changeset
|
2431 |
filter_type_args_in_iterm thy monom_constrs type_enc |
46437 | 2432 |
val ho_term_of = |
2433 |
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
|
2434 |
val name1 as (s1, _) = |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2435 |
base_name |> ary - 1 > base_ary ? aliased_uncurried (ary - 1) |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2436 |
val name2 as (s2, _) = base_name |> aliased_uncurried ary |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2437 |
val (arg_Ts, _) = chop_fun ary T |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2438 |
val bound_names = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2439 |
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
|
2440 |
val bounds = bound_names ~~ arg_Ts |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2441 |
val (first_bounds, last_bound) = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2442 |
bounds |> map (fn (name, T) => IConst (name, T, [])) |> split_last |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2443 |
val tm1 = |
46410
78ff6a886b50
made sure to filter type args also for "uncurried alias" equations
blanchet
parents:
46409
diff
changeset
|
2444 |
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
|
2445 |
last_bound |
46437 | 2446 |
|> filter_ty_args |
2447 |
val tm2 = |
|
2448 |
list_app (do_const name2) (first_bounds @ [last_bound]) |
|
2449 |
|> filter_ty_args |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2450 |
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
|
2451 |
val eq = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2452 |
eq_formula type_enc (atomic_types_of T) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2453 |
(map (apsnd do_bound_type) bounds) false |
46437 | 2454 |
(ho_term_of tm1) (ho_term_of tm2) |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2455 |
in |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2456 |
([tm1, tm2], |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2457 |
[Formula (uncurried_alias_eq_prefix ^ s2, kind, eq |> maybe_negate, |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2458 |
NONE, isabelle_info format spec_eqN helper_rank)]) |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2459 |
|> (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
|
2460 |
else pair_append (do_alias (ary - 1))) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2461 |
end |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2462 |
in do_alias end |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2463 |
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
|
2464 |
type_enc sym_tab0 sym_tab |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2465 |
(s, {min_ary, types, in_conj, ...} : sym_info) = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2466 |
case unprefix_and_unascii const_prefix s of |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2467 |
SOME mangled_s => |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2468 |
if String.isSubstring uncurried_alias_sep mangled_s then |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2469 |
let |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2470 |
val base_s0 = mangled_s |> unmangled_const_name |> hd |> invert_const |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2471 |
in |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2472 |
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
|
2473 |
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
|
2474 |
end |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2475 |
else |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2476 |
([], []) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2477 |
| NONE => ([], []) |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2478 |
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
|
2479 |
mono type_enc uncurried_aliases sym_tab0 sym_tab = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2480 |
([], []) |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2481 |
|> uncurried_aliases |
46402
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
blanchet
parents:
46400
diff
changeset
|
2482 |
? Symtab.fold_rev |
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
blanchet
parents:
46400
diff
changeset
|
2483 |
(pair_append |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2484 |
o uncurried_alias_lines_for_sym ctxt monom_constrs format |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2485 |
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
|
2486 |
|
44768 | 2487 |
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
|
2488 |
Config.get ctxt type_tag_idempotence andalso |
44768 | 2489 |
is_type_level_monotonicity_based level andalso |
2490 |
poly <> Mangled_Monomorphic |
|
44496
c1884789ff80
added config options to control two aspects of the translation, for evaluation purposes
blanchet
parents:
44495
diff
changeset
|
2491 |
| needs_type_tag_idempotence _ _ = false |
42831
c9b0968484fb
more work on "shallow" encoding + adjustments to other encodings
blanchet
parents:
42830
diff
changeset
|
2492 |
|
42998
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2493 |
val implicit_declsN = "Should-be-implicit typings" |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
2494 |
val explicit_declsN = "Explicit typings" |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2495 |
val uncurried_alias_eqsN = "Uncurried aliases" |
41157 | 2496 |
val factsN = "Relevant facts" |
2497 |
val class_relsN = "Class relationships" |
|
42543
f9d402d144d4
declare TFF types so that SNARK can be used with types
blanchet
parents:
42542
diff
changeset
|
2498 |
val aritiesN = "Arities" |
41157 | 2499 |
val helpersN = "Helper facts" |
2500 |
val conjsN = "Conjectures" |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
2501 |
val free_typesN = "Type variables" |
41157 | 2502 |
|
45828
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2503 |
(* 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
|
2504 |
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
|
2505 |
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
|
2506 |
|
45875
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
blanchet
parents:
45831
diff
changeset
|
2507 |
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
|
2508 |
let |
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2509 |
val ind = |
45875
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
blanchet
parents:
45831
diff
changeset
|
2510 |
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
|
2511 |
Simple_Types _ => |
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
blanchet
parents:
45831
diff
changeset
|
2512 |
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
|
2513 |
else individual_atype |
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
blanchet
parents:
45831
diff
changeset
|
2514 |
| _ => 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
|
2515 |
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
|
2516 |
| 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
|
2517 |
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
|
2518 |
|
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2519 |
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
|
2520 |
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
|
2521 |
|
45875
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
blanchet
parents:
45831
diff
changeset
|
2522 |
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
|
2523 |
let |
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2524 |
val declared = declared_syms_in_problem problem |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2525 |
fun do_sym (name as (s, _)) ty = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2526 |
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
|
2527 |
AList.default (op =) (name, ty) |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2528 |
else |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2529 |
I |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2530 |
fun do_type (AType (name, tys)) = |
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2531 |
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
|
2532 |
#> 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
|
2533 |
| 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
|
2534 |
| 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
|
2535 |
fun do_term pred_sym (ATerm (name as (s, _), tms)) = |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2536 |
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
|
2537 |
#> 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
|
2538 |
| 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
|
2539 |
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
|
2540 |
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
|
2541 |
| 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
|
2542 |
| 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
|
2543 |
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
|
2544 |
| 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
|
2545 |
in |
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2546 |
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
|
2547 |
|> 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
|
2548 |
end |
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2549 |
|
45875
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
blanchet
parents:
45831
diff
changeset
|
2550 |
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
|
2551 |
let |
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2552 |
val decls = |
3b8606fba2dd
correctly declare implicit TFF1 types that appear first as type arguments with "$tType" and not "$i
blanchet
parents:
45780
diff
changeset
|
2553 |
problem |
45875
6d3533fd26ea
make sure that all symbols are declared in untyped SPASS DFG output (broken since 3b8606fba2dd)
blanchet
parents:
45831
diff
changeset
|
2554 |
|> 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
|
2555 |
|> 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
|
2556 |
|> 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
|
2557 |
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
|
2558 |
|
45945
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2559 |
fun exists_subdtype P = |
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2560 |
let |
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2561 |
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
|
2562 |
(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
|
2563 |
in ex end |
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2564 |
|
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2565 |
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
|
2566 |
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
|
2567 |
|
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2568 |
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
|
2569 |
Symtab.fold (snd |
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2570 |
#> #descr |
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2571 |
#> maps (snd #> #3) |
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2572 |
#> (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
|
2573 |
(Datatype.get_all thy) [] |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2574 |
|> List.partition is_poly_constr |
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2575 |
|> pairself (map fst) |
45945
aa8100cc02dc
no need for type arguments for monomorphic constructors of polymorphic datatypes (e.g. "Nil")
blanchet
parents:
45939
diff
changeset
|
2576 |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2577 |
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
|
2578 |
|
44397
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44396
diff
changeset
|
2579 |
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
|
2580 |
lam_trans uncurried_aliases readable_names preproc |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2581 |
hyp_ts concl_t facts = |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
2582 |
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
|
2583 |
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
|
2584 |
val type_enc = type_enc |> adjust_type_enc format |
46389 | 2585 |
(* Forcing explicit applications is expensive for polymorphic encodings, |
2586 |
because it takes only one existential variable ranging over "'a => 'b" to |
|
2587 |
ruin everything. Hence we do it only if there are few facts (which is |
|
2588 |
normally the case for "metis" and the minimizer). *) |
|
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2589 |
val app_op_level = |
46370 | 2590 |
if polymorphism_of_type_enc type_enc = Polymorphic andalso |
46392
676a4b4b6e73
implemented partial application aliases (for SPASS mainly)
blanchet
parents:
46389
diff
changeset
|
2591 |
length facts >= app_op_threshold then |
46389 | 2592 |
Incomplete_Apply |
46370 | 2593 |
else |
46389 | 2594 |
Sufficient_Apply |
45514 | 2595 |
val lam_trans = |
45520 | 2596 |
if lam_trans = keep_lamsN andalso |
2597 |
not (is_type_enc_higher_order type_enc) then |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45516
diff
changeset
|
2598 |
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
|
2599 |
\encoding.") |
3693baa6befb
move lambda-lifting code to ATP encoding, so it can be used by Metis
blanchet
parents:
44003
diff
changeset
|
2600 |
else |
45514 | 2601 |
lam_trans |
45508 | 2602 |
val (fact_names, classes, conjs, facts, class_rel_clauses, arity_clauses, |
2603 |
lifted) = |
|
45514 | 2604 |
translate_formulas ctxt format prem_kind type_enc lam_trans preproc hyp_ts |
2605 |
concl_t facts |
|
46402
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
blanchet
parents:
46400
diff
changeset
|
2606 |
val sym_tab0 = sym_table_for_facts ctxt type_enc app_op_level conjs facts |
44399 | 2607 |
val mono = conjs @ facts |> mononotonicity_info_for_facts ctxt type_enc |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2608 |
val (polym_constrs, monom_constrs) = |
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2609 |
all_constrs_of_polymorphic_datatypes thy |
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2610 |
|>> map (make_fixed_const (SOME format)) |
46400
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
blanchet
parents:
46399
diff
changeset
|
2611 |
fun firstorderize in_helper = |
46402
ef8d65f64f77
change 9ce354a77908 wasn't quite right -- here's an improvement
blanchet
parents:
46400
diff
changeset
|
2612 |
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
|
2613 |
(uncurried_aliases andalso not in_helper) |
46400
9ce354a77908
don't introduce new symbols in helpers -- makes problems unprovable
blanchet
parents:
46399
diff
changeset
|
2614 |
val (conjs, facts) = (conjs, facts) |> pairself (map (firstorderize false)) |
46389 | 2615 |
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
|
2616 |
val helpers = |
44773
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
blanchet
parents:
44772
diff
changeset
|
2617 |
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
|
2618 |
|> map (firstorderize true) |
44501 | 2619 |
val mono_Ts = |
45937 | 2620 |
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
|
2621 |
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
|
2622 |
val (uncurried_alias_eq_tms, uncurried_alias_eq_lines) = |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2623 |
uncurried_alias_lines_for_sym_table ctxt monom_constrs format |
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2624 |
conj_sym_kind mono type_enc uncurried_aliases sym_tab0 sym_tab |
42680 | 2625 |
val sym_decl_lines = |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2626 |
(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
|
2627 |
|> sym_decl_table_for_facts ctxt format type_enc sym_tab |
44399 | 2628 |
|> problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono |
44501 | 2629 |
type_enc mono_Ts |
46406 | 2630 |
val num_facts = length facts |
2631 |
val fact_lines = |
|
2632 |
map (formula_line_for_fact ctxt polym_constrs format fact_prefix |
|
2633 |
ascii_of (not exporter) (not exporter) mono type_enc |
|
2634 |
(rank_of_fact_num num_facts)) |
|
2635 |
(0 upto num_facts - 1 ~~ facts) |
|
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42879
diff
changeset
|
2636 |
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
|
2637 |
0 upto length helpers - 1 ~~ helpers |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2638 |
|> map (formula_line_for_fact ctxt polym_constrs format helper_prefix I |
46406 | 2639 |
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
|
2640 |
|> (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
|
2641 |
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
|
2642 |
else |
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents:
43139
diff
changeset
|
2643 |
I) |
42522
413b56894f82
close ATP formulas universally earlier, so that we can add type predicates
blanchet
parents:
42521
diff
changeset
|
2644 |
(* Reordering these might confuse the proof reconstruction code or the SPASS |
43039 | 2645 |
FLOTTER hack. *) |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
2646 |
val problem = |
44595
444d111bde7d
generate properly typed TFF1 (PFF) problems in the presence of type class predicates
blanchet
parents:
44594
diff
changeset
|
2647 |
[(explicit_declsN, class_decl_lines @ sym_decl_lines), |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46406
diff
changeset
|
2648 |
(uncurried_alias_eqsN, uncurried_alias_eq_lines), |
46406 | 2649 |
(factsN, fact_lines), |
44593
ccf40af26ae9
implement more of the polymorphic simply typed format TFF(1)
blanchet
parents:
44591
diff
changeset
|
2650 |
(class_relsN, |
45301
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45299
diff
changeset
|
2651 |
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
|
2652 |
class_rel_clauses), |
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45299
diff
changeset
|
2653 |
(aritiesN, |
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45299
diff
changeset
|
2654 |
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
|
2655 |
(helpersN, helper_lines), |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46642
diff
changeset
|
2656 |
(free_typesN, formula_lines_for_free_types type_enc (facts @ conjs)), |
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42956
diff
changeset
|
2657 |
(conjsN, |
45947
7eccf8147f57
treat polymorphic constructors specially in @? encodings
blanchet
parents:
45946
diff
changeset
|
2658 |
map (formula_line_for_conjecture ctxt polym_constrs format mono |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46642
diff
changeset
|
2659 |
type_enc) conjs)] |
42543
f9d402d144d4
declare TFF types so that SNARK can be used with types
blanchet
parents:
42542
diff
changeset
|
2660 |
val problem = |
42561
23ddc4e3d19c
have properly type-instantiated helper facts (combinators and If)
blanchet
parents:
42560
diff
changeset
|
2661 |
problem |
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43086
diff
changeset
|
2662 |
|> (case format of |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43086
diff
changeset
|
2663 |
CNF => ensure_cnf_problem |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43086
diff
changeset
|
2664 |
| CNF_UEQ => filter_cnf_ueq_problem |
44589
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents:
44587
diff
changeset
|
2665 |
| FOF => I |
44754 | 2666 |
| TFF (_, TPTP_Implicit) => I |
2667 |
| 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
|
2668 |
| _ => 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
|
2669 |
val (problem, pool) = problem |> nice_atp_problem readable_names format |
44772 | 2670 |
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
|
2671 |
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
|
2672 |
in |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
2673 |
(problem, |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
2674 |
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
|
2675 |
fact_names |> Vector.fromList, |
45508 | 2676 |
lifted, |
44772 | 2677 |
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
|
2678 |
end |
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
2679 |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
2680 |
(* FUDGE *) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
2681 |
val conj_weight = 0.0 |
41770 | 2682 |
val hyp_weight = 0.1 |
2683 |
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
|
2684 |
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
|
2685 |
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
|
2686 |
|
46442
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2687 |
(* 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
|
2688 |
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
|
2689 |
let |
46442
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2690 |
fun add_term_weights weight (ATerm (s, tms)) = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2691 |
is_tptp_user_symbol s ? Symtab.default (s, weight) |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2692 |
#> fold (add_term_weights weight) tms |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2693 |
| 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
|
2694 |
fun add_line_weights weight (Formula (_, _, phi, _, _)) = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2695 |
formula_fold NONE (K (add_term_weights weight)) phi |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2696 |
| add_line_weights _ _ = I |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2697 |
fun add_conjectures_weights [] = I |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2698 |
| add_conjectures_weights conjs = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2699 |
let val (hyps, conj) = split_last conjs in |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2700 |
add_line_weights conj_weight conj |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2701 |
#> fold (add_line_weights hyp_weight) hyps |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2702 |
end |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2703 |
fun add_facts_weights facts = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2704 |
let |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2705 |
val num_facts = length facts |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2706 |
fun weight_of j = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2707 |
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
|
2708 |
/ Real.fromInt num_facts |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2709 |
in |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2710 |
map weight_of (0 upto num_facts - 1) ~~ facts |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2711 |
|> fold (uncurry add_line_weights) |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2712 |
end |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2713 |
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
|
2714 |
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
|
2715 |
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
|
2716 |
|> 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
|
2717 |
|> add_facts_weights (get factsN) |
46442
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2718 |
|> 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
|
2719 |
[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
|
2720 |
|> 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
|
2721 |
|> 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
|
2722 |
end |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41211
diff
changeset
|
2723 |
|
46443
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2724 |
fun have_head_rolling (ATerm (s, args)) = |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2725 |
(* ugly hack: may make innocent victims *) |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2726 |
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
|
2727 |
have_head_rolling (hd args) ||> append (tl args) |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2728 |
else |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2729 |
(s, args) |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2730 |
| have_head_rolling _ = ("", []) |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2731 |
|
46450 | 2732 |
val max_kbo_weight = 2147483647 |
2733 |
||
46442
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2734 |
fun atp_problem_kbo_weights problem = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2735 |
let |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2736 |
fun add_term_deps head (ATerm (s, args)) = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2737 |
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
|
2738 |
#> fold (add_term_deps head) args |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2739 |
| 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
|
2740 |
fun add_eq_deps (ATerm (s, [lhs as _, rhs])) = |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2741 |
if is_tptp_equal s then |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2742 |
let val (head, rest) = have_head_rolling lhs in |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2743 |
fold (add_term_deps head) (rhs :: rest) |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2744 |
end |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2745 |
else |
c86276014571
improved KBO weights -- beware of explicit applications
blanchet
parents:
46442
diff
changeset
|
2746 |
I |
46442
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2747 |
| add_eq_deps _ = I |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2748 |
fun add_line_deps _ (Decl (_, s, ty)) = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2749 |
is_function_type ty ? Graph.default_node (s, ()) |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2750 |
| add_line_deps status (Formula (_, _, phi, _, info)) = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2751 |
if extract_isabelle_status info = SOME status then |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2752 |
formula_fold NONE (K add_eq_deps) phi |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2753 |
else |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2754 |
I |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2755 |
val graph = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2756 |
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
|
2757 |
|> fold (fold (add_line_deps simpN) o snd) problem |
46450 | 2758 |
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
|
2759 |
fun add_weights _ [] = I |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2760 |
| add_weights weight syms = |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2761 |
fold (AList.update (op =) o rpair weight) syms |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2762 |
#> add_weights (next_weight weight) |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2763 |
(fold (union (op =) o Graph.immediate_succs graph) syms []) |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2764 |
in |
46446 | 2765 |
[] |> 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
|
2766 |
end |
1e07620d724c
added possibility of generating KBO weights to DFG problems
blanchet
parents:
46437
diff
changeset
|
2767 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
diff
changeset
|
2768 |
end; |