| author | haftmann | 
| Mon, 26 Apr 2010 11:34:15 +0200 | |
| changeset 36348 | 89c54f51f55a | 
| parent 36293 | e6db3ba0b61d | 
| child 36369 | d2cd0d04b8e6 | 
| permissions | -rw-r--r-- | 
| 35826 | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML | 
| 33310 | 2 | Author: Lawrence C Paulson and Claire Quigley, Cambridge University Computer Laboratory | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 3 | |
| 33310 | 4 | Transfer of proofs from external provers. | 
| 5 | *) | |
| 6 | ||
| 35826 | 7 | signature SLEDGEHAMMER_PROOF_RECONSTRUCT = | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 8 | sig | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 9 | type minimize_command = string list -> string | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 10 | |
| 25492 
4cc7976948ac
Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
 paulson parents: 
25414diff
changeset | 11 | val chained_hint: string | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 12 | val invert_const: string -> string | 
| 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 13 | val invert_type_const: string -> string | 
| 33243 | 14 | val num_typargs: theory -> string -> int | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 15 | val make_tvar: string -> typ | 
| 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 16 | val strip_prefix: string -> string -> string option | 
| 35865 | 17 | val is_proof_well_formed: string -> bool | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 18 | val metis_line: int -> int -> string list -> string | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 19 | val metis_proof_text: | 
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 20 | minimize_command * string * string vector * thm * int | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 21 | -> string * string list | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 22 | val isar_proof_text: | 
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 23 | bool -> int -> bool -> Proof.context | 
| 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 24 | -> minimize_command * string * string vector * thm * int | 
| 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 25 | -> string * string list | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 26 | val proof_text: | 
| 36288 | 27 | bool -> bool -> int -> bool -> Proof.context | 
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 28 | -> minimize_command * string * string vector * thm * int | 
| 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 29 | -> string * string list | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 30 | end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 31 | |
| 35826 | 32 | structure Sledgehammer_Proof_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 33 | struct | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 34 | |
| 35865 | 35 | open Sledgehammer_FOL_Clause | 
| 36 | open Sledgehammer_Fact_Preprocessor | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 37 | |
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 38 | type minimize_command = string list -> string | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 39 | |
| 35865 | 40 | val trace_proof_path = Path.basic "atp_trace"; | 
| 41 | ||
| 42 | fun trace_proof_msg f = | |
| 43 | if !trace then File.append (File.tmp_path trace_proof_path) (f ()) else (); | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 44 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31866diff
changeset | 45 | fun string_of_thm ctxt = PrintMode.setmp [] (Display.string_of_thm ctxt); | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 46 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 47 | fun is_ident_char c = Char.isAlphaNum c orelse c = #"_" | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 48 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 49 | fun is_axiom thm_names line_no = line_no <= Vector.length thm_names | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 50 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 51 | (**** PARSING OF TSTP FORMAT ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 52 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 53 | (* Syntax trees, either term list or formulae *) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 54 | datatype stree = Int of int | Br of string * stree list; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 55 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 56 | fun atom x = Br(x,[]); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 57 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 58 | fun scons (x,y) = Br("cons", [x,y]);
 | 
