author | blanchet |
Thu, 22 Apr 2010 14:47:52 +0200 | |
changeset 36287 | 96f45c5ffb36 |
parent 36285 | d868b34d604c |
child 36288 | 156e4f179bb0 |
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:
24387
diff
changeset
|
8 |
sig |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
9 |
type minimize_command = string list -> string |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
10 |
|
25492
4cc7976948ac
Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents:
25414
diff
changeset
|
11 |
val chained_hint: string |
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
12 |
val invert_const: string -> string |
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
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:
24387
diff
changeset
|
15 |
val make_tvar: string -> typ |
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
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:
36001
diff
changeset
|
18 |
val metis_line: int -> int -> string list -> string |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
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:
36285
diff
changeset
|
20 |
minimize_command * string * string vector * thm * int |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
21 |
-> string * string list |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
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:
36285
diff
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:
36285
diff
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:
36285
diff
changeset
|
25 |
-> string * string list |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
26 |
val 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:
36285
diff
changeset
|
27 |
bool -> bool -> 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:
36285
diff
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:
36285
diff
changeset
|
29 |
-> string * string list |
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
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:
36231
diff
changeset
|
38 |
type minimize_command = string list -> string |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
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:
31866
diff
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:
25999
diff
changeset
|
46 |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
47 |
(**** PARSING OF TSTP FORMAT ****) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
48 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
49 |
(*Syntax trees, either termlist or formulae*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
50 |
datatype stree = Int of int | Br of string * stree list; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
51 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
52 |
fun atom x = Br(x,[]); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
53 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
54 |
fun scons (x,y) = Br("cons", [x,y]); |
30190 | 55 |
val listof = List.foldl scons (atom "nil"); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
56 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
57 |
(*Strings enclosed in single quotes, e.g. filenames*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
58 |
val quoted = $$"'" |-- Scan.repeat (~$$"'") --| $$"'" >> implode; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
59 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
60 |
(*Intended for $true and $false*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
61 |
fun tf s = "c_" ^ str (Char.toUpper (String.sub(s,0))) ^ String.extract(s,1,NONE); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
62 |
val truefalse = $$"$" |-- Symbol.scan_id >> (atom o tf); |
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 |
(*Integer constants, typically proof line numbers*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
65 |
fun is_digit s = Char.isDigit (String.sub(s,0)); |
33035 | 66 |
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
|
67 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
68 |
(*Generalized FO terms, which include filenames, numbers, etc.*) |
25999 | 69 |
fun termlist x = (term ::: Scan.repeat ($$"," |-- term)) x |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
70 |
and term x = (quoted >> atom || integer>>Int || truefalse || |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
71 |
Symbol.scan_id -- Scan.optional ($$"(" |-- termlist --| $$")") [] >> Br || |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
72 |
$$"(" |-- term --| $$")" || |
24547
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
73 |
$$"[" |-- Scan.optional termlist [] --| $$"]" >> listof) x; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
74 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
75 |
fun negate t = Br("c_Not", [t]); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
76 |
fun equate (t1,t2) = Br("c_equal", [t1,t2]); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
77 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
78 |
(*Apply equal or not-equal to a term*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
79 |
fun syn_equal (t, NONE) = t |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
80 |
| syn_equal (t1, SOME (NONE, t2)) = equate (t1,t2) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
81 |
| syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2)); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
82 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
83 |
(*Literals can involve negation, = and !=.*) |
24547
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
84 |
fun literal x = ($$"~" |-- literal >> negate || |
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
85 |
(term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal)) x; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
86 |
|
25999 | 87 |
val literals = literal ::: Scan.repeat ($$"|" |-- literal); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
88 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
89 |
(*Clause: a list of literals separated by the disjunction sign*) |
24547
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
90 |
val clause = $$"(" |-- literals --| $$")" || Scan.single literal; |
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 |
val annotations = $$"," |-- term -- Scan.option ($$"," |-- termlist); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
93 |
|
25718 | 94 |
(*<cnf_annotated> ::= cnf(<name>,<formula_role>,<cnf_formula><annotations>). |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
95 |
The <name> could be an identifier, but we assume integers.*) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
96 |
val tstp_line = (Scan.this_string "cnf" -- $$"(") |-- |
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
97 |
integer --| $$"," -- Symbol.scan_id --| $$"," -- |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
98 |
clause -- Scan.option annotations --| $$ ")"; |
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 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
101 |
(**** INTERPRETATION OF TSTP SYNTAX TREES ****) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
102 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
103 |
exception STREE of stree; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
104 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
105 |
(*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:
23083
diff
changeset
|
106 |
fun strip_prefix s1 s = |
31038 | 107 |
if String.isPrefix s1 s |
35865 | 108 |
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
|
109 |
else NONE; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
110 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
111 |
(*Invert the table of translations between Isabelle and ATPs*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
112 |
val type_const_trans_table_inv = |
35865 | 113 |
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
|
114 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
115 |
fun invert_type_const c = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
116 |
case Symtab.lookup type_const_trans_table_inv c of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
117 |
SOME c' => c' |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
118 |
| NONE => c; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
119 |
|
36285
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
blanchet
parents:
36283
diff
changeset
|
120 |
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:
36283
diff
changeset
|
121 |
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
|
122 |
fun make_var (b,T) = Var((b,0),T); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
123 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
124 |
(*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
|
125 |
by information from type literals, or by type inference.*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
126 |
fun type_of_stree t = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
127 |
case t of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
128 |
Int _ => raise STREE t |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
129 |
| Br (a,ts) => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
130 |
let val Ts = map type_of_stree ts |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
131 |
in |
35865 | 132 |
case strip_prefix tconst_prefix a of |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
133 |
SOME b => Type(invert_type_const b, Ts) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
134 |
| NONE => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
135 |
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:
23083
diff
changeset
|
136 |
else |
35865 | 137 |
case strip_prefix tfree_prefix a of |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
138 |
SOME b => TFree("'" ^ b, HOLogic.typeS) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
139 |
| NONE => |
35865 | 140 |
case strip_prefix tvar_prefix a of |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
141 |
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:
36283
diff
changeset
|
142 |
| 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
|
143 |
end; |
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 |
(*Invert the table of translations between Isabelle and ATPs*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
146 |
val const_trans_table_inv = |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
147 |
Symtab.update ("fequal", "op =") |
35865 | 148 |
(Symtab.make (map swap (Symtab.dest const_trans_table))); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
149 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
150 |
fun invert_const c = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
151 |
case Symtab.lookup const_trans_table_inv c of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
152 |
SOME c' => c' |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
153 |
| NONE => c; |
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 |
(*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
|
156 |
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
|
157 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
158 |
(*Generates a constant, given its type arguments*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
159 |
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
|
160 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
161 |
(*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
|
162 |
them to be inferred.*) |
22428 | 163 |
fun term_of_stree args thy t = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
164 |
case t of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
165 |
Int _ => raise STREE t |
22428 | 166 |
| Br ("hBOOL",[t]) => term_of_stree [] thy t (*ignore hBOOL*) |
167 |
| Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
168 |
| Br (a,ts) => |
35865 | 169 |
case strip_prefix const_prefix a of |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
170 |
SOME "equal" => |
35865 | 171 |
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:
23083
diff
changeset
|
172 |
| SOME b => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
173 |
let val c = invert_const b |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
174 |
val nterms = length ts - num_typargs thy c |
22428 | 175 |
val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args) |
176 |
(*Extra args from hAPP come AFTER any arguments given directly to the |
|
177 |
constant.*) |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
178 |
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
|
179 |
in list_comb(const_of thy (c, Ts), us) end |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
180 |
| NONE => (*a variable, not a constant*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
181 |
let val T = HOLogic.typeT |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
182 |
val opr = (*a Free variable is typically a Skolem function*) |
35865 | 183 |
case strip_prefix fixed_var_prefix a of |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
184 |
SOME b => Free(b,T) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
185 |
| NONE => |
35865 | 186 |
case strip_prefix schematic_var_prefix a of |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
187 |
SOME b => make_var (b,T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
188 |
| NONE => make_var (a,T) (*Variable from the ATP, say X1*) |
23519 | 189 |
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
|
190 |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
191 |
(*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
|
192 |
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
|
193 |
| constraint_of_stree pol t = case t of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
194 |
Int _ => raise STREE t |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
195 |
| Br (a,ts) => |
35865 | 196 |
(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
|
197 |
(SOME b, [T]) => (pol, b, T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
198 |
| _ => raise STREE t); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
199 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
200 |
(** Accumulate type constraints in a clause: negative type literals **) |
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 |
fun addix (key,z) = Vartab.map_default (key,[]) (cons z); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
203 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
204 |
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
|
205 |
| 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
|
206 |
| add_constraint (_, vt) = vt; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
207 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
208 |
(*False literals (which E includes in its proofs) are deleted*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
209 |
val nofalses = filter (not o equal HOLogic.false_const); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
210 |
|
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
211 |
(*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:
22470
diff
changeset
|
212 |
fun finish [] = HOLogic.true_const (*No "real" literals means only type information*) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
213 |
| finish lits = |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
214 |
case nofalses lits of |
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
215 |
[] => 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:
22470
diff
changeset
|
216 |
| xs => foldr1 HOLogic.mk_disj (rev xs); |
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
217 |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
218 |
(*Accumulate sort constraints in vt, with "real" literals in lits.*) |
32994 | 219 |
fun lits_of_strees _ (vt, lits) [] = (vt, finish lits) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
220 |
| 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:
21999
diff
changeset
|
221 |
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:
23083
diff
changeset
|
222 |
handle STREE _ => |
22428 | 223 |
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
|
224 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
225 |
(*Update TVars/TFrees with detected sort constraints.*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
226 |
fun fix_sorts vt = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
227 |
let fun tysubst (Type (a, Ts)) = Type (a, map tysubst Ts) |
33035 | 228 |
| tysubst (TVar (xi, s)) = TVar (xi, the_default s (Vartab.lookup vt xi)) |
229 |
| 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
|
230 |
fun tmsubst (Const (a, T)) = Const (a, tysubst T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
231 |
| tmsubst (Free (a, T)) = Free (a, tysubst T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
232 |
| tmsubst (Var (xi, T)) = Var (xi, tysubst T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
233 |
| tmsubst (t as Bound _) = t |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
234 |
| 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
|
235 |
| 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:
36283
diff
changeset
|
236 |
in not (Vartab.is_empty vt) ? tmsubst end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
237 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
238 |
(*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
|
239 |
vt0 holds the initial sort constraints, from the conjecture clauses.*) |
23519 | 240 |
fun clause_of_strees ctxt vt0 ts = |
22728 | 241 |
let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts in |
24680 | 242 |
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:
36283
diff
changeset
|
243 |
end |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
244 |
|
29268 | 245 |
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
|
246 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
247 |
fun ints_of_stree_aux (Int n, ns) = n::ns |
30190 | 248 |
| ints_of_stree_aux (Br(_,ts), ns) = List.foldl ints_of_stree_aux ns ts; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
249 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
250 |
fun ints_of_stree t = ints_of_stree_aux (t, []); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
251 |
|
25086
729f9aad1f50
Improving the propagation of type constraints for Frees
paulson
parents:
24958
diff
changeset
|
252 |
fun decode_tstp vt0 (name, role, ts, annots) ctxt = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
253 |
let val deps = case annots of NONE => [] | SOME (source,_) => ints_of_stree source |
25086
729f9aad1f50
Improving the propagation of type constraints for Frees
paulson
parents:
24958
diff
changeset
|
254 |
val cl = clause_of_strees ctxt vt0 ts |
29268 | 255 |
in ((name, role, cl, deps), fold Variable.declare_term (OldTerm.term_frees cl) ctxt) end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
256 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
257 |
fun dest_tstp ((((name, role), ts), annots), chs) = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
258 |
case chs of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
259 |
"."::_ => (name, role, ts, annots) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
260 |
| _ => error ("TSTP line not terminated by \".\": " ^ implode chs); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
261 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
262 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
263 |
(** 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
|
264 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
265 |
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
|
266 |
| add_tfree_constraint (_, vt) = vt; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
267 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
268 |
fun tfree_constraints_of_clauses vt [] = vt |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
269 |
| tfree_constraints_of_clauses vt ([lit]::tss) = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
270 |
(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
|
271 |
handle STREE _ => (*not a positive type constraint: ignore*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
272 |
tfree_constraints_of_clauses vt tss) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
273 |
| 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
|
274 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
275 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
276 |
(**** Translation of TSTP files to Isar Proofs ****) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
277 |
|
22012
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents:
21999
diff
changeset
|
278 |
fun decode_tstp_list ctxt tuples = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
279 |
let val vt0 = tfree_constraints_of_clauses Vartab.empty (map #3 tuples) |
25086
729f9aad1f50
Improving the propagation of type constraints for Frees
paulson
parents:
24958
diff
changeset
|
280 |
in #1 (fold_map (decode_tstp vt0) tuples ctxt) end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
281 |
|
23519 | 282 |
(** Finding a matching assumption. The literals may be permuted, and variable names |
31038 | 283 |
may disagree. We have to try all combinations of literals (quadratic!) and |
23519 | 284 |
match up the variable names consistently. **) |
285 |
||
35865 | 286 |
fun strip_alls_aux n (Const(@{const_name all}, _)$Abs(a,T,t)) = |
23519 | 287 |
strip_alls_aux (n+1) (subst_bound (Var ((a,n), T), t)) |
288 |
| strip_alls_aux _ t = t; |
|
289 |
||
290 |
val strip_alls = strip_alls_aux 0; |
|
291 |
||
292 |
exception MATCH_LITERAL; |
|
22012
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents:
21999
diff
changeset
|
293 |
|
23519 | 294 |
(*Ignore types: they are not to be trusted...*) |
295 |
fun match_literal (t1$u1) (t2$u2) env = |
|
296 |
match_literal t1 t2 (match_literal u1 u2 env) |
|
31038 | 297 |
| match_literal (Abs (_,_,t1)) (Abs (_,_,t2)) env = |
23519 | 298 |
match_literal t1 t2 env |
31038 | 299 |
| match_literal (Bound i1) (Bound i2) env = |
23519 | 300 |
if i1=i2 then env else raise MATCH_LITERAL |
31038 | 301 |
| match_literal (Const(a1,_)) (Const(a2,_)) env = |
23519 | 302 |
if a1=a2 then env else raise MATCH_LITERAL |
31038 | 303 |
| match_literal (Free(a1,_)) (Free(a2,_)) env = |
23519 | 304 |
if a1=a2 then env else raise MATCH_LITERAL |
305 |
| match_literal (Var(ix1,_)) (Var(ix2,_)) env = insert (op =) (ix1,ix2) env |
|
32994 | 306 |
| match_literal _ _ _ = raise MATCH_LITERAL; |
23519 | 307 |
|
308 |
(*Checking that all variable associations are unique. The list env contains no |
|
309 |
repetitions, but does it contain say (x,y) and (y,y)? *) |
|
31038 | 310 |
fun good env = |
23519 | 311 |
let val (xs,ys) = ListPair.unzip env |
312 |
in not (has_duplicates (op=) xs orelse has_duplicates (op=) ys) end; |
|
313 |
||
314 |
(*Match one list of literals against another, ignoring types and the order of |
|
315 |
literals. Sorting is unreliable because we don't have types or variable names.*) |
|
316 |
fun matches_aux _ [] [] = true |
|
317 |
| matches_aux env (lit::lits) ts = |
|
318 |
let fun match1 us [] = false |
|
319 |
| match1 us (t::ts) = |
|
320 |
let val env' = match_literal lit t env |
|
31038 | 321 |
in (good env' andalso matches_aux env' lits (us@ts)) orelse |
322 |
match1 (t::us) ts |
|
23519 | 323 |
end |
324 |
handle MATCH_LITERAL => match1 (t::us) ts |
|
31038 | 325 |
in match1 [] ts end; |
23519 | 326 |
|
327 |
(*Is this length test useful?*) |
|
31038 | 328 |
fun matches (lits1,lits2) = |
329 |
length lits1 = length lits2 andalso |
|
23519 | 330 |
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:
21979
diff
changeset
|
331 |
|
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
332 |
fun permuted_clause t = |
24958 | 333 |
let val lits = HOLogic.disjuncts t |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
334 |
fun perm [] = NONE |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
335 |
| perm (ctm::ctms) = |
24958 | 336 |
if matches (lits, HOLogic.disjuncts (HOLogic.dest_Trueprop (strip_alls ctm))) |
23519 | 337 |
then SOME ctm else perm ctms |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
338 |
in perm end; |
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
339 |
|
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
340 |
(*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:
21979
diff
changeset
|
341 |
ATP may have their literals reordered.*) |
36064
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
blanchet
parents:
36063
diff
changeset
|
342 |
fun isar_proof_body ctxt sorts ctms = |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
343 |
let |
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
344 |
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:
36063
diff
changeset
|
345 |
val string_of_term = |
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
blanchet
parents:
36063
diff
changeset
|
346 |
PrintMode.setmp (filter (curry (op =) Symbol.xsymbolsN) |
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
blanchet
parents:
36063
diff
changeset
|
347 |
(print_mode_value ())) |
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
blanchet
parents:
36063
diff
changeset
|
348 |
(Syntax.string_of_term ctxt) |
35966
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents:
35963
diff
changeset
|
349 |
fun have_or_show "show" _ = " show \"" |
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents:
35963
diff
changeset
|
350 |
| have_or_show have lname = " " ^ have ^ " " ^ lname ^ ": \"" |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
351 |
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:
36283
diff
changeset
|
352 |
(* No depedencies: it's a conjecture clause, with no proof. *) |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
353 |
(case permuted_clause t ctms of |
35966
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents:
35963
diff
changeset
|
354 |
SOME u => " assume " ^ lname ^ ": \"" ^ string_of_term u ^ "\"\n" |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
355 |
| NONE => raise TERM ("Sledgehammer_Proof_Reconstruct.isar_proof_body", |
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
356 |
[t])) |
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
357 |
| do_line have (lname, t, deps) = |
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
358 |
have_or_show have lname ^ |
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
359 |
string_of_term (gen_all_vars (HOLogic.mk_Trueprop t)) ^ |
35966
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents:
35963
diff
changeset
|
360 |
"\"\n by (metis " ^ space_implode " " deps ^ ")\n" |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
361 |
fun do_lines [(lname, t, deps)] = [do_line "show" (lname, t, deps)] |
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
362 |
| do_lines ((lname, t, deps) :: lines) = |
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
363 |
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:
36063
diff
changeset
|
364 |
in setmp_CRITICAL show_sorts sorts do_lines end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
365 |
|
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
366 |
fun unequal t (_, t', _) = not (t aconv t'); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
367 |
|
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
368 |
(*No "real" literals means only type information*) |
23519 | 369 |
fun eq_types t = t aconv HOLogic.true_const; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
370 |
|
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
371 |
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
|
372 |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
373 |
fun replace_deps (old:int, new) (lno, t, deps) = |
33042 | 374 |
(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
|
375 |
|
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
376 |
(*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:
22470
diff
changeset
|
377 |
only in type information.*) |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
378 |
fun add_prfline ((lno, "axiom", t, []), lines) = (*axioms are not proof lines*) |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
379 |
if eq_types t (*must be clsrel/clsarity: type information, so delete refs to it*) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
380 |
then map (replace_deps (lno, [])) lines |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
381 |
else |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
382 |
(case take_prefix (unequal t) lines of |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
383 |
(_,[]) => lines (*no repetition of proof line*) |
32994 | 384 |
| (pre, (lno', _, _) :: post) => (*repetition: replace later line by earlier one*) |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
385 |
pre @ map (replace_deps (lno', [lno])) post) |
32994 | 386 |
| add_prfline ((lno, _, t, []), lines) = (*no deps: conjecture clause*) |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
387 |
(lno, t, []) :: lines |
32994 | 388 |
| add_prfline ((lno, _, t, deps), lines) = |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
389 |
if eq_types t then (lno, t, deps) :: lines |
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
390 |
(*Type information will be deleted later; skip repetition test.*) |
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
391 |
else (*FIXME: Doesn't this code risk conflating proofs involving different types??*) |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
392 |
case take_prefix (unequal t) lines of |
22044
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
393 |
(_,[]) => (lno, t, deps) :: lines (*no repetition of proof line*) |
32994 | 394 |
| (pre, (lno', t', _) :: post) => |
22044
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
395 |
(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:
22012
diff
changeset
|
396 |
(pre @ map (replace_deps (lno', [lno])) post); |
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
397 |
|
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
398 |
(*Recursively delete empty lines (type information) from the proof.*) |
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
399 |
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:
22470
diff
changeset
|
400 |
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:
23083
diff
changeset
|
401 |
then delete_dep lno lines |
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
402 |
else (lno, t, []) :: lines |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
403 |
| add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines |
30190 | 404 |
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:
22428
diff
changeset
|
405 |
|
35865 | 406 |
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:
22728
diff
changeset
|
407 |
| bad_free _ = false; |
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
408 |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
409 |
(*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:
22470
diff
changeset
|
410 |
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:
22470
diff
changeset
|
411 |
counts the number of proof lines processed so far. |
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
412 |
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:
22012
diff
changeset
|
413 |
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:
36063
diff
changeset
|
414 |
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:
22470
diff
changeset
|
415 |
(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:
36063
diff
changeset
|
416 |
| add_wanted_prfline ctxt modulus ((lno, t, deps), (nlines, lines)) = |
29272 | 417 |
if eq_types t orelse not (null (Term.add_tvars t [])) orelse |
29268 | 418 |
exists_subterm bad_free t orelse |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24920
diff
changeset
|
419 |
(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:
36063
diff
changeset
|
420 |
(length deps < 2 orelse nlines mod modulus <> 0)) |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
421 |
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:
22470
diff
changeset
|
422 |
else (nlines+1, (lno, t, deps) :: lines); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
423 |
|
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
424 |
(*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
|
425 |
fun stringify_deps thm_names deps_map [] = [] |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
426 |
| stringify_deps thm_names deps_map ((lno, t, deps) :: lines) = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
427 |
if lno <= Vector.length thm_names (*axiom*) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
428 |
then (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:
21978
diff
changeset
|
429 |
else let val lname = Int.toString (length deps_map) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
430 |
fun fix lno = if lno <= Vector.length thm_names |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
431 |
then SOME(Vector.sub(thm_names,lno-1)) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
432 |
else AList.lookup op= deps_map lno; |
32952 | 433 |
in (lname, t, map_filter fix (distinct (op=) deps)) :: |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
434 |
stringify_deps thm_names ((lno,lname)::deps_map) lines |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
435 |
end; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
436 |
|
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
437 |
fun isar_proof_start i = |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
438 |
(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:
36001
diff
changeset
|
439 |
"proof (neg_clausify)\n"; |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
440 |
fun isar_fixes [] = "" |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
441 |
| isar_fixes ts = " fix " ^ space_implode " " ts ^ "\n"; |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
442 |
fun isar_proof_end 1 = "qed" |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
443 |
| isar_proof_end _ = "next" |
21979
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents:
21978
diff
changeset
|
444 |
|
36064
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
blanchet
parents:
36063
diff
changeset
|
445 |
fun isar_proof_from_tstp_file cnfs modulus sorts ctxt goal i thm_names = |
35868
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
446 |
let |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
447 |
val _ = trace_proof_msg (K "\nisar_proof_from_tstp_file: start\n") |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
448 |
val tuples = map (dest_tstp o tstp_line o explode) cnfs |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
449 |
val _ = trace_proof_msg (fn () => |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
450 |
Int.toString (length tuples) ^ " tuples extracted\n") |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
451 |
val ctxt = ProofContext.set_mode ProofContext.mode_schematic ctxt |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
452 |
val raw_lines = List.foldr add_prfline [] (decode_tstp_list ctxt tuples) |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
453 |
val _ = trace_proof_msg (fn () => |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
454 |
Int.toString (length raw_lines) ^ " raw_lines extracted\n") |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
455 |
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:
35865
diff
changeset
|
456 |
val _ = trace_proof_msg (fn () => |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
457 |
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:
36063
diff
changeset
|
458 |
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:
35865
diff
changeset
|
459 |
val _ = trace_proof_msg (fn () => |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
460 |
Int.toString (length lines) ^ " lines extracted\n") |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
461 |
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:
35865
diff
changeset
|
462 |
val _ = trace_proof_msg (fn () => |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
463 |
Int.toString (length ccls) ^ " conjecture clauses\n") |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
464 |
val ccls = map forall_intr_vars ccls |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
465 |
val _ = app (fn th => trace_proof_msg |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
466 |
(fn () => "\nccl: " ^ string_of_thm ctxt th)) ccls |
36064
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
blanchet
parents:
36063
diff
changeset
|
467 |
val body = isar_proof_body ctxt sorts (map prop_of ccls) |
35869
cac366550624
start work on direct proof reconstruction for Sledgehammer
blanchet
parents:
35868
diff
changeset
|
468 |
(stringify_deps thm_names [] lines) |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
469 |
val n = Logic.count_prems (prop_of goal) |
35868
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
470 |
val _ = trace_proof_msg (K "\nisar_proof_from_tstp_file: finishing\n") |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
471 |
in |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
472 |
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:
36001
diff
changeset
|
473 |
isar_proof_end n ^ "\n" |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
474 |
end |
35868
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
475 |
handle STREE _ => error "Could not extract proof (ATP output malformed?)"; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
476 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
477 |
|
33310 | 478 |
(*=== EXTRACTING PROOF-TEXT === *) |
31866 | 479 |
|
35865 | 480 |
val begin_proof_strs = ["# SZS output start CNFRefutation.", |
33310 | 481 |
"=========== Refutation ==========", |
31866 | 482 |
"Here is a proof"]; |
33310 | 483 |
|
35865 | 484 |
val end_proof_strs = ["# SZS output end CNFRefutation", |
33310 | 485 |
"======= End of refutation =======", |
31866 | 486 |
"Formulae used in the proof"]; |
33310 | 487 |
|
488 |
fun get_proof_extract proof = |
|
489 |
let |
|
31866 | 490 |
(*splits to_split by the first possible of a list of splitters*) |
491 |
val (begin_string, end_string) = |
|
35865 | 492 |
(find_first (fn s => String.isSubstring s proof) begin_proof_strs, |
493 |
find_first (fn s => String.isSubstring s proof) end_proof_strs) |
|
33310 | 494 |
in |
495 |
if is_none begin_string orelse is_none end_string |
|
496 |
then error "Could not extract proof (no substring indicating a proof)" |
|
497 |
else proof |> first_field (the begin_string) |> the |> snd |
|
498 |
|> first_field (the end_string) |> the |> fst |
|
499 |
end; |
|
31866 | 500 |
|
35865 | 501 |
(* ==== CHECK IF PROOF WAS SUCCESSFUL === *) |
31866 | 502 |
|
35865 | 503 |
fun is_proof_well_formed proof = |
504 |
exists (fn s => String.isSubstring s proof) begin_proof_strs andalso |
|
505 |
exists (fn s => String.isSubstring s proof) end_proof_strs |
|
31866 | 506 |
|
33310 | 507 |
(* === EXTRACTING LEMMAS === *) |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
508 |
(* A list consisting of the first number in each line is returned. |
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
509 |
TPTP: Interesting lines have the form "cnf(108, axiom, ...)", where the |
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
510 |
number (108) is extracted. |
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
511 |
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:
36140
diff
changeset
|
512 |
extracted. *) |
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
513 |
fun get_step_nums proof_extract = |
35865 | 514 |
let |
515 |
val toks = String.tokens (not o Char.isAlphaNum) |
|
516 |
fun inputno ("cnf" :: ntok :: "axiom" :: _) = Int.fromString ntok |
|
517 |
| inputno ("cnf" :: ntok :: "negated" :: "conjecture" :: _) = |
|
518 |
Int.fromString ntok |
|
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
519 |
| inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok (* DFG *) |
35865 | 520 |
| inputno _ = NONE |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
521 |
val lines = split_lines proof_extract |
35865 | 522 |
in map_filter (inputno o toks) lines end |
33310 | 523 |
|
524 |
(*Used to label theorems chained into the sledgehammer call*) |
|
525 |
val chained_hint = "CHAINED"; |
|
35865 | 526 |
val kill_chained = filter_out (curry (op =) chained_hint) |
527 |
||
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
528 |
fun apply_command _ 1 = "by " |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
529 |
| apply_command 1 _ = "apply " |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
530 |
| 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:
36001
diff
changeset
|
531 |
fun metis_command i n [] = |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
532 |
apply_command i n ^ "metis" |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
533 |
| metis_command i n xs = |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
534 |
apply_command i n ^ "(metis " ^ space_implode " " xs ^ ")" |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
535 |
fun metis_line i n xs = |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
536 |
"Try this command: " ^ |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
537 |
Markup.markup Markup.sendback (metis_command i n xs) ^ ".\n" |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
538 |
fun minimize_line _ [] = "" |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
539 |
| minimize_line minimize_command facts = |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
540 |
case minimize_command facts of |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
541 |
"" => "" |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
542 |
| command => |
36065 | 543 |
"To minimize the number of lemmas, try this command: " ^ |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
544 |
Markup.markup Markup.sendback command ^ ".\n" |
31840
beeaa1ed1f47
check if conjectures have been used in proof
immler@in.tum.de
parents:
31479
diff
changeset
|
545 |
|
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:
36285
diff
changeset
|
546 |
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:
36001
diff
changeset
|
547 |
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:
36225
diff
changeset
|
548 |
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:
36225
diff
changeset
|
549 |
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:
36225
diff
changeset
|
550 |
|> get_step_nums |
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:
36225
diff
changeset
|
551 |
|> filter (fn i => i <= Vector.length thm_names) |
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:
36225
diff
changeset
|
552 |
|> 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:
36225
diff
changeset
|
553 |
|> 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:
36225
diff
changeset
|
554 |
|> sort_distinct string_ord |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
555 |
val n = Logic.count_prems (prop_of goal) |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
556 |
val xs = kill_chained lemmas |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36001
diff
changeset
|
557 |
in |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36231
diff
changeset
|
558 |
(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:
36140
diff
changeset
|
559 |
end |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
30874
diff
changeset
|
560 |
|
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:
36285
diff
changeset
|
561 |
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:
36285
diff
changeset
|
562 |
(minimize_command, proof, thm_names, goal, i) = |
33310 | 563 |
let |
36064
48aec67c284f
added "modulus" and "sorts" options to control Sledgehammer's Isar proof output
blanchet
parents:
36063
diff
changeset
|
564 |
(* We could use "split_lines", but it can return blank lines. *) |
35865 | 565 |
val lines = String.tokens (equal #"\n"); |
35868
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
566 |
val kill_spaces = |
491a97039ce1
renamed "e_full" and "vampire_full" to "e_isar" and "vampire_isar";
blanchet
parents:
35865
diff
changeset
|
567 |
String.translate (fn c => if Char.isSpace c then "" else str c) |
35865 | 568 |
val extract = get_proof_extract proof |
569 |
val cnfs = filter (String.isPrefix "cnf(") (map kill_spaces (lines extract)) |
|
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
570 |
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:
36285
diff
changeset
|
571 |
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:
35865
diff
changeset
|
572 |
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:
36281
diff
changeset
|
573 |
fun 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:
36281
diff
changeset
|
574 |
case isar_proof_from_tstp_file cnfs modulus sorts ctxt goal i thm_names of |
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:
36281
diff
changeset
|
575 |
"" => "" |
36285
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
blanchet
parents:
36283
diff
changeset
|
576 |
| isar_proof => |
d868b34d604c
Isar proof reconstruction: type variables introduced by the ATP should be treated as type parameters for type inference purposes;
blanchet
parents:
36283
diff
changeset
|
577 |
"\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:
35865
diff
changeset
|
578 |
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:
36281
diff
changeset
|
579 |
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:
36281
diff
changeset
|
580 |
"" |
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:
36281
diff
changeset
|
581 |
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:
36281
diff
changeset
|
582 |
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:
36281
diff
changeset
|
583 |
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:
36281
diff
changeset
|
584 |
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:
36285
diff
changeset
|
585 |
|> 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:
36281
diff
changeset
|
586 |
in (one_line_proof ^ isar_proof, lemma_names) end |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
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:
36285
diff
changeset
|
588 |
fun proof_text supports_isar_proofs isar_proof debug modulus sorts ctxt params = |
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:
36285
diff
changeset
|
589 |
if isar_proof then |
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:
36285
diff
changeset
|
590 |
if supports_isar_proofs then |
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:
36285
diff
changeset
|
591 |
isar_proof_text debug modulus sorts ctxt params |
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:
36285
diff
changeset
|
592 |
else |
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:
36285
diff
changeset
|
593 |
metis_proof_text params |
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:
36285
diff
changeset
|
594 |
|>> suffix "(Isar proof output is not supported for this ATP.)\n" |
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:
36285
diff
changeset
|
595 |
else |
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:
36285
diff
changeset
|
596 |
metis_proof_text params |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36140
diff
changeset
|
597 |
|
31038 | 598 |
end; |