| author | wenzelm | 
| Wed, 02 May 2012 21:23:14 +0200 | |
| changeset 47861 | c7144da6abfb | 
| parent 47774 | 6d9a51a00a6a | 
| child 47918 | 81ae96996223 | 
| 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 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 6 | 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 | 
| 43678 | 11 |   type ('a, 'b) ho_term = ('a, 'b) ATP_Problem.ho_term
 | 
| 43127 | 12 |   type ('a, 'b, 'c) formula = ('a, 'b, 'c) ATP_Problem.formula
 | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 13 | type 'a proof = 'a ATP_Proof.proof | 
| 46340 | 14 | type stature = ATP_Problem_Generate.stature | 
| 43033 | 15 | |
| 16 | datatype reconstructor = | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 17 | Metis of string * string | | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 18 | SMT | 
| 43033 | 19 | |
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 20 | datatype play = | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 21 | Played of reconstructor * Time.time | | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 22 | Trust_Playable of reconstructor * Time.time option | | 
| 43166 | 23 | Failed_to_Play of reconstructor | 
| 43033 | 24 | |
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 25 | type minimize_command = string list -> string | 
| 43033 | 26 | type one_line_params = | 
| 46340 | 27 | play * string * (string * stature) list * minimize_command * int * int | 
| 38818 
61cf050f8b2e
improve SPASS hack, when a clause comes from several facts
 blanchet parents: 
38752diff
changeset | 28 | type isar_params = | 
| 46340 | 29 | bool * int * string Symtab.table * (string * stature) list vector | 
| 45551 | 30 | * int Symtab.table * string proof * thm | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 31 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 32 | val metisN : string | 
| 45520 | 33 | val smtN : string | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 34 | val full_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 35 | val partial_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 36 | val no_typesN : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 37 | val really_full_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 38 | val full_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 39 | val partial_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 40 | val no_type_enc : string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 41 | val full_type_encs : string list | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 42 | val partial_type_encs : string list | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 43 | val metis_default_lam_trans : string | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 44 | val metis_call : string -> string -> string | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 45 | val string_for_reconstructor : reconstructor -> string | 
| 42451 
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
 blanchet parents: 
42449diff
changeset | 46 | val used_facts_in_atp_proof : | 
| 46340 | 47 | Proof.context -> (string * stature) list vector -> string proof | 
| 48 | -> (string * stature) list | |
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 49 | val lam_trans_from_atp_proof : string proof -> string -> string | 
| 45590 | 50 | val is_typed_helper_used_in_atp_proof : string proof -> bool | 
| 42876 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 51 | val used_facts_in_unsound_atp_proof : | 
| 46340 | 52 | Proof.context -> (string * stature) list vector -> 'a proof | 
| 45551 | 53 | -> string list option | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 54 | val unalias_type_enc : string -> string list | 
| 43033 | 55 | val one_line_proof_text : one_line_params -> string | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 56 | val make_tvar : string -> typ | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 57 | val make_tfree : Proof.context -> string -> typ | 
| 43094 | 58 | val term_from_atp : | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 59 | Proof.context -> bool -> int Symtab.table -> typ option | 
| 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 60 | -> (string, string) ho_term -> term | 
| 43127 | 61 | val prop_from_atp : | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 62 | Proof.context -> bool -> int Symtab.table | 
| 43678 | 63 | -> (string, string, (string, string) ho_term) formula -> term | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 64 | val isar_proof_text : | 
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 65 | Proof.context -> bool -> isar_params -> one_line_params -> string | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 66 | val proof_text : | 
| 43033 | 67 | Proof.context -> bool -> isar_params -> one_line_params -> string | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 68 | end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 69 | |
| 46320 | 70 | structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 71 | struct | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 72 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 73 | open ATP_Util | 
| 38028 | 74 | open ATP_Problem | 
| 39452 | 75 | open ATP_Proof | 
| 46320 | 76 | open ATP_Problem_Generate | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 77 | |
| 46320 | 78 | structure String_Redirect = ATP_Proof_Redirect( | 
| 45887 | 79 | type key = step_name | 
| 80 | val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s') | |
| 81 | val string_of = fst) | |
| 82 | ||
| 45882 | 83 | open String_Redirect | 
| 84 | ||
| 43033 | 85 | datatype reconstructor = | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 86 | Metis of string * string | | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 87 | SMT | 
| 43033 | 88 | |
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 89 | datatype play = | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 90 | Played of reconstructor * Time.time | | 
| 43033 | 91 | Trust_Playable of reconstructor * Time.time option | | 
| 43166 | 92 | Failed_to_Play of reconstructor | 
| 43033 | 93 | |
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 94 | type minimize_command = string list -> string | 
| 43033 | 95 | type one_line_params = | 
| 46340 | 96 | play * string * (string * stature) list * minimize_command * int * int | 
| 38818 
61cf050f8b2e
improve SPASS hack, when a clause comes from several facts
 blanchet parents: 