| 30190 | 59 | val listof = List.foldl scons (atom "nil"); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 60 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 61 | (*Strings enclosed in single quotes, e.g. filenames*) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 62 | val quoted = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 63 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 64 | (*Intended for $true and $false*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 65 | fun tf s = "c_" ^ str (Char.toUpper (String.sub(s,0))) ^ String.extract(s,1,NONE); | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 66 | val truefalse = $$ "$" |-- Symbol.scan_id >> (atom o tf); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 67 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 68 | (*Integer constants, typically proof line numbers*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 69 | fun is_digit s = Char.isDigit (String.sub(s,0)); | 
| 33035 | 70 | val integer = Scan.many1 is_digit >> (the o Int.fromString o implode); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 71 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 72 | (* needed for SPASS's nonstandard output format *) | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 73 | fun fix_symbol "equal" = "c_equal" | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 74 | | fix_symbol s = s | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 75 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 76 | (*Generalized FO terms, which include filenames, numbers, etc.*) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 77 | fun term x = (quoted >> atom || integer >> Int || truefalse | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 78 | || (Symbol.scan_id >> fix_symbol) | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 79 |                  -- Scan.optional ($$ "(" |-- terms --| $$ ")") [] >> Br
 | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 80 |               || $$ "(" |-- term --| $$ ")"
 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 81 | || $$ "[" |-- Scan.optional terms [] --| $$ "]" >> listof) x | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 82 | and terms x = (term ::: Scan.repeat ($$ "," |-- term)) x | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 83 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 84 | fun negate t = Br ("c_Not", [t])
 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 85 | fun equate t1 t2 = Br ("c_equal", [t1, t2]);
 | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 86 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 87 | (*Apply equal or not-equal to a term*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 88 | fun syn_equal (t, NONE) = t | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 89 | | syn_equal (t1, SOME (NONE, t2)) = equate t1 t2 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 90 | | syn_equal (t1, SOME (SOME _, t2)) = negate (equate t1 t2) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 91 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 92 | (*Literals can involve negation, = and !=.*) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 93 | fun literal x = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 94 | ($$ "~" |-- literal >> negate | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 95 | || (term -- Scan.option (Scan.option ($$ "!") --| $$ "=" -- term) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 96 | >> syn_equal)) x | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 97 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 98 | val literals = literal ::: Scan.repeat ($$ "|" |-- literal); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 99 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 100 | (*Clause: a list of literals separated by the disjunction sign*) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 101 | val clause = $$ "(" |-- literals --| $$ ")" || Scan.single literal;
 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 102 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 103 | fun ints_of_stree (Int n) = cons n | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 104 | | ints_of_stree (Br (_, ts)) = fold ints_of_stree ts | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 105 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 106 | val tstp_annotations = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 107 | Scan.optional ($$ "," |-- term --| Scan.option ($$ "," |-- terms) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 108 | >> (fn source => ints_of_stree source [])) [] | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 109 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 110 | fun retuple_tstp_line ((name, ts), deps) = (name, ts, deps) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 111 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 112 | (* <cnf_annotated> ::= cnf(<name>, <formula_role>, <cnf_formula> <annotations>). | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 113 | The <name> could be an identifier, but we assume integers. *) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 114 | val parse_tstp_line = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 115 |   (Scan.this_string "cnf" -- $$ "(") |-- integer --| $$ "," --| Symbol.scan_id
 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 116 | --| $$ "," -- clause -- tstp_annotations --| $$ ")" --| $$ "." | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 117 | >> retuple_tstp_line | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 118 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 119 | (**** PARSING OF SPASS OUTPUT ****) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 120 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 121 | val dot_name = integer --| $$ "." --| integer | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 122 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 123 | val spass_annotations = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 124 | Scan.optional ($$ ":" |-- Scan.repeat (dot_name --| Scan.option ($$ ","))) [] | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 125 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 126 | val starred_literal = literal --| Scan.repeat ($$ "*" || $$ " ") | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 127 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 128 | val horn_clause = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 129 | Scan.repeat starred_literal --| $$ "-" --| $$ ">" | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 130 | -- Scan.repeat starred_literal | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 131 | >> (fn ([], []) => [atom (tf "false")] | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 132 | | (clauses1, clauses2) => map negate clauses1 @ clauses2) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 133 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 134 | fun retuple_spass_proof_line ((name, deps), ts) = (name, ts, deps) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 135 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 136 | (* Syntax: <name>[0:<inference><annotations>] || -> <cnf_formula> **. *) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 137 | val parse_spass_proof_line = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 138 | integer --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 139 | -- spass_annotations --| $$ "]" --| $$ "|" --| $$ "|" -- horn_clause | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 140 | --| $$ "." | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 141 | >> retuple_spass_proof_line | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 142 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 143 | val parse_proof_line = fst o (parse_tstp_line || parse_spass_proof_line) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 144 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 145 | (**** INTERPRETATION OF TSTP SYNTAX TREES ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 146 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 147 | exception STREE of stree; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 148 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 149 | (*If string s has the prefix s1, return the result of deleting it.*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 150 | fun strip_prefix s1 s = | 
| 31038 | 151 | if String.isPrefix s1 s | 
| 35865 | 152 | then SOME (undo_ascii_of (String.extract (s, size s1, NONE))) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 153 | else NONE; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 154 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 155 | (*Invert the table of translations between Isabelle and ATPs*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 156 | val type_const_trans_table_inv = | 
| 35865 | 157 | Symtab.make (map swap (Symtab.dest type_const_trans_table)); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 158 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 159 | fun invert_type_const c = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 160 | case Symtab.lookup type_const_trans_table_inv c of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 161 | SOME c' => c' | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 162 | | NONE => c; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 163 | |
| 36285 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 164 | fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS);
 | 
| 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 165 | fun make_tparam s = TypeInfer.param 0 (s, HOLogic.typeS) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 166 | fun make_var (b,T) = Var((b,0),T); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 167 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 168 | (*Type variables are given the basic sort, HOL.type. Some will later be constrained | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 169 | by information from type literals, or by type inference.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 170 | fun type_of_stree t = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 171 | case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 172 | Int _ => raise STREE t | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 173 | | Br (a,ts) => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 174 | let val Ts = map type_of_stree ts | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 175 | in | 
| 35865 | 176 | case strip_prefix tconst_prefix a of | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 177 | SOME b => Type(invert_type_const b, Ts) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 178 | | NONE => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 179 | if not (null ts) then raise STREE t (*only tconsts have type arguments*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 180 | else | 
| 35865 | 181 | case strip_prefix tfree_prefix a of | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 182 |                     SOME b => TFree("'" ^ b, HOLogic.typeS)
 | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 183 | | NONE => | 
| 35865 | 184 | case strip_prefix tvar_prefix a of | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 185 | SOME b => make_tvar b | 
| 36285 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 186 | | NONE => make_tparam a (* Variable from the ATP, say "X1" *) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 187 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 188 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 189 | (*Invert the table of translations between Isabelle and ATPs*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 190 | val const_trans_table_inv = | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 191 |       Symtab.update ("fequal", "op =")
 | 
| 35865 | 192 | (Symtab.make (map swap (Symtab.dest const_trans_table))); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 193 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 194 | fun invert_const c = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 195 | case Symtab.lookup const_trans_table_inv c of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 196 | SOME c' => c' | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 197 | | NONE => c; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 198 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 199 | (*The number of type arguments of a constant, zero if it's monomorphic*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 200 | fun num_typargs thy s = length (Sign.const_typargs thy (s, Sign.the_const_type thy s)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 201 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 202 | (*Generates a constant, given its type arguments*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 203 | fun const_of thy (a,Ts) = Const(a, Sign.const_instance thy (a,Ts)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 204 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 205 | (*First-order translation. No types are known for variables. HOLogic.typeT should allow | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 206 | them to be inferred.*) | 
| 22428 | 207 | fun term_of_stree args thy t = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 208 | case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 209 | Int _ => raise STREE t | 
| 22428 | 210 |     | Br ("hBOOL",[t]) => term_of_stree [] thy t  (*ignore hBOOL*)
 | 
| 211 |     | Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t
 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 212 | | Br (a,ts) => | 
| 35865 | 213 | case strip_prefix const_prefix a of | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 214 | SOME "equal" => | 
| 35865 | 215 |               list_comb(Const (@{const_name "op ="}, HOLogic.typeT), List.map (term_of_stree [] thy) ts)
 | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 216 | | SOME b => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 217 | let val c = invert_const b | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 218 | val nterms = length ts - num_typargs thy c | 
| 22428 | 219 | val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args) | 
| 220 | (*Extra args from hAPP come AFTER any arguments given directly to the | |
| 221 | constant.*) | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 222 | val Ts = List.map type_of_stree (List.drop(ts,nterms)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 223 | in list_comb(const_of thy (c, Ts), us) end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 224 | | NONE => (*a variable, not a constant*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 225 | let val T = HOLogic.typeT | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 226 | val opr = (*a Free variable is typically a Skolem function*) | 
| 35865 | 227 | case strip_prefix fixed_var_prefix a of | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 228 | SOME b => Free(b,T) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 229 | | NONE => | 
| 35865 | 230 | case strip_prefix schematic_var_prefix a of | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 231 | SOME b => make_var (b,T) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 232 | | NONE => make_var (a,T) (* Variable from the ATP, say "X1" *) | 
| 23519 | 233 | in list_comb (opr, List.map (term_of_stree [] thy) (ts@args)) end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 234 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 235 | (*Type class literal applied to a type. Returns triple of polarity, class, type.*) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 236 | fun constraint_of_stree pol (Br("c_Not",[t])) = constraint_of_stree (not pol) t
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 237 | | constraint_of_stree pol t = case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 238 | Int _ => raise STREE t | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 239 | | Br (a,ts) => | 
| 35865 | 240 | (case (strip_prefix class_prefix a, map type_of_stree ts) of | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 241 | (SOME b, [T]) => (pol, b, T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 242 | | _ => raise STREE t); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 243 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 244 | (** Accumulate type constraints in a clause: negative type literals **) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 245 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 246 | fun addix (key,z) = Vartab.map_default (key,[]) (cons z); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 247 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 248 | fun add_constraint ((false, cl, TFree(a,_)), vt) = addix ((a,~1),cl) vt | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 249 | | add_constraint ((false, cl, TVar(ix,_)), vt) = addix (ix,cl) vt | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 250 | | add_constraint (_, vt) = vt; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 251 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 252 | (*False literals (which E includes in its proofs) are deleted*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 253 | val nofalses = filter (not o equal HOLogic.false_const); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 254 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 255 | (*Final treatment of the list of "real" literals from a clause.*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 256 | fun finish [] = HOLogic.true_const (*No "real" literals means only type information*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 257 | | finish lits = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 258 | case nofalses lits of | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 259 | [] => HOLogic.false_const (*The empty clause, since we started with real literals*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 260 | | xs => foldr1 HOLogic.mk_disj (rev xs); | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 261 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 262 | (*Accumulate sort constraints in vt, with "real" literals in lits.*) | 
| 32994 | 263 | fun lits_of_strees _ (vt, lits) [] = (vt, finish lits) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 264 | | lits_of_strees ctxt (vt, lits) (t::ts) = | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 265 | lits_of_strees ctxt (add_constraint (constraint_of_stree true t, vt), lits) ts | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 266 | handle STREE _ => | 
| 22428 | 267 | lits_of_strees ctxt (vt, term_of_stree [] (ProofContext.theory_of ctxt) t :: lits) ts; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 268 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 269 | (*Update TVars/TFrees with detected sort constraints.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 270 | fun fix_sorts vt = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 271 | let fun tysubst (Type (a, Ts)) = Type (a, map tysubst Ts) | 
| 33035 | 272 | | tysubst (TVar (xi, s)) = TVar (xi, the_default s (Vartab.lookup vt xi)) | 
| 273 | | tysubst (TFree (x, s)) = TFree (x, the_default s (Vartab.lookup vt (x, ~1))) | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 274 | fun tmsubst (Const (a, T)) = Const (a, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 275 | | tmsubst (Free (a, T)) = Free (a, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 276 | | tmsubst (Var (xi, T)) = Var (xi, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 277 | | tmsubst (t as Bound _) = t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 278 | | tmsubst (Abs (a, T, t)) = Abs (a, tysubst T, tmsubst t) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 279 | | tmsubst (t $ u) = tmsubst t $ tmsubst u; | 
| 36285 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 280 | in not (Vartab.is_empty vt) ? tmsubst end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 281 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 282 | (*Interpret a list of syntax trees as a clause, given by "real" literals and sort constraints. | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 283 | vt0 holds the initial sort constraints, from the conjecture clauses.*) | 
| 23519 | 284 | fun clause_of_strees ctxt vt0 ts = | 
| 22728 | 285 | let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts in | 
| 24680 | 286 | singleton (Syntax.check_terms ctxt) (TypeInfer.constrain HOLogic.boolT (fix_sorts vt dt)) | 
| 36285 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 287 | end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 288 | |
| 29268 | 289 | fun gen_all_vars t = fold_rev Logic.all (OldTerm.term_vars t) t; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 290 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 291 | fun decode_proof_step vt0 (name, ts, deps) ctxt = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 292 | let val cl = clause_of_strees ctxt vt0 ts in | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 293 | ((name, cl, deps), fold Variable.declare_term (OldTerm.term_frees cl) ctxt) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 294 | end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 295 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 296 | (** Global sort constraints on TFrees (from tfree_tcs) are positive unit clauses. **) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 297 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 298 | fun add_tfree_constraint ((true, cl, TFree(a,_)), vt) = addix ((a,~1),cl) vt | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 299 | | add_tfree_constraint (_, vt) = vt; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 300 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 301 | fun tfree_constraints_of_clauses vt [] = vt | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 302 | | tfree_constraints_of_clauses vt ([lit]::tss) = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 303 | (tfree_constraints_of_clauses (add_tfree_constraint (constraint_of_stree true lit, vt)) tss | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 304 | handle STREE _ => (*not a positive type constraint: ignore*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 305 | tfree_constraints_of_clauses vt tss) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 306 | | tfree_constraints_of_clauses vt (_::tss) = tfree_constraints_of_clauses vt tss; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 307 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 308 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 309 | (**** Translation of TSTP files to Isar Proofs ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 310 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 311 | fun decode_proof_steps ctxt tuples = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 312 | let val vt0 = tfree_constraints_of_clauses Vartab.empty (map #2 tuples) in | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 313 | #1 (fold_map (decode_proof_step vt0) tuples ctxt) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 314 | end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 315 | |
| 23519 | 316 | (** Finding a matching assumption. The literals may be permuted, and variable names | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 317 | may disagree. We must try all combinations of literals (quadratic!) and | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 318 | match the variable names consistently. **) | 
| 23519 | 319 | |
| 35865 | 320 | fun strip_alls_aux n (Const(@{const_name all}, _)$Abs(a,T,t))  =
 | 
| 23519 | 321 | strip_alls_aux (n+1) (subst_bound (Var ((a,n), T), t)) | 
| 322 | | strip_alls_aux _ t = t; | |
| 323 | ||
| 324 | val strip_alls = strip_alls_aux 0; | |
| 325 | ||
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 326 | exception MATCH_LITERAL of unit | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 327 | |
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 328 | (* Remark 1: Ignore types. They are not to be trusted. | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 329 | Remark 2: Ignore order of arguments for equality. SPASS sometimes swaps | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 330 | them for no apparent reason. *) | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 331 | fun match_literal (Const (@{const_name "op ="}, _) $ t1 $ u1)
 | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 332 |                   (Const (@{const_name "op ="}, _) $ t2 $ u2) env =
 | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 333 | (env |> match_literal t1 t2 |> match_literal u1 u2 | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 334 | handle MATCH_LITERAL () => | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 335 | env |> match_literal t1 u2 |> match_literal u1 t2) | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 336 | | match_literal (t1 $ u1) (t2 $ u2) env = | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 337 | env |> match_literal t1 t2 |> match_literal u1 u2 | 
| 31038 | 338 | | match_literal (Abs (_,_,t1)) (Abs (_,_,t2)) env = | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 339 | match_literal t1 t2 env | 
| 31038 | 340 | | match_literal (Bound i1) (Bound i2) env = | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 341 | if i1=i2 then env else raise MATCH_LITERAL () | 
| 31038 | 342 | | match_literal (Const(a1,_)) (Const(a2,_)) env = | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 343 | if a1=a2 then env else raise MATCH_LITERAL () | 
| 31038 | 344 | | match_literal (Free(a1,_)) (Free(a2,_)) env = | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 345 | if a1=a2 then env else raise MATCH_LITERAL () | 
| 23519 | 346 | | match_literal (Var(ix1,_)) (Var(ix2,_)) env = insert (op =) (ix1,ix2) env | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 347 | | match_literal _ _ _ = raise MATCH_LITERAL () | 
| 23519 | 348 | |
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 349 | (* Checking that all variable associations are unique. The list "env" contains | 
| 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 350 | no repetitions, but does it contain say (x, y) and (y, y)? *) | 
| 31038 | 351 | fun good env = | 
| 23519 | 352 | let val (xs,ys) = ListPair.unzip env | 
| 353 | in not (has_duplicates (op=) xs orelse has_duplicates (op=) ys) end; | |
| 354 | ||
| 355 | (*Match one list of literals against another, ignoring types and the order of | |
| 356 | literals. Sorting is unreliable because we don't have types or variable names.*) | |
| 357 | fun matches_aux _ [] [] = true | |
| 358 | | matches_aux env (lit::lits) ts = | |
| 359 | let fun match1 us [] = false | |
| 360 | | match1 us (t::ts) = | |
| 361 | let val env' = match_literal lit t env | |
| 31038 | 362 | in (good env' andalso matches_aux env' lits (us@ts)) orelse | 
| 363 | match1 (t::us) ts | |
| 23519 | 364 | end | 
| 36293 
e6db3ba0b61d
handle SPASS's equality predicate correctly in Isar proof reconstruction
 blanchet parents: 
36291diff
changeset | 365 | handle MATCH_LITERAL () => match1 (t::us) ts | 
| 31038 | 366 | in match1 [] ts end; | 
| 23519 | 367 | |
| 368 | (*Is this length test useful?*) | |
| 31038 | 369 | fun matches (lits1,lits2) = | 
| 370 | length lits1 = length lits2 andalso | |
| 23519 | 371 | matches_aux [] (map Envir.eta_contract lits1) (map Envir.eta_contract lits2); | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 372 | |
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 373 | fun permuted_clause t = | 
| 24958 | 374 | let val lits = HOLogic.disjuncts t | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 375 | fun perm [] = NONE | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 376 | | perm (ctm::ctms) = | 
| 24958 | 377 | if matches (lits, HOLogic.disjuncts (HOLogic.dest_Trueprop (strip_alls ctm))) | 
| 23519 | 378 | then SOME ctm else perm ctms | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 379 | in perm end; | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 380 | |
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 381 | (*ctms is a list of conjecture clauses as yielded by Isabelle. Those returned by the | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 382 | ATP may have their literals reordered.*) | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 383 | fun isar_proof_body ctxt sorts ctms = | 
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 384 | let | 
| 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 385 | val _ = trace_proof_msg (K "\n\nisar_proof_body: start\n") | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 386 | val string_of_term = | 
| 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 387 | PrintMode.setmp (filter (curry (op =) Symbol.xsymbolsN) | 
| 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 388 | (print_mode_value ())) | 
| 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 389 | (Syntax.string_of_term ctxt) | 
| 35966 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 blanchet parents: 
35963diff
changeset | 390 | fun have_or_show "show" _ = " show \"" | 
| 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 blanchet parents: 
35963diff
changeset | 391 | | have_or_show have lname = " " ^ have ^ " " ^ lname ^ ": \"" | 
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 392 | fun do_line _ (lname, t, []) = | 
| 36285 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 393 | (* No depedencies: it's a conjecture clause, with no proof. *) | 
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 394 | (case permuted_clause t ctms of | 
| 35966 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 blanchet parents: 
35963diff
changeset | 395 | SOME u => " assume " ^ lname ^ ": \"" ^ string_of_term u ^ "\"\n" | 
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 396 |         | NONE => raise TERM ("Sledgehammer_Proof_Reconstruct.isar_proof_body",
 | 
| 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 397 | [t])) | 
| 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 398 | | do_line have (lname, t, deps) = | 
| 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 399 | have_or_show have lname ^ | 
| 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 400 | string_of_term (gen_all_vars (HOLogic.mk_Trueprop t)) ^ | 
| 35966 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 blanchet parents: 
35963diff
changeset | 401 | "\"\n by (metis " ^ space_implode " " deps ^ ")\n" | 
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 402 | fun do_lines [(lname, t, deps)] = [do_line "show" (lname, t, deps)] | 
| 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 403 | | do_lines ((lname, t, deps) :: lines) = | 
| 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 404 | do_line "have" (lname, t, deps) :: do_lines lines | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 405 | in setmp_CRITICAL show_sorts sorts do_lines end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 406 | |
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 407 | fun unequal t (_, t', _) = not (t aconv t'); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 408 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 409 | (*No "real" literals means only type information*) | 
| 23519 | 410 | fun eq_types t = t aconv HOLogic.true_const; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 411 | |
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 412 | fun replace_dep (old:int, new) dep = if dep=old then new else [dep]; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 413 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 414 | fun replace_deps (old:int, new) (lno, t, deps) = | 
| 33042 | 415 | (lno, t, List.foldl (uncurry (union (op =))) [] (map (replace_dep (old, new)) deps)); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 416 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 417 | (*Discard axioms; consolidate adjacent lines that prove the same clause, since they differ | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 418 | only in type information.*) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 419 | fun add_proof_line thm_names (lno, t, []) lines = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 420 | (* No dependencies: axiom or conjecture clause *) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 421 | if is_axiom thm_names lno then | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 422 | (* Axioms are not proof lines *) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 423 | if eq_types t then | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 424 | (* Must be clsrel/clsarity: type information, so delete refs to it *) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 425 | map (replace_deps (lno, [])) lines | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 426 | else | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 427 | (case take_prefix (unequal t) lines of | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 428 | (_,[]) => lines (*no repetition of proof line*) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 429 | | (pre, (lno', _, _) :: post) => (*repetition: replace later line by earlier one*) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 430 | pre @ map (replace_deps (lno', [lno])) post) | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 431 | else | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 432 | (lno, t, []) :: lines | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 433 | | add_proof_line _ (lno, t, deps) lines = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 434 | if eq_types t then (lno, t, deps) :: lines | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 435 | (*Type information will be deleted later; skip repetition test.*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 436 | else (*FIXME: Doesn't this code risk conflating proofs involving different types??*) | 
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 437 | case take_prefix (unequal t) lines of | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 438 | (_,[]) => (lno, t, deps) :: lines (*no repetition of proof line*) | 
| 32994 | 439 | | (pre, (lno', t', _) :: post) => | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 440 | (lno, t', deps) :: (*repetition: replace later line by earlier one*) | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 441 | (pre @ map (replace_deps (lno', [lno])) post); | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 442 | |
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 443 | (*Recursively delete empty lines (type information) from the proof.*) | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 444 | fun add_nonnull_prfline ((lno, t, []), lines) = (*no dependencies, so a conjecture clause*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 445 | if eq_types t (*must be type information, tfree_tcs, clsrel, clsarity: delete refs to it*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 446 | then delete_dep lno lines | 
| 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 447 | else (lno, t, []) :: lines | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 448 | | add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines | 
| 30190 | 449 | and delete_dep lno lines = List.foldr add_nonnull_prfline [] (map (replace_deps (lno, [])) lines); | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 450 | |
| 35865 | 451 | fun bad_free (Free (a,_)) = String.isPrefix skolem_prefix a | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 452 | | bad_free _ = false; | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 453 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 454 | (*TVars are forbidden in goals. Also, we don't want lines with <2 dependencies. | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 455 | To further compress proofs, setting modulus:=n deletes every nth line, and nlines | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 456 | counts the number of proof lines processed so far. | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 457 | Deleted lines are replaced by their own dependencies. Note that the "add_nonnull_prfline" | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 458 | phase may delete some dependencies, hence this phase comes later.*) | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 459 | fun add_wanted_prfline ctxt _ ((lno, t, []), (nlines, lines)) = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 460 | (nlines, (lno, t, []) :: lines) (*conjecture clauses must be kept*) | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 461 | | add_wanted_prfline ctxt modulus ((lno, t, deps), (nlines, lines)) = | 
| 29272 | 462 | if eq_types t orelse not (null (Term.add_tvars t [])) orelse | 
| 29268 | 463 | exists_subterm bad_free t orelse | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24920diff
changeset | 464 | (not (null lines) andalso (*final line can't be deleted for these reasons*) | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 465 | (length deps < 2 orelse nlines mod modulus <> 0)) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 466 | then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 467 | else (nlines+1, (lno, t, deps) :: lines); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 468 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 469 | (*Replace numeric proof lines by strings, either from thm_names or sequential line numbers*) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 470 | fun stringify_deps thm_names deps_map [] = [] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 471 | | stringify_deps thm_names deps_map ((lno, t, deps) :: lines) = | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 472 | if is_axiom thm_names lno then | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 473 | (Vector.sub(thm_names,lno-1), t, []) :: stringify_deps thm_names deps_map lines | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 474 | else let val lname = Int.toString (length deps_map) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 475 | fun fix lno = if is_axiom thm_names lno | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 476 | then SOME(Vector.sub(thm_names,lno-1)) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 477 | else AList.lookup (op =) deps_map lno; | 
| 32952 | 478 | in (lname, t, map_filter fix (distinct (op=) deps)) :: | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 479 | stringify_deps thm_names ((lno,lname)::deps_map) lines | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 480 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 481 | |
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 482 | fun isar_proof_start i = | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 483 | (if i = 1 then "" else "prefer " ^ string_of_int i ^ "\n") ^ | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 484 | "proof (neg_clausify)\n"; | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 485 | fun isar_fixes [] = "" | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 486 | | isar_fixes ts = " fix " ^ space_implode " " ts ^ "\n"; | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 487 | fun isar_proof_end 1 = "qed" | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 488 | | isar_proof_end _ = "next" | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 489 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 490 | fun isar_proof_from_atp_proof cnfs modulus sorts ctxt goal i thm_names = | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 491 | let | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 492 | val _ = trace_proof_msg (K "\nisar_proof_from_atp_proof: start\n") | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 493 | val tuples = map (parse_proof_line o explode) cnfs | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 494 | val _ = trace_proof_msg (fn () => | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 495 | Int.toString (length tuples) ^ " tuples extracted\n") | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 496 | val ctxt = ProofContext.set_mode ProofContext.mode_schematic ctxt | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 497 | val raw_lines = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 498 | fold_rev (add_proof_line thm_names) (decode_proof_steps ctxt tuples) [] | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 499 | val _ = trace_proof_msg (fn () => | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 500 | Int.toString (length raw_lines) ^ " raw_lines extracted\n") | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 501 | val nonnull_lines = List.foldr add_nonnull_prfline [] raw_lines | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 502 | val _ = trace_proof_msg (fn () => | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 503 | Int.toString (length nonnull_lines) ^ " nonnull_lines extracted\n") | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 504 | val (_, lines) = List.foldr (add_wanted_prfline ctxt modulus) (0,[]) nonnull_lines | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 505 | val _ = trace_proof_msg (fn () => | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 506 | Int.toString (length lines) ^ " lines extracted\n") | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 507 | val (ccls, fixes) = neg_conjecture_clauses ctxt goal i | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 508 | val _ = trace_proof_msg (fn () => | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 509 | Int.toString (length ccls) ^ " conjecture clauses\n") | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 510 | val ccls = map forall_intr_vars ccls | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 511 | val _ = app (fn th => trace_proof_msg | 
| 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 512 | (fn () => "\nccl: " ^ string_of_thm ctxt th)) ccls | 
| 36064 
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
 blanchet parents: 
36063diff
changeset | 513 | val body = isar_proof_body ctxt sorts (map prop_of ccls) | 
| 35869 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 blanchet parents: 
35868diff
changeset | 514 | (stringify_deps thm_names [] lines) | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 515 | val n = Logic.count_prems (prop_of goal) | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 516 | val _ = trace_proof_msg (K "\nisar_proof_from_atp_proof: finishing\n") | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 517 | in | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 518 | isar_proof_start i ^ isar_fixes (map #1 fixes) ^ implode body ^ | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 519 | isar_proof_end n ^ "\n" | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 520 | end | 
| 36288 | 521 | handle STREE _ => raise Fail "Cannot parse ATP output"; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 522 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 523 | |
| 33310 | 524 | (*=== EXTRACTING PROOF-TEXT === *) | 
| 31866 | 525 | |
| 35865 | 526 | val begin_proof_strs = ["# SZS output start CNFRefutation.", | 
| 33310 | 527 | "=========== Refutation ==========", | 
| 31866 | 528 | "Here is a proof"]; | 
| 33310 | 529 | |
| 35865 | 530 | val end_proof_strs = ["# SZS output end CNFRefutation", | 
| 33310 | 531 | "======= End of refutation =======", | 
| 31866 | 532 | "Formulae used in the proof"]; | 
| 33310 | 533 | |
| 534 | fun get_proof_extract proof = | |
| 36288 | 535 | (* Splits by the first possible of a list of splitters. *) | 
| 536 | case pairself (find_first (fn s => String.isSubstring s proof)) | |
| 537 | (begin_proof_strs, end_proof_strs) of | |
| 538 | (SOME begin_string, SOME end_string) => | |
| 539 | proof |> first_field begin_string |> the |> snd | |
| 540 | |> first_field end_string |> the |> fst | |
| 541 | | _ => raise Fail "Cannot extract proof" | |
| 31866 | 542 | |
| 35865 | 543 | (* ==== CHECK IF PROOF WAS SUCCESSFUL === *) | 
| 31866 | 544 | |
| 35865 | 545 | fun is_proof_well_formed proof = | 
| 36288 | 546 | forall (exists (fn s => String.isSubstring s proof)) | 
| 547 | [begin_proof_strs, end_proof_strs] | |
| 31866 | 548 | |
| 33310 | 549 | (* === EXTRACTING LEMMAS === *) | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 550 | (* A list consisting of the first number in each line is returned. | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 551 | TPTP: Interesting lines have the form "cnf(108, axiom, ...)", where the | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 552 | number (108) is extracted. | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 553 | DFG: Lines have the form "108[0:Inp] ...", where the first number (108) is | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 554 | extracted. *) | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 555 | fun get_step_nums proof_extract = | 
| 35865 | 556 | let | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 557 | val toks = String.tokens (not o is_ident_char) | 
| 35865 | 558 |     fun inputno ("cnf" :: ntok :: "axiom" :: _) = Int.fromString ntok
 | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 559 |       | inputno ("cnf" :: ntok :: "negated_conjecture" :: _) =
 | 
| 35865 | 560 | Int.fromString ntok | 
| 36288 | 561 | | inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok (* DFG format *) | 
| 35865 | 562 | | inputno _ = NONE | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 563 | val lines = split_lines proof_extract | 
| 35865 | 564 | in map_filter (inputno o toks) lines end | 
| 33310 | 565 | |
| 566 | (*Used to label theorems chained into the sledgehammer call*) | |
| 567 | val chained_hint = "CHAINED"; | |
| 35865 | 568 | val kill_chained = filter_out (curry (op =) chained_hint) | 
| 569 | ||
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 570 | fun apply_command _ 1 = "by " | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 571 | | apply_command 1 _ = "apply " | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 572 | | apply_command i _ = "prefer " ^ string_of_int i ^ " apply " | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 573 | fun metis_command i n [] = | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 574 | apply_command i n ^ "metis" | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 575 | | metis_command i n xs = | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 576 | apply_command i n ^ "(metis " ^ space_implode " " xs ^ ")" | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 577 | fun metis_line i n xs = | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 578 | "Try this command: " ^ | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 579 | Markup.markup Markup.sendback (metis_command i n xs) ^ ".\n" | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 580 | fun minimize_line _ [] = "" | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 581 | | minimize_line minimize_command facts = | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 582 | case minimize_command facts of | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 583 | "" => "" | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 584 | | command => | 
| 36065 | 585 | "To minimize the number of lemmas, try this command: " ^ | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 586 | Markup.markup Markup.sendback command ^ ".\n" | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 587 | |
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 588 | fun metis_proof_text (minimize_command, proof, thm_names, goal, i) = | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 589 | let | 
| 36231 
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
 blanchet parents: 
36225diff
changeset | 590 | val lemmas = | 
| 
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
 blanchet parents: 
36225diff
changeset | 591 | proof |> get_proof_extract | 
| 
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
 blanchet parents: 
36225diff
changeset | 592 | |> get_step_nums | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 593 | |> filter (is_axiom thm_names) | 
| 36231 
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
 blanchet parents: 
36225diff
changeset | 594 | |> map (fn i => Vector.sub (thm_names, i - 1)) | 
| 
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
 blanchet parents: 
36225diff
changeset | 595 | |> filter (fn x => x <> "??.unknown") | 
| 
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
 blanchet parents: 
36225diff
changeset | 596 | |> sort_distinct string_ord | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 597 | val n = Logic.count_prems (prop_of goal) | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 598 | val xs = kill_chained lemmas | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36001diff
changeset | 599 | in | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36231diff
changeset | 600 | (metis_line i n xs ^ minimize_line minimize_command xs, kill_chained lemmas) | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 601 | end | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30874diff
changeset | 602 | |
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 603 | val is_proof_line = String.isPrefix "cnf(" orf String.isSubstring "||"
 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 604 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 605 | fun do_space c = if Char.isSpace c then "" else str c | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 606 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 607 | fun strip_spaces_in_list [] = "" | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 608 | | strip_spaces_in_list [c1] = do_space c1 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 609 | | strip_spaces_in_list [c1, c2] = do_space c1 ^ do_space c2 | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 610 | | strip_spaces_in_list (c1 :: c2 :: c3 :: cs) = | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 611 | if Char.isSpace c1 then | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 612 | strip_spaces_in_list (c2 :: c3 :: cs) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 613 | else if Char.isSpace c2 then | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 614 | if Char.isSpace c3 then | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 615 | strip_spaces_in_list (c1 :: c3 :: cs) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 616 | else | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 617 | str c1 ^ | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 618 | (if is_ident_char c1 andalso is_ident_char c3 then " " else "") ^ | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 619 | strip_spaces_in_list (c3 :: cs) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 620 | else | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 621 | str c1 ^ strip_spaces_in_list (c2 :: c3 :: cs) | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 622 | |
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 623 | val strip_spaces = strip_spaces_in_list o String.explode | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 624 | |
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 625 | fun isar_proof_text debug modulus sorts ctxt | 
| 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 626 | (minimize_command, proof, thm_names, goal, i) = | 
| 33310 | 627 | let | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 628 | val cnfs = proof |> get_proof_extract |> split_lines |> map strip_spaces | 
| 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 629 | |> filter is_proof_line | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 630 | val (one_line_proof, lemma_names) = | 
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 631 | metis_proof_text (minimize_command, proof, thm_names, goal, i) | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 632 | val tokens = String.tokens (fn c => c = #" ") one_line_proof | 
| 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 | 633 | fun isar_proof_for () = | 
| 36291 
b4c2043cc96c
added Isar proof reconstruction support for SPASS -- which means all provers can now yield Isar proofs;
 blanchet parents: 
36288diff
changeset | 634 | case isar_proof_from_atp_proof cnfs modulus sorts ctxt goal i thm_names of | 
| 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 | 635 | "" => "" | 
| 36285 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 636 | | isar_proof => | 
| 
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
 blanchet parents: 
36283diff
changeset | 637 | "\nStructured proof:\n" ^ Markup.markup Markup.sendback isar_proof | 
| 35868 
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
 blanchet parents: 
35865diff
changeset | 638 | val isar_proof = | 
| 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 | 639 | if member (op =) tokens chained_hint then | 
| 
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 | 640 | "" | 
| 
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 | 641 | else if debug then | 
| 
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 | 642 | 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 | 643 | else | 
| 
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 | 644 | try isar_proof_for () | 
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36285diff
changeset | 645 | |> the_default "Warning: The Isar proof construction failed.\n" | 
| 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 | 646 | in (one_line_proof ^ isar_proof, lemma_names) end | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 647 | |
| 36288 | 648 | fun proof_text isar_proof debug modulus sorts ctxt = | 
| 649 | if isar_proof then isar_proof_text debug modulus sorts ctxt | |
| 650 | else metis_proof_text | |
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36140diff
changeset | 651 | |
| 31038 | 652 | end; |