| author | wenzelm | 
| Sun, 09 Jun 2024 15:11:07 +0200 | |
| changeset 80305 | 95b51df1382e | 
| parent 79734 | 0fa4bebbdd75 | 
| child 81254 | d3c0734059ee | 
| permissions | -rw-r--r-- | 
| 46320 | 1 | (* Title: HOL/Tools/ATP/atp_proof_reconstruct.ML | 
| 38027 | 2 | Author: Lawrence C. Paulson, Cambridge University Computer Laboratory | 
| 3 | Author: Claire Quigley, Cambridge University Computer Laboratory | |
| 36392 | 4 | Author: Jasmin Blanchette, TU Muenchen | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 5 | |
| 49914 | 6 | Basic proof reconstruction from ATP proofs. | 
| 33310 | 7 | *) | 
| 8 | ||
| 46320 | 9 | signature ATP_PROOF_RECONSTRUCT = | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 10 | sig | 
| 54811 | 11 | type 'a atp_type = 'a ATP_Problem.atp_type | 
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
52758diff
changeset | 12 |   type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term
 | 
| 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
52758diff
changeset | 13 |   type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula
 | 
| 54500 | 14 | type stature = ATP_Problem_Generate.stature | 
| 54772 | 15 | type atp_step_name = ATP_Proof.atp_step_name | 
| 54499 | 16 |   type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
 | 
| 17 | type 'a atp_proof = 'a ATP_Proof.atp_proof | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 18 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 19 | val metisN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 20 | val full_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 21 | val partial_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 22 | val no_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 23 | val really_full_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 24 | val full_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 25 | val partial_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 26 | val no_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 27 | val full_type_encs : string list | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 28 | val partial_type_encs : string list | 
| 54500 | 29 | val default_metis_lam_trans : string | 
| 54772 | 30 | |
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 31 | val leo2_extcnf_equal_neg_rule : string | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 32 | val satallax_tab_rule_prefix : string | 
| 79734 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 33 | val zipperposition_cnf_rule : string | 
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 34 | |
| 49983 
33e18e9916a8
use metaquantification when possible in Isar proofs
 blanchet parents: 
49914diff
changeset | 35 | val forall_of : term -> term -> term | 
| 
33e18e9916a8
use metaquantification when possible in Isar proofs
 blanchet parents: 
49914diff
changeset | 36 | val exists_of : term -> term -> term | 
| 57767 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 37 | val simplify_bool : term -> term | 
| 57792 | 38 | val var_name_of_typ : typ -> string | 
| 57765 | 39 | val rename_bound_vars : term -> term | 
| 55285 | 40 | val type_enc_aliases : (string * string list) list | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 41 | val unalias_type_enc : string -> string list | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 42 | val term_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc -> | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 43 | bool -> int Symtab.table -> typ option -> (string, string atp_type) atp_term -> term | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 44 | val prop_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc -> | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 45 | bool -> int Symtab.table -> | 
| 54811 | 46 | (string, string, (string, string atp_type) atp_term, string) atp_formula -> term | 
| 54499 | 47 | |
| 77432 | 48 | val is_conjecture_used_in_proof : string atp_proof -> bool | 
| 57263 | 49 | val used_facts_in_atp_proof : Proof.context -> (string * stature) list -> string atp_proof -> | 
| 50 | (string * stature) list | |
| 55257 | 51 | val atp_proof_prefers_lifting : string atp_proof -> bool | 
| 54500 | 52 | val is_typed_helper_used_in_atp_proof : string atp_proof -> bool | 
| 54772 | 53 |   val replace_dependencies_in_line : atp_step_name * atp_step_name list -> ('a, 'b) atp_step ->
 | 
| 54 |     ('a, 'b) atp_step
 | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 55 | val termify_atp_proof : Proof.context -> string -> ATP_Problem.atp_format -> | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 56 | ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list -> | 
| 54499 | 57 | int Symtab.table -> string atp_proof -> (term, string) atp_step list | 
| 57267 
8b87114357bd
integrated new Waldmeister code with 'sledgehammer' command
 blanchet parents: 
57263diff
changeset | 58 | val repair_waldmeister_endgame : (term, 'a) atp_step list -> (term, 'a) atp_step list | 
| 57699 | 59 | val infer_formulas_types : Proof.context -> term list -> term list | 
| 72401 
2783779b7dd3
removed obsolete unmaintained experimental prover Pirate
 blanchet parents: 
