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