| author | blanchet | 
| Mon, 06 Jun 2011 20:36:35 +0200 | |
| changeset 43178 | b5862142d378 | 
| parent 43176 | 29a3a1a7794d | 
| child 43182 | 649bada59658 | 
| permissions | -rw-r--r-- | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 1 | (* Title: HOL/Tools/ATP/atp_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 | ||
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 9 | signature ATP_RECONSTRUCT = | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 10 | sig | 
| 43094 | 11 | type 'a fo_term = 'a ATP_Problem.fo_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 | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 14 | type locality = ATP_Translate.locality | 
| 43102 | 15 | type type_sys = ATP_Translate.type_sys | 
| 43033 | 16 | |
| 17 | datatype reconstructor = | |
| 18 | Metis | | |
| 19 | MetisFT | | |
| 43168 
467d5b34e1f5
temporarily added "MetisX" as reconstructor and minimizer
 blanchet parents: 
43166diff
changeset | 20 | MetisX | | 
| 43033 | 21 | SMT of string | 
| 22 | ||
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 23 | datatype play = | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 24 | Played of reconstructor * Time.time | | 
| 43033 | 25 | Trust_Playable of reconstructor * Time.time option| | 
| 43166 | 26 | Failed_to_Play of reconstructor | 
| 43033 | 27 | |
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 28 | type minimize_command = string list -> string | 
| 43033 | 29 | type one_line_params = | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 30 | play * string * (string * locality) list * minimize_command * int * int | 
| 38818 
61cf050f8b2e
improve SPASS hack, when a clause comes from several facts
 blanchet parents: 
38752diff
changeset | 31 | type isar_params = | 
| 43102 | 32 | bool * bool * int * type_sys * string Symtab.table * int list list * int | 
| 33 | * (string * locality) list vector * int Symtab.table * string proof * thm | |
| 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 | 34 | val repair_conjecture_shape_and_fact_names : | 
| 43102 | 35 | type_sys -> string -> int list list -> int | 
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 36 | -> (string * locality) list vector -> int list | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 37 | -> int list list * int * (string * locality) list vector * int list | 
| 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 | 38 | val used_facts_in_atp_proof : | 
| 43102 | 39 | Proof.context -> type_sys -> int -> (string * locality) list vector | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 40 | -> string proof -> (string * locality) list | 
| 42876 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 41 | val used_facts_in_unsound_atp_proof : | 
| 43102 | 42 | Proof.context -> type_sys -> int list list -> int | 
| 43033 | 43 | -> (string * locality) list vector -> 'a proof -> string list option | 
| 44 | val uses_typed_helpers : int list -> 'a proof -> bool | |
| 43051 | 45 | val reconstructor_name : reconstructor -> string | 
| 43033 | 46 | 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 | 47 | 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 | 48 | val make_tfree : Proof.context -> string -> typ | 
| 43094 | 49 | val term_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 | 50 | Proof.context -> bool -> int Symtab.table -> typ option -> string fo_term | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 51 | -> term | 
| 43127 | 52 | 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 | 53 | Proof.context -> bool -> int Symtab.table | 
| 43127 | 54 | -> (string, string, string fo_term) formula -> term | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 55 | val isar_proof_text : | 
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 56 | 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 | 57 | val proof_text : | 
| 43033 | 58 | 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 | 59 | end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 60 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 61 | structure ATP_Reconstruct : ATP_RECONSTRUCT = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 62 | struct | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 63 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 64 | open ATP_Util | 
| 38028 | 65 | open ATP_Problem | 
| 39452 | 66 | open ATP_Proof | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 67 | open ATP_Translate | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 68 | |
| 43033 | 69 | datatype reconstructor = | 
| 70 | Metis | | |
| 71 | MetisFT | | |
| 43168 
467d5b34e1f5
temporarily added "MetisX" as reconstructor and minimizer
 blanchet parents: 
43166diff
changeset | 72 | MetisX | | 
| 43033 | 73 | SMT of string | 
| 74 | ||
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 75 | datatype play = | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 76 | Played of reconstructor * Time.time | | 
| 43033 | 77 | Trust_Playable of reconstructor * Time.time option | | 
| 43166 | 78 | Failed_to_Play of reconstructor | 
| 43033 | 79 | |
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 80 | type minimize_command = string list -> string | 
| 43033 | 81 | type one_line_params = | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 82 | play * string * (string * locality) list * minimize_command * int * int | 
| 38818 
61cf050f8b2e
improve SPASS hack, when a clause comes from several facts
 blanchet parents: 
38752diff
changeset | 83 | type isar_params = | 
| 43102 | 84 | bool * bool * int * type_sys * string Symtab.table * int list list * int | 
| 43037 | 85 | * (string * locality) list vector * int Symtab.table * string proof * thm | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 86 | |
| 39500 | 87 | fun is_head_digit s = Char.isDigit (String.sub (s, 0)) | 
| 88 | val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode) | |
| 89 | ||
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 90 | val is_typed_helper_name = | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 91 | String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 92 | |
| 39500 | 93 | fun find_first_in_list_vector vec key = | 
| 94 | Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key | |
| 95 | | (_, value) => value) NONE vec | |
| 96 | ||
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 97 | |
| 43039 | 98 | (** SPASS's FLOTTER hack **) | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 99 | |
| 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 | 100 | (* This is a hack required for keeping track of facts after they have been | 
| 43039 | 101 | clausified by SPASS's FLOTTER preprocessor. The "ATP/scripts/spass" script is | 
| 102 | also part of this hack. *) | |
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 103 | |
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 104 | val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation" | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 105 | |
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 106 | fun extract_clause_sequence output = | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 107 | let | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 108 | val tokens_of = String.tokens (not o Char.isAlphaNum) | 
| 42587 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 blanchet parents: 
42579diff
changeset | 109 |     fun extract_num ("clause" :: (ss as _ :: _)) = Int.fromString (List.last ss)
 | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 110 | | extract_num _ = NONE | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 111 | in output |> split_lines |> map_filter (extract_num o tokens_of) end | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 112 | |
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 113 | val parse_clause_formula_pair = | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 114 |   $$ "(" |-- scan_integer --| $$ ","
 | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 115 | -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")" | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 116 | --| Scan.option ($$ ",") | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 117 | val parse_clause_formula_relation = | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 118 |   Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
 | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 119 | |-- Scan.repeat parse_clause_formula_pair | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 120 | val extract_clause_formula_relation = | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 121 | Substring.full #> Substring.position set_ClauseFormulaRelationN | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 122 | #> snd #> Substring.position "." #> fst #> Substring.string | 
| 40627 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 wenzelm parents: 
40204diff
changeset | 123 | #> raw_explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 124 | #> fst | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 125 | |
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 126 | fun maybe_unprefix_fact_number type_sys = | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 127 | polymorphism_of_type_sys type_sys <> Polymorphic | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 128 | ? (space_implode "_" o tl o space_explode "_") | 
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 129 | |
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 130 | fun repair_conjecture_shape_and_fact_names type_sys output conjecture_shape | 
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 131 | fact_offset fact_names typed_helpers = | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 132 | if String.isSubstring set_ClauseFormulaRelationN output then | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 133 | let | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 134 | val j0 = hd (hd conjecture_shape) | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 135 | val seq = extract_clause_sequence output | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 136 | val name_map = extract_clause_formula_relation output | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 137 | fun renumber_conjecture j = | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 138 | conjecture_prefix ^ string_of_int (j - j0) | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 139 | |> AList.find (fn (s, ss) => member (op =) ss s) name_map | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 140 | |> map (fn s => find_index (curry (op =) s) seq + 1) | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 141 | fun names_for_number j = | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 142 | j |> AList.lookup (op =) name_map |> these | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 143 | |> map_filter (try (unascii_of o maybe_unprefix_fact_number type_sys | 
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 144 | o unprefix fact_prefix)) | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 145 | |> map (fn name => | 
| 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 | 146 | (name, name |> find_first_in_list_vector fact_names |> the) | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 147 | handle Option.Option => | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 148 |                             error ("No such fact: " ^ quote name ^ "."))
 | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 149 | in | 
| 42587 
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
 blanchet parents: 
42579diff
changeset | 150 | (conjecture_shape |> map (maps renumber_conjecture), 0, | 
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 151 | seq |> map names_for_number |> Vector.fromList, | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 152 | name_map |> filter (forall is_typed_helper_name o snd) |> map fst) | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 153 | end | 
| 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 154 | else | 
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 155 | (conjecture_shape, fact_offset, fact_names, typed_helpers) | 
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39455diff
changeset | 156 | |
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 157 | val vampire_step_prefix = "f" (* grrr... *) | 
| 41203 
1393514094d7
fixed more issues with the Vampire output parser, and added support for Vampire's TSTP output (--proof tptp)
 blanchet parents: 
41151diff
changeset | 158 | |
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 159 | val extract_step_number = | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 160 | Int.fromString o perhaps (try (unprefix vampire_step_prefix)) | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 161 | |
| 42751 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 162 | fun resolve_fact type_sys _ fact_names (_, SOME s) = | 
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 163 | (case try (unprefix fact_prefix) s of | 
| 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 164 | SOME s' => | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 165 | let val s' = s' |> maybe_unprefix_fact_number type_sys |> unascii_of in | 
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 166 | case find_first_in_list_vector fact_names s' of | 
| 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 167 | SOME x => [(s', x)] | 
| 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 168 | | NONE => [] | 
| 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 169 | end | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 170 | | NONE => []) | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 171 | | resolve_fact _ facts_offset fact_names (num, NONE) = | 
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 172 | (case extract_step_number num of | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 173 | SOME j => | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 174 | let val j = j - facts_offset in | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 175 | if j > 0 andalso j <= Vector.length fact_names then | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 176 | Vector.sub (fact_names, j - 1) | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 177 | else | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 178 | [] | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 179 | end | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 180 | | NONE => []) | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 181 | |
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 182 | fun is_fact type_sys conjecture_shape = | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 183 | not o null o resolve_fact type_sys 0 conjecture_shape | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 184 | |
| 42751 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 185 | fun resolve_conjecture _ (_, SOME s) = | 
| 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 186 | (case try (unprefix conjecture_prefix) s of | 
| 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 187 | SOME s' => | 
| 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 188 | (case Int.fromString s' of | 
| 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 189 | SOME j => [j] | 
| 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 190 | | NONE => []) | 
| 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 191 | | NONE => []) | 
| 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 192 | | resolve_conjecture conjecture_shape (num, NONE) = | 
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 193 | case extract_step_number num of | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 194 | SOME i => (case find_index (exists (curry (op =) i)) conjecture_shape of | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 195 | ~1 => [] | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 196 | | j => [j]) | 
| 42751 
f42fd9754724
fixed conjecture resolution bug for Vampire 1.0's TSTP output
 blanchet parents: 
42680diff
changeset | 197 | | NONE => [] | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 198 | |
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 199 | fun is_conjecture conjecture_shape = | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 200 | not o null o resolve_conjecture conjecture_shape | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 201 | |
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 202 | fun is_typed_helper _ (_, SOME s) = is_typed_helper_name s | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 203 | | is_typed_helper typed_helpers (num, NONE) = | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 204 | (case extract_step_number num of | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 205 | SOME i => member (op =) typed_helpers i | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 206 | | NONE => false) | 
| 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 207 | |
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 208 | val leo2_ext = "extcnf_equal_neg" | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 209 | val isa_ext = Thm.get_name_hint @{thm ext}
 | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 210 | val isa_short_ext = Long_Name.base_name isa_ext | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 211 | |
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 212 | fun ext_name ctxt = | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 213 |   if Thm.eq_thm_prop (@{thm ext},
 | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 214 | singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 215 | isa_short_ext | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 216 | else | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 217 | isa_ext | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 218 | |
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 219 | fun add_fact _ type_sys facts_offset fact_names (Inference (name, _, [])) = | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 220 | union (op =) (resolve_fact type_sys facts_offset fact_names name) | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 221 | | add_fact ctxt _ _ _ (Inference (_, _, deps)) = | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 222 | if AList.defined (op =) deps leo2_ext then | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 223 | insert (op =) (ext_name ctxt, General (* or Chained... *)) | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 224 | else | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 225 | I | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 226 | | add_fact _ _ _ _ _ = I | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 227 | |
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 228 | fun used_facts_in_atp_proof ctxt type_sys facts_offset fact_names atp_proof = | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 229 | if null atp_proof then Vector.foldl (op @) [] fact_names | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 230 | else fold (add_fact ctxt type_sys facts_offset fact_names) atp_proof [] | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 231 | |
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 232 | fun is_conjecture_referred_to_in_proof conjecture_shape = | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 233 | exists (fn Inference (name, _, []) => is_conjecture conjecture_shape name | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 234 | | _ => false) | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 235 | |
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 236 | fun used_facts_in_unsound_atp_proof ctxt type_sys conjecture_shape facts_offset | 
| 42876 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 237 | fact_names atp_proof = | 
| 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 238 | let | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 239 | val used_facts = | 
| 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 240 | used_facts_in_atp_proof ctxt type_sys facts_offset fact_names atp_proof | 
| 42876 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 241 | in | 
| 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 242 | if forall (is_locality_global o snd) used_facts andalso | 
| 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 243 | not (is_conjecture_referred_to_in_proof conjecture_shape atp_proof) then | 
| 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 244 | SOME (map fst used_facts) | 
| 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 245 | else | 
| 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 246 | NONE | 
| 
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
 blanchet parents: 
42761diff
changeset | 247 | end | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 248 | |
| 43033 | 249 | fun uses_typed_helpers typed_helpers = | 
| 250 | exists (fn Inference (name, _, []) => is_typed_helper typed_helpers name | |
| 251 | | _ => false) | |
| 252 | ||
| 253 | ||
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 254 | (** Soft-core proof reconstruction: Metis one-liner **) | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 255 | |
| 43033 | 256 | fun reconstructor_name Metis = "metis" | 
| 257 | | reconstructor_name MetisFT = "metisFT" | |
| 43168 
467d5b34e1f5
temporarily added "MetisX" as reconstructor and minimizer
 blanchet parents: 
43166diff
changeset | 258 | | reconstructor_name MetisX = "metisX" | 
| 43033 | 259 | | reconstructor_name (SMT _) = "smt" | 
| 260 | ||
| 261 | fun reconstructor_settings (SMT settings) = settings | |
| 262 | | reconstructor_settings _ = "" | |
| 263 | ||
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 264 | 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 | 265 | |
| 43033 | 266 | fun show_time NONE = "" | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 267 |   | show_time (SOME ext_time) = " (" ^ string_from_ext_time ext_time ^ ")"
 | 
| 43033 | 268 | |
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 269 | fun set_settings "" = "" | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 270 | | set_settings settings = "using [[" ^ settings ^ "]] " | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 271 | fun apply_on_subgoal settings _ 1 = set_settings settings ^ "by " | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 272 | | apply_on_subgoal settings 1 _ = set_settings settings ^ "apply " | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 273 | | apply_on_subgoal settings i n = | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 274 | "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal settings 1 n | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 275 | fun command_call name [] = name | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 276 |   | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
 | 
| 43033 | 277 | fun try_command_line banner time command = | 
| 278 | banner ^ ": " ^ Markup.markup Markup.sendback command ^ show_time time ^ "." | |
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 279 | fun using_labels [] = "" | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 280 | | using_labels ls = | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 281 | "using " ^ space_implode " " (map string_for_label ls) ^ " " | 
| 43033 | 282 | fun reconstructor_command reconstructor i n (ls, ss) = | 
| 283 | using_labels ls ^ | |
| 284 | apply_on_subgoal (reconstructor_settings reconstructor) i n ^ | |
| 285 | command_call (reconstructor_name reconstructor) ss | |
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 286 | fun minimize_line _ [] = "" | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 287 | | minimize_line minimize_command ss = | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 288 | case minimize_command ss of | 
| 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 289 | "" => "" | 
| 43006 | 290 | | command => "\nTo minimize: " ^ Markup.markup Markup.sendback command ^ "." | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 291 | |
| 40723 
a82badd0e6ef
put facts found by SMT solver in alphabetical order and omit chained facts, as was done already for ATP proofs
 blanchet parents: 
40627diff
changeset | 292 | val split_used_facts = | 
| 
a82badd0e6ef
put facts found by SMT solver in alphabetical order and omit chained facts, as was done already for ATP proofs
 blanchet parents: 
40627diff
changeset | 293 | List.partition (curry (op =) Chained o snd) | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 294 | #> pairself (sort_distinct (string_ord o pairself fst)) | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 295 | |
| 43033 | 296 | fun one_line_proof_text (preplay, banner, used_facts, minimize_command, | 
| 43037 | 297 | subgoal, subgoal_count) = | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 298 | let | 
| 43033 | 299 | val (chained, extra) = split_used_facts used_facts | 
| 43166 | 300 | val (failed, reconstructor, ext_time) = | 
| 43033 | 301 | case preplay of | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43039diff
changeset | 302 | Played (reconstructor, time) => | 
| 43166 | 303 | (false, reconstructor, (SOME (false, time))) | 
| 43033 | 304 | | Trust_Playable (reconstructor, time) => | 
| 43166 | 305 | (false, reconstructor, | 
| 43033 | 306 | case time of | 
| 307 | NONE => NONE | |
| 308 | | SOME time => | |
| 309 | if time = Time.zeroTime then NONE else SOME (true, time)) | |
| 43166 | 310 | | Failed_to_Play reconstructor => (true, reconstructor, NONE) | 
| 43033 | 311 | val try_line = | 
| 43166 | 312 | ([], map fst extra) | 
| 313 | |> reconstructor_command reconstructor subgoal subgoal_count | |
| 314 | |> (if failed then enclose "One-line proof reconstruction failed: " "." | |
| 315 | else try_command_line banner ext_time) | |
| 43033 | 316 | 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 | 317 | |
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 318 | (** 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 | 319 | |
| 38014 | 320 | (* Simple simplifications to ensure that sort annotations don't leave a trail of | 
| 321 | spurious "True"s. *) | |
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 322 | fun s_not (Const (@{const_name All}, T) $ Abs (s, T', t')) =
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 323 |     Const (@{const_name Ex}, T) $ Abs (s, T', s_not t')
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 324 |   | s_not (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 325 |     Const (@{const_name All}, T) $ Abs (s, T', s_not t')
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 326 |   | s_not (@{const HOL.implies} $ t1 $ t2) = @{const HOL.conj} $ t1 $ s_not t2
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 327 |   | s_not (@{const HOL.conj} $ t1 $ t2) =
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 328 |     @{const HOL.disj} $ s_not t1 $ s_not t2
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 329 |   | s_not (@{const HOL.disj} $ t1 $ t2) =
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 330 |     @{const HOL.conj} $ s_not t1 $ s_not t2
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 331 |   | s_not (@{const False}) = @{const True}
 | 
| 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 332 |   | s_not (@{const True}) = @{const False}
 | 
| 38014 | 333 |   | s_not (@{const Not} $ t) = t
 | 
| 334 |   | s_not t = @{const Not} $ t
 | |
| 335 | fun s_conj (@{const True}, t2) = t2
 | |
| 336 |   | s_conj (t1, @{const True}) = t1
 | |
| 337 | | s_conj p = HOLogic.mk_conj p | |
| 338 | fun s_disj (@{const False}, t2) = t2
 | |
| 339 |   | s_disj (t1, @{const False}) = t1
 | |
| 340 | | s_disj p = HOLogic.mk_disj p | |
| 341 | fun s_imp (@{const True}, t2) = t2
 | |
| 342 |   | s_imp (t1, @{const False}) = s_not t1
 | |
| 343 | | s_imp p = HOLogic.mk_imp p | |
| 344 | fun s_iff (@{const True}, t2) = t2
 | |
| 345 |   | s_iff (t1, @{const True}) = t1
 | |
| 346 | | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2 | |
| 347 | ||
| 39425 | 348 | fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t | 
| 349 | fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t | |
| 350 | ||
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 351 | 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 | 352 | 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 | 353 | 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 | 354 | 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 | 355 | end | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 356 | |
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 357 | val indent_size = 2 | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 358 | val no_label = ("", ~1)
 | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 359 | |
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 360 | val raw_prefix = "X" | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 361 | val assum_prefix = "A" | 
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 362 | val have_prefix = "F" | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 363 | |
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 364 | fun raw_label_for_name conjecture_shape name = | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 365 | case resolve_conjecture conjecture_shape name of | 
| 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 366 | [j] => (conjecture_prefix, j) | 
| 39455 | 367 | | _ => case Int.fromString (fst name) of | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 368 | SOME j => (raw_prefix, j) | 
| 39455 | 369 | | NONE => (raw_prefix ^ fst name, 0) | 
| 39453 
1740a2d6bef9
use the same TSTP/Vampire/SPASS parser for one-liners as for Isar proofs
 blanchet parents: 
39452diff
changeset | 370 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 371 | (**** INTERPRETATION OF TSTP SYNTAX TREES ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 372 | |
| 37991 | 373 | exception FO_TERM of string fo_term list | 
| 42531 
a462dbaa584f
added more rudimentary type support to Sledgehammer's ATP encoding
 blanchet parents: 
42526diff
changeset | 374 | exception FORMULA of (string, string, string fo_term) formula list | 
| 37991 | 375 | exception SAME of unit | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 376 | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 377 | (* Type variables are given the basic sort "HOL.type". Some will later be | 
| 37991 | 378 | 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 | 379 | 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 | 380 | let val Ts = map (typ_from_atp ctxt) us in | 
| 38748 | 381 | case strip_prefix_and_unascii type_const_prefix a of | 
| 37991 | 382 | SOME b => Type (invert_const b, Ts) | 
| 383 | | NONE => | |
| 384 | if not (null us) then | |
| 385 | raise FO_TERM [u] (* only "tconst"s have type arguments *) | |
| 38748 | 386 | else case strip_prefix_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 | 387 | SOME b => make_tfree ctxt b | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 388 | | NONE => | 
| 38748 | 389 | case strip_prefix_and_unascii tvar_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 | 390 | SOME b => make_tvar b | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 391 | | NONE => | 
| 37991 | 392 | (* Variable from the ATP, say "X1" *) | 
| 393 | Type_Infer.param 0 (a, HOLogic.typeS) | |
| 394 | end | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 395 | |
| 38014 | 396 | (* Type class literal applied to a type. Returns triple of polarity, class, | 
| 397 | 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 | 398 | fun type_constraint_from_term 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 | 399 | case (strip_prefix_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 | 400 | (SOME b, [T]) => (b, T) | 
| 38014 | 401 | | _ => raise FO_TERM [u] | 
| 402 | ||
| 38085 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 blanchet parents: 
38066diff
changeset | 403 | (** Accumulate type constraints in a formula: negative type literals **) | 
| 38014 | 404 | 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 | 405 | 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 | 406 | | 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 | 407 | | add_type_constraint _ _ = I | 
| 38014 | 408 | |
| 43094 | 409 | fun repair_variable_name f s = | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 410 | let | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 411 | fun subscript_name s n = s ^ nat_subscript n | 
| 38488 | 412 | val s = String.map f s | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 413 | in | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 414 | case space_explode "_" s of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 415 | [_] => (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 | 416 | (cs1 as _ :: _, cs2 as _ :: _) => | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 417 | subscript_name (String.implode cs1) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 418 | (the (Int.fromString (String.implode cs2))) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 419 | | (_, _) => s) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 420 | | [s1, s2] => (case Int.fromString s2 of | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 421 | SOME n => subscript_name s1 n | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 422 | | NONE => s) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 423 | | _ => s | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 424 | end | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 425 | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 426 | (* First-order translation. No types are known for variables. "HOLogic.typeT" | 
| 38014 | 427 | 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 | 428 | 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 | 429 | 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 | 430 | val thy = Proof_Context.theory_of ctxt | 
| 43094 | 431 | (* see also "mk_var" in "Metis_Reconstruct" *) | 
| 432 | 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 | 433 | 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 | 434 | case u of | 
| 42539 
f6ba908b8b27
generate typing for "hBOOL" in "Many_Typed" mode + tuning
 blanchet parents: 
42531diff
changeset | 435 | ATerm (a, us) => | 
| 42962 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42943diff
changeset | 436 | if String.isPrefix simple_type_prefix a then | 
| 
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
 blanchet parents: 
42943diff
changeset | 437 |           @{const True} (* ignore TPTP type information *)
 | 
| 43000 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 438 | else if a = tptp_equal then | 
| 43093 | 439 | let val ts = map (do_term [] NONE) us in | 
| 43094 | 440 | if textual andalso length ts = 2 andalso | 
| 441 | hd ts aconv List.last ts then | |
| 39106 | 442 | (* Vampire is keen on producing these. *) | 
| 443 |               @{const True}
 | |
| 444 | else | |
| 445 |               list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
 | |
| 446 | end | |
| 43000 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 447 | else case strip_prefix_and_unascii const_prefix a of | 
| 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 448 | SOME s => | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 449 | let | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 450 | val ((s', s), mangled_us) = s |> unmangled_const |>> `invert_const | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 451 | 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 | 452 | 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 | 453 | 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 | 454 | [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 | 455 | do_term extra_ts (SOME (typ_from_atp ctxt typ_u)) term_u | 
| 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 | 456 | | _ => raise FO_TERM us | 
| 42966 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 blanchet parents: 
42962diff
changeset | 457 | else if s' = predicator_name then | 
| 43093 | 458 |               do_term [] (SOME @{typ bool}) (hd us)
 | 
| 42966 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 blanchet parents: 
42962diff
changeset | 459 | else if s' = app_op_name then | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 460 | 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 | 461 | do_term (extra_t :: extra_ts) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 462 | (case opt_T of | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 463 | SOME T => SOME (fastype_of extra_t --> T) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 464 | | NONE => NONE) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 465 | (nth us (length us - 2)) | 
| 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 466 | end | 
| 42966 
4e2d6c1e5392
more work on parsing LEO-II proofs without lambdas
 blanchet parents: 
42962diff
changeset | 467 | else if s' = type_pred_name then | 
| 42551 
cd99d6d3027a
reconstruct TFF type predicates correctly for ToFoF
 blanchet parents: 
42549diff
changeset | 468 |               @{const True} (* ignore type predicates *)
 | 
| 42549 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 469 | else | 
| 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 470 | let | 
| 42755 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 blanchet parents: 
42751diff
changeset | 471 | 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 | 472 | 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 | 473 | val (type_us, term_us) = | 
| 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 474 | chop num_ty_args us |>> append mangled_us | 
| 43093 | 475 | 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 | 476 | val T = | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 477 | if not (null type_us) andalso | 
| 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43062diff
changeset | 478 | num_type_args thy s' = length type_us 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 | 479 | (s', map (typ_from_atp ctxt) type_us) | 
| 43093 | 480 | |> Sign.const_instance thy | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 481 | else case opt_T of | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 482 | SOME T => map fastype_of term_ts ---> T | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 483 | | NONE => HOLogic.typeT | 
| 42755 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 blanchet parents: 
42751diff
changeset | 484 | val s' = s' |> unproxify_const | 
| 
4603154a3018
robustly detect how many type args were passed to the ATP, even if some of them were omitted
 blanchet parents: 
42751diff
changeset | 485 | in list_comb (Const (s', T), term_ts @ extra_ts) end | 
| 42549 
b9754f26c7bc
handle special constants correctly in Isar proof reconstruction code, especially type predicates
 blanchet parents: 
42544diff
changeset | 486 | end | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 487 | | NONE => (* a free or schematic variable *) | 
| 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 488 | let | 
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 489 | val ts = map (do_term [] NONE) us @ extra_ts | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 490 | val T = map fastype_of ts ---> HOLogic.typeT | 
| 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 491 | val t = | 
| 38748 | 492 | case strip_prefix_and_unascii fixed_var_prefix a of | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 493 | SOME b => Free (b, T) | 
| 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 494 | | NONE => | 
| 38748 | 495 | case strip_prefix_and_unascii schematic_var_prefix a of | 
| 43094 | 496 | SOME b => Var ((b, var_index), T) | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 497 | | NONE => | 
| 43095 | 498 | Var ((a |> textual ? repair_variable_name Char.toLower, | 
| 499 | var_index), T) | |
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 500 | in list_comb (t, ts) end | 
| 43093 | 501 | in do_term [] end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 502 | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 503 | fun term_from_atom ctxt textual sym_tab pos (u as ATerm (s, _)) = | 
| 38014 | 504 | 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 | 505 | add_type_constraint pos (type_constraint_from_term ctxt u) | 
| 38014 | 506 |     #> pair @{const True}
 | 
| 507 | 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 | 508 |     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 | 509 | |
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 510 | val combinator_table = | 
| 39953 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 blanchet parents: 
39710diff
changeset | 511 |   [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def_raw}),
 | 
| 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 blanchet parents: 
39710diff
changeset | 512 |    (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def_raw}),
 | 
| 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 blanchet parents: 
39710diff
changeset | 513 |    (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def_raw}),
 | 
| 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 blanchet parents: 
39710diff
changeset | 514 |    (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def_raw}),
 | 
| 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 blanchet parents: 
39710diff
changeset | 515 |    (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def_raw})]
 | 
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 516 | |
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 517 | 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 | 518 | let | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 519 | 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 | 520 | | 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 | 521 | | 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 | 522 | (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 | 523 | 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 | 524 | |> 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 | 525 | | NONE => t) | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 526 | | aux t = t | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 527 | in aux end | 
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 528 | |
| 37991 | 529 | (* Update schematic type variables with detected sort constraints. It's not | 
| 42563 | 530 | totally clear whether this code is necessary. *) | 
| 38014 | 531 | 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 | 532 | let | 
| 37991 | 533 | fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) | 
| 534 | | do_type (TVar (xi, s)) = | |
| 535 | TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) | |
| 536 | | do_type (TFree z) = TFree z | |
| 537 | fun do_term (Const (a, T)) = Const (a, do_type T) | |
| 538 | | do_term (Free (a, T)) = Free (a, do_type T) | |
| 539 | | do_term (Var (xi, T)) = Var (xi, do_type T) | |
| 540 | | do_term (t as Bound _) = t | |
| 541 | | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) | |
| 542 | | do_term (t1 $ t2) = do_term t1 $ do_term t2 | |
| 543 | in t |> not (Vartab.is_empty tvar_tab) ? do_term end | |
| 544 | ||
| 39425 | 545 | fun quantify_over_var quant_of var_s t = | 
| 546 | let | |
| 547 | val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s) | |
| 548 | |> map Var | |
| 549 | in fold_rev quant_of vars t end | |
| 37991 | 550 | |
| 38085 
cc44e887246c
avoid "clause" and "cnf" terminology where it no longer makes sense
 blanchet parents: 
38066diff
changeset | 551 | (* 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 | 552 | appear in the formula. *) | 
| 43136 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 553 | fun raw_prop_from_atp ctxt textual sym_tab phi = | 
| 38014 | 554 | let | 
| 555 | fun do_formula pos phi = | |
| 37991 | 556 | case phi of | 
| 38014 | 557 | AQuant (_, [], phi) => do_formula pos phi | 
| 42526 | 558 | | AQuant (q, (s, _) :: xs, phi') => | 
| 38014 | 559 | do_formula pos (AQuant (q, xs, phi')) | 
| 42526 | 560 | (* FIXME: TFF *) | 
| 39425 | 561 | #>> quantify_over_var (case q of | 
| 562 | AForall => forall_of | |
| 563 | | AExists => exists_of) | |
| 43095 | 564 | (s |> textual ? repair_variable_name Char.toLower) | 
| 38014 | 565 | | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not | 
| 37991 | 566 | | AConn (c, [phi1, phi2]) => | 
| 38014 | 567 | do_formula (pos |> c = AImplies ? not) phi1 | 
| 568 | ##>> do_formula pos phi2 | |
| 569 | #>> (case c of | |
| 570 | AAnd => s_conj | |
| 571 | | AOr => s_disj | |
| 572 | | AImplies => s_imp | |
| 38038 | 573 | | AIff => s_iff | 
| 43163 | 574 | | 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 | 575 | | AAtom tm => term_from_atom ctxt textual sym_tab pos tm | 
| 37991 | 576 | | _ => raise FORMULA [phi] | 
| 38014 | 577 | in repair_tvar_sorts (do_formula true phi Vartab.empty) end | 
| 37991 | 578 | |
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 579 | fun infer_formula_types ctxt = | 
| 39288 | 580 | 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 | 581 | #> Syntax.check_term | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 582 | (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 583 | |
| 43136 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 584 | fun prop_from_atp ctxt textual sym_tab = | 
| 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 585 | let val thy = Proof_Context.theory_of ctxt in | 
| 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 586 | raw_prop_from_atp ctxt textual sym_tab | 
| 43176 
29a3a1a7794d
only uncombine combinators in textual Isar proofs, not in Metis
 blanchet parents: 
43168diff
changeset | 587 | #> 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 | 588 | end | 
| 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 589 | |
| 43093 | 590 | (**** Translation of TSTP files to Isar proofs ****) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 591 | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 592 | 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 | 593 |   | 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 | 594 | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 595 | fun decode_line sym_tab (Definition (name, phi1, phi2)) ctxt = | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 596 | let | 
| 42361 | 597 | val thy = Proof_Context.theory_of ctxt | 
| 43136 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 598 | (* FIXME: prop or term? *) | 
| 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 599 | val t1 = raw_prop_from_atp ctxt true sym_tab phi1 | 
| 36551 | 600 | val vars = snd (strip_comb t1) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 601 | val frees = map unvarify_term vars | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 602 | val unvarify_args = subst_atomic (vars ~~ frees) | 
| 43136 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 603 | val t2 = raw_prop_from_atp ctxt true sym_tab phi2 | 
| 36551 | 604 | val (t1, t2) = | 
| 605 | HOLogic.eq_const HOLogic.typeT $ t1 $ t2 | |
| 43131 
9e9420122f91
fixed interaction between type tags and hAPP in reconstruction code
 blanchet parents: 
43130diff
changeset | 606 | |> unvarify_args |> uncombine_term thy |> infer_formula_types ctxt | 
| 36555 
8ff45c2076da
expand combinators in Isar proofs constructed by Sledgehammer;
 blanchet parents: 
36551diff
changeset | 607 | |> HOLogic.dest_eq | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 608 | in | 
| 39368 | 609 | (Definition (name, t1, t2), | 
| 36551 | 610 | fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 611 | 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 | 612 | | decode_line sym_tab (Inference (name, u, deps)) ctxt = | 
| 43136 
cf5cda219058
handle lightweight tags sym theorems gracefully in the presence of TVars with interesting type classes
 blanchet parents: 
43135diff
changeset | 613 | let val t = u |> prop_from_atp ctxt true sym_tab in | 
| 39368 | 614 | (Inference (name, t, deps), | 
| 36551 | 615 | fold Variable.declare_term (OldTerm.term_frees t) ctxt) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 616 | 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 | 617 | 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 | 618 | fst (fold_map (decode_line sym_tab) lines ctxt) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 619 | |
| 38035 | 620 | fun is_same_inference _ (Definition _) = false | 
| 621 | | is_same_inference t (Inference (_, t', _)) = t aconv t' | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 622 | |
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 623 | (* 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 | 624 | clsarity). *) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 625 | val is_only_type_information = curry (op aconv) HOLogic.true_const | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 626 | |
| 39373 | 627 | 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 | 628 | if is_same_atp_step dep old then new else [dep] | 
| 39373 | 629 | fun replace_dependencies_in_line _ (line as Definition _) = line | 
| 630 | | replace_dependencies_in_line p (Inference (name, t, deps)) = | |
| 631 | Inference (name, t, fold (union (op =) o replace_one_dependency p) deps []) | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 632 | |
| 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 | 633 | (* 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 | 634 | they differ only in type information.*) | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 635 | fun add_line _ _ _ (line as Definition _) lines = line :: lines | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 636 | | add_line type_sys conjecture_shape fact_names (Inference (name, t, [])) | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 637 | 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 | 638 | (* 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 | 639 | definitions. *) | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 640 | if is_fact type_sys fact_names name 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 | 641 | (* Facts are not proof lines. *) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 642 | if is_only_type_information t then | 
| 39373 | 643 | map (replace_dependencies_in_line (name, [])) lines | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 644 | (* Is there a repetition? If so, replace later line by earlier one. *) | 
| 38035 | 645 | else case take_prefix (not o is_same_inference t) lines of | 
| 39373 | 646 | (_, []) => lines (* no repetition of proof line *) | 
| 39368 | 647 | | (pre, Inference (name', _, _) :: post) => | 
| 39373 | 648 | pre @ map (replace_dependencies_in_line (name', [name])) post | 
| 40069 | 649 | | _ => raise Fail "unexpected inference" | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 650 | else if is_conjecture conjecture_shape name then | 
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 651 | Inference (name, s_not t, []) :: lines | 
| 36551 | 652 | else | 
| 39373 | 653 | map (replace_dependencies_in_line (name, [])) lines | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 654 | | add_line _ _ _ (Inference (name, t, deps)) lines = | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 655 | (* 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 | 656 | if is_only_type_information t then | 
| 39368 | 657 | Inference (name, t, deps) :: lines | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 658 | (* Is there a repetition? If so, replace later line by earlier one. *) | 
| 38035 | 659 | 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 | 660 | (* FIXME: Doesn't this code risk conflating proofs involving different | 
| 38035 | 661 | types? *) | 
| 39368 | 662 | (_, []) => Inference (name, t, deps) :: lines | 
| 663 | | (pre, Inference (name', t', _) :: post) => | |
| 664 | Inference (name, t', deps) :: | |
| 39373 | 665 | pre @ map (replace_dependencies_in_line (name', [name])) post | 
| 40069 | 666 | | _ => raise Fail "unexpected inference" | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 667 | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 668 | (* Recursively delete empty lines (type information) from the proof. *) | 
| 39368 | 669 | fun add_nontrivial_line (Inference (name, t, [])) lines = | 
| 39373 | 670 | if is_only_type_information t then delete_dependency name lines | 
| 39368 | 671 | else Inference (name, t, []) :: lines | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 672 | | add_nontrivial_line line lines = line :: lines | 
| 39373 | 673 | and delete_dependency name lines = | 
| 674 | fold_rev add_nontrivial_line | |
| 675 | (map (replace_dependencies_in_line (name, [])) lines) [] | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 676 | |
| 37323 | 677 | (* ATPs sometimes reuse free variable names in the strangest ways. Removing | 
| 678 | offending lines often does the trick. *) | |
| 36560 
45c1870f234f
fixed definition of "bad frees" so that it actually works
 blanchet parents: 
36559diff
changeset | 679 | 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 | 680 | | is_bad_free _ _ = false | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 681 | |
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 682 | fun add_desired_line _ _ _ _ _ (line as Definition (name, _, _)) (j, lines) = | 
| 39373 | 683 | (j, line :: map (replace_dependencies_in_line (name, [])) lines) | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 684 | | add_desired_line type_sys isar_shrink_factor conjecture_shape fact_names | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 685 | frees (Inference (name, t, deps)) (j, lines) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 686 | (j + 1, | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 687 | if is_fact type_sys fact_names name orelse | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 688 | is_conjecture conjecture_shape name orelse | 
| 39373 | 689 | (* the last line must be kept *) | 
| 690 | 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 | 691 | (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 | 692 | 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 | 693 | not (exists_subterm (is_bad_free frees) t) andalso | 
| 39373 | 694 | length deps >= 2 andalso j mod isar_shrink_factor = 0 andalso | 
| 695 | (* kill next to last line, which usually results in a trivial step *) | |
| 696 | j <> 1) then | |
| 39368 | 697 | Inference (name, t, deps) :: lines (* keep line *) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 698 | else | 
| 39373 | 699 | 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 | 700 | |
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 701 | (** Isar proof construction and manipulation **) | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 702 | |
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 703 | fun merge_fact_sets (ls1, ss1) (ls2, ss2) = | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 704 | (union (op =) ls1 ls2, union (op =) ss1 ss2) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 705 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 706 | type label = string * int | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 707 | type facts = label list * string list | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 708 | |
| 39452 | 709 | 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 | 710 | |
| 39452 | 711 | datatype isar_step = | 
| 36478 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 712 | Fix of (string * typ) list | | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 713 | Let of term * term | | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 714 | Assume of label * term | | 
| 39452 | 715 | Have 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 | 716 | and byline = | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 717 | ByMetis of facts | | 
| 39452 | 718 | CaseSplit 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 | 719 | |
| 36574 | 720 | fun smart_case_split [] facts = ByMetis facts | 
| 721 | | smart_case_split proofs facts = CaseSplit (proofs, facts) | |
| 722 | ||
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 723 | fun add_fact_from_dependency type_sys conjecture_shape facts_offset fact_names | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 724 | name = | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 725 | if is_fact type_sys fact_names name then | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 726 | apsnd (union (op =) | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 727 | (map fst (resolve_fact type_sys facts_offset fact_names name))) | 
| 36475 
05209b869a6b
new Isar proof construction code: stringfy axiom names correctly
 blanchet parents: 
36474diff
changeset | 728 | else | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 729 | apfst (insert (op =) (raw_label_for_name conjecture_shape name)) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 730 | |
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 731 | fun step_for_line _ _ _ _ _ (Definition (_, t1, t2)) = Let (t1, t2) | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 732 | | step_for_line _ conjecture_shape _ _ _ (Inference (name, t, [])) = | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 733 | Assume (raw_label_for_name conjecture_shape name, t) | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 734 | | step_for_line type_sys conjecture_shape facts_offset | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 735 | fact_names j (Inference (name, t, deps)) = | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 736 | Have (if j = 1 then [Show] else [], | 
| 39425 | 737 | raw_label_for_name conjecture_shape name, | 
| 738 | fold_rev forall_of (map Var (Term.add_vars t [])) t, | |
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 739 | ByMetis (fold (add_fact_from_dependency type_sys conjecture_shape | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 740 | facts_offset fact_names) | 
| 39373 | 741 | deps ([], []))) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 742 | |
| 39454 | 743 | fun repair_name "$true" = "c_True" | 
| 744 | | repair_name "$false" = "c_False" | |
| 43000 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 745 | | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *) | 
| 39454 | 746 | | repair_name s = | 
| 43000 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 747 | if is_tptp_equal s orelse | 
| 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 748 | (* seen in Vampire proofs *) | 
| 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 749 | (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then | 
| 
bd424c3dde46
cleaner handling of equality and proxies (esp. for THF)
 blanchet parents: 
42998diff
changeset | 750 | tptp_equal | 
| 39454 | 751 | else | 
| 752 | s | |
| 753 | ||
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 754 | fun isar_proof_from_atp_proof pool ctxt type_sys isar_shrink_factor | 
| 43033 | 755 | conjecture_shape facts_offset fact_names sym_tab params frees | 
| 756 | atp_proof = | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 757 | let | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 758 | val lines = | 
| 42449 
494e4ac5b0f8
detect some unsound proofs before showing them to the user
 blanchet parents: 
42361diff
changeset | 759 | atp_proof | 
| 42968 
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
 blanchet parents: 
42966diff
changeset | 760 | |> clean_up_atp_proof_dependencies | 
| 39454 | 761 | |> nasty_atp_proof pool | 
| 762 | |> map_term_names_in_atp_proof repair_name | |
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 763 | |> decode_lines ctxt sym_tab | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 764 | |> rpair [] |-> fold_rev (add_line type_sys conjecture_shape fact_names) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 765 | |> rpair [] |-> fold_rev add_nontrivial_line | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 766 | |> rpair (0, []) | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 767 | |-> fold_rev (add_desired_line type_sys isar_shrink_factor | 
| 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 768 | conjecture_shape fact_names frees) | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 769 | |> snd | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 770 | in | 
| 36909 
7d5587f6d5f7
made Sledgehammer's full-typed proof reconstruction work for the first time;
 blanchet parents: 
36607diff
changeset | 771 | (if null params then [] else [Fix params]) @ | 
| 42647 
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
 blanchet parents: 
42613diff
changeset | 772 | map2 (step_for_line type_sys conjecture_shape facts_offset fact_names) | 
| 42541 
8938507b2054
move type declarations to the front, for TFF-compliance
 blanchet parents: 
42539diff
changeset | 773 | (length lines downto 1) lines | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 774 | end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 775 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 776 | (* When redirecting proofs, we keep information about the labels seen so far in | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 777 | the "backpatches" data structure. The first component indicates which facts | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 778 | should be associated with forthcoming proof steps. The second component is a | 
| 37322 
0347b1a16682
fix bug in direct Isar proofs, which was exhibited by the "BigO" example
 blanchet parents: 
37319diff
changeset | 779 |    pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
 | 
| 
0347b1a16682
fix bug in direct Isar proofs, which was exhibited by the "BigO" example
 blanchet parents: 
37319diff
changeset | 780 | become assumptions and "drop_ls" are the labels that should be dropped in a | 
| 
0347b1a16682
fix bug in direct Isar proofs, which was exhibited by the "BigO" example
 blanchet parents: 
37319diff
changeset | 781 | case split. *) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 782 | type backpatches = (label * facts) list * (label list * label list) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 783 | |
| 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 | 784 | fun used_labels_of_step (Have (_, _, _, by)) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 785 | (case by of | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 786 | ByMetis (ls, _) => ls | 
| 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 | 787 | | CaseSplit (proofs, (ls, _)) => | 
| 
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 | 788 | fold (union (op =) o used_labels_of) proofs ls) | 
| 
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 | 789 | | used_labels_of_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 | 790 | 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 | 791 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 792 | fun new_labels_of_step (Fix _) = [] | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 793 | | new_labels_of_step (Let _) = [] | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 794 | | new_labels_of_step (Assume (l, _)) = [l] | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 795 | | new_labels_of_step (Have (_, l, _, _)) = [l] | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 796 | val new_labels_of = maps new_labels_of_step | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 797 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 798 | val join_proofs = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 799 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 800 | fun aux _ [] = NONE | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 801 | | aux proof_tail (proofs as (proof1 :: _)) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 802 | if exists null proofs then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 803 | NONE | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 804 | else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 805 | aux (hd proof1 :: proof_tail) (map tl proofs) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 806 | else case hd proof1 of | 
| 37498 
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
 blanchet parents: 
37479diff
changeset | 807 | Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 808 | if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t') | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 809 | | _ => false) (tl proofs) andalso | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 810 | not (exists (member (op =) (maps new_labels_of proofs)) | 
| 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 | 811 | (used_labels_of proof_tail)) then | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 812 | SOME (l, t, map rev proofs, proof_tail) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 813 | else | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 814 | NONE | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 815 | | _ => NONE | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 816 | in aux [] o map rev end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 817 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 818 | fun case_split_qualifiers proofs = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 819 | case length proofs of | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 820 | 0 => [] | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 821 | | 1 => [Then] | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 822 | | _ => [Ultimately] | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 823 | |
| 39372 
2fd8a9a7080d
first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
 blanchet parents: 
39370diff
changeset | 824 | fun redirect_proof hyp_ts concl_t proof = | 
| 33310 | 825 | let | 
| 37324 
d77250dd2416
fix bugs in Sledgehammer's Isar proof "redirection" code
 blanchet parents: 
37323diff
changeset | 826 | (* The first pass outputs those steps that are independent of the negated | 
| 
d77250dd2416
fix bugs in Sledgehammer's Isar proof "redirection" code
 blanchet parents: 
37323diff
changeset | 827 | conjecture. The second pass flips the proof by contradiction to obtain a | 
| 
d77250dd2416
fix bugs in Sledgehammer's Isar proof "redirection" code
 blanchet parents: 
37323diff
changeset | 828 | direct proof, introducing case splits when an inference depends on | 
| 
d77250dd2416
fix bugs in Sledgehammer's Isar proof "redirection" code
 blanchet parents: 
37323diff
changeset | 829 | several facts that depend on the negated conjecture. *) | 
| 39372 
2fd8a9a7080d
first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
 blanchet parents: 
39370diff
changeset | 830 | val concl_l = (conjecture_prefix, length hyp_ts) | 
| 38040 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 831 | fun first_pass ([], contra) = ([], contra) | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 832 | | first_pass ((step as Fix _) :: proof, contra) = | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 833 | first_pass (proof, contra) |>> cons step | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 834 | | first_pass ((step as Let _) :: proof, contra) = | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 835 | first_pass (proof, contra) |>> cons step | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 836 | | first_pass ((step as Assume (l as (_, j), _)) :: proof, contra) = | 
| 39372 
2fd8a9a7080d
first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
 blanchet parents: 
39370diff
changeset | 837 | if l = concl_l then first_pass (proof, contra ||> cons step) | 
| 
2fd8a9a7080d
first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
 blanchet parents: 
39370diff
changeset | 838 | else first_pass (proof, contra) |>> cons (Assume (l, nth hyp_ts j)) | 
| 38040 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 839 | | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) = | 
| 39372 
2fd8a9a7080d
first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
 blanchet parents: 
39370diff
changeset | 840 | let val step = Have (qs, l, t, ByMetis (ls, ss)) in | 
| 38040 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 841 | if exists (member (op =) (fst contra)) ls then | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 842 | first_pass (proof, contra |>> cons l ||> cons step) | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 843 | else | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 844 | first_pass (proof, contra) |>> cons step | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 845 | end | 
| 
174568533593
fix bug in the SPASS Flotter hack, when a conjecture FOF is translated to several CNF clauses
 blanchet parents: 
38039diff
changeset | 846 | | first_pass _ = raise Fail "malformed proof" | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 847 | val (proof_top, (contra_ls, contra_proof)) = | 
| 39372 
2fd8a9a7080d
first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
 blanchet parents: 
39370diff
changeset | 848 | first_pass (proof, ([concl_l], [])) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 849 | val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 850 | fun backpatch_labels patches ls = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 851 | fold merge_fact_sets (map (backpatch_label patches) ls) ([], []) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 852 | fun second_pass end_qs ([], assums, patches) = | 
| 37324 
d77250dd2416
fix bugs in Sledgehammer's Isar proof "redirection" code
 blanchet parents: 
37323diff
changeset | 853 | ([Have (end_qs, no_label, concl_t, | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 854 | ByMetis (backpatch_labels patches (map snd assums)))], patches) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 855 | | second_pass end_qs (Assume (l, t) :: proof, assums, patches) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 856 | second_pass end_qs (proof, (t, l) :: assums, patches) | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 857 | | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums, | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 858 | patches) = | 
| 39373 | 859 | (if member (op =) (snd (snd patches)) l andalso | 
| 860 | not (member (op =) (fst (snd patches)) l) andalso | |
| 861 | not (AList.defined (op =) (fst patches) l) then | |
| 862 | second_pass end_qs (proof, assums, patches ||> apsnd (append ls)) | |
| 863 | else case List.partition (member (op =) contra_ls) ls of | |
| 864 | ([contra_l], co_ls) => | |
| 865 | if member (op =) qs Show then | |
| 866 | second_pass end_qs (proof, assums, | |
| 867 | patches |>> cons (contra_l, (co_ls, ss))) | |
| 868 | else | |
| 869 | second_pass end_qs | |
| 870 | (proof, assums, | |
| 871 | patches |>> cons (contra_l, (l :: co_ls, ss))) | |
| 872 | |>> cons (if member (op =) (fst (snd patches)) l then | |
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 873 | Assume (l, s_not t) | 
| 39373 | 874 | else | 
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 875 | Have (qs, l, s_not t, | 
| 39373 | 876 | ByMetis (backpatch_label patches l))) | 
| 877 | | (contra_ls as _ :: _, co_ls) => | |
| 878 | let | |
| 879 | val proofs = | |
| 880 | map_filter | |
| 881 | (fn l => | |
| 882 | if l = concl_l then | |
| 883 | NONE | |
| 884 | else | |
| 885 | let | |
| 886 | val drop_ls = filter (curry (op <>) l) contra_ls | |
| 887 | in | |
| 888 | second_pass [] | |
| 889 | (proof, assums, | |
| 890 | patches ||> apfst (insert (op =) l) | |
| 891 | ||> apsnd (union (op =) drop_ls)) | |
| 892 | |> fst |> SOME | |
| 893 | end) contra_ls | |
| 894 | val (assumes, facts) = | |
| 895 | if member (op =) (fst (snd patches)) l then | |
| 42606 
0c76cf483899
show sorts not just types in Isar proofs + tuning
 blanchet parents: 
42595diff
changeset | 896 | ([Assume (l, s_not t)], (l :: co_ls, ss)) | 
| 39373 | 897 | else | 
| 898 | ([], (co_ls, ss)) | |
| 899 | in | |
| 900 | (case join_proofs proofs of | |
| 901 | SOME (l, t, proofs, proof_tail) => | |
| 902 | Have (case_split_qualifiers proofs @ | |
| 903 | (if null proof_tail then end_qs else []), l, t, | |
| 904 | smart_case_split proofs facts) :: proof_tail | |
| 905 | | NONE => | |
| 906 | [Have (case_split_qualifiers proofs @ end_qs, no_label, | |
| 907 | concl_t, smart_case_split proofs facts)], | |
| 908 | patches) | |
| 909 | |>> append assumes | |
| 910 | end | |
| 911 | | _ => raise Fail "malformed proof") | |
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 912 | | second_pass _ _ = raise Fail "malformed proof" | 
| 36486 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 913 | val proof_bottom = | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 914 | second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 915 | in proof_top @ proof_bottom end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 916 | |
| 38490 | 917 | (* FIXME: Still needed? Probably not. *) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 918 | val kill_duplicate_assumptions_in_proof = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 919 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 920 | fun relabel_facts subst = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 921 | 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 | 922 | 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 | 923 | (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 | 924 | 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 | 925 | | NONE => (step :: proof, subst, (t, l) :: assums)) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 926 | | do_step (Have (qs, l, t, by)) (proof, subst, assums) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 927 | (Have (qs, l, t, | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 928 | case by of | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 929 | ByMetis facts => ByMetis (relabel_facts subst facts) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 930 | | CaseSplit (proofs, facts) => | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 931 | CaseSplit (map do_proof proofs, relabel_facts subst facts)) :: | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 932 | 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 | 933 | | 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 | 934 | 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 | 935 | in do_proof end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 936 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 937 | val then_chain_proof = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 938 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 939 | fun aux _ [] = [] | 
| 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 | 940 | | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 941 | | aux l' (Have (qs, l, t, by) :: proof) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 942 | (case by of | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 943 | ByMetis (ls, ss) => | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 944 | Have (if member (op =) ls l' then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 945 | (Then :: qs, l, t, | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 946 | ByMetis (filter_out (curry (op =) l') ls, ss)) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 947 | else | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 948 | (qs, l, t, ByMetis (ls, ss))) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 949 | | CaseSplit (proofs, facts) => | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 950 | Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) :: | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 951 | aux l proof | 
| 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 | 952 | | aux _ (step :: proof) = step :: aux no_label proof | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 953 | in aux no_label end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 954 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 955 | 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 | 956 | 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 | 957 | 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 | 958 | 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 | 959 | fun do_step (Assume (l, t)) = Assume (do_label l, t) | 
| 
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 | 960 | | do_step (Have (qs, l, t, by)) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 961 | Have (qs, do_label l, t, | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 962 | case by of | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 963 | CaseSplit (proofs, facts) => | 
| 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 | 964 | CaseSplit (map (map do_step) proofs, facts) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 965 | | _ => 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 | 966 | | 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 | 967 | 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 | 968 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 969 | 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 | 970 | |
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 971 | val relabel_proof = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 972 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 973 | fun aux _ _ _ [] = [] | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 974 | | 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 | 975 | if l = no_label then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 976 | 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 | 977 | else | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 978 | 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 | 979 | Assume (l', t) :: | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 980 | 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 | 981 | end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 982 | | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 983 | let | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 984 | val (l', subst, next_fact) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 985 | if l = no_label then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 986 | (l, subst, next_fact) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 987 | else | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 988 | let | 
| 42180 
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
 blanchet parents: 
41742diff
changeset | 989 | 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 | 990 | 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 | 991 | val relabel_facts = | 
| 39370 
f8292d3020db
use same hack as in "Async_Manager" to work around Proof General bug
 blanchet parents: 
39368diff
changeset | 992 | 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 | 993 | val by = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 994 | case by of | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 995 | ByMetis facts => ByMetis (relabel_facts facts) | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 996 | | CaseSplit (proofs, facts) => | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 997 | CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs, | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 998 | relabel_facts facts) | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 999 | in | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1000 | Have (qs, l', t, by) :: | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1001 | 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 | 1002 | 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 | 1003 | | 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 | 1004 | 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 | 1005 | in aux [] 0 (1, 1) end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1006 | |
| 42881 
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
 blanchet parents: 
42876diff
changeset | 1007 | fun string_for_proof ctxt0 full_types i n = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1008 | let | 
| 42761 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 1009 | val ctxt = | 
| 
8ea9c6fa8b53
fixed several bugs in Isar proof reconstruction, in particular w.r.t. mangled types and hAPP
 blanchet parents: 
42759diff
changeset | 1010 | 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 | 1011 | |> 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 | 1012 | |> Config.put show_sorts true | 
| 37319 
42268dc7d6c4
show types in Isar proofs, but not for free variables;
 blanchet parents: 
37172diff
changeset | 1013 | fun fix_print_mode f x = | 
| 39134 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
39115diff
changeset | 1014 | Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) | 
| 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
39115diff
changeset | 1015 | (print_mode_value ())) f x | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1016 | 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 | 1017 | fun do_free (s, T) = | 
| 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 1018 | maybe_quote s ^ " :: " ^ | 
| 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 1019 | 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 | 1020 | 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 | 1021 | fun do_have qs = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1022 | (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 | 1023 | (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 | 1024 | (if member (op =) qs Then then | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1025 | 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 | 1026 | else | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1027 | 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 | 1028 | val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt) | 
| 43168 
467d5b34e1f5
temporarily added "MetisX" as reconstructor and minimizer
 blanchet parents: 
43166diff
changeset | 1029 | val reconstructor = if full_types then MetisX else Metis | 
| 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 | 1030 | fun do_facts (ls, ss) = | 
| 43033 | 1031 | reconstructor_command reconstructor 1 1 | 
| 1032 | (ls |> sort_distinct (prod_ord string_ord int_ord), | |
| 1033 | ss |> sort_distinct string_ord) | |
| 36478 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 1034 | and do_step ind (Fix xs) = | 
| 
1aba777a367f
fix types of "fix" variables to help proof reconstruction and aid readability
 blanchet parents: 
36477diff
changeset | 1035 | 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 | 1036 | | do_step ind (Let (t1, t2)) = | 
| 
c2d7e2dff59e
support Vampire definitions of constants as "let" constructs in Isar proofs
 blanchet parents: 
36485diff
changeset | 1037 | 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 | 1038 | | do_step ind (Assume (l, t)) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1039 | do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n" | 
| 36564 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 1040 | | do_step ind (Have (qs, l, t, ByMetis facts)) = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1041 | do_indent ind ^ do_have qs ^ " " ^ | 
| 36479 | 1042 | do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n" | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1043 | | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) = | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1044 | space_implode (do_indent ind ^ "moreover\n") | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1045 | (map (do_block ind) proofs) ^ | 
| 36479 | 1046 | 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 | 1047 | do_facts facts ^ "\n" | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1048 | 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 | 1049 | 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 | 1050 | 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 | 1051 | String.extract (s, ind * indent_size, | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1052 | 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 | 1053 | suffix ^ "\n" | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1054 | end | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1055 |     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 | 1056 | (* 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 | 1057 | directly. *) | 
| 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 1058 | and do_proof [Have (_, _, _, ByMetis _)] = "" | 
| 
96f767f546e7
be more discriminate: some one-line Isar proofs are silly
 blanchet parents: 
36563diff
changeset | 1059 | | 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 | 1060 | (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ | 
| 39452 | 1061 | do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^ | 
| 1062 | (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 | 1063 | in do_proof end | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1064 | |
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1065 | fun isar_proof_text ctxt isar_proof_requested | 
| 43033 | 1066 | (debug, full_types, isar_shrink_factor, type_sys, pool, | 
| 43037 | 1067 | conjecture_shape, facts_offset, fact_names, sym_tab, atp_proof, goal) | 
| 1068 | (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 | 1069 | let | 
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1070 | val isar_shrink_factor = | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1071 | (if isar_proof_requested then 1 else 2) * isar_shrink_factor | 
| 43037 | 1072 | 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 | 1073 | val frees = fold Term.add_frees (concl_t :: hyp_ts) [] | 
| 43033 | 1074 | val one_line_proof = one_line_proof_text one_line_params | 
| 36283 
25e69e93954d
failure of reconstructing the Isar proof (e.g., exception) should not prevent Sledgehammer from showing the short one-liner proof -- but in "debug" mode we do want to know what the exception is
 blanchet parents: 
36281diff
changeset | 1075 | fun isar_proof_for () = | 
| 43033 | 1076 | case atp_proof | 
| 43135 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 1077 | |> isar_proof_from_atp_proof pool ctxt type_sys isar_shrink_factor | 
| 
8c32a0160b0d
more uniform handling of tfree sort inference in ATP reconstruction code, based on what Metis always has done
 blanchet parents: 
43131diff
changeset | 1078 | conjecture_shape facts_offset fact_names sym_tab params frees | 
| 39372 
2fd8a9a7080d
first step in generalizing to nonnumeric proof step names (e.g. remote Vampire 0.6)
 blanchet parents: 
39370diff
changeset | 1079 | |> redirect_proof hyp_ts concl_t | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1080 | |> kill_duplicate_assumptions_in_proof | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1081 | |> then_chain_proof | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1082 | |> kill_useless_labels_in_proof | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1083 | |> relabel_proof | 
| 43037 | 1084 | |> string_for_proof ctxt full_types subgoal subgoal_count of | 
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1085 | "" => | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1086 | if isar_proof_requested then | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1087 | "\nNo structured proof available (proof too short)." | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1088 | else | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1089 | "" | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1090 | | proof => | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1091 | "\n\n" ^ (if isar_proof_requested then "Structured proof" | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1092 | else "Perhaps this will work") ^ | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1093 | ":\n" ^ Markup.markup Markup.sendback proof | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 1094 | val isar_proof = | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36396diff
changeset | 1095 | if debug then | 
| 36283 
25e69e93954d
failure of reconstructing the Isar proof (e.g., exception) should not prevent Sledgehammer from showing the short one-liner proof -- but in "debug" mode we do want to know what the exception is
 blanchet parents: 
36281diff
changeset | 1096 | isar_proof_for () | 
| 
25e69e93954d
failure of reconstructing the Isar proof (e.g., exception) should not prevent Sledgehammer from showing the short one-liner proof -- but in "debug" mode we do want to know what the exception is
 blanchet parents: 
36281diff
changeset | 1097 | else | 
| 43062 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1098 | case try isar_proof_for () of | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1099 | SOME s => s | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1100 | | NONE => if isar_proof_requested then | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1101 | "\nWarning: The Isar proof construction failed." | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1102 | else | 
| 
2834548a7a48
nicer failure message when one-line proof reconstruction fails
 blanchet parents: 
43051diff
changeset | 1103 | "" | 
| 43033 | 1104 | in one_line_proof ^ isar_proof end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 1105 | |
| 43033 | 1106 | fun proof_text ctxt isar_proof isar_params | 
| 1107 | (one_line_params as (preplay, _, _, _, _, _)) = | |
| 43166 | 1108 | (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 | 1109 | isar_proof_text ctxt isar_proof isar_params | 
| 43033 | 1110 | else | 
| 1111 | one_line_proof_text) one_line_params | |
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 1112 | |
| 31038 | 1113 | end; |