38752diff
changeset | 97 | type isar_params = | 
| 46340 | 98 | bool * int * string Symtab.table * (string * stature) list vector | 
| 45551 | 99 | * int Symtab.table * string proof * thm | 
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 100 | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 101 | val metisN = "metis" | 
| 45520 | 102 | val smtN = "smt" | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 103 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 104 | val full_typesN = "full_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 105 | val partial_typesN = "partial_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 106 | val no_typesN = "no_types" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 107 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 108 | val really_full_type_enc = "mono_tags" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 109 | val full_type_enc = "poly_guards_query" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 110 | val partial_type_enc = "poly_args" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 111 | val no_type_enc = "erased" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 112 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 113 | val full_type_encs = [full_type_enc, really_full_type_enc] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 114 | val partial_type_encs = partial_type_enc :: full_type_encs | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 115 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 116 | val type_enc_aliases = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 117 | [(full_typesN, full_type_encs), | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 118 | (partial_typesN, partial_type_encs), | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 119 | (no_typesN, [no_type_enc])] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 120 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 121 | fun unalias_type_enc s = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 122 | AList.lookup (op =) type_enc_aliases s |> the_default [s] | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 123 | |
| 46365 | 124 | val metis_default_lam_trans = combsN | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 125 | |
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 126 | fun metis_call type_enc lam_trans = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 127 | let | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 128 | val type_enc = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 129 | case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 130 | type_enc of | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 131 | [alias] => alias | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 132 | | _ => type_enc | 
| 45522 | 133 | val opts = [] |> type_enc <> partial_typesN ? cons type_enc | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 134 | |> lam_trans <> metis_default_lam_trans ? cons lam_trans | 
| 45520 | 135 |   in metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") end
 | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 136 | |
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 137 | fun string_for_reconstructor (Metis (type_enc, lam_trans)) = | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45511diff
changeset | 138 | metis_call type_enc lam_trans | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 139 | | string_for_reconstructor SMT = smtN | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 140 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 141 | fun find_first_in_list_vector vec key = | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 142 | Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 143 | | (_, value) => value) NONE vec | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 144 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 145 | val unprefix_fact_number = space_implode "_" o tl o space_explode "_" | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 146 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 147 | fun resolve_one_named_fact fact_names s = | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 148 | case try (unprefix fact_prefix) s of | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 149 | SOME s' => | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 150 | let val s' = s' |> unprefix_fact_number |> unascii_of in | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 151 | s' |> find_first_in_list_vector fact_names |> Option.map (pair s') | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 152 | end | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 153 | | NONE => NONE | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 154 | fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names) | 
| 45579 | 155 | fun is_fact fact_names = not o null o resolve_fact fact_names | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 156 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 157 | fun resolve_one_named_conjecture s = | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 158 | case try (unprefix conjecture_prefix) s of | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 159 | SOME s' => Int.fromString s' | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 160 | | NONE => NONE | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 161 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 162 | val resolve_conjecture = map_filter resolve_one_named_conjecture | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 163 | val is_conjecture = not o null o resolve_conjecture | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 164 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 165 | fun is_axiom_used_in_proof pred = | 
| 47774 | 166 | exists (fn Inference_Step ((_, ss), _, _, []) => exists pred ss | _ => false) | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 167 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 168 | val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix) | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 169 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 170 | val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 171 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 172 | (* overapproximation (good enough) *) | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 173 | fun is_lam_lifted s = | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 174 | String.isPrefix fact_prefix s andalso | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 175 | String.isSubstring ascii_of_lam_fact_prefix s | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 176 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 177 | fun lam_trans_from_atp_proof atp_proof default = | 
| 46365 | 178 | case (is_axiom_used_in_proof is_combinator_def atp_proof, | 
| 179 | is_axiom_used_in_proof is_lam_lifted atp_proof) of | |
| 180 | (false, false) => default | |
| 181 | | (false, true) => liftingN | |
| 46367 | 182 | (* | (true, true) => combs_and_liftingN -- not supported by "metis" *) | 
| 183 | | (true, _) => combsN | |
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 184 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 185 | val is_typed_helper_name = | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 186 | String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix | 
| 45590 | 187 | fun is_typed_helper_used_in_atp_proof atp_proof = | 
| 188 | is_axiom_used_in_proof is_typed_helper_name atp_proof | |
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 189 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 190 | val leo2_ext = "extcnf_equal_neg" | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 191 | val isa_ext = Thm.get_name_hint @{thm ext}
 | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 192 | val isa_short_ext = Long_Name.base_name isa_ext | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 193 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 194 | fun ext_name ctxt = | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 195 |   if Thm.eq_thm_prop (@{thm ext},
 | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 196 | singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 197 | isa_short_ext | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 198 | else | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 199 | isa_ext | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 200 | |
| 47774 | 201 | fun add_fact _ fact_names (Inference_Step ((_, ss), _, _, [])) = | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 202 | union (op =) (resolve_fact fact_names ss) | 
| 47774 | 203 | | add_fact ctxt _ (Inference_Step (_, _, rule, _)) = | 
| 46340 | 204 | if rule = leo2_ext then insert (op =) (ext_name ctxt, (Global, General)) | 
| 205 | else I | |
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 206 | | add_fact _ _ _ = I | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 207 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 208 | fun used_facts_in_atp_proof ctxt fact_names atp_proof = | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 209 | if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 210 | else fold (add_fact ctxt fact_names) atp_proof [] | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 211 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 212 | fun used_facts_in_unsound_atp_proof _ _ [] = NONE | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 213 | | used_facts_in_unsound_atp_proof ctxt fact_names atp_proof = | 
| 46427 | 214 | let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in | 
| 46340 | 215 | if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 216 | not (is_axiom_used_in_proof (is_conjecture o single) atp_proof) then | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 217 | SOME (map fst used_facts) | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 218 | else | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 219 | NONE | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 220 | end | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 221 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 222 | |
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 223 | (** Soft-core proof reconstruction: one-liners **) | 
| 43033 | 224 | |
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 225 | fun string_for_label (s, num) = s ^ string_of_int num | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 226 | |
| 43033 | 227 | fun show_time NONE = "" | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 228 |   | show_time (SOME ext_time) = " (" ^ string_from_ext_time ext_time ^ ")"
 | 
| 43033 | 229 | |
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 230 | fun apply_on_subgoal _ 1 = "by " | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 231 | | apply_on_subgoal 1 _ = "apply " | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 232 | | apply_on_subgoal i n = | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 233 | "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n | 
| 44002 | 234 | fun command_call name [] = | 
| 235 |     name |> not (Lexicon.is_identifier name) ? enclose "(" ")"
 | |
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 236 |   | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
 | 
| 43033 | 237 | fun try_command_line banner time command = | 
| 45666 | 238 | banner ^ ": " ^ Markup.markup Isabelle_Markup.sendback command ^ show_time time ^ "." | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 239 | fun using_labels [] = "" | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 240 | | using_labels ls = | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 241 | "using " ^ space_implode " " (map string_for_label ls) ^ " " | 
| 45520 | 242 | fun reconstructor_command reconstr i n (ls, ss) = | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 243 | using_labels ls ^ apply_on_subgoal i n ^ | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 244 | command_call (string_for_reconstructor reconstr) ss | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 245 | fun minimize_line _ [] = "" | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 246 | | minimize_line minimize_command ss = | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 247 | case minimize_command ss of | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 248 | "" => "" | 
| 46340 | 249 | | command => | 
| 250 | "\nTo minimize: " ^ Markup.markup Isabelle_Markup.sendback command ^ "." | |
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 251 | |
| 46342 | 252 | fun split_used_facts facts = | 
| 253 | facts |> List.partition (fn (_, (sc, _)) => sc = Chained) | |
| 254 | |> pairself (sort_distinct (string_ord o pairself fst)) | |
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 255 | |
| 43033 | 256 | fun one_line_proof_text (preplay, banner, used_facts, minimize_command, | 
| 43037 | 257 | subgoal, subgoal_count) = | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 258 | let | 
| 43033 | 259 | val (chained, extra) = split_used_facts used_facts | 
| 45520 | 260 | val (failed, reconstr, ext_time) = | 
| 43033 | 261 | case preplay of | 
| 45520 | 262 | Played (reconstr, time) => (false, reconstr, (SOME (false, time))) | 
| 263 | | Trust_Playable (reconstr, time) => | |
| 264 | (false, reconstr, | |
| 43033 | 265 | case time of | 
| 266 | NONE => NONE | |
| 267 | | SOME time => | |
| 268 | if time = Time.zeroTime then NONE else SOME (true, time)) | |
| 45520 | 269 | | Failed_to_Play reconstr => (true, reconstr, NONE) | 
| 43033 | 270 | val try_line = | 
| 43166 | 271 | ([], map fst extra) | 
| 45520 | 272 | |> reconstructor_command reconstr subgoal subgoal_count | 
| 47147 | 273 | |> (if failed then | 
| 274 | enclose "One-line proof reconstruction failed: " | |
| 275 | ".\n(Invoking \"sledgehammer\" with \"[strict]\" might \ | |
| 276 | \solve this.)" | |
| 277 | else | |
| 278 | try_command_line banner ext_time) | |
| 43033 | 279 | in try_line ^ minimize_line minimize_command (map fst (extra @ chained)) end | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 280 | |
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 281 | (** Hard-core proof reconstruction: structured Isar proofs **) | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 282 | |
| 39425 | 283 | fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t | 
| 284 | fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t | |
| 285 | ||
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 286 | fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
 | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 287 | 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 | 288 | let val ww = "'" ^ w in | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 289 | TFree (ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1))) | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 290 | end | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 291 | |
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 292 | val indent_size = 2 | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 293 | val no_label = ("", ~1)
 | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 294 | |
