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