author | wenzelm |
Sat, 06 Jun 2009 21:46:36 +0200 | |
changeset 31479 | 08e2a70d002a |
parent 31410 | c231efe693ce |
child 31840 | beeaa1ed1f47 |
permissions | -rw-r--r-- |
29268 | 1 |
(* Author: L C Paulson and Claire Quigley, Cambridge University Computer Laboratory *) |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
2 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
3 |
(***************************************************************************) |
21979
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents:
21978
diff
changeset
|
4 |
(* Code to deal with the transfer of proofs from a prover process *) |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
5 |
(***************************************************************************) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
6 |
signature RES_RECONSTRUCT = |
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
7 |
sig |
25492
4cc7976948ac
Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents:
25414
diff
changeset
|
8 |
val chained_hint: string |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
9 |
|
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
10 |
val fix_sorts: sort Vartab.table -> term -> term |
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
11 |
val invert_const: string -> string |
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
12 |
val invert_type_const: string -> string |
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
13 |
val num_typargs: Context.theory -> string -> int |
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
14 |
val make_tvar: string -> typ |
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
15 |
val strip_prefix: string -> string -> string option |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25492
diff
changeset
|
16 |
val setup: Context.theory -> Context.theory |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
17 |
(* extracting lemma list*) |
30874
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
18 |
val find_failure: string -> string option |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
30874
diff
changeset
|
19 |
val lemma_list_dfg: string -> string * string vector * Proof.context * Thm.thm * int -> string |
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
30874
diff
changeset
|
20 |
val lemma_list_tstp: string -> string * string vector * Proof.context * Thm.thm * int -> string |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
21 |
(* structured proofs *) |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
30874
diff
changeset
|
22 |
val structured_proof: string -> string * string vector * Proof.context * Thm.thm * int -> string |
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
23 |
end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
24 |
|
24425
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
paulson
parents:
24387
diff
changeset
|
25 |
structure ResReconstruct : RES_RECONSTRUCT = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
26 |
struct |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
27 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
28 |
val trace_path = Path.basic "atp_trace"; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
29 |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
30 |
fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
31 |
else (); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
32 |
|
26928 | 33 |
val string_of_thm = PrintMode.setmp [] Display.string_of_thm; |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
34 |
|
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25492
diff
changeset
|
35 |
(*For generating structured proofs: keep every nth proof line*) |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
36 |
val (modulus, modulus_setup) = Attrib.config_int "sledgehammer_modulus" 1; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
37 |
|
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25492
diff
changeset
|
38 |
(*Indicates whether to include sort information in generated proofs*) |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
39 |
val (recon_sorts, recon_sorts_setup) = Attrib.config_bool "sledgehammer_sorts" true; |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
40 |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
41 |
(*Indicated whether to generate full proofs or just lemma lists - now via setup of atps*) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
42 |
(* val (full_proofs, full_proofs_setup) = Attrib.config_bool "sledgehammer_full" false; *) |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
43 |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
44 |
val setup = modulus_setup #> recon_sorts_setup; |
22044
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
45 |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
46 |
(**** PARSING OF TSTP FORMAT ****) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
47 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
48 |
(*Syntax trees, either termlist or formulae*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
49 |
datatype stree = Int of int | Br of string * stree list; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
50 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
51 |
fun atom x = Br(x,[]); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
52 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
53 |
fun scons (x,y) = Br("cons", [x,y]); |
30190 | 54 |
val listof = List.foldl scons (atom "nil"); |
21978
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 |
(*Strings enclosed in single quotes, e.g. filenames*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
57 |
val quoted = $$"'" |-- Scan.repeat (~$$"'") --| $$"'" >> implode; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
58 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
59 |
(*Intended for $true and $false*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
60 |
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
|
61 |
val truefalse = $$"$" |-- Symbol.scan_id >> (atom o tf); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
62 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
63 |
(*Integer constants, typically proof line numbers*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
64 |
fun is_digit s = Char.isDigit (String.sub(s,0)); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
65 |
val integer = Scan.many1 is_digit >> (valOf o Int.fromString o implode); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
66 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
67 |
(*Generalized FO terms, which include filenames, numbers, etc.*) |
25999 | 68 |
fun termlist x = (term ::: Scan.repeat ($$"," |-- term)) x |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
69 |
and term x = (quoted >> atom || integer>>Int || truefalse || |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
70 |
Symbol.scan_id -- Scan.optional ($$"(" |-- termlist --| $$")") [] >> Br || |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
71 |
$$"(" |-- term --| $$")" || |
24547
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
72 |
$$"[" |-- Scan.optional termlist [] --| $$"]" >> listof) x; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
73 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
74 |
fun negate t = Br("c_Not", [t]); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
75 |
fun equate (t1,t2) = Br("c_equal", [t1,t2]); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
76 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
77 |
(*Apply equal or not-equal to a term*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
78 |
fun syn_equal (t, NONE) = t |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
79 |
| syn_equal (t1, SOME (NONE, t2)) = equate (t1,t2) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
80 |
| syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2)); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
81 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
82 |
(*Literals can involve negation, = and !=.*) |
24547
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
83 |
fun literal x = ($$"~" |-- literal >> negate || |
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
84 |
(term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal)) x; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
85 |
|
25999 | 86 |
val literals = literal ::: Scan.repeat ($$"|" |-- literal); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
87 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
88 |
(*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
|
89 |
val clause = $$"(" |-- literals --| $$")" || Scan.single literal; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
90 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
91 |
val annotations = $$"," |-- term -- Scan.option ($$"," |-- termlist); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
92 |
|
25718 | 93 |
(*<cnf_annotated> ::= cnf(<name>,<formula_role>,<cnf_formula><annotations>). |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
94 |
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
|
95 |
val tstp_line = (Scan.this_string "cnf" -- $$"(") |-- |
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
96 |
integer --| $$"," -- Symbol.scan_id --| $$"," -- |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
97 |
clause -- Scan.option annotations --| $$ ")"; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
98 |
|
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 |
(**** INTERPRETATION OF TSTP SYNTAX TREES ****) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
101 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
102 |
exception STREE of stree; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
103 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
104 |
(*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
|
105 |
fun strip_prefix s1 s = |
31038 | 106 |
if String.isPrefix s1 s |
24182 | 107 |
then SOME (ResClause.undo_ascii_of (String.extract (s, size s1, NONE))) |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
108 |
else NONE; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
109 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
110 |
(*Invert the table of translations between Isabelle and ATPs*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
111 |
val type_const_trans_table_inv = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
112 |
Symtab.make (map swap (Symtab.dest ResClause.type_const_trans_table)); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
113 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
114 |
fun invert_type_const c = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
115 |
case Symtab.lookup type_const_trans_table_inv c of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
116 |
SOME c' => c' |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
117 |
| NONE => c; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
118 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
119 |
fun make_tvar b = TVar(("'" ^ b, 0), HOLogic.typeS); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
120 |
fun make_var (b,T) = Var((b,0),T); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
121 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
122 |
(*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
|
123 |
by information from type literals, or by type inference.*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
124 |
fun type_of_stree t = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
125 |
case t of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
126 |
Int _ => raise STREE t |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
127 |
| Br (a,ts) => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
128 |
let val Ts = map type_of_stree ts |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
129 |
in |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
130 |
case strip_prefix ResClause.tconst_prefix a of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
131 |
SOME b => Type(invert_type_const b, Ts) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
132 |
| NONE => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
133 |
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
|
134 |
else |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
135 |
case strip_prefix ResClause.tfree_prefix a of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
136 |
SOME b => TFree("'" ^ b, HOLogic.typeS) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
137 |
| NONE => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
138 |
case strip_prefix ResClause.tvar_prefix a of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
139 |
SOME b => make_tvar b |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
140 |
| NONE => make_tvar a (*Variable from the ATP, say X1*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
141 |
end; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
142 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
143 |
(*Invert the table of translations between Isabelle and ATPs*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
144 |
val const_trans_table_inv = |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
145 |
Symtab.update ("fequal", "op =") |
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
146 |
(Symtab.make (map swap (Symtab.dest ResClause.const_trans_table))); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
147 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
148 |
fun invert_const c = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
149 |
case Symtab.lookup const_trans_table_inv c of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
150 |
SOME c' => c' |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
151 |
| NONE => c; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
152 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
153 |
(*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
|
154 |
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
|
155 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
156 |
(*Generates a constant, given its type arguments*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
157 |
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
|
158 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
159 |
(*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
|
160 |
them to be inferred.*) |
22428 | 161 |
fun term_of_stree args thy t = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
162 |
case t of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
163 |
Int _ => raise STREE t |
22428 | 164 |
| Br ("hBOOL",[t]) => term_of_stree [] thy t (*ignore hBOOL*) |
165 |
| 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
|
166 |
| Br (a,ts) => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
167 |
case strip_prefix ResClause.const_prefix a of |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
168 |
SOME "equal" => |
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
169 |
list_comb(Const ("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
|
170 |
| SOME b => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
171 |
let val c = invert_const b |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
172 |
val nterms = length ts - num_typargs thy c |
22428 | 173 |
val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args) |
174 |
(*Extra args from hAPP come AFTER any arguments given directly to the |
|
175 |
constant.*) |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
176 |
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
|
177 |
in list_comb(const_of thy (c, Ts), us) end |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
178 |
| NONE => (*a variable, not a constant*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
179 |
let val T = HOLogic.typeT |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
180 |
val opr = (*a Free variable is typically a Skolem function*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
181 |
case strip_prefix ResClause.fixed_var_prefix a of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
182 |
SOME b => Free(b,T) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
183 |
| NONE => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
184 |
case strip_prefix ResClause.schematic_var_prefix a of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
185 |
SOME b => make_var (b,T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
186 |
| NONE => make_var (a,T) (*Variable from the ATP, say X1*) |
23519 | 187 |
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
|
188 |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
189 |
(*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
|
190 |
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
|
191 |
| constraint_of_stree pol t = case t of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
192 |
Int _ => raise STREE t |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
193 |
| Br (a,ts) => |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
194 |
(case (strip_prefix ResClause.class_prefix a, map type_of_stree ts) of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
195 |
(SOME b, [T]) => (pol, b, T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
196 |
| _ => raise STREE t); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
197 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
198 |
(** Accumulate type constraints in a clause: negative type literals **) |
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 |
fun addix (key,z) = Vartab.map_default (key,[]) (cons z); |
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 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
|
203 |
| 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
|
204 |
| add_constraint (_, vt) = vt; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
205 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
206 |
(*False literals (which E includes in its proofs) are deleted*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
207 |
val nofalses = filter (not o equal HOLogic.false_const); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
208 |
|
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
209 |
(*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
|
210 |
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
|
211 |
| finish lits = |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
212 |
case nofalses lits of |
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
213 |
[] => 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
|
214 |
| 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
|
215 |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
216 |
(*Accumulate sort constraints in vt, with "real" literals in lits.*) |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
217 |
fun lits_of_strees ctxt (vt, lits) [] = (vt, finish lits) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
218 |
| 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
|
219 |
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
|
220 |
handle STREE _ => |
22428 | 221 |
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
|
222 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
223 |
(*Update TVars/TFrees with detected sort constraints.*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
224 |
fun fix_sorts vt = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
225 |
let fun tysubst (Type (a, Ts)) = Type (a, map tysubst Ts) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
226 |
| tysubst (TVar (xi, s)) = TVar (xi, getOpt (Vartab.lookup vt xi, s)) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
227 |
| tysubst (TFree (x, s)) = TFree (x, getOpt (Vartab.lookup vt (x,~1), s)) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
228 |
fun tmsubst (Const (a, T)) = Const (a, tysubst T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
229 |
| tmsubst (Free (a, T)) = Free (a, tysubst T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
230 |
| tmsubst (Var (xi, T)) = Var (xi, tysubst T) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
231 |
| tmsubst (t as Bound _) = t |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
232 |
| 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
|
233 |
| tmsubst (t $ u) = tmsubst t $ tmsubst u; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
234 |
in fn t => if Vartab.is_empty vt then t else tmsubst t end; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
235 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
236 |
(*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
|
237 |
vt0 holds the initial sort constraints, from the conjecture clauses.*) |
23519 | 238 |
fun clause_of_strees ctxt vt0 ts = |
22728 | 239 |
let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts in |
24680 | 240 |
singleton (Syntax.check_terms ctxt) (TypeInfer.constrain HOLogic.boolT (fix_sorts vt dt)) |
22728 | 241 |
end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
242 |
|
29268 | 243 |
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
|
244 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
245 |
fun ints_of_stree_aux (Int n, ns) = n::ns |
30190 | 246 |
| 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
|
247 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
248 |
fun ints_of_stree t = ints_of_stree_aux (t, []); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
249 |
|
25086
729f9aad1f50
Improving the propagation of type constraints for Frees
paulson
parents:
24958
diff
changeset
|
250 |
fun decode_tstp vt0 (name, role, ts, annots) ctxt = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
251 |
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
|
252 |
val cl = clause_of_strees ctxt vt0 ts |
29268 | 253 |
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
|
254 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
255 |
fun dest_tstp ((((name, role), ts), annots), chs) = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
256 |
case chs of |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
257 |
"."::_ => (name, role, ts, annots) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
258 |
| _ => error ("TSTP line not terminated by \".\": " ^ implode chs); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
259 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
260 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
261 |
(** 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
|
262 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
263 |
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
|
264 |
| add_tfree_constraint (_, vt) = vt; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
265 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
266 |
fun tfree_constraints_of_clauses vt [] = vt |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
267 |
| tfree_constraints_of_clauses vt ([lit]::tss) = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
268 |
(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
|
269 |
handle STREE _ => (*not a positive type constraint: ignore*) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
270 |
tfree_constraints_of_clauses vt tss) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
271 |
| 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
|
272 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
273 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
274 |
(**** Translation of TSTP files to Isar Proofs ****) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
275 |
|
22012
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents:
21999
diff
changeset
|
276 |
fun decode_tstp_list ctxt tuples = |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
277 |
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
|
278 |
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
|
279 |
|
23519 | 280 |
(** Finding a matching assumption. The literals may be permuted, and variable names |
31038 | 281 |
may disagree. We have to try all combinations of literals (quadratic!) and |
23519 | 282 |
match up the variable names consistently. **) |
283 |
||
31038 | 284 |
fun strip_alls_aux n (Const("all",_)$Abs(a,T,t)) = |
23519 | 285 |
strip_alls_aux (n+1) (subst_bound (Var ((a,n), T), t)) |
286 |
| strip_alls_aux _ t = t; |
|
287 |
||
288 |
val strip_alls = strip_alls_aux 0; |
|
289 |
||
290 |
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
|
291 |
|
23519 | 292 |
(*Ignore types: they are not to be trusted...*) |
293 |
fun match_literal (t1$u1) (t2$u2) env = |
|
294 |
match_literal t1 t2 (match_literal u1 u2 env) |
|
31038 | 295 |
| match_literal (Abs (_,_,t1)) (Abs (_,_,t2)) env = |
23519 | 296 |
match_literal t1 t2 env |
31038 | 297 |
| match_literal (Bound i1) (Bound i2) env = |
23519 | 298 |
if i1=i2 then env else raise MATCH_LITERAL |
31038 | 299 |
| match_literal (Const(a1,_)) (Const(a2,_)) env = |
23519 | 300 |
if a1=a2 then env else raise MATCH_LITERAL |
31038 | 301 |
| match_literal (Free(a1,_)) (Free(a2,_)) env = |
23519 | 302 |
if a1=a2 then env else raise MATCH_LITERAL |
303 |
| match_literal (Var(ix1,_)) (Var(ix2,_)) env = insert (op =) (ix1,ix2) env |
|
304 |
| match_literal _ _ env = raise MATCH_LITERAL; |
|
305 |
||
306 |
(*Checking that all variable associations are unique. The list env contains no |
|
307 |
repetitions, but does it contain say (x,y) and (y,y)? *) |
|
31038 | 308 |
fun good env = |
23519 | 309 |
let val (xs,ys) = ListPair.unzip env |
310 |
in not (has_duplicates (op=) xs orelse has_duplicates (op=) ys) end; |
|
311 |
||
312 |
(*Match one list of literals against another, ignoring types and the order of |
|
313 |
literals. Sorting is unreliable because we don't have types or variable names.*) |
|
314 |
fun matches_aux _ [] [] = true |
|
315 |
| matches_aux env (lit::lits) ts = |
|
316 |
let fun match1 us [] = false |
|
317 |
| match1 us (t::ts) = |
|
318 |
let val env' = match_literal lit t env |
|
31038 | 319 |
in (good env' andalso matches_aux env' lits (us@ts)) orelse |
320 |
match1 (t::us) ts |
|
23519 | 321 |
end |
322 |
handle MATCH_LITERAL => match1 (t::us) ts |
|
31038 | 323 |
in match1 [] ts end; |
23519 | 324 |
|
325 |
(*Is this length test useful?*) |
|
31038 | 326 |
fun matches (lits1,lits2) = |
327 |
length lits1 = length lits2 andalso |
|
23519 | 328 |
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
|
329 |
|
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
330 |
fun permuted_clause t = |
24958 | 331 |
let val lits = HOLogic.disjuncts t |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
332 |
fun perm [] = NONE |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
333 |
| perm (ctm::ctms) = |
24958 | 334 |
if matches (lits, HOLogic.disjuncts (HOLogic.dest_Trueprop (strip_alls ctm))) |
23519 | 335 |
then SOME ctm else perm ctms |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
336 |
in perm end; |
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
337 |
|
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
338 |
fun have_or_show "show " lname = "show \"" |
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
339 |
| have_or_show have lname = have ^ lname ^ ": \"" |
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
340 |
|
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
341 |
(*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
|
342 |
ATP may have their literals reordered.*) |
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
343 |
fun isar_lines ctxt ctms = |
28572 | 344 |
let val string_of = PrintMode.setmp [] (fn term => Syntax.string_of_term ctxt term) |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
345 |
val _ = trace ("\n\nisar_lines: start\n") |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
346 |
fun doline have (lname, t, []) = (*No deps: it's a conjecture clause, with no proof.*) |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
347 |
(case permuted_clause t ctms of |
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
348 |
SOME u => "assume " ^ lname ^ ": \"" ^ string_of u ^ "\"\n" |
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
349 |
| NONE => "assume? " ^ lname ^ ": \"" ^ string_of t ^ "\"\n") (*no match!!*) |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
350 |
| doline have (lname, t, deps) = |
23519 | 351 |
have_or_show have lname ^ string_of (gen_all_vars (HOLogic.mk_Trueprop t)) ^ |
22372 | 352 |
"\"\n by (metis " ^ space_implode " " deps ^ ")\n" |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
353 |
fun dolines [(lname, t, deps)] = [doline "show " (lname, t, deps)] |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
354 |
| dolines ((lname, t, deps)::lines) = doline "have " (lname, t, deps) :: dolines lines |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25492
diff
changeset
|
355 |
in setmp show_sorts (Config.get ctxt recon_sorts) dolines end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
356 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
357 |
fun notequal t (_,t',_) = not (t aconv t'); |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
358 |
|
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
359 |
(*No "real" literals means only type information*) |
23519 | 360 |
fun eq_types t = t aconv HOLogic.true_const; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
361 |
|
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
362 |
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
|
363 |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
364 |
fun replace_deps (old:int, new) (lno, t, deps) = |
30190 | 365 |
(lno, t, List.foldl (op union_int) [] (map (replace_dep (old, new)) deps)); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
366 |
|
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
367 |
(*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
|
368 |
only in type information.*) |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
369 |
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
|
370 |
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
|
371 |
then map (replace_deps (lno, [])) lines |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
372 |
else |
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
373 |
(case take_prefix (notequal t) lines of |
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
374 |
(_,[]) => lines (*no repetition of proof line*) |
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
375 |
| (pre, (lno',t',deps')::post) => (*repetition: replace later line by earlier one*) |
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
376 |
pre @ map (replace_deps (lno', [lno])) post) |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
377 |
| add_prfline ((lno, role, t, []), lines) = (*no deps: conjecture clause*) |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
378 |
(lno, t, []) :: lines |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
379 |
| add_prfline ((lno, role, t, deps), lines) = |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
380 |
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
|
381 |
(*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
|
382 |
else (*FIXME: Doesn't this code risk conflating proofs involving different types??*) |
22044
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
383 |
case take_prefix (notequal t) lines of |
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
384 |
(_,[]) => (lno, t, deps) :: lines (*no repetition of proof line*) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
385 |
| (pre, (lno',t',deps')::post) => |
22044
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents:
22012
diff
changeset
|
386 |
(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
|
387 |
(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
|
388 |
|
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
389 |
(*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
|
390 |
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
|
391 |
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
|
392 |
then delete_dep lno lines |
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
393 |
else (lno, t, []) :: lines |
22470
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
paulson
parents:
22428
diff
changeset
|
394 |
| add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines |
30190 | 395 |
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
|
396 |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24920
diff
changeset
|
397 |
fun bad_free (Free (a,_)) = String.isPrefix "sko_" a |
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
398 |
| bad_free _ = false; |
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
399 |
|
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
400 |
(*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
|
401 |
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
|
402 |
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
|
403 |
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
|
404 |
phase may delete some dependencies, hence this phase comes later.*) |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25492
diff
changeset
|
405 |
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
|
406 |
(nlines, (lno, t, []) :: lines) (*conjecture clauses must be kept*) |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25492
diff
changeset
|
407 |
| add_wanted_prfline ctxt ((lno, t, deps), (nlines, lines)) = |
29272 | 408 |
if eq_types t orelse not (null (Term.add_tvars t [])) orelse |
29268 | 409 |
exists_subterm bad_free t orelse |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24920
diff
changeset
|
410 |
(not (null lines) andalso (*final line can't be deleted for these reasons*) |
31038 | 411 |
(length deps < 2 orelse nlines mod (Config.get ctxt modulus) <> 0)) |
22491
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents:
22470
diff
changeset
|
412 |
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
|
413 |
else (nlines+1, (lno, t, deps) :: lines); |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
414 |
|
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
415 |
(*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
|
416 |
fun stringify_deps thm_names deps_map [] = [] |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
417 |
| stringify_deps thm_names deps_map ((lno, t, deps) :: lines) = |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
418 |
if lno <= Vector.length thm_names (*axiom*) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
419 |
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
|
420 |
else let val lname = Int.toString (length deps_map) |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
421 |
fun fix lno = if lno <= Vector.length thm_names |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
422 |
then SOME(Vector.sub(thm_names,lno-1)) |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
423 |
else AList.lookup op= deps_map lno; |
22731
abfdccaed085
trying to make single-step proofs work better, especially if they contain
paulson
parents:
22728
diff
changeset
|
424 |
in (lname, t, List.mapPartial fix (distinct (op=) deps)) :: |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
425 |
stringify_deps thm_names ((lno,lname)::deps_map) lines |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
426 |
end; |
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
427 |
|
24547
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
428 |
val proofstart = "proof (neg_clausify)\n"; |
21979
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents:
21978
diff
changeset
|
429 |
|
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents:
21978
diff
changeset
|
430 |
fun isar_header [] = proofstart |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
431 |
| isar_header ts = proofstart ^ "fix " ^ space_implode " " ts ^ "\n"; |
21979
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents:
21978
diff
changeset
|
432 |
|
22012
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents:
21999
diff
changeset
|
433 |
fun decode_tstp_file cnfs ctxt th sgno thm_names = |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
434 |
let val _ = trace "\ndecode_tstp_file: start\n" |
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
435 |
val tuples = map (dest_tstp o tstp_line o explode) cnfs |
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
436 |
val _ = trace (Int.toString (length tuples) ^ " tuples extracted\n") |
24552 | 437 |
val ctxt = ProofContext.set_mode ProofContext.mode_schematic ctxt |
30190 | 438 |
val raw_lines = List.foldr add_prfline [] (decode_tstp_list ctxt tuples) |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
439 |
val _ = trace (Int.toString (length raw_lines) ^ " raw_lines extracted\n") |
30190 | 440 |
val nonnull_lines = List.foldr add_nonnull_prfline [] raw_lines |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
441 |
val _ = trace (Int.toString (length nonnull_lines) ^ " nonnull_lines extracted\n") |
30190 | 442 |
val (_,lines) = List.foldr (add_wanted_prfline ctxt) (0,[]) nonnull_lines |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
443 |
val _ = trace (Int.toString (length lines) ^ " lines extracted\n") |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
444 |
val (ccls,fixes) = ResAxioms.neg_conjecture_clauses th sgno |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
445 |
val _ = trace (Int.toString (length ccls) ^ " conjecture clauses\n") |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21979
diff
changeset
|
446 |
val ccls = map forall_intr_vars ccls |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
447 |
val _ = if !Output.debugging then app (fn th => trace ("\nccl: " ^ string_of_thm th)) ccls |
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
448 |
else () |
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
449 |
val ilines = isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines) |
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
450 |
val _ = trace "\ndecode_tstp_file: finishing\n" |
23139
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
wenzelm
parents:
23083
diff
changeset
|
451 |
in |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
452 |
isar_header (map #1 fixes) ^ String.concat ilines ^ "qed\n" |
24547
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
453 |
end |
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents:
24493
diff
changeset
|
454 |
handle e => (*FIXME: exn handler is too general!*) |
31479 | 455 |
let val msg = "Translation of TSTP raised an exception: " ^ ML_Compiler.exn_message e |
26333
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
paulson
parents:
25999
diff
changeset
|
456 |
in trace msg; msg end; |
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
457 |
|
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
458 |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
459 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
460 |
(* ==== CHECK IF PROOF OF E OR VAMPIRE WAS SUCCESSFUL === *) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
461 |
|
30874
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
462 |
val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable", |
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
463 |
"SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"]; |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
464 |
val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"]; |
30874
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
465 |
val failure_strings_SPASS = ["SPASS beiseite: Completion found.", |
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
466 |
"SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."]; |
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
467 |
val failure_strings_remote = ["Remote-script could not extract proof"]; |
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
468 |
fun find_failure proof = |
29597 | 469 |
let val failures = |
31038 | 470 |
map_filter (fn s => if String.isSubstring s proof then SOME s else NONE) |
30874
34927a1e0ae8
reverted to explicitly check the presence of a refutation
immler@in.tum.de
parents:
30857
diff
changeset
|
471 |
(failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote) |
29597 | 472 |
in if null failures then NONE else SOME (hd failures) end; |
473 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
474 |
(*=== EXTRACTING PROOF-TEXT === *) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
475 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
476 |
val begin_proof_strings = ["# SZS output start CNFRefutation.", |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
477 |
"=========== Refutation ==========", |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
478 |
"Here is a proof"]; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
479 |
val end_proof_strings = ["# SZS output end CNFRefutation", |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
480 |
"======= End of refutation =======", |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
481 |
"Formulae used in the proof"]; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
482 |
fun get_proof_extract proof = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
483 |
let |
31038 | 484 |
(*splits to_split by the first possible of a list of splitters*) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
485 |
fun first_field_any [] to_split = NONE |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
486 |
| first_field_any (splitter::splitters) to_split = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
487 |
let |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
488 |
val result = (first_field splitter to_split) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
489 |
in |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
490 |
if isSome result then result else first_field_any splitters to_split |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
491 |
end |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
492 |
val (a:string, b:string) = valOf (first_field_any begin_proof_strings proof) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
493 |
val (proofextract:string, c:string) = valOf (first_field_any end_proof_strings b) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
494 |
in proofextract end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
495 |
|
31038 | 496 |
(* === EXTRACTING LEMMAS === *) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
497 |
(* lines have the form "cnf(108, axiom, ...", |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
498 |
the number (108) has to be extracted)*) |
31038 | 499 |
fun get_step_nums_tstp proofextract = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
500 |
let val toks = String.tokens (not o Char.isAlphaNum) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
501 |
fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
502 |
| inputno _ = NONE |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
503 |
val lines = split_lines proofextract |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
504 |
in List.mapPartial (inputno o toks) lines end |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
505 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
506 |
(*String contains multiple lines. We want those of the form |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
507 |
"253[0:Inp] et cetera..." |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
508 |
A list consisting of the first number in each line is returned. *) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
509 |
fun get_step_nums_dfg proofextract = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
510 |
let val toks = String.tokens (not o Char.isAlphaNum) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
511 |
fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
512 |
| inputno _ = NONE |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
513 |
val lines = split_lines proofextract |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
514 |
in List.mapPartial (inputno o toks) lines end |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
515 |
|
31038 | 516 |
(*extracting lemmas from tstp-output between the lines from above*) |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
30874
diff
changeset
|
517 |
fun extract_lemmas get_step_nums (proof, thm_names, _, _, _) = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
518 |
let |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
519 |
(* get the names of axioms from their numbers*) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
520 |
fun get_axiom_names thm_names step_nums = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
521 |
let |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
522 |
fun is_axiom n = n <= Vector.length thm_names |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
523 |
fun getname i = Vector.sub(thm_names, i-1) |
31038 | 524 |
in |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
525 |
sort_distinct string_ord (filter (fn x => x <> "??.unknown") (map getname (filter is_axiom step_nums))) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
526 |
end |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
527 |
val proofextract = get_proof_extract proof |
31038 | 528 |
in |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
529 |
get_axiom_names thm_names (get_step_nums proofextract) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
530 |
end; |
31410 | 531 |
|
532 |
(*Used to label theorems chained into the sledgehammer call*) |
|
533 |
val chained_hint = "CHAINED"; |
|
534 |
val nochained = filter_out (fn y => y = chained_hint) |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
27865
diff
changeset
|
535 |
|
31038 | 536 |
(* metis-command *) |
537 |
fun metis_line [] = "apply metis" |
|
538 |
| metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")" |
|
539 |
||
540 |
(* atp_minimize [atp=<prover>] <lemmas> *) |
|
541 |
fun minimize_line _ [] = "" |
|
542 |
| minimize_line name lemmas = "For minimizing the number of lemmas try this command:\n" ^ |
|
31410 | 543 |
(Markup.markup Markup.sendback) ("atp_minimize [atp=" ^ name ^ "] " ^ |
544 |
space_implode " " (nochained lemmas)) |
|
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
30874
diff
changeset
|
545 |
|
31038 | 546 |
fun sendback_metis_nochained lemmas = |
31410 | 547 |
(Markup.markup Markup.sendback o metis_line) (nochained lemmas) |
31038 | 548 |
fun lemma_list_tstp name result = |
549 |
let val lemmas = extract_lemmas get_step_nums_tstp result |
|
550 |
in sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas end; |
|
551 |
fun lemma_list_dfg name result = |
|
552 |
let val lemmas = extract_lemmas get_step_nums_dfg result |
|
553 |
in sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas end; |
|
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
30874
diff
changeset
|
554 |
|
31038 | 555 |
(* === Extracting structured Isar-proof === *) |
556 |
fun structured_proof name (result as (proof, thm_names, ctxt, goal, subgoalno)) = |
|
557 |
let |
|
558 |
(*Could use split_lines, but it can return blank lines...*) |
|
559 |
val lines = String.tokens (equal #"\n"); |
|
560 |
val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c) |
|
561 |
val proofextract = get_proof_extract proof |
|
562 |
val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract)) |
|
563 |
val one_line_proof = lemma_list_tstp name result |
|
564 |
val structured = if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then "" |
|
565 |
else decode_tstp_file cnfs ctxt goal subgoalno thm_names |
|
566 |
in |
|
567 |
one_line_proof ^ "\n\n" ^ (Markup.markup Markup.sendback) structured |
|
568 |
end |
|
21978
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff
changeset
|
569 |
|
31038 | 570 |
end; |