| 45882 | 295 | val raw_prefix = "x" | 
| 296 | val assum_prefix = "a" | |
| 297 | val have_prefix = "f" | |
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 298 | |
| 45552 | 299 | fun raw_label_for_name (num, ss) = | 
| 300 | case resolve_conjecture ss of | |
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 301 | [j] => (conjecture_prefix, j) | 
| 45552 | 302 | | _ => case Int.fromString num of | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 303 | SOME j => (raw_prefix, j) | 
| 45552 | 304 | | NONE => (raw_prefix ^ num, 0) | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 305 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 306 | (**** INTERPRETATION OF TSTP SYNTAX TREES ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 307 | |
| 43678 | 308 | exception HO_TERM of (string, string) ho_term list | 
| 309 | exception FORMULA of (string, string, (string, string) ho_term) formula list | |
| 37991 | 310 | exception SAME of unit | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 311 | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 312 | (* Type variables are given the basic sort "HOL.type". Some will later be | 
| 37991 | 313 | constrained by information from type literals, or by type inference. *) | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 314 | fun typ_from_atp ctxt (u as ATerm (a, us)) = | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 315 | let val Ts = map (typ_from_atp ctxt) us in | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 316 | case unprefix_and_unascii type_const_prefix a of | 
| 37991 | 317 | SOME b => Type (invert_const b, Ts) | 
| 318 | | NONE => | |
| 319 | if not (null us) then | |
| 43678 | 320 | raise HO_TERM [u] (* only "tconst"s have type arguments *) | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 321 | else case unprefix_and_unascii tfree_prefix a of | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 322 | SOME b => make_tfree ctxt b | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 323 | | NONE => | 
| 43302 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 blanchet parents: 
43296diff
changeset | 324 | (* Could be an Isabelle variable or a variable from the ATP, say "X1" | 
| 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 blanchet parents: 
43296diff
changeset | 325 | or "_5018". Sometimes variables from the ATP are indistinguishable | 
| 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 blanchet parents: 
43296diff
changeset | 326 | from Isabelle variables, which forces us to use a type parameter in | 
| 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 blanchet parents: 
43296diff
changeset | 327 | all cases. *) | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 328 | (a |> perhaps (unprefix_and_unascii tvar_prefix), HOLogic.typeS) | 
| 43302 
566f970006e5
improve sort inference in Metis proofs -- in some rare cases Metis steals Isabelle's variable names and the sorts must then be inferred as well
 blanchet parents: 
43296diff
changeset | 329 | |> Type_Infer.param 0 | 
| 37991 | 330 | end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 331 | |
| 38014 | 332 | (* Type class literal applied to a type. Returns triple of polarity, class, | 
| 333 | type. *) | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 334 | fun type_constraint_from_term ctxt (u as ATerm (a, us)) = | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 335 | case (unprefix_and_unascii class_prefix a, map (typ_from_atp ctxt) us) of | 
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 336 | (SOME b, [T]) => (b, T) | 
| 43678 | 337 | | _ => raise HO_TERM [u] | 
| 38014 | 338 | |
| 43907 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 339 | (* Accumulate type constraints in a formula: negative type literals. *) | 
| 38014 | 340 | 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 | 341 | 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 | 342 | | 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 | 343 | | add_type_constraint _ _ = I | 
| 38014 | 344 | |
| 43094 | 345 | fun repair_variable_name f s = | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 346 | let | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 347 | fun subscript_name s n = s ^ nat_subscript n | 
| 38488 | 348 | val s = String.map f s | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 349 | in | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 350 | case space_explode "_" s of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 351 | [_] => (case take_suffix Char.isDigit (String.explode s) of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 352 | (cs1 as _ :: _, cs2 as _ :: _) => | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 353 | subscript_name (String.implode cs1) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 354 | (the (Int.fromString (String.implode cs2))) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 355 | | (_, _) => s) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 356 | | [s1, s2] => (case Int.fromString s2 of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 357 | SOME n => subscript_name s1 n | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 358 | | NONE => s) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 359 | | _ => s | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 360 | end | 
| 43182 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 361 | |
| 43907 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 362 | (* The number of type arguments of a constant, zero if it's monomorphic. For | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 363 | (instances of) Skolem pseudoconstants, this information is encoded in the | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 364 | constant name. *) | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 365 | fun num_type_args thy s = | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 366 | 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 | 367 | 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 | 368 | 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 | 369 | 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 | 370 | else | 
| 
073ab5379842
pass type arguments to lambda-lifted Frees, to account for polymorphism
 blanchet parents: 
43863diff
changeset | 371 | (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 | 372 | |
| 43182 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 373 | fun slack_fastype_of t = fastype_of t handle TERM _ => HOLogic.typeT | 
| 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 374 | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 375 | (* First-order translation. No types are known for variables. "HOLogic.typeT" | 
| 38014 | 376 | should allow them to be inferred. *) | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 377 | fun term_from_atp ctxt textual sym_tab = | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 378 | 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 | 379 | val thy = Proof_Context.theory_of ctxt | 
| 43212 | 380 | (* For Metis, we use 1 rather than 0 because variable references in clauses | 
| 381 | may otherwise conflict with variable constraints in the goal. At least, | |
| 382 | type inference often fails otherwise. See also "axiom_inference" in | |
| 383 | "Metis_Reconstruct". *) | |
| 43094 | 384 | val var_index = if textual then 0 else 1 | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 385 | fun do_term extra_ts opt_T u = | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 386 | case u of | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 387 | ATerm (s, us) => | 
| 46435 | 388 | if String.isPrefix native_type_prefix s then | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42943diff
changeset | 389 |           @{const True} (* ignore TPTP type information *)
 | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 390 | else if s = tptp_equal then | 
| 43093 | 391 | let val ts = map (do_term [] NONE) us in | 
| 43094 | 392 | if textual andalso length ts = 2 andalso | 
| 393 | hd ts aconv List.last ts then | |
| 39106 | 394 | (* Vampire is keen on producing these. *) | 
| 395 |               @{const True}
 | |
| 396 | else | |
| 397 |               list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
 | |
| 398 | end | |
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 399 | else case unprefix_and_unascii const_prefix s of | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 400 | SOME s' => | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 401 | let | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 402 | val ((s', s''), mangled_us) = | 
| 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 403 | s' |> unmangled_const |>> `invert_const | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 404 | in | 
| 42755 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 blanchet parents: 
42751diff
changeset | 405 | if s' = type_tag_name then | 
| 42589 
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
 blanchet parents: 
42587diff
changeset | 406 | case mangled_us @ us of | 
| 
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
 blanchet parents: 
42587diff
changeset | 407 | [typ_u, term_u] => | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 408 | do_term extra_ts (SOME (typ_from_atp ctxt typ_u)) term_u | 
| 43678 | 409 | | _ => raise HO_TERM us | 
| 42966 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 blanchet parents: 
42962diff
changeset | 410 | else if s' = predicator_name then | 
| 43093 | 411 |               do_term [] (SOME @{typ bool}) (hd us)
 | 
| 42966 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 blanchet parents: 
42962diff
changeset | 412 | else if s' = app_op_name then | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 413 | let val extra_t = do_term [] NONE (List.last us) in | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 414 | do_term (extra_t :: extra_ts) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 415 | (case opt_T of | 
| 43182 
649bada59658
slacker version of "fastype_of", in case a function has dummy type
 blanchet parents: 
43176diff
changeset | 416 | SOME T => SOME (slack_fastype_of extra_t --> T) | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 417 | | NONE => NONE) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 418 | (nth us (length us - 2)) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 419 | end | 
| 44396 
66b9b3fcd608
started cleaning up polymorphic monotonicity-based encodings, based on discussions with Nick Smallbone
 blanchet parents: 