70940diff
changeset | 60 | val introduce_spass_skolems : (term, string) atp_step list -> (term, string) atp_step list | 
| 79734 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 61 | val regroup_zipperposition_skolems : (term, string) atp_step list -> (term, string) atp_step list | 
| 57263 | 62 | val factify_atp_proof : (string * 'a) list -> term list -> term -> (term, string) atp_step list -> | 
| 63 | (term, string) atp_step list | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 64 | val termify_atp_abduce_candidate : Proof.context -> string -> ATP_Problem.atp_format -> | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 65 | ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list -> | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 66 | int Symtab.table -> (string, string, (string, string atp_type) atp_term, string) atp_formula -> | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 67 | term | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 68 | val top_abduce_candidates : int -> term list -> term list | 
| 77602 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 69 | val sort_propositions_by_provability : Proof.context -> term list -> term list | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 70 | end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 71 | |
| 46320 | 72 | structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 73 | struct | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 74 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 75 | open ATP_Util | 
| 38028 | 76 | open ATP_Problem | 
| 39452 | 77 | open ATP_Proof | 
| 46320 | 78 | open ATP_Problem_Generate | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 79 | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 80 | val metisN = "metis" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 81 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 82 | val full_typesN = "full_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 83 | val partial_typesN = "partial_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 84 | val no_typesN = "no_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 85 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 86 | val really_full_type_enc = "mono_tags" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 87 | val full_type_enc = "poly_guards_query" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 88 | val partial_type_enc = "poly_args" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 89 | val no_type_enc = "erased" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 90 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 91 | val full_type_encs = [full_type_enc, really_full_type_enc] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 92 | val partial_type_encs = partial_type_enc :: full_type_encs | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 93 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 94 | val type_enc_aliases = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 95 | [(full_typesN, full_type_encs), | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 96 | (partial_typesN, partial_type_encs), | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 97 | (no_typesN, [no_type_enc])] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 98 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 99 | fun unalias_type_enc s = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 100 | AList.lookup (op =) type_enc_aliases s |> the_default [s] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 101 | |
| 54500 | 102 | val default_metis_lam_trans = combsN | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 103 | |
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 104 | val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg" | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 105 | val satallax_tab_rule_prefix = "tab_" | 
| 79734 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 106 | val zipperposition_cnf_rule = "cnf" | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 107 | val pseudo_zipperposition_conjunct = "conjunct" | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 108 | |
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 109 | val pseudo_zipperposition_skolemization_suffix = ".skolem" | 
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 110 | |
| 50670 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
49983diff
changeset | 111 | fun term_name' (Var ((s, _), _)) = perhaps (try Name.dest_skolem) s | 
| 52034 
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
 blanchet parents: 
52031diff
changeset | 112 | | term_name' _ = "" | 
| 50670 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
49983diff
changeset | 113 | |
| 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
49983diff
changeset | 114 | fun lambda' v = Term.lambda_name (term_name' v, v) | 
| 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
49983diff
changeset | 115 | |
| 78693 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 116 | fun If_const T = Const (\<^const_name>\<open>If\<close>, HOLogic.boolT --> T --> T --> T) | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 117 | |
| 50670 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
49983diff
changeset | 118 | fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda' v t | 
| 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
49983diff
changeset | 119 | fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda' v t | 
| 39425 | 120 | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 121 | fun make_tfree ctxt w = | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 122 | let val ww = "'" ^ w in | 
| 69593 | 123 | TFree (ww, the_default \<^sort>\<open>type\<close> (Variable.def_sort ctxt (ww, ~1))) | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 124 | end | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 125 | |
| 69593 | 126 | fun simplify_bool ((all as Const (\<^const_name>\<open>All\<close>, _)) $ Abs (s, T, t)) = | 
| 57767 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 127 | let val t' = simplify_bool t in | 
| 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 128 | if loose_bvar1 (t', 0) then all $ Abs (s, T, t') else t' | 
| 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 129 | end | 
| 69593 | 130 | | simplify_bool (Const (\<^const_name>\<open>Not\<close>, _) $ t) = s_not (simplify_bool t) | 
| 131 | | simplify_bool (Const (\<^const_name>\<open>conj\<close>, _) $ t $ u) = | |
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 132 | s_conj (simplify_bool t, simplify_bool u) | 
| 69593 | 133 | | simplify_bool (Const (\<^const_name>\<open>disj\<close>, _) $ t $ u) = | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 134 | s_disj (simplify_bool t, simplify_bool u) | 
| 69593 | 135 | | simplify_bool (Const (\<^const_name>\<open>implies\<close>, _) $ t $ u) = | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 136 | s_imp (simplify_bool t, simplify_bool u) | 
| 69593 | 137 | | simplify_bool ((t as Const (\<^const_name>\<open>HOL.eq\<close>, _)) $ u $ v) = | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 138 | (case (u, v) of | 
| 69593 | 139 | (Const (\<^const_name>\<open>True\<close>, _), _) => v | 
| 140 | | (u, Const (\<^const_name>\<open>True\<close>, _)) => u | |
| 141 | | (Const (\<^const_name>\<open>False\<close>, _), v) => s_not v | |
| 142 | | (u, Const (\<^const_name>\<open>False\<close>, _)) => s_not u | |
| 74379 | 143 | | _ => if u aconv v then \<^Const>\<open>True\<close> else t $ simplify_bool u $ simplify_bool v) | 
| 57767 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 144 | | simplify_bool (t $ u) = simplify_bool t $ simplify_bool u | 
| 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 145 | | simplify_bool (Abs (s, T, t)) = Abs (s, T, simplify_bool t) | 
| 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 146 | | simplify_bool t = t | 
| 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 147 | |
| 57765 | 148 | fun single_letter upper s = | 
| 149 | let val s' = if String.isPrefix "o" s orelse String.isPrefix "O" s then "z" else s in | |
| 150 | String.extract (Name.desymbolize (SOME upper) (Long_Name.base_name s'), 0, SOME 1) | |
| 151 | end | |
| 152 | ||
| 69593 | 153 | fun var_name_of_typ (Type (\<^type_name>\<open>fun\<close>, [_, T])) = | 
| 57765 | 154 | if body_type T = HOLogic.boolT then "p" else "f" | 
| 69593 | 155 | | var_name_of_typ (Type (\<^type_name>\<open>set\<close>, [T])) = | 
| 58600 | 156 | let fun default () = single_letter true (var_name_of_typ T) in | 
| 58599 | 157 | (case T of | 
| 158 | Type (s, [T1, T2]) => if String.isSuffix "prod" s andalso T1 = T2 then "r" else default () | |
| 159 | | _ => default ()) | |
| 160 | end | |
| 57765 | 161 | | var_name_of_typ (Type (s, Ts)) = | 
| 162 | if String.isSuffix "list" s then var_name_of_typ (the_single Ts) ^ "s" | |
| 163 | else single_letter false (Long_Name.base_name s) | |
| 164 | | var_name_of_typ (TFree (s, _)) = single_letter false (perhaps (try (unprefix "'")) s) | |
| 165 | | var_name_of_typ (TVar ((s, _), T)) = var_name_of_typ (TFree (s, T)) | |
| 166 | ||
| 167 | fun rename_bound_vars (t $ u) = rename_bound_vars t $ rename_bound_vars u | |
| 168 | | rename_bound_vars (Abs (_, T, t)) = Abs (var_name_of_typ T, T, rename_bound_vars t) | |
| 169 | | rename_bound_vars t = t | |
| 170 | ||
| 54820 | 171 | exception ATP_CLASS of string list | 
| 54818 | 172 | exception ATP_TYPE of string atp_type list | 
| 54811 | 173 | exception ATP_TERM of (string, string atp_type) atp_term list | 
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
52758diff
changeset | 174 | exception ATP_FORMULA of | 
| 54811 | 175 | (string, string, (string, string atp_type) atp_term, string) atp_formula list | 
| 37991 | 176 | exception SAME of unit | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 177 | |
| 54820 | 178 | fun class_of_atp_class cls = | 
| 179 | (case unprefix_and_unascii class_prefix cls of | |
| 180 | SOME s => s | |
| 181 | | NONE => raise ATP_CLASS [cls]) | |
| 182 | ||
| 57697 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 183 | fun atp_type_of_atp_term (ATerm ((s, _), us)) = | 
| 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 184 | let val tys = map atp_type_of_atp_term us in | 
| 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 185 | if s = tptp_fun_type then | 
| 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 186 | (case tys of | 
| 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 187 | [ty1, ty2] => AFun (ty1, ty2) | 
| 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 188 | | _ => raise ATP_TYPE tys) | 
| 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 189 | else | 
| 
44341963ade3
correctly translate THF functions from terms to types
 blanchet parents: 
57635diff
changeset | 190 | AType ((s, []), tys) | 
| 37991 | 191 | end | 
| 54818 | 192 | |
| 58477 | 193 | (* Type variables are given the basic sort "HOL.type". Some will later be constrained by information | 
| 194 | from type literals, or by type inference. *) | |
| 195 | fun typ_of_atp_type ctxt (ty as AType ((a, clss), tys)) = | |
| 196 | let val Ts = map (typ_of_atp_type ctxt) tys in | |
| 197 | (case unprefix_and_unascii native_type_prefix a of | |
| 198 | SOME b => typ_of_atp_type ctxt (atp_type_of_atp_term (unmangled_type b)) | |
| 199 | | NONE => | |
| 200 | (case unprefix_and_unascii type_const_prefix a of | |
| 201 | SOME b => Type (invert_const b, Ts) | |
| 202 | | NONE => | |
| 203 | if not (null tys) then | |
| 204 | raise ATP_TYPE [ty] (* only "tconst"s have type arguments *) | |
| 205 | else | |
| 206 | (case unprefix_and_unascii tfree_prefix a of | |
| 207 | SOME b => make_tfree ctxt b | |
| 208 | | NONE => | |
| 209 | (* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018". | |
| 210 | Sometimes variables from the ATP are indistinguishable from Isabelle variables, which | |
| 211 | forces us to use a type parameter in all cases. *) | |
| 212 |               Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a,
 | |
| 69593 | 213 | (if null clss then \<^sort>\<open>type\<close> else map class_of_atp_class clss))))) | 
| 58477 | 214 | end | 
| 215 | | typ_of_atp_type ctxt (AFun (ty1, ty2)) = typ_of_atp_type ctxt ty1 --> typ_of_atp_type ctxt ty2 | |
| 216 | ||
| 54818 | 217 | fun typ_of_atp_term ctxt = typ_of_atp_type ctxt o atp_type_of_atp_term | 
| 218 | ||
| 54789 | 219 | (* Type class literal applied to a type. Returns triple of polarity, class, type. *) | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51878diff
changeset | 220 | fun type_constraint_of_term ctxt (u as ATerm ((a, _), us)) = | 
| 54818 | 221 | (case (unprefix_and_unascii class_prefix a, map (typ_of_atp_term ctxt) us) of | 
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 222 | (SOME b, [T]) => (b, T) | 
| 54789 | 223 | | _ => raise ATP_TERM [u]) | 
| 38014 | 224 | |
| 43907 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 225 | (* Accumulate type constraints in a formula: negative type literals. *) | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 226 | fun add_var (key, z) = Vartab.map_default (key, []) (cons z) | 
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 227 | fun add_type_constraint false (cl, TFree (a ,_)) = add_var ((a, ~1), cl) | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 228 | | add_type_constraint false (cl, TVar (ix, _)) = add_var (ix, cl) | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 229 | | add_type_constraint _ _ = I | 
| 38014 | 230 | |
| 57789 
a73255cffb5b
don't rename variables which will be renamed later anyway
 blanchet parents: 
57788diff
changeset | 231 | fun repair_var_name s = | 
| 55185 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 blanchet parents: 
54822diff
changeset | 232 | (case unprefix_and_unascii schematic_var_prefix s of | 
| 57789 
a73255cffb5b
don't rename variables which will be renamed later anyway
 blanchet parents: 
57788diff
changeset | 233 | SOME s' => s' | 
| 
a73255cffb5b
don't rename variables which will be renamed later anyway
 blanchet parents: 
57788diff
changeset | 234 | | NONE => s) | 
| 55185 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 blanchet parents: 
54822diff
changeset | 235 | |
| 54789 | 236 | (* The number of type arguments of a constant, zero if it's monomorphic. For (instances of) Skolem | 
| 237 | pseudoconstants, this information is encoded in the constant name. *) | |
| 52125 
ac7830871177
improved handling of free variables' types in Isar proofs
 blanchet parents: 
52034diff
changeset | 238 | fun robust_const_num_type_args thy s = | 
| 43907 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 239 | if String.isPrefix skolem_const_prefix s then | 
| 46711 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 wenzelm parents: 
46435diff
changeset | 240 | s |> Long_Name.explode |> List.last |> Int.fromString |> the | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 241 | 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: 
45553diff
changeset | 242 | if String.isPrefix lam_lifted_poly_prefix s then 2 else 0 | 
| 43907 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 243 | else | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 244 | (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 245 | |
| 69593 | 246 | fun slack_fastype_of t = fastype_of t handle TERM _ => Type_Infer.anyT \<^sort>\<open>type\<close> | 
| 43182 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 247 | |
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 248 | val spass_skolem_prefixes = ["skc", "skf"] | 
| 79734 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 249 | val zipperposition_skolem_prefix = "sk" | 
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 250 | |
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 251 | fun is_spass_skolem s = | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 252 | exists (fn pre => String.isPrefix pre s) spass_skolem_prefixes | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 253 | |
| 79734 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 254 | val is_zipperposition_skolem = String.isPrefix zipperposition_skolem_prefix | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 255 | |
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 256 | fun is_reverse_skolem s = | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 257 | exists (fn pre => String.isPrefix pre s) (zipperposition_skolem_prefix :: spass_skolem_prefixes) | 
| 50703 
76a2e506c125
swap Vampire's Skolem arguments to bring them in line with what E and metis's new skolemizer do (helps Isar proof reconstruction in some cases)
 blanchet parents: 
50693diff
changeset | 258 | |
| 78693 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 259 | val zip_internal_ite_prefix = "zip_internal_ite_" | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 260 | |
| 57257 | 261 | fun var_index_of_textual textual = if textual then 0 else 1 | 
| 262 | ||
| 263 | fun quantify_over_var textual quant_of var_s var_T t = | |
| 264 | let | |
| 57790 | 265 | val vars = | 
| 266 | ((var_s, var_index_of_textual textual), var_T) :: | |
| 267 | filter (fn ((s, _), _) => s = var_s) (Term.add_vars t []) | |
| 57257 | 268 | val normTs = vars |> AList.group (op =) |> map (apsnd hd) | 
| 269 | fun norm_var_types (Var (x, T)) = | |
| 270 | Var (x, the_default T (AList.lookup (op =) normTs x)) | |
| 271 | | norm_var_types t = t | |
| 272 | in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end | |
| 273 | ||
| 78697 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 274 | (* This assumes that distinct names are mapped to distinct names by | 
| 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 275 | "Variable.variant_frees". This does not hold in general but should hold for | 
| 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 276 | ATP-generated Skolem function names, since these end with a digit and | 
| 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 277 | "variant_frees" appends letters. *) | 
| 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 278 | fun desymbolize_and_fresh_up ctxt s = | 
| 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 279 | fst (hd (Variable.variant_frees ctxt [] [(Name.desymbolize (SOME false) s, ())])) | 
| 57790 | 280 | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 281 | (* Higher-order translation. Variables are typed (although we don't use that information). Lambdas | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 282 | are typed. The code is similar to "term_of_atp_fo". *) | 
| 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 283 | fun term_of_atp_ho ctxt sym_tab = | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 284 | let | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 285 | val thy = Proof_Context.theory_of ctxt | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 286 | val var_index = var_index_of_textual true | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 287 | |
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 288 | fun do_term opt_T u = | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 289 | (case u of | 
| 78696 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 blanchet parents: 
78693diff
changeset | 290 | AAbs (((var, ty), term), us) => | 
| 57257 | 291 | let | 
| 292 | val typ = typ_of_atp_type ctxt ty | |
| 57789 
a73255cffb5b
don't rename variables which will be renamed later anyway
 blanchet parents: 
57788diff
changeset | 293 | val var_name = repair_var_name var | 
| 57257 | 294 | val tm = do_term NONE term | 
| 78696 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 blanchet parents: 
78693diff
changeset | 295 | val lam = quantify_over_var true lambda' var_name typ tm | 
| 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 blanchet parents: 
78693diff
changeset | 296 | val args = map (do_term NONE) us | 
| 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 blanchet parents: 
78693diff
changeset | 297 | in | 
| 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 blanchet parents: 
78693diff
changeset | 298 | list_comb (lam, args) | 
| 
ef89f1beee95
parse applie lambdas correctly plus deal gracefully with lambda-lifting in Zipperposition
 blanchet parents: 
78693diff
changeset | 299 | end | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 300 | | ATerm ((s, tys), us) => | 
| 74050 | 301 | if s = "" (* special marker generated on parse error *) then | 
| 302 | error "Isar proof reconstruction failed because the ATP proof contains unparsable \ | |
| 303 | \material" | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 304 | else if s = tptp_equal then | 
| 69593 | 305 | list_comb (Const (\<^const_name>\<open>HOL.eq\<close>, Type_Infer.anyT \<^sort>\<open>type\<close>), | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 306 | map (do_term NONE) us) | 
| 75052 | 307 | else if s = tptp_not_equal andalso length us = 2 then | 
| 75048 
e926618f9474
handle TPTP '!=' more gracefully in Isar proof reconstruction
 blanchet parents: 
74379diff
changeset | 308 | \<^const>\<open>HOL.Not\<close> $ list_comb (Const (\<^const_name>\<open>HOL.eq\<close>, Type_Infer.anyT \<^sort>\<open>type\<close>), | 
| 
e926618f9474
handle TPTP '!=' more gracefully in Isar proof reconstruction
 blanchet parents: 
74379diff
changeset | 309 | map (do_term NONE) us) | 
| 57256 | 310 | else if not (null us) then | 
| 311 | let | |
| 57820 | 312 | val args = map (do_term NONE) us | 
| 57256 | 313 | val opt_T' = SOME (map slack_fastype_of args ---> the_default dummyT opt_T) | 
| 314 | val func = do_term opt_T' (ATerm ((s, tys), [])) | |
| 315 | in foldl1 (op $) (func :: args) end | |
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 316 | else if s = tptp_or then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 317 | HOLogic.disj | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 318 | else if s = tptp_and then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 319 | HOLogic.conj | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 320 | else if s = tptp_implies then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 321 | HOLogic.imp | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 322 | else if s = tptp_iff orelse s = tptp_equal then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 323 | HOLogic.eq_const dummyT | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 324 | else if s = tptp_not_iff orelse s = tptp_not_equal then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 325 | \<^term>\<open>\<lambda>x y. x \<noteq> y\<close> | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 326 | else if s = tptp_if then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 327 | \<^term>\<open>\<lambda>P Q. Q \<longrightarrow> P\<close> | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 328 | else if s = tptp_not_and then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 329 | \<^term>\<open>\<lambda>P Q. \<not> (P \<and> Q)\<close> | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 330 | else if s = tptp_not_or then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 331 | \<^term>\<open>\<lambda>P Q. \<not> (P \<or> Q)\<close> | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 332 | else if s = tptp_not then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 333 | HOLogic.Not | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 334 | else if s = tptp_ho_forall then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 335 | HOLogic.all_const dummyT | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 336 | else if s = tptp_ho_exists then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 337 | HOLogic.exists_const dummyT | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 338 | else if s = tptp_hilbert_choice then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 339 | HOLogic.choice_const dummyT | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 340 | else if s = tptp_hilbert_the then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 341 | \<^term>\<open>The\<close> | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 342 | else if String.isPrefix zip_internal_ite_prefix s then | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 343 | If_const dummyT | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 344 | else | 
| 57256 | 345 | (case unprefix_and_unascii const_prefix s of | 
| 346 | SOME s' => | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 347 | let | 
| 57547 
677b07d777c3
don't generate TPTP THF 'Definition's, because they complicate reconstruction for AgsyHOL and Satallax
 blanchet parents: 
57267diff
changeset | 348 | val ((s', _), mangled_us) = s' |> unmangled_const |>> `invert_const | 
| 57256 | 349 | val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) | 
| 350 | val (type_us, term_us) = chop num_ty_args us |>> append mangled_us | |
| 351 | val term_ts = map (do_term NONE) term_us | |
| 352 | val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us | |
| 353 | val T = | |
| 354 | (if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then | |
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 355 | try (Sign.const_instance thy) (s', Ts) | 
| 57256 | 356 | else | 
| 357 | NONE) | |
| 358 | |> (fn SOME T => T | |
| 359 | | NONE => | |
| 360 | map slack_fastype_of term_ts ---> | |
| 69593 | 361 | the_default (Type_Infer.anyT \<^sort>\<open>type\<close>) opt_T) | 
| 57256 | 362 | val t = Const (unproxify_const s', T) | 
| 363 | in list_comb (t, term_ts) end | |
| 364 | | NONE => (* a free or schematic variable *) | |
| 365 | let | |
| 366 | val ts = map (do_term NONE) us | |
| 367 | val T = | |
| 58500 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 368 | (case tys of | 
| 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 369 | [ty] => typ_of_atp_type ctxt ty | 
| 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 370 | | _ => | 
| 57256 | 371 | map slack_fastype_of ts ---> | 
| 58500 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 372 | (case opt_T of | 
| 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 373 | SOME T => T | 
| 69593 | 374 | | NONE => Type_Infer.anyT \<^sort>\<open>type\<close>)) | 
| 57256 | 375 | val t = | 
| 376 | (case unprefix_and_unascii fixed_var_prefix s of | |
| 377 | SOME s => Free (s, T) | |
| 378 | | NONE => | |
| 78697 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 379 | if not (is_tptp_variable s) then Free (desymbolize_and_fresh_up ctxt s, T) | 
| 57789 
a73255cffb5b
don't rename variables which will be renamed later anyway
 blanchet parents: 
57788diff
changeset | 380 | else Var ((repair_var_name s, var_index), T)) | 
| 57256 | 381 | in list_comb (t, ts) end)) | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 382 | in do_term end | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 383 | |
| 56254 | 384 | (* First-order translation. No types are known for variables. "Type_Infer.anyT @{sort type}"
 | 
| 385 | should allow them to be inferred. *) | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 386 | fun term_of_atp_fo ctxt textual sym_tab = | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 387 | let | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 388 | val thy = Proof_Context.theory_of ctxt | 
| 54789 | 389 | (* For Metis, we use 1 rather than 0 because variable references in clauses may otherwise | 
| 390 | conflict with variable constraints in the goal. At least, type inference often fails | |
| 391 | otherwise. See also "axiom_inference" in "Metis_Reconstruct". *) | |
| 57257 | 392 | val var_index = var_index_of_textual textual | 
| 57199 | 393 | |
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 394 | fun do_term extra_ts opt_T u = | 
| 54789 | 395 | (case u of | 
| 54818 | 396 | ATerm ((s, tys), us) => | 
| 74050 | 397 | if s = "" (* special marker generated on parse error *) then | 
| 57199 | 398 | error "Isar proof reconstruction failed because the ATP proof contains unparsable \ | 
| 63692 | 399 | \material" | 
| 51878 | 400 | else if String.isPrefix native_type_prefix s then | 
| 74379 | 401 | \<^Const>\<open>True\<close> (* ignore TPTP type information (needed?) *) | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 402 | else if s = tptp_equal then | 
| 69593 | 403 | list_comb (Const (\<^const_name>\<open>HOL.eq\<close>, Type_Infer.anyT \<^sort>\<open>type\<close>), | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 404 | map (do_term [] NONE) us) | 
| 54789 | 405 | else | 
| 406 | (case unprefix_and_unascii const_prefix s of | |
| 407 | SOME s' => | |
| 54818 | 408 | let val ((s', s''), mangled_us) = s' |> unmangled_const |>> `invert_const in | 
| 54789 | 409 | if s' = type_tag_name then | 
| 410 | (case mangled_us @ us of | |
| 54818 | 411 | [typ_u, term_u] => do_term extra_ts (SOME (typ_of_atp_term ctxt typ_u)) term_u | 
| 54789 | 412 | | _ => raise ATP_TERM us) | 
| 413 | else if s' = predicator_name then | |
| 69593 | 414 | do_term [] (SOME \<^typ>\<open>bool\<close>) (hd us) | 
| 54789 | 415 | else if s' = app_op_name then | 
| 416 | let val extra_t = do_term [] NONE (List.last us) in | |
| 417 | do_term (extra_t :: extra_ts) | |
| 57199 | 418 | (case opt_T of SOME T => SOME (slack_fastype_of extra_t --> T) | NONE => NONE) | 
| 419 | (nth us (length us - 2)) | |
| 54789 | 420 | end | 
| 421 | else if s' = type_guard_name then | |
| 74379 | 422 | \<^Const>\<open>True\<close> (* ignore type predicates *) | 
| 54789 | 423 | else | 
| 424 | let | |
| 425 | val new_skolem = String.isPrefix new_skolem_const_prefix s'' | |
| 54818 | 426 | val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) | 
| 427 | val (type_us, term_us) = chop num_ty_args us |>> append mangled_us | |
| 54789 | 428 | val term_ts = map (do_term [] NONE) term_us | 
| 54818 | 429 | |
| 430 | val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us | |
| 54789 | 431 | val T = | 
| 54818 | 432 | (if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then | 
| 57199 | 433 | if new_skolem then SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) | 
| 434 | else if textual then try (Sign.const_instance thy) (s', Ts) | |
| 435 | else NONE | |
| 54789 | 436 | else | 
| 437 | NONE) | |
| 438 | |> (fn SOME T => T | |
| 439 | | NONE => | |
| 56254 | 440 | map slack_fastype_of term_ts ---> | 
| 69593 | 441 | the_default (Type_Infer.anyT \<^sort>\<open>type\<close>) opt_T) | 
| 54789 | 442 | val t = | 
| 443 | if new_skolem then Var ((new_skolem_var_name_of_const s'', var_index), T) | |
| 444 | else Const (unproxify_const s', T) | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 445 | in | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 446 | list_comb (t, term_ts @ extra_ts) | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 447 | end | 
| 54789 | 448 | end | 
| 449 | | NONE => (* a free or schematic variable *) | |
| 450 | let | |
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 451 | val term_ts = map (do_term [] NONE) us |> is_reverse_skolem s ? rev | 
| 54789 | 452 | val ts = term_ts @ extra_ts | 
| 453 | val T = | |
| 58500 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 454 | (case tys of | 
| 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 455 | [ty] => typ_of_atp_type ctxt ty | 
| 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 456 | | _ => | 
| 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 457 | (case opt_T of | 
| 
430306aa03b1
proper types for applied variables, for typed formats (TFF0, DFG)
 blanchet parents: 
58484diff
changeset | 458 | SOME T => map slack_fastype_of term_ts ---> T | 
| 69593 | 459 | | NONE => map slack_fastype_of ts ---> Type_Infer.anyT \<^sort>\<open>type\<close>)) | 
| 54789 | 460 | val t = | 
| 461 | (case unprefix_and_unascii fixed_var_prefix s of | |
| 462 | SOME s => Free (s, T) | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 463 | | NONE => | 
| 55185 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 blanchet parents: 
54822diff
changeset | 464 | if textual andalso not (is_tptp_variable s) then | 
| 78697 
8ca71c0ae31f
avoid legacy binding errors in Sledgehammer Isar proofs
 blanchet parents: 
78696diff
changeset | 465 | Free (desymbolize_and_fresh_up ctxt s, T) | 
| 55185 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 blanchet parents: 
54822diff
changeset | 466 | else | 
| 57789 
a73255cffb5b
don't rename variables which will be renamed later anyway
 blanchet parents: 
57788diff
changeset | 467 | Var ((repair_var_name s, var_index), T)) | 
| 54789 | 468 | in list_comb (t, ts) end)) | 
| 43093 | 469 | in do_term [] end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 470 | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 471 | fun term_of_atp ctxt (ATP_Problem.THF _) type_enc = | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 472 | if ATP_Problem_Generate.is_type_enc_higher_order type_enc then K (term_of_atp_ho ctxt) | 
| 63692 | 473 | else error "Unsupported Isar reconstruction" | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 474 | | term_of_atp ctxt _ type_enc = | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 475 | if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc) then term_of_atp_fo ctxt | 
| 63692 | 476 | else error "Unsupported Isar reconstruction" | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 477 | |
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 478 | fun term_of_atom ctxt format type_enc textual sym_tab pos (u as ATerm ((s, _), _)) = | 
| 38014 | 479 | if String.isPrefix class_prefix s then | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51878diff
changeset | 480 | add_type_constraint pos (type_constraint_of_term ctxt u) | 
| 74379 | 481 | #> pair \<^Const>\<open>True\<close> | 
| 38014 | 482 | else | 
| 69593 | 483 | pair (term_of_atp ctxt format type_enc textual sym_tab (SOME \<^typ>\<open>bool\<close>) u) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 484 | |
| 37991 | 485 | (* Update schematic type variables with detected sort constraints. It's not | 
| 42563 | 486 | totally clear whether this code is necessary. *) | 
| 38014 | 487 | fun repair_tvar_sorts (t, tvar_tab) = | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 488 | let | 
| 37991 | 489 | fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) | 
| 490 | | do_type (TVar (xi, s)) = | |
| 491 | TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) | |
| 492 | | do_type (TFree z) = TFree z | |
| 493 | fun do_term (Const (a, T)) = Const (a, do_type T) | |
| 494 | | do_term (Free (a, T)) = Free (a, do_type T) | |
| 495 | | do_term (Var (xi, T)) = Var (xi, do_type T) | |
| 496 | | do_term (t as Bound _) = t | |
| 497 | | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) | |
| 498 | | do_term (t1 $ t2) = do_term t1 $ do_term t2 | |
| 499 | in t |> not (Vartab.is_empty tvar_tab) ? do_term end | |
| 500 | ||
| 54811 | 501 | (* Interpret an ATP formula as a HOL term, extracting sort constraints as they appear in the | 
| 502 | formula. *) | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 503 | fun prop_of_atp ctxt format type_enc textual sym_tab phi = | 
| 38014 | 504 | let | 
| 505 | fun do_formula pos phi = | |
| 54789 | 506 | (case phi of | 
| 38014 | 507 | AQuant (_, [], phi) => do_formula pos phi | 
| 42526 | 508 | | AQuant (q, (s, _) :: xs, phi') => | 
| 38014 | 509 | do_formula pos (AQuant (q, xs, phi')) | 
| 42526 | 510 | (* FIXME: TFF *) | 
| 57257 | 511 | #>> quantify_over_var textual (case q of AForall => forall_of | AExists => exists_of) | 
| 57789 
a73255cffb5b
don't rename variables which will be renamed later anyway
 blanchet parents: 
57788diff
changeset | 512 | (repair_var_name s) dummyT | 
| 38014 | 513 | | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not | 
| 37991 | 514 | | AConn (c, [phi1, phi2]) => | 
| 38014 | 515 | do_formula (pos |> c = AImplies ? not) phi1 | 
| 516 | ##>> do_formula pos phi2 | |
| 517 | #>> (case c of | |
| 54811 | 518 | AAnd => s_conj | 
| 519 | | AOr => s_disj | |
| 520 | | AImplies => s_imp | |
| 521 | | AIff => s_iff | |
| 522 | | ANot => raise Fail "impossible connective") | |
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 523 | | AAtom tm => term_of_atom ctxt format type_enc textual sym_tab pos tm | 
| 54789 | 524 | | _ => raise ATP_FORMULA [phi]) | 
| 525 | in | |
| 526 | repair_tvar_sorts (do_formula true phi Vartab.empty) | |
| 527 | end | |
| 37991 | 528 | |
| 54500 | 529 | val unprefix_fact_number = space_implode "_" o tl o space_explode "_" | 
| 530 | ||
| 57263 | 531 | fun resolve_fact facts s = | 
| 54789 | 532 | (case try (unprefix fact_prefix) s of | 
| 54500 | 533 | SOME s' => | 
| 534 | let val s' = s' |> unprefix_fact_number |> unascii_of in | |
| 57263 | 535 | AList.lookup (op =) facts s' |> Option.map (pair s') | 
| 54500 | 536 | end | 
| 54789 | 537 | | NONE => NONE) | 
| 54506 | 538 | |
| 57263 | 539 | fun resolve_conjecture s = | 
| 54789 | 540 | (case try (unprefix conjecture_prefix) s of | 
| 54500 | 541 | SOME s' => Int.fromString s' | 
| 54789 | 542 | | NONE => NONE) | 
| 54500 | 543 | |
| 57263 | 544 | fun resolve_facts facts = map_filter (resolve_fact facts) | 
| 545 | val resolve_conjectures = map_filter resolve_conjecture | |
| 54500 | 546 | |
| 547 | fun is_axiom_used_in_proof pred = | |
| 548 | exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false) | |
| 549 | ||
| 77432 | 550 | val is_conjecture_used_in_proof = is_axiom_used_in_proof (is_some o resolve_conjecture) | 
| 551 | ||
| 70940 
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
 blanchet parents: 
69597diff
changeset | 552 | fun add_fact ctxt facts ((num, ss), _, _, rule, deps) = | 
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 553 | (if member (op =) [agsyhol_core_rule, leo2_extcnf_equal_neg_rule] rule orelse | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58652diff
changeset | 554 | String.isPrefix satallax_tab_rule_prefix rule then | 
| 56104 
fd6e132ee4fb
correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
 blanchet parents: 
55285diff
changeset | 555 | insert (op =) (short_thm_name ctxt ext, (Global, General)) | 
| 54500 | 556 | else | 
| 557 | I) | |
| 70940 
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
 blanchet parents: 
69597diff
changeset | 558 | #> (if null deps then union (op =) (resolve_facts facts (num :: ss)) else I) | 
| 54500 | 559 | |
| 57263 | 560 | fun used_facts_in_atp_proof ctxt facts atp_proof = | 
| 561 | if null atp_proof then facts else fold (add_fact ctxt facts) atp_proof [] | |
| 54500 | 562 | |
| 563 | val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix | |
| 564 | ||
| 565 | (* overapproximation (good enough) *) | |
| 566 | fun is_lam_lifted s = | |
| 567 | String.isPrefix fact_prefix s andalso | |
| 568 | String.isSubstring ascii_of_lam_fact_prefix s | |
| 569 | ||
| 570 | val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix) | |
| 571 | ||
| 55257 | 572 | fun atp_proof_prefers_lifting atp_proof = | 
| 573 | (is_axiom_used_in_proof is_combinator_def atp_proof, | |
| 574 | is_axiom_used_in_proof is_lam_lifted atp_proof) = (false, true) | |
| 54500 | 575 | |
| 576 | val is_typed_helper_name = | |
| 577 | String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix | |
| 578 | ||
| 579 | fun is_typed_helper_used_in_atp_proof atp_proof = | |
| 580 | is_axiom_used_in_proof is_typed_helper_name atp_proof | |
| 581 | ||
| 54772 | 582 | fun replace_one_dependency (old, new) dep = if is_same_atp_step dep old then new else [dep] | 
| 583 | fun replace_dependencies_in_line old_new (name, role, t, rule, deps) = | |
| 584 | (name, role, t, rule, fold (union (op =) o replace_one_dependency old_new) deps []) | |
| 585 | ||
| 54499 | 586 | fun repair_name "$true" = "c_True" | 
| 587 | | repair_name "$false" = "c_False" | |
| 588 | | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *) | |
| 589 | | repair_name s = | |
| 590 | if is_tptp_equal s orelse | |
| 591 | (* seen in Vampire proofs *) | |
| 592 | (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then | |
| 593 | tptp_equal | |
| 594 | else | |
| 595 | s | |
| 596 | ||
| 57699 | 597 | fun set_var_index j = map_aterms (fn Var ((s, 0), T) => Var ((s, j), T) | t => t) | 
| 57635 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 blanchet parents: 
57547diff
changeset | 598 | |
| 57699 | 599 | fun infer_formulas_types ctxt = | 
| 57635 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 blanchet parents: 
57547diff
changeset | 600 | map_index (uncurry (fn j => set_var_index j #> Type.constraint HOLogic.boolT)) | 
| 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 blanchet parents: 
57547diff
changeset | 601 | #> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) | 
| 54499 | 602 | |
| 603 | val combinator_table = | |
| 69593 | 604 |   [(\<^const_name>\<open>Meson.COMBI\<close>, @{thm Meson.COMBI_def [abs_def]}),
 | 
| 605 |    (\<^const_name>\<open>Meson.COMBK\<close>, @{thm Meson.COMBK_def [abs_def]}),
 | |
| 606 |    (\<^const_name>\<open>Meson.COMBB\<close>, @{thm Meson.COMBB_def [abs_def]}),
 | |
| 607 |    (\<^const_name>\<open>Meson.COMBC\<close>, @{thm Meson.COMBC_def [abs_def]}),
 | |
| 608 |    (\<^const_name>\<open>Meson.COMBS\<close>, @{thm Meson.COMBS_def [abs_def]})]
 | |
| 54499 | 609 | |
| 610 | fun uncombine_term thy = | |
| 611 | let | |
| 57717 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 612 | fun uncomb (t1 $ t2) = betapply (uncomb t1, uncomb t2) | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 613 | | uncomb (Abs (s, T, t)) = Abs (s, T, uncomb t) | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 614 | | uncomb (t as Const (x as (s, _))) = | 
| 54499 | 615 | (case AList.lookup (op =) combinator_table s of | 
| 59582 | 616 | SOME thm => thm |> Thm.prop_of |> specialize_type thy x |> Logic.dest_equals |> snd | 
| 54756 | 617 | | NONE => t) | 
| 57717 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 618 | | uncomb t = t | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 619 | in uncomb end | 
| 54499 | 620 | |
| 57717 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 621 | fun unlift_aterm lifted (t as Const (s, _)) = | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 622 | if String.isPrefix lam_lifted_prefix s then | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 623 | (* FIXME: do something about the types *) | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 624 | (case AList.lookup (op =) lifted s of | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 625 | SOME t' => unlift_term lifted t' | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 626 | | NONE => t) | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 627 | else | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 628 | t | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 629 | | unlift_aterm _ t = t | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 630 | and unlift_term lifted = | 
| 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 631 | map_aterms (unlift_aterm lifted) | 
| 54499 | 632 | |
| 57256 | 633 | fun termify_line _ _ _ _ _ (_, Type_Role, _, _, _) = NONE | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 634 | | termify_line ctxt format type_enc lifted sym_tab (name, role, u, rule, deps) = | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 635 | let | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 636 | val thy = Proof_Context.theory_of ctxt | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 637 | val t = u | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 638 | |> prop_of_atp ctxt format type_enc true sym_tab | 
| 57717 
949838aba487
unlift before uncombine, because the definition of a lambda-lifted symbol might have an SK combinator in it (in hybrid encodings)
 blanchet parents: 
57713diff
changeset | 639 | |> unlift_term lifted | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 640 | |> uncombine_term thy | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 641 | |> simplify_bool | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 642 | in | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 643 | SOME (name, role, t, rule, deps) | 
| 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 644 | end | 
| 54499 | 645 | |
| 646 | val waldmeister_conjecture_num = "1.0.0.0" | |
| 647 | ||
| 54756 | 648 | fun repair_waldmeister_endgame proof = | 
| 54499 | 649 | let | 
| 74379 | 650 | fun repair_tail (name, _, \<^Const>\<open>Trueprop for t\<close>, rule, deps) = | 
| 651 | (name, Negated_Conjecture, \<^Const>\<open>Trueprop for \<open>s_not t\<close>\<close>, rule, deps) | |
| 54756 | 652 | fun repair_body [] = [] | 
| 653 | | repair_body ((line as ((num, _), _, _, _, _)) :: lines) = | |
| 654 | if num = waldmeister_conjecture_num then map repair_tail (line :: lines) | |
| 655 | else line :: repair_body lines | |
| 656 | in | |
| 657 | repair_body proof | |
| 658 | end | |
| 54499 | 659 | |
| 57635 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 blanchet parents: 
57547diff
changeset | 660 | fun map_proof_terms f (lines : ('a * 'b * 'c * 'd * 'e) list) =
 | 
| 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 blanchet parents: 
57547diff
changeset | 661 | map2 (fn c => fn (a, b, _, d, e) => (a, b, c, d, e)) (f (map #3 lines)) lines | 
| 
97adb86619a4
more robust handling of types for skolems (modeled as Frees)
 blanchet parents: 
57547diff
changeset | 662 | |
| 57258 
67d85a8aa6cc
Moving the remote prefix deleting on Sledgehammer's side
 fleury parents: 
57257diff
changeset | 663 | fun termify_atp_proof ctxt local_prover format type_enc pool lifted sym_tab = | 
| 57267 
8b87114357bd
integrated new Waldmeister code with 'sledgehammer' command
 blanchet parents: 
57263diff
changeset | 664 | nasty_atp_proof pool | 
| 54499 | 665 | #> map_term_names_in_atp_proof repair_name | 
| 57255 
488046fdda59
add support for Isar reconstruction for thf1 ATP provers like Leo-II.
 fleury parents: 
57199diff
changeset | 666 | #> map_filter (termify_line ctxt format type_enc lifted sym_tab) | 
| 57699 | 667 | #> map_proof_terms (infer_formulas_types ctxt #> map HOLogic.mk_Trueprop) | 
| 57258 
67d85a8aa6cc
Moving the remote prefix deleting on Sledgehammer's side
 fleury parents: 
57257diff
changeset | 668 | #> local_prover = waldmeisterN ? repair_waldmeister_endgame | 
| 54499 | 669 | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 670 | fun unskolemize_spass_term skos = | 
| 55192 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 blanchet parents: 
55185diff
changeset | 671 | let | 
| 58597 | 672 | val is_skolem_name = member (op =) skos | 
| 55192 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 blanchet parents: 
55185diff
changeset | 673 | |
| 58091 | 674 | fun find_argless_skolem (Free _ $ Var _) = NONE | 
| 58597 | 675 | | find_argless_skolem (Free (x as (s, _))) = if is_skolem_name s then SOME x else NONE | 
| 58091 | 676 | | find_argless_skolem (t $ u) = | 
| 677 | (case find_argless_skolem t of NONE => find_argless_skolem u | sk => sk) | |
| 678 | | find_argless_skolem (Abs (_, _, t)) = find_argless_skolem t | |
| 679 | | find_argless_skolem _ = NONE | |
| 55192 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 blanchet parents: 
55185diff
changeset | 680 | |
| 58597 | 681 | fun find_skolem_arg (Free (s, _) $ Var v) = if is_skolem_name s then SOME v else NONE | 
| 58091 | 682 | | find_skolem_arg (t $ u) = (case find_skolem_arg t of NONE => find_skolem_arg u | v => v) | 
| 683 | | find_skolem_arg (Abs (_, _, t)) = find_skolem_arg t | |
| 684 | | find_skolem_arg _ = NONE | |
| 685 | ||
| 58597 | 686 | fun kill_skolem_arg (t as Free (s, T) $ Var _) = | 
| 687 | if is_skolem_name s then Free (s, range_type T) else t | |
| 688 | | kill_skolem_arg (t $ u) = kill_skolem_arg t $ kill_skolem_arg u | |
| 689 | | kill_skolem_arg (Abs (s, T, t)) = Abs (s, T, kill_skolem_arg t) | |
| 690 | | kill_skolem_arg t = t | |
| 691 | ||
| 58091 | 692 | fun find_var (Var v) = SOME v | 
| 693 | | find_var (t $ u) = (case find_var t of NONE => find_var u | v => v) | |
| 694 | | find_var (Abs (_, _, t)) = find_var t | |
| 695 | | find_var _ = NONE | |
| 696 | ||
| 697 | val safe_abstract_over = abstract_over o apsnd (incr_boundvars 1) | |
| 55192 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 blanchet parents: 
55185diff
changeset | 698 | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 699 | fun unskolem_inner t = | 
| 58091 | 700 | (case find_argless_skolem t of | 
| 701 | SOME (x as (s, T)) => | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 702 | HOLogic.exists_const T $ Abs (s, T, unskolem_inner (safe_abstract_over (Free x, t))) | 
| 58091 | 703 | | NONE => | 
| 58597 | 704 | (case find_skolem_arg t of | 
| 58091 | 705 | SOME (v as ((s, _), T)) => | 
| 706 | let | |
| 707 | val (haves, have_nots) = | |
| 708 | HOLogic.disjuncts t | |
| 58597 | 709 | |> List.partition (exists_subterm (curry (op =) (Var v))) | 
| 69593 | 710 | |> apply2 (fn lits => fold (curry s_disj) lits \<^term>\<open>False\<close>) | 
| 58091 | 711 | in | 
| 712 | s_disj (HOLogic.all_const T | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 713 | $ Abs (s, T, unskolem_inner (safe_abstract_over (Var v, kill_skolem_arg haves))), | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 714 | unskolem_inner have_nots) | 
| 58091 | 715 | end | 
| 58597 | 716 | | NONE => | 
| 717 | (case find_var t of | |
| 718 | SOME (v as ((s, _), T)) => | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 719 | HOLogic.all_const T $ Abs (s, T, unskolem_inner (safe_abstract_over (Var v, t))) | 
| 58597 | 720 | | NONE => t))) | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 721 | |
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 722 |     fun unskolem_outer (@{const Trueprop} $ t) = @{const Trueprop} $ unskolem_outer t
 | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 723 | | unskolem_outer t = unskolem_inner t | 
| 55195 | 724 | in | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 725 | unskolem_outer | 
| 55192 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 blanchet parents: 
55185diff
changeset | 726 | end | 
| 
b75b52c7cf94
unskolemize SPASS formula to ensure that the variables are in the right order for 'metis's skolemizer
 blanchet parents: 
55185diff
changeset | 727 | |
| 57788 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 728 | fun rename_skolem_args t = | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 729 | let | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 730 | fun add_skolem_args (Abs (_, _, t)) = add_skolem_args t | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 731 | | add_skolem_args t = | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 732 | (case strip_comb t of | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 733 | (Free (s, _), args as _ :: _) => | 
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 734 | if is_spass_skolem s then insert (op =) (s, take_prefix is_Var args) | 
| 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 735 | else fold add_skolem_args args | 
| 57788 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 736 | | (u, args as _ :: _) => fold add_skolem_args (u :: args) | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 737 | | _ => I) | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 738 | |
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 739 | fun subst_of_skolem (sk, args) = | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 740 | map_index (fn (j, Var (z, T)) => (z, Var ((sk ^ "_" ^ string_of_int j, 0), T))) args | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 741 | |
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 742 | val subst = maps subst_of_skolem (add_skolem_args t []) | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 743 | in | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 744 | subst_vars ([], subst) t | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 745 | end | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 746 | |
| 72401 
2783779b7dd3
removed obsolete unmaintained experimental prover Pirate
 blanchet parents: 
70940diff
changeset | 747 | fun introduce_spass_skolems proof = | 
| 57787 | 748 | let | 
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 749 | fun add_skolem (Free (s, _)) = is_spass_skolem s ? insert (op =) s | 
| 57788 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 750 | | add_skolem _ = I | 
| 54772 | 751 | |
| 57787 | 752 | (* union-find would be faster *) | 
| 753 | fun add_cycle [] = I | |
| 754 | | add_cycle ss = | |
| 755 | fold (fn s => Graph.default_node (s, ())) ss | |
| 756 | #> fold Graph.add_edge (ss ~~ tl ss @ [hd ss]) | |
| 54772 | 757 | |
| 57787 | 758 | val (input_steps, other_steps) = List.partition (null o #5) proof | 
| 54772 | 759 | |
| 58597 | 760 | (* The names of the formulas are added to the Skolem constants, to ensure that formulas giving | 
| 761 | rise to several clauses are skolemized together. *) | |
| 762 | val skoXss = map (fn ((_, ss), _, t, _, _) => Term.fold_aterms add_skolem t ss) input_steps | |
| 58601 
85fa90262807
avoid creating needless skolemization steps for SPASS
 blanchet parents: 
58600diff
changeset | 763 | val groups0 = Graph.strong_conn (fold add_cycle skoXss Graph.empty) | 
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 764 | val groups = filter (exists is_spass_skolem) groups0 | 
| 58601 
85fa90262807
avoid creating needless skolemization steps for SPASS
 blanchet parents: 
58600diff
changeset | 765 | |
| 
85fa90262807
avoid creating needless skolemization steps for SPASS
 blanchet parents: 
58600diff
changeset | 766 | val skoXss_input_steps = skoXss ~~ input_steps | 
| 54772 | 767 | |
| 58597 | 768 | fun step_name_of_group skoXs = (implode skoXs, []) | 
| 57787 | 769 | fun in_group group = member (op =) group o hd | 
| 58601 
85fa90262807
avoid creating needless skolemization steps for SPASS
 blanchet parents: 
58600diff
changeset | 770 | fun group_of skoX = find_first (fn group => in_group group skoX) groups | 
| 54772 | 771 | |
| 58652 | 772 | fun new_steps (skoXss_steps : (string list * (term, 'a) atp_step) list) group = | 
| 57787 | 773 | let | 
| 774 | val name = step_name_of_group group | |
| 775 | val name0 = name |>> prefix "0" | |
| 57788 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 776 | val t = | 
| 58597 | 777 | (case map (snd #> #3) skoXss_steps of | 
| 57788 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 778 | [t] => t | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 779 | | ts => ts | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 780 | |> map (HOLogic.dest_Trueprop #> rename_skolem_args) | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 781 | |> Library.foldr1 s_conj | 
| 
0a38c47ebb69
normalize skolem argument variable names so that they coincide when taking the conjunction
 blanchet parents: 
57787diff
changeset | 782 | |> HOLogic.mk_Trueprop) | 
| 79730 
4031aafc2dda
deal with new-style Vampire skolemization in reconstructed Isar proofs
 blanchet parents: 
78697diff
changeset | 783 | val skos = fold (union (op =) o filter is_spass_skolem o fst) skoXss_steps [] | 
| 58597 | 784 | val deps = map (snd #> #1) skoXss_steps | 
| 57787 | 785 | in | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 786 | [(name0, Unknown, unskolemize_spass_term skos t, spass_pre_skolemize_rule, deps), | 
| 57787 | 787 | (name, Unknown, t, spass_skolemize_rule, [name0])] | 
| 788 | end | |
| 54772 | 789 | |
| 57787 | 790 | val sko_steps = | 
| 58597 | 791 | maps (fn group => new_steps (filter (in_group group o fst) skoXss_input_steps) group) groups | 
| 54772 | 792 | |
| 57787 | 793 | val old_news = | 
| 58601 
85fa90262807
avoid creating needless skolemization steps for SPASS
 blanchet parents: 
58600diff
changeset | 794 | map_filter (fn (skoXs, (name, _, _, _, _)) => | 
| 
85fa90262807
avoid creating needless skolemization steps for SPASS
 blanchet parents: 
58600diff
changeset | 795 | Option.map (pair name o single o step_name_of_group) (group_of skoXs)) | 
| 58597 | 796 | skoXss_input_steps | 
| 57787 | 797 | val repair_deps = fold replace_dependencies_in_line old_news | 
| 798 | in | |
| 799 | input_steps @ sko_steps @ map repair_deps other_steps | |
| 800 | end | |
| 54772 | 801 | |
| 79734 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 802 | fun regroup_zipperposition_skolems steps = | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 803 | let | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 804 | val contains_skolem = | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 805 | Term.exists_subterm (fn Free (sk, _) => is_zipperposition_skolem sk | _ => false) | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 806 | |
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 807 | fun is_skolem_step (_, _, t, rule, _) = | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 808 | rule = zipperposition_cnf_rule andalso contains_skolem t | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 809 | |
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 810 | val dep_to_skolem_steps = steps | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 811 | |> filter is_skolem_step | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 812 | |> map_filter (fn step as (_, _, _, _, [dep]) => SOME (fst dep, step) | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 813 | | _ => NONE) | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 814 | |> AList.group (op =) | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 815 | |
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 816 | val prop_of_skolem_steps = | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 817 | map (#3 #> perhaps (try HOLogic.dest_Trueprop)) | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 818 | #> Library.foldr1 HOLogic.mk_conj | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 819 | #> HOLogic.mk_Trueprop | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 820 | |
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 821 | fun insert_back_regrouped_skolem_steps (step as (name, _, _, _, _)) = | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 822 | (case AList.lookup (op =) dep_to_skolem_steps (fst name) of | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 823 | SOME skolem_steps => | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 824 | [step, ((fst name ^ pseudo_zipperposition_skolemization_suffix, []), Plain, | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 825 | prop_of_skolem_steps skolem_steps, zipperposition_cnf_rule, [name])] | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 826 | | NONE => [step]) | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 827 | |
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 828 | fun adjust_dependencies (step as (name, role, t, _, [dep])) = | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 829 | if is_skolem_step step then | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 830 | (name, role, t, pseudo_zipperposition_conjunct, | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 831 | [(fst dep ^ pseudo_zipperposition_skolemization_suffix, [])]) | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 832 | else | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 833 | step | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 834 | | adjust_dependencies step = step | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 835 | in | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 836 | steps | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 837 | |> map adjust_dependencies | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 838 | |> maps insert_back_regrouped_skolem_steps | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 839 | end | 
| 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 840 | |
| 78693 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 841 | val zf_stmt_prefix = "zf_stmt_" | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 842 | |
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 843 | fun is_if_True_or_False_axiom true_or_false t = | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 844 | (case t of | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 845 |     @{const Trueprop} $
 | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 846 |      (Const (@{const_name HOL.eq}, _) $
 | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 847 |        (Const (@{const_name If}, _) $ cond $ Var _ $ Var _) $
 | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 848 | Var _) => | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 849 | cond aconv true_or_false | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 850 | | _ => false) | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 851 | |
| 57263 | 852 | fun factify_atp_proof facts hyp_ts concl_t atp_proof = | 
| 54505 | 853 | let | 
| 57785 | 854 | fun factify_step ((num, ss), role, t, rule, deps) = | 
| 54505 | 855 | let | 
| 856 | val (ss', role', t') = | |
| 57263 | 857 | (case resolve_conjectures ss of | 
| 54505 | 858 | [j] => | 
| 58484 
b4c0e2b00036
support hypotheses with schematics in Isar proofs
 blanchet parents: 
58477diff
changeset | 859 | if j = length hyp_ts then ([], Conjecture, concl_t) | 
| 
b4c0e2b00036
support hypotheses with schematics in Isar proofs
 blanchet parents: 
58477diff
changeset | 860 | else ([], Hypothesis, close_form (nth hyp_ts j)) | 
| 55185 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 blanchet parents: 
54822diff
changeset | 861 | | _ => | 
| 70940 
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
 blanchet parents: 
69597diff
changeset | 862 | (case resolve_facts facts (num :: ss) of | 
| 78693 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 863 | [] => | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 864 | if role = Axiom andalso String.isPrefix zf_stmt_prefix num | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 865 |                  andalso is_if_True_or_False_axiom @{const True} t then
 | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 866 | (["if_True"], Axiom, t) | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 867 | else if role = Axiom andalso String.isPrefix zf_stmt_prefix num | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 868 |                  andalso is_if_True_or_False_axiom @{const False} t then
 | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 869 | (["if_False"], Axiom, t) | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 870 | else | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 871 | (ss, | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 872 | if role = Definition then Definition | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 873 | else if role = Lemma then Lemma | 
| 
1c0576840bf4
reconstruct Zipperposition's if-then-else in Sledgehammer Isar proofs
 blanchet parents: 
77602diff
changeset | 874 | else Plain, t) | 
| 55185 
a0bd8d6414e6
centralized & repaired handling of bound variables in intermediate data structure (viva de Bruijn)
 blanchet parents: 
54822diff
changeset | 875 | | facts => (map fst facts, Axiom, t))) | 
| 54505 | 876 | in | 
| 79734 
0fa4bebbdd75
support Zipperposition's skolemization in generated Isar proofs
 blanchet parents: 
79730diff
changeset | 877 | ((num, ss' |> distinct (op =)), role', t', rule, deps) | 
| 54505 | 878 | end | 
| 879 | ||
| 880 | val atp_proof = map factify_step atp_proof | |
| 881 | val names = map #1 atp_proof | |
| 882 | ||
| 883 | fun repair_dep (num, ss) = (num, the_default ss (AList.lookup (op =) names num)) | |
| 884 | fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps) | |
| 54772 | 885 | in | 
| 886 | map repair_deps atp_proof | |
| 887 | end | |
| 54505 | 888 | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 889 | fun termify_atp_abduce_candidate ctxt local_prover format type_enc pool lifted sym_tab phi = | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 890 | let | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 891 |     val proof = [(("", []), Conjecture, mk_anot phi, "", [])]
 | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 892 | val new_proof = termify_atp_proof ctxt local_prover format type_enc pool lifted sym_tab proof | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 893 | in | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 894 | (case new_proof of | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 895 | [(_, _, phi', _, _)] => phi' | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 896 | | _ => error "Impossible case in termify_atp_abduce_candidate") | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 897 | end | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 898 | |
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 899 | fun sort_top n scored_items = | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 900 | if n <= 0 orelse null scored_items then | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 901 | [] | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 902 | else | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 903 | let | 
| 77421 | 904 | fun split_min accum [] (_, best_item) = (best_item, List.rev accum) | 
| 905 | | split_min accum ((score, item) :: scored_items) (best_score, best_item) = | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 906 | if score < best_score then | 
| 77421 | 907 | split_min ((best_score, best_item) :: accum) scored_items (score, item) | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 908 | else | 
| 77421 | 909 | split_min ((score, item) :: accum) scored_items (best_score, best_item) | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 910 | |
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 911 | val (min_item, other_scored_items) = split_min [] (tl scored_items) (hd scored_items) | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 912 | in | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 913 | min_item :: sort_top (n - 1) (filter_out (equal min_item o snd) other_scored_items) | 
| 77421 | 914 | |> distinct (op aconv) | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 915 | end | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 916 | |
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 917 | fun top_abduce_candidates max_candidates candidates = | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 918 | let | 
| 77421 | 919 | (* We prefer free variables to other constructs, so that e.g. "x \<le> y" is | 
| 77576 
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
 blanchet parents: 
77432diff
changeset | 920 | prioritized over "x \<le> 5". *) | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 921 | fun score t = | 
| 77425 | 922 | Term.fold_aterms (fn t => fn score => score + (case t of Free _ => 1 | _ => 2)) t 0 | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 923 | |
| 77421 | 924 | (* Equations of the form "x = ..." or "... = x" or similar are too specific | 
| 925 | to be useful. Quantified formulas are also filtered out. As for "True", | |
| 926 | it may seem an odd choice for abduction, but it sometimes arises in | |
| 927 | conjunction with type class constraints, which are removed by the | |
| 928 | termifier. *) | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 929 | fun maybe_score t = | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 930 | (case t of | 
| 77421 | 931 |         @{prop True} => NONE
 | 
| 932 |       | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ Free _ $ _) => NONE
 | |
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 933 |       | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Free _) => NONE
 | 
| 77421 | 934 |       | @{const Trueprop} $ (@{const less(nat)} $ _ $ @{const zero_class.zero(nat)}) => NONE
 | 
| 935 |       | @{const Trueprop} $ (@{const less_eq(nat)} $ _ $ @{const zero_class.zero(nat)}) => NONE
 | |
| 936 |       | @{const Trueprop} $ (@{const less(nat)} $ _ $ @{const one_class.one(nat)}) => NONE
 | |
| 937 |       | @{const Trueprop} $ (@{const Not} $
 | |
| 938 |           (@{const less(nat)} $ @{const zero_class.zero(nat)} $ _)) => NONE
 | |
| 939 |       | @{const Trueprop} $ (@{const Not} $
 | |
| 940 |           (@{const less_eq(nat)} $ @{const zero_class.zero(nat)} $ _)) => NONE
 | |
| 941 |       | @{const Trueprop} $ (@{const Not} $
 | |
| 77425 | 942 |           (@{const less_eq(nat)} $ @{const one_class.one(nat)} $ _)) => NONE
 | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 943 |       | @{const Trueprop} $ (Const (@{const_name All}, _) $ _) => NONE
 | 
| 77421 | 944 |       | @{const Trueprop} $ (Const (@{const_name Ex}, _) $ _) => NONE
 | 
| 77576 
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
 blanchet parents: 
77432diff
changeset | 945 | | _ => | 
| 
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
 blanchet parents: 
77432diff
changeset | 946 | (* We require the presence of at least one free variable. A "missing | 
| 
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
 blanchet parents: 
77432diff
changeset | 947 | assumption" that does not talk about any free variable is likely | 
| 
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
 blanchet parents: 
77432diff
changeset | 948 | spurious. *) | 
| 
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
 blanchet parents: 
77432diff
changeset | 949 | if exists_subterm (fn Free _ => true | _ => false) t then SOME (score t, t) | 
| 
80bcebe6cf33
require the presence of free variables to do abduction in Sledgehammer
 blanchet parents: 
77432diff
changeset | 950 | else NONE) | 
| 77418 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 951 | in | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 952 | sort_top max_candidates (map_filter maybe_score candidates) | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 953 | end | 
| 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 blanchet parents: 
77272diff
changeset | 954 | |
| 77602 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 955 | fun provability_status ctxt t = | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 956 | let | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 957 | val res = Timeout.apply (seconds 0.1) | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 958 | (Thm.term_of o Thm.rhs_of o Simplifier.full_rewrite ctxt) (Thm.cterm_of ctxt t) | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 959 | in | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 960 |     if res aconv @{prop True} then SOME true
 | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 961 |     else if res aconv @{prop False} then SOME false
 | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 962 | else NONE | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 963 | end | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 964 | handle Timeout.TIMEOUT _ => NONE | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 965 | |
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 966 | (* Put propositions that simplify to "True" first, and filter out propositions | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 967 | that simplify to "False". *) | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 968 | fun sort_propositions_by_provability ctxt ts = | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 969 | let | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 970 | val statuses_ts = map (`(provability_status ctxt)) ts | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 971 | in | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 972 | maps (fn (SOME true, t) => [t] | _ => []) statuses_ts @ | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 973 | maps (fn (NONE, t) => [t] | _ => []) statuses_ts | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 974 | end | 
| 
7c25451ae2c1
use simplifier to classify the missing assumptions in Sledgehammer's abduction mechanism
 blanchet parents: 
77576diff
changeset | 975 | |
| 31038 | 976 | end; |