44121diff
changeset | 420 | else if s' = type_guard_name then | 
| 42551 
cd99d6d3027a
reconstruct TFF type predicates correctly for ToFoF
 blanchet parents: 
42549diff
changeset | 421 |               @{const True} (* ignore type predicates *)
 | 
| 42549 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 422 | else | 
| 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 423 | let | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 424 | val new_skolem = String.isPrefix new_skolem_const_prefix s'' | 
| 42755 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 blanchet parents: 
42751diff
changeset | 425 | val num_ty_args = | 
| 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 blanchet parents: 
42751diff
changeset | 426 | length us - the_default 0 (Symtab.lookup sym_tab s) | 
| 42549 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 427 | val (type_us, term_us) = | 
| 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 428 | chop num_ty_args us |>> append mangled_us | 
| 43093 | 429 | val term_ts = map (do_term [] NONE) term_us | 
| 42549 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 430 | val T = | 
| 43183 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 431 | (if not (null type_us) andalso | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 432 | num_type_args thy s' = length type_us then | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 433 | let val Ts = type_us |> map (typ_from_atp ctxt) in | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 434 | if new_skolem then | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 435 | SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) | 
| 43200 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 blanchet parents: 
43199diff
changeset | 436 | else if textual then | 
| 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 blanchet parents: 
43199diff
changeset | 437 | try (Sign.const_instance thy) (s', Ts) | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 438 | else | 
| 43200 
ca7b0a48515d
imported patch metis_reconstr_give_type_infer_a_chance
 blanchet parents: 
43199diff
changeset | 439 | NONE | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 440 | end | 
| 43183 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 441 | else | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 442 | NONE) | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 443 | |> (fn SOME T => T | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 444 | | NONE => map slack_fastype_of term_ts ---> | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 445 | (case opt_T of | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 446 | SOME T => T | 
| 
faece9668bce
don't throw exception on unknown constants (e.g. skolems), and give more precise type to applied functions
 blanchet parents: 
43182diff
changeset | 447 | | NONE => HOLogic.typeT)) | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 448 | val t = | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 449 | if new_skolem then | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 450 | Var ((new_skolem_var_name_from_const s'', var_index), T) | 
| 43191 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 451 | else | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 452 | Const (unproxify_const s', T) | 
| 
0a72c0527111
fixed reconstruction of new Skolem constants in new Metis
 blanchet parents: 
43184diff
changeset | 453 | in list_comb (t, term_ts @ extra_ts) end | 
| 42549 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 454 | end | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 455 | | NONE => (* a free or schematic variable *) | 
| 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 456 | let | 
| 45042 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 457 | val term_ts = map (do_term [] NONE) us | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 458 | val ts = term_ts @ extra_ts | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 459 | val T = | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 460 | case opt_T of | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 461 | SOME T => map slack_fastype_of term_ts ---> T | 
| 
89341b897412
better type reconstruction -- prevents ill-instantiations in proof replay
 blanchet parents: 
44783diff
changeset | 462 | | NONE => map slack_fastype_of ts ---> HOLogic.typeT | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 463 | val t = | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 464 | case unprefix_and_unascii fixed_var_prefix s of | 
| 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 465 | SOME s => Free (s, T) | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 466 | | NONE => | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45379diff
changeset | 467 | case unprefix_and_unascii schematic_var_prefix s of | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 468 | SOME s => Var ((s, var_index), T) | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 469 | | NONE => | 
| 44773 
e701dabbfe37
perform mangling before computing symbol arity, to avoid needless "hAPP"s and "hBOOL"s
 blanchet parents: 
44399diff
changeset | 470 | Var ((s |> textual ? repair_variable_name Char.toLower, | 
| 43095 | 471 | var_index), T) | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 472 | in list_comb (t, ts) end | 
| 43093 | 473 | in do_term [] end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 474 | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 475 | fun term_from_atom ctxt textual sym_tab pos (u as ATerm (s, _)) = | 
| 38014 | 476 | if String.isPrefix class_prefix s then | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 477 | add_type_constraint pos (type_constraint_from_term ctxt u) | 
| 38014 | 478 |     #> pair @{const True}
 | 
| 479 | else | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 480 |     pair (term_from_atp ctxt textual sym_tab (SOME @{typ bool}) u)
 | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 481 | |
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 482 | val combinator_table = | 
| 46904 | 483 |   [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
 | 
| 484 |    (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
 | |
| 485 |    (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
 | |
| 486 |    (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
 | |
| 487 |    (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
 | |
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 488 | |
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 489 | fun uncombine_term thy = | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 490 | let | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 491 | fun aux (t1 $ t2) = betapply (pairself aux (t1, t2)) | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 492 | | aux (Abs (s, T, t')) = Abs (s, T, aux t') | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 493 | | aux (t as Const (x as (s, _))) = | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 494 | (case AList.lookup (op =) combinator_table s of | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 495 | SOME thm => thm |> prop_of |> specialize_type thy x | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 496 | |> Logic.dest_equals |> snd | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 497 | | NONE => t) | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 498 | | aux t = t | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 499 | in aux end | 
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 500 | |
| 37991 | 501 | (* Update schematic type variables with detected sort constraints. It's not | 
| 42563 | 502 | totally clear whether this code is necessary. *) | 
| 38014 | 503 | 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 | 504 | let | 
| 37991 | 505 | fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) | 
| 506 | | do_type (TVar (xi, s)) = | |
| 507 | TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) | |
| 508 | | do_type (TFree z) = TFree z | |
| 509 | fun do_term (Const (a, T)) = Const (a, do_type T) | |
| 510 | | do_term (Free (a, T)) = Free (a, do_type T) | |
| 511 | | do_term (Var (xi, T)) = Var (xi, do_type T) | |
| 512 | | do_term (t as Bound _) = t | |
| 513 | | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) | |
| 514 | | do_term (t1 $ t2) = do_term t1 $ do_term t2 | |
| 515 | in t |> not (Vartab.is_empty tvar_tab) ? do_term end | |
| 516 | ||
| 39425 | 517 | fun quantify_over_var quant_of var_s t = | 
| 518 | let | |
| 519 | val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s) | |
| 520 | |> map Var | |
| 521 | in fold_rev quant_of vars t end | |
| 37991 | 522 | |
| 38085 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 blanchet parents: 
38066diff
changeset | 523 | (* Interpret an ATP formula as a HOL term, extracting sort constraints as they | 
| 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 blanchet parents: 
38066diff
changeset | 524 | appear in the formula. *) | 
| 43184 | 525 | fun prop_from_atp ctxt textual sym_tab phi = | 
| 38014 | 526 | let | 
| 527 | fun do_formula pos phi = | |
| 37991 | 528 | case phi of | 
| 38014 | 529 | AQuant (_, [], phi) => do_formula pos phi | 
| 42526 | 530 | | AQuant (q, (s, _) :: xs, phi') => | 
| 38014 | 531 | do_formula pos (AQuant (q, xs, phi')) | 
| 42526 | 532 | (* FIXME: TFF *) | 
| 39425 | 533 | #>> quantify_over_var (case q of | 
| 534 | AForall => forall_of | |
| 535 | | AExists => exists_of) | |
| 43095 | 536 | (s |> textual ? repair_variable_name Char.toLower) | 
| 38014 | 537 | | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not | 
| 37991 | 538 | | AConn (c, [phi1, phi2]) => | 
| 38014 | 539 | do_formula (pos |> c = AImplies ? not) phi1 | 
| 540 | ##>> do_formula pos phi2 | |
| 541 | #>> (case c of | |
| 542 | AAnd => s_conj | |
| 543 | | AOr => s_disj | |
| 544 | | AImplies => s_imp | |
| 38038 | 545 | | AIff => s_iff | 
| 43163 | 546 | | ANot => raise Fail "impossible connective") | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 547 | | AAtom tm => term_from_atom ctxt textual sym_tab pos tm | 
| 37991 | 548 | | _ => raise FORMULA [phi] | 
| 38014 | 549 | in repair_tvar_sorts (do_formula true phi Vartab.empty) end | 
| 37991 | 550 | |
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 551 | fun infer_formula_types ctxt = | 
| 39288 | 552 | Type.constraint HOLogic.boolT | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 553 | #> Syntax.check_term | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 554 | (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 555 | |
| 43184 | 556 | fun uncombined_etc_prop_from_atp ctxt textual sym_tab = | 
| 43136 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 557 | let val thy = Proof_Context.theory_of ctxt in | 
| 43184 | 558 | prop_from_atp ctxt textual sym_tab | 
| 43176 
29a3a1a7794d
only uncombine combinators in textual Isar proofs, not in Metis
 blanchet parents: 
43168diff
changeset | 559 | #> textual ? uncombine_term thy #> infer_formula_types ctxt | 
| 43136 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 560 | end | 
| 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 561 | |
| 43093 | 562 | (**** Translation of TSTP files to Isar proofs ****) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 563 | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 564 | fun unvarify_term (Var ((s, 0), T)) = Free (s, T) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 565 |   | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
 | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 566 | |
| 47774 | 567 | fun decode_line sym_tab (Definition_Step (name, phi1, phi2)) ctxt = | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 568 | let | 
| 42361 | 569 | val thy = Proof_Context.theory_of ctxt | 
| 43184 | 570 | val t1 = prop_from_atp ctxt true sym_tab phi1 | 
| 36551 | 571 | val vars = snd (strip_comb t1) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 572 | val frees = map unvarify_term vars | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 573 | val unvarify_args = subst_atomic (vars ~~ frees) | 
| 43184 | 574 | val t2 = prop_from_atp ctxt true sym_tab phi2 | 
| 36551 | 575 | val (t1, t2) = | 
| 576 | HOLogic.eq_const HOLogic.typeT $ t1 $ t2 | |
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 577 | |> unvarify_args |> uncombine_term thy |> infer_formula_types ctxt | 
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 578 | |> HOLogic.dest_eq | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 579 | in | 
| 47774 | 580 | (Definition_Step (name, t1, t2), | 
| 44121 | 581 | fold Variable.declare_term (maps Misc_Legacy.term_frees [t1, t2]) ctxt) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 582 | end | 
| 47774 | 583 | | decode_line sym_tab (Inference_Step (name, u, rule, deps)) ctxt = | 
| 43184 | 584 | let val t = u |> uncombined_etc_prop_from_atp ctxt true sym_tab in | 
| 47774 | 585 | (Inference_Step (name, t, rule, deps), | 
| 44121 | 586 | fold Variable.declare_term (Misc_Legacy.term_frees t) ctxt) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 587 | end | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 588 | fun decode_lines ctxt sym_tab lines = | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 589 | fst (fold_map (decode_line sym_tab) lines ctxt) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 590 | |
| 47774 | 591 | fun is_same_inference _ (Definition_Step _) = false | 
| 592 | | is_same_inference t (Inference_Step (_, t', _, _)) = t aconv t' | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 593 | |
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 594 | (* No "real" literals means only type information (tfree_tcs, clsrel, or | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 595 | clsarity). *) | 
| 45740 | 596 | val is_only_type_information = curry (op aconv) @{term True}
 | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 597 | |
| 39373 | 598 | fun replace_one_dependency (old, new) dep = | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 599 | if is_same_atp_step dep old then new else [dep] | 
| 47774 | 600 | fun replace_dependencies_in_line _ (line as Definition_Step _) = line | 
| 601 | | replace_dependencies_in_line p (Inference_Step (name, t, rule, deps)) = | |
| 602 | Inference_Step (name, t, rule, | |
| 603 | fold (union (op =) o replace_one_dependency p) deps []) | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 604 | |
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40114diff
changeset | 605 | (* Discard facts; consolidate adjacent lines that prove the same formula, since | 
| 38085 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 blanchet parents: 
38066diff
changeset | 606 | they differ only in type information.*) | 
| 47774 | 607 | fun add_line _ (line as Definition_Step _) lines = line :: lines | 
| 608 | | add_line fact_names (Inference_Step (name as (_, ss), t, rule, [])) lines = | |
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40114diff
changeset | 609 | (* No dependencies: fact, conjecture, or (for Vampire) internal facts or | 
| 38085 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 blanchet parents: 
38066diff
changeset | 610 | definitions. *) | 
| 45552 | 611 | if is_fact fact_names ss then | 
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40114diff
changeset | 612 | (* Facts are not proof lines. *) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 613 | if is_only_type_information t then | 
| 39373 | 614 | map (replace_dependencies_in_line (name, [])) lines | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 615 | (* Is there a repetition? If so, replace later line by earlier one. *) | 
| 38035 | 616 | else case take_prefix (not o is_same_inference t) lines of | 
| 39373 | 617 | (_, []) => lines (* no repetition of proof line *) | 
| 47774 | 618 | | (pre, Inference_Step (name', _, _, _) :: post) => | 
| 39373 | 619 | pre @ map (replace_dependencies_in_line (name', [name])) post | 
| 40069 | 620 | | _ => raise Fail "unexpected inference" | 
| 45552 | 621 | else if is_conjecture ss then | 
| 47774 | 622 | Inference_Step (name, s_not t, rule, []) :: lines | 
| 36551 | 623 | else | 
| 39373 | 624 | map (replace_dependencies_in_line (name, [])) lines | 
| 47774 | 625 | | add_line _ (Inference_Step (name, t, rule, deps)) lines = | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 626 | (* Type information will be deleted later; skip repetition test. *) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 627 | if is_only_type_information t then | 
| 47774 | 628 | Inference_Step (name, t, rule, deps) :: lines | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 629 | (* Is there a repetition? If so, replace later line by earlier one. *) | 
| 38035 | 630 | else case take_prefix (not o is_same_inference t) lines of | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 631 | (* FIXME: Doesn't this code risk conflating proofs involving different | 
| 38035 | 632 | types? *) | 
| 47774 | 633 | (_, []) => Inference_Step (name, t, rule, deps) :: lines | 
| 634 | | (pre, Inference_Step (name', t', rule, _) :: post) => | |
| 635 | Inference_Step (name, t', rule, deps) :: | |
| 39373 | 636 | pre @ map (replace_dependencies_in_line (name', [name])) post | 
| 40069 | 637 | | _ => raise Fail "unexpected inference" | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 638 | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 639 | (* Recursively delete empty lines (type information) from the proof. *) | 
| 47774 | 640 | fun add_nontrivial_line (line as Inference_Step (name, t, _, [])) lines = | 
| 39373 | 641 | if is_only_type_information t then delete_dependency name lines | 
| 45209 | 642 | else line :: lines | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 643 | | add_nontrivial_line line lines = line :: lines | 
| 39373 | 644 | and delete_dependency name lines = | 
| 645 | fold_rev add_nontrivial_line | |
| 646 | (map (replace_dependencies_in_line (name, [])) lines) [] | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 647 | |
| 37323 | 648 | (* ATPs sometimes reuse free variable names in the strangest ways. Removing | 
| 649 | offending lines often does the trick. *) | |
| 36560 
45c1870f234f
fixed definition of "bad frees" so that it actually works
 blanchet parents: 
36559diff
changeset | 650 | fun is_bad_free frees (Free x) = not (member (op =) frees x) | 
| 
45c1870f234f
fixed definition of "bad frees" so that it actually works
 blanchet parents: 
36559diff
changeset | 651 | | is_bad_free _ _ = false | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 652 | |
| 47774 | 653 | fun add_desired_line _ _ _ (line as Definition_Step (name, _, _)) (j, lines) = | 
| 39373 | 654 | (j, line :: map (replace_dependencies_in_line (name, [])) lines) | 
| 45551 | 655 | | add_desired_line isar_shrink_factor fact_names frees | 
| 47774 | 656 | (Inference_Step (name as (_, ss), t, rule, deps)) (j, lines) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 657 | (j + 1, | 
| 45552 | 658 | if is_fact fact_names ss orelse | 
| 659 | is_conjecture ss orelse | |
| 39373 | 660 | (* the last line must be kept *) | 
| 661 | j = 0 orelse | |
| 36570 
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
 blanchet parents: 
36567diff
changeset | 662 | (not (is_only_type_information t) andalso | 
| 
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
 blanchet parents: 
36567diff
changeset | 663 | null (Term.add_tvars t []) andalso | 
| 
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
 blanchet parents: 
36567diff
changeset | 664 | not (exists_subterm (is_bad_free frees) t) andalso | 
| 39373 | 665 | length deps >= 2 andalso j mod isar_shrink_factor = 0 andalso | 
| 666 | (* kill next to last line, which usually results in a trivial step *) | |
| 667 | j <> 1) then | |
| 47774 | 668 | Inference_Step (name, t, rule, deps) :: lines (* keep line *) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 669 | else | 
| 39373 | 670 | map (replace_dependencies_in_line (name, deps)) lines) (* drop line *) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 671 | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 672 | (** Isar proof construction and manipulation **) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 673 | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 674 | type label = string * int | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 675 | type facts = label list * string list | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 676 | |
| 39452 | 677 | datatype isar_qualifier = Show | Then | Moreover | Ultimately | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 678 | |
| 39452 | 679 | datatype isar_step = | 
| 36478 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 680 | Fix of (string * typ) list | | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 681 | Let of term * term | | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 682 | Assume of label * term | | 
| 45882 | 683 | Prove of isar_qualifier list * label * term * byline | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 684 | and byline = | 
| 45882 | 685 | By_Metis of facts | | 
| 686 | Case_Split of isar_step list list * facts | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 687 | |
| 45552 | 688 | fun add_fact_from_dependency fact_names (name as (_, ss)) = | 
| 689 | if is_fact fact_names ss then | |
| 690 | apsnd (union (op =) (map fst (resolve_fact fact_names ss))) | |
| 36475 
05209b869a6b
new Isar proof construction code: stringfy axiom names correctly
 blanchet parents: 
36474diff
changeset | 691 | else | 
| 45551 | 692 | apfst (insert (op =) (raw_label_for_name name)) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 693 | |
| 39454 | 694 | fun repair_name "$true" = "c_True" | 
| 695 | | repair_name "$false" = "c_False" | |
| 43000 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 696 | | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *) | 
| 39454 | 697 | | repair_name s = | 
| 43000 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 698 | if is_tptp_equal s orelse | 
| 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 699 | (* seen in Vampire proofs *) | 
| 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 700 | (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then | 
| 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 701 | tptp_equal | 
| 39454 | 702 | else | 
| 703 | s | |
| 704 | ||
| 45883 | 705 | (* FIXME: Still needed? Try with SPASS proofs perhaps. *) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 706 | val kill_duplicate_assumptions_in_proof = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 707 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 708 | fun relabel_facts subst = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 709 | apfst (map (fn l => AList.lookup (op =) subst l |> the_default l)) | 
| 36491 
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
 blanchet parents: 
36488diff
changeset | 710 | fun do_step (step as Assume (l, t)) (proof, subst, assums) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 711 | (case AList.lookup (op aconv) assums t of | 
| 36967 
3c804030474b
fix bug in Isar proof reconstruction step relabeling + don't try to infer the sorts of TVars, since this often fails miserably
 blanchet parents: 
36924diff
changeset | 712 | SOME l' => (proof, (l, l') :: subst, assums) | 
| 36491 
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
 blanchet parents: 
36488diff
changeset | 713 | | NONE => (step :: proof, subst, (t, l) :: assums)) | 
| 45882 | 714 | | do_step (Prove (qs, l, t, by)) (proof, subst, assums) = | 
| 715 | (Prove (qs, l, t, | |
| 716 | case by of | |
| 717 | By_Metis facts => By_Metis (relabel_facts subst facts) | |
| 718 | | Case_Split (proofs, facts) => | |
| 719 | Case_Split (map do_proof proofs, | |
| 720 | relabel_facts subst facts)) :: | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 721 | proof, subst, assums) | 
| 36491 
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
 blanchet parents: 
36488diff
changeset | 722 | | do_step step (proof, subst, assums) = (step :: proof, subst, assums) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 723 | and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 724 | in do_proof end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 725 | |
| 45883 | 726 | fun used_labels_of_step (Prove (_, _, _, by)) = | 
| 727 | (case by of | |
| 728 | By_Metis (ls, _) => ls | |
| 729 | | Case_Split (proofs, (ls, _)) => | |
| 730 | fold (union (op =) o used_labels_of) proofs ls) | |
| 731 | | used_labels_of_step _ = [] | |
| 732 | and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof [] | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 733 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 734 | fun kill_useless_labels_in_proof proof = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 735 | let | 
| 36556 
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
 blanchet parents: 
36555diff
changeset | 736 | val used_ls = used_labels_of proof | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 737 | fun do_label l = if member (op =) used_ls l then l else no_label | 
| 36556 
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
 blanchet parents: 
36555diff
changeset | 738 | fun do_step (Assume (l, t)) = Assume (do_label l, t) | 
| 45882 | 739 | | do_step (Prove (qs, l, t, by)) = | 
| 740 | Prove (qs, do_label l, t, | |
| 741 | case by of | |
| 742 | Case_Split (proofs, facts) => | |
| 743 | Case_Split (map (map do_step) proofs, facts) | |
| 744 | | _ => by) | |
| 36556 
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
 blanchet parents: 
36555diff
changeset | 745 | | do_step step = step | 
| 
81dc2c20f052
use readable names in "debug" mode for type vars + don't pipe facts using "using" but rather give them directly to metis (works better with type variables)
 blanchet parents: 
36555diff
changeset | 746 | in map do_step proof end | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 747 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 748 | fun prefix_for_depth n = replicate_string (n + 1) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 749 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 750 | val relabel_proof = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 751 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 752 | fun aux _ _ _ [] = [] | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 753 | | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 754 | if l = no_label then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 755 | Assume (l, t) :: aux subst depth (next_assum, next_fact) proof | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 756 | else | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 757 | let val l' = (prefix_for_depth depth assum_prefix, next_assum) in | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 758 | Assume (l', t) :: | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 759 | aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 760 | end | 
| 45882 | 761 | | aux subst depth (next_assum, next_fact) | 
| 762 | (Prove (qs, l, t, by) :: proof) = | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 763 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 764 | val (l', subst, next_fact) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 765 | if l = no_label then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 766 | (l, subst, next_fact) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 767 | else | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 768 | let | 
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 769 | val l' = (prefix_for_depth depth have_prefix, next_fact) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 770 | in (l', (l, l') :: subst, next_fact + 1) end | 
| 36570 
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
 blanchet parents: 
36567diff
changeset | 771 | val relabel_facts = | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 772 | apfst (maps (the_list o AList.lookup (op =) subst)) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 773 | val by = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 774 | case by of | 
| 45882 | 775 | By_Metis facts => By_Metis (relabel_facts facts) | 
| 776 | | Case_Split (proofs, facts) => | |
| 777 | Case_Split (map (aux subst (depth + 1) (1, 1)) proofs, | |
| 778 | relabel_facts facts) | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 779 | in | 
| 45882 | 780 | Prove (qs, l', t, by) :: aux subst depth (next_assum, next_fact) proof | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 781 | end | 
| 36491 
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
 blanchet parents: 
36488diff
changeset | 782 | | aux subst depth nextp (step :: proof) = | 
| 
6769f8eacaac
unskolemize formulas in proof reconstruction + detect newer SPASS versions to avoid truncating identifiers if not necessary (truncating confuses proof reconstruction)
 blanchet parents: 
36488diff
changeset | 783 | step :: aux subst depth nextp proof | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 784 | in aux [] 0 (1, 1) end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 785 | |
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 786 | fun string_for_proof ctxt0 type_enc lam_trans i n = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 787 | let | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 788 | val ctxt = | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 789 | ctxt0 |> Config.put show_free_types false | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 790 | |> Config.put show_types true | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 791 | |> Config.put show_sorts true | 
| 37319 
42268dc7d6c4
show types in Isar proofs, but not for free variables;
 blanchet parents: 
37172diff
changeset | 792 | fun fix_print_mode f x = | 
| 39134 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
39115diff
changeset | 793 | Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) | 
| 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
39115diff
changeset | 794 | (print_mode_value ())) f x | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 795 | fun do_indent ind = replicate_string (ind * indent_size) " " | 
| 36478 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 796 | fun do_free (s, T) = | 
| 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 797 | maybe_quote s ^ " :: " ^ | 
| 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 798 | maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T) | 
| 36570 
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
 blanchet parents: 
36567diff
changeset | 799 | fun do_label l = if l = no_label then "" else string_for_label l ^ ": " | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 800 | fun do_have qs = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 801 | (if member (op =) qs Moreover then "moreover " else "") ^ | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 802 | (if member (op =) qs Ultimately then "ultimately " else "") ^ | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 803 | (if member (op =) qs Then then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 804 | if member (op =) qs Show then "thus" else "hence" | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 805 | else | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 806 | if member (op =) qs Show then "show" else "have") | 
| 36478 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 807 | val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt) | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 808 | val reconstr = Metis (type_enc, lam_trans) | 
| 36570 
9bebcb40599f
identify axioms/conjectures more reliably in ATP proofs (an empty dependency list doesn't always indicate an axiom or conjecture!)
 blanchet parents: 
36567diff
changeset | 809 | fun do_facts (ls, ss) = | 
| 45520 | 810 | reconstructor_command reconstr 1 1 | 
| 43033 | 811 | (ls |> sort_distinct (prod_ord string_ord int_ord), | 
| 812 | ss |> sort_distinct string_ord) | |
| 36478 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 813 | and do_step ind (Fix xs) = | 
| 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 814 | do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n" | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 815 | | do_step ind (Let (t1, t2)) = | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 816 | do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n" | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 817 | | do_step ind (Assume (l, t)) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 818 | do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n" | 
| 45882 | 819 | | do_step ind (Prove (qs, l, t, By_Metis facts)) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 820 | do_indent ind ^ do_have qs ^ " " ^ | 
| 36479 | 821 | do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n" | 
| 45882 | 822 | | do_step ind (Prove (qs, l, t, Case_Split (proofs, facts))) = | 
| 823 | implode (map (prefix (do_indent ind ^ "moreover\n") o do_block ind) | |
| 824 | proofs) ^ | |
| 36479 | 825 | do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^ | 
| 36478 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 826 | do_facts facts ^ "\n" | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 827 | and do_steps prefix suffix ind steps = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 828 | let val s = implode (map (do_step ind) steps) in | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 829 | replicate_string (ind * indent_size - size prefix) " " ^ prefix ^ | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 830 | String.extract (s, ind * indent_size, | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 831 | SOME (size s - ind * indent_size - 1)) ^ | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 832 | suffix ^ "\n" | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 833 | end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 834 |     and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
 | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 835 | (* One-step proofs are pointless; better use the Metis one-liner | 
| 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 836 | directly. *) | 
| 45882 | 837 | and do_proof [Prove (_, _, _, By_Metis _)] = "" | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 838 | | do_proof proof = | 
| 36480 
1e01a7162435
polish Isar proofs: don't mention facts twice, and don't show one-liner "structured" proofs
 blanchet parents: 
36479diff
changeset | 839 | (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ | 
| 39452 | 840 | do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^ | 
| 841 | (if n <> 1 then "next" else "qed") | |
| 36488 
32c92af68ec9
remove Sledgehammer's "sorts" option to annotate variables with sorts in proof;
 blanchet parents: 
36486diff
changeset | 842 | in do_proof end | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 843 | |
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 844 | fun isar_proof_text ctxt isar_proof_requested | 
| 45552 | 845 | (debug, isar_shrink_factor, pool, fact_names, sym_tab, atp_proof, goal) | 
| 43037 | 846 | (one_line_params as (_, _, _, _, subgoal, subgoal_count)) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 847 | let | 
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 848 | val isar_shrink_factor = | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 849 | (if isar_proof_requested then 1 else 2) * isar_shrink_factor | 
| 43037 | 850 | val (params, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 851 | val frees = fold Term.add_frees (concl_t :: hyp_ts) [] | 
| 43033 | 852 | val one_line_proof = one_line_proof_text one_line_params | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 853 | val type_enc = | 
| 45590 | 854 | if is_typed_helper_used_in_atp_proof atp_proof then full_typesN | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 855 | else partial_typesN | 
| 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 856 | val lam_trans = lam_trans_from_atp_proof atp_proof metis_default_lam_trans | 
| 45882 | 857 | |
| 45883 | 858 | fun isar_proof_of () = | 
| 859 | let | |
| 860 | val atp_proof = | |
| 861 | atp_proof | |
| 862 | |> clean_up_atp_proof_dependencies | |
| 863 | |> nasty_atp_proof pool | |
| 864 | |> map_term_names_in_atp_proof repair_name | |
| 865 | |> decode_lines ctxt sym_tab | |
| 866 | |> rpair [] |-> fold_rev (add_line fact_names) | |
| 867 | |> rpair [] |-> fold_rev add_nontrivial_line | |
| 868 | |> rpair (0, []) | |
| 869 | |-> fold_rev (add_desired_line isar_shrink_factor fact_names frees) | |
| 870 | |> snd | |
| 871 | val conj_name = conjecture_prefix ^ string_of_int (length hyp_ts) | |
| 872 | val conjs = | |
| 873 | atp_proof | |
| 47774 | 874 | |> map_filter (fn Inference_Step (name as (_, ss), _, _, []) => | 
| 45883 | 875 | if member (op =) ss conj_name then SOME name else NONE | 
| 876 | | _ => NONE) | |
| 47774 | 877 | fun dep_of_step (Definition_Step _) = NONE | 
| 878 | | dep_of_step (Inference_Step (name, _, _, from)) = SOME (from, name) | |
| 45883 | 879 | val ref_graph = atp_proof |> map_filter dep_of_step |> make_ref_graph | 
| 880 | val axioms = axioms_of_ref_graph ref_graph conjs | |
| 881 | val tainted = tainted_atoms_of_ref_graph ref_graph conjs | |
| 882 | val props = | |
| 883 | Symtab.empty | |
| 47774 | 884 | |> fold (fn Definition_Step _ => I (* FIXME *) | 
| 885 | | Inference_Step ((s, _), t, _, _) => | |
| 45883 | 886 | Symtab.update_new (s, | 
| 887 | t |> member (op = o apsnd fst) tainted s ? s_not)) | |
| 888 | atp_proof | |
| 889 | (* FIXME: add "fold_rev forall_of (map Var (Term.add_vars t []))"? *) | |
| 890 | fun prop_of_clause c = | |
| 891 | fold (curry s_disj) (map_filter (Symtab.lookup props o fst) c) | |
| 892 |                @{term False}
 | |
| 893 | fun label_of_clause c = (space_implode "___" (map fst c), 0) | |
| 894 | fun maybe_show outer c = | |
| 895 | (outer andalso length c = 1 andalso subset (op =) (c, conjs)) | |
| 896 | ? cons Show | |
| 897 | fun do_have outer qs (gamma, c) = | |
| 898 | Prove (maybe_show outer c qs, label_of_clause c, prop_of_clause c, | |
| 899 | By_Metis (fold (add_fact_from_dependency fact_names | |
| 900 | o the_single) gamma ([], []))) | |
| 901 | fun do_inf outer (Have z) = do_have outer [] z | |
| 902 | | do_inf outer (Hence z) = do_have outer [Then] z | |
| 903 | | do_inf outer (Cases cases) = | |
| 904 | let val c = succedent_of_cases cases in | |
| 905 | Prove (maybe_show outer c [Ultimately], label_of_clause c, | |
| 906 | prop_of_clause c, | |
| 907 | Case_Split (map (do_case false) cases, ([], []))) | |
| 908 | end | |
| 909 | and do_case outer (c, infs) = | |
| 910 | Assume (label_of_clause c, prop_of_clause c) :: | |
| 911 | map (do_inf outer) infs | |
| 912 | val isar_proof = | |
| 913 | (if null params then [] else [Fix params]) @ | |
| 914 | (ref_graph | |
| 915 | |> redirect_graph axioms tainted | |
| 916 | |> chain_direct_proof | |
| 917 | |> map (do_inf true) | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 918 | |> kill_duplicate_assumptions_in_proof | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 919 | |> kill_useless_labels_in_proof | 
| 45883 | 920 | |> relabel_proof) | 
| 921 | |> string_for_proof ctxt type_enc lam_trans subgoal subgoal_count | |
| 922 | in | |
| 923 | case isar_proof of | |
| 924 | "" => | |
| 925 | if isar_proof_requested then | |
| 926 | "\nNo structured proof available (proof too short)." | |
| 927 | else | |
| 928 | "" | |
| 929 | | _ => | |
| 930 | "\n\n" ^ (if isar_proof_requested then "Structured proof" | |
| 931 | else "Perhaps this will work") ^ | |
| 932 | ":\n" ^ Markup.markup Isabelle_Markup.sendback isar_proof | |
| 933 | end | |
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 934 | val isar_proof = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 935 | if debug then | 
| 45883 | 936 | isar_proof_of () | 
| 937 | else case try isar_proof_of () of | |
| 45882 | 938 | SOME s => s | 
| 939 | | NONE => if isar_proof_requested then | |
| 940 | "\nWarning: The Isar proof construction failed." | |
| 941 | else | |
| 942 | "" | |
| 43033 | 943 | in one_line_proof ^ isar_proof end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 944 | |
| 43033 | 945 | fun proof_text ctxt isar_proof isar_params | 
| 946 | (one_line_params as (preplay, _, _, _, _, _)) = | |
| 43166 | 947 | (if case preplay of Failed_to_Play _ => true | _ => isar_proof then | 
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 948 | isar_proof_text ctxt isar_proof isar_params | 
| 43033 | 949 | else | 
| 950 | one_line_proof_text) one_line_params | |
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 951 | |
| 31038 | 952 | end; |