| author | paulson | 
| Tue, 09 Jan 2007 18:12:59 +0100 | |
| changeset 22044 | 6c0702a96076 | 
| parent 22012 | adf68479ae1b | 
| child 22130 | 0906fd95e0b5 | 
| permissions | -rwxr-xr-x | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
1  | 
(* ID: $Id$  | 
| 
21979
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
2  | 
Author: L C Paulson and Claire Quigley  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Copyright 2004 University of Cambridge  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
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  | 
(***************************************************************************)  | 
| 
21979
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
7  | 
(* 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
 | 
8  | 
(***************************************************************************)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
9  | 
signature RES_RECONSTRUCT =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
11  | 
val checkEProofFound:  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
12  | 
TextIO.instream * TextIO.outstream * Posix.Process.pid *  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
13  | 
string * Proof.context * thm * int * string Vector.vector -> bool  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
14  | 
val checkVampProofFound:  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
15  | 
TextIO.instream * TextIO.outstream * Posix.Process.pid *  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
16  | 
string * Proof.context * thm * int * string Vector.vector -> bool  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
17  | 
val checkSpassProofFound:  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
18  | 
TextIO.instream * TextIO.outstream * Posix.Process.pid *  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
19  | 
string * Proof.context * thm * int * string Vector.vector -> bool  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
20  | 
val signal_parent:  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
21  | 
TextIO.outstream * Posix.Process.pid * string * string -> unit  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
23  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
25  | 
structure ResReconstruct =  | 
| 
 
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  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
30  | 
fun trace s = if !Output.show_debug_msgs then File.append (File.tmp_path trace_path) s  | 
| 
 
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  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
33  | 
(*Full proof reconstruction wanted*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
34  | 
val full = ref true;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
35  | 
|
| 
22044
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
36  | 
val min_deps = ref 2; (*consolidate proof lines containing fewer dependencies*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
37  | 
|
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
38  | 
(**** PARSING OF TSTP FORMAT ****)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
40  | 
(*Syntax trees, either termlist or formulae*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
41  | 
datatype stree = Int of int | Br of string * stree list;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
43  | 
fun atom x = Br(x,[]);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
45  | 
fun scons (x,y) = Br("cons", [x,y]);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
46  | 
val listof = foldl scons (atom "nil");  | 
| 
 
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  | 
(*Strings enclosed in single quotes, e.g. filenames*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
49  | 
val quoted = $$"'" |-- Scan.repeat (~$$"'") --| $$"'" >> implode;  | 
| 
 
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  | 
(*Intended for $true and $false*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
52  | 
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
 | 
53  | 
val truefalse = $$"$" |-- Symbol.scan_id >> (atom o tf);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
55  | 
(*Integer constants, typically proof line numbers*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
56  | 
fun is_digit s = Char.isDigit (String.sub(s,0));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
57  | 
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
 | 
58  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
59  | 
(*Generalized FO terms, which include filenames, numbers, etc.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
60  | 
fun termlist x = (term -- Scan.repeat ($$"," |-- term) >> op::) x  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
61  | 
and term x = (quoted >> atom || integer>>Int || truefalse ||  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
62  | 
              Symbol.scan_id -- Scan.optional ($$"(" |-- termlist --| $$")") [] >> Br ||
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
63  | 
              $$"(" |-- term --| $$")" ||
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
64  | 
$$"[" |-- termlist --| $$"]" >> listof) x;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
66  | 
fun negate t = Br("c_Not", [t]);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
67  | 
fun equate (t1,t2) = Br("c_equal", [t1,t2]);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
69  | 
(*Apply equal or not-equal to a term*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
70  | 
fun syn_equal (t, NONE) = t  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
71  | 
| syn_equal (t1, SOME (NONE, t2)) = equate (t1,t2)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
72  | 
| syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2));  | 
| 
 
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  | 
(*Literals can involve negation, = and !=.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
75  | 
val literal = $$"~" |-- term >> negate ||  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
76  | 
(term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal) ;  | 
| 
 
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  | 
val literals = literal -- Scan.repeat ($$"|" |-- literal) >> op:: ;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
79  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
80  | 
(*Clause: a list of literals separated by the disjunction sign*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
81  | 
val clause = $$"(" |-- literals --| $$")";
 | 
| 
 
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  | 
val annotations = $$"," |-- term -- Scan.option ($$"," |-- termlist);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
85  | 
(*<cnf_annotated> ::=Êcnf(<name>,<formula_role>,<cnf_formula><annotations>).  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
86  | 
The <name> could be an identifier, but we assume integers.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
87  | 
val tstp_line = (Scan.this_string "cnf" -- $$"(") |-- 
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
88  | 
integer --| $$"," -- Symbol.scan_id --| $$"," --  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
89  | 
clause -- Scan.option annotations --| $$ ")";  | 
| 
 
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  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
92  | 
(**** DUPLICATE of Susanto's code to remove ASCII armouring from names in proof files ****)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
94  | 
(*original file: Isabelle_ext.sml*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
95  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
96  | 
val A_min_spc = Char.ord #"A" - Char.ord #" ";  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
98  | 
fun cList2int chs = getOpt (Int.fromString (String.implode (rev chs)), 0);  | 
| 
 
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  | 
(*why such a tiny range?*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
101  | 
fun check_valid_int x =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
102  | 
let val val_x = cList2int x  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
103  | 
in (length x = 3) andalso (val_x >= 123) andalso (val_x <= 126)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
104  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
106  | 
fun normalise_s s [] st_ sti =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
107  | 
String.implode(rev(  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
108  | 
if st_  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
109  | 
then if null sti  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
110  | 
then (#"_" :: s)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
111  | 
else if check_valid_int sti  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
112  | 
then (Char.chr (cList2int sti) :: s)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
113  | 
else (sti @ (#"_" :: s))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
114  | 
else s))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
115  | 
| normalise_s s (#"_"::cs) st_ sti =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
116  | 
if st_  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
117  | 
then let val s' = if null sti  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
118  | 
then (#"_"::s)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
119  | 
else if check_valid_int sti  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
120  | 
then (Char.chr (cList2int sti) :: s)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
121  | 
else (sti @ (#"_" :: s))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
122  | 
in normalise_s s' cs false []  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
123  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
124  | 
else normalise_s s cs true []  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
125  | 
| normalise_s s (c::cs) true sti =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
126  | 
if (Char.isDigit c)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
127  | 
then normalise_s s cs true (c::sti)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
128  | 
else let val s' = if null sti  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
129  | 
then if ((c >= #"A") andalso (c<= #"P"))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
130  | 
then ((Char.chr(Char.ord c - A_min_spc))::s)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
131  | 
else (c :: (#"_" :: s))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
132  | 
else if check_valid_int sti  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
133  | 
then (Char.chr (cList2int sti) :: s)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
134  | 
else (sti @ (#"_" :: s))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
135  | 
in normalise_s s' cs false []  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
136  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
137  | 
| normalise_s s (c::cs) _ _ = normalise_s (c::s) cs false [];  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
139  | 
(*This version does not look for standard prefixes first.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
140  | 
fun normalise_string s = normalise_s [] (String.explode s) false [];  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
141  | 
|
| 
 
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  | 
(**** INTERPRETATION OF TSTP SYNTAX TREES ****)  | 
| 
 
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  | 
exception STREE of stree;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
146  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
147  | 
(*If string s has the prefix s1, return the result of deleting it.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
148  | 
fun strip_prefix s1 s =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
149  | 
if String.isPrefix s1 s then SOME (normalise_string (String.extract (s, size s1, NONE)))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
150  | 
else NONE;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
152  | 
(*Invert the table of translations between Isabelle and ATPs*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
153  | 
val type_const_trans_table_inv =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
154  | 
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
 | 
155  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
156  | 
fun invert_type_const c =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
157  | 
case Symtab.lookup type_const_trans_table_inv c of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
158  | 
SOME c' => c'  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
159  | 
| NONE => c;  | 
| 
 
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  | 
fun make_tvar b = TVar(("'" ^ b, 0), HOLogic.typeS);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
162  | 
fun make_var (b,T) = Var((b,0),T);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
163  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
164  | 
(*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
 | 
165  | 
by information from type literals, or by type inference.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
166  | 
fun type_of_stree t =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
167  | 
case t of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
168  | 
Int _ => raise STREE t  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
169  | 
| Br (a,ts) =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
170  | 
let val Ts = map type_of_stree ts  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
171  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
172  | 
case strip_prefix ResClause.tconst_prefix a of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
173  | 
SOME b => Type(invert_type_const b, Ts)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
174  | 
| NONE =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
175  | 
if not (null ts) then raise STREE t (*only tconsts have type arguments*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
176  | 
else  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
177  | 
case strip_prefix ResClause.tfree_prefix a of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
178  | 
                    SOME b => TFree("'" ^ b, HOLogic.typeS)
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
179  | 
| NONE =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
180  | 
case strip_prefix ResClause.tvar_prefix a of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
181  | 
SOME b => make_tvar b  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
182  | 
| NONE => make_tvar a (*Variable from the ATP, say X1*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
183  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
185  | 
(*Invert the table of translations between Isabelle and ATPs*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
186  | 
val const_trans_table_inv =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
187  | 
Symtab.make (map swap (Symtab.dest ResClause.const_trans_table));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
188  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
189  | 
fun invert_const c =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
190  | 
case Symtab.lookup const_trans_table_inv c of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
191  | 
SOME c' => c'  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
192  | 
| NONE => c;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
193  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
194  | 
(*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
 | 
195  | 
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
 | 
196  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
197  | 
(*Generates a constant, given its type arguments*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
198  | 
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
 | 
199  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
200  | 
(*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
 | 
201  | 
them to be inferred.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
202  | 
fun term_of_stree thy t =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
203  | 
case t of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
204  | 
Int _ => raise STREE t  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
205  | 
| Br (a,ts) =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
206  | 
case strip_prefix ResClause.const_prefix a of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
207  | 
SOME "equal" =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
208  | 
if length ts = 2 then  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
209  | 
                list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree thy) ts)
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
210  | 
else raise STREE t (*equality needs two arguments*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
211  | 
| SOME b =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
212  | 
let val c = invert_const b  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
213  | 
val nterms = length ts - num_typargs thy c  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
214  | 
val us = List.map (term_of_stree thy) (List.take(ts,nterms))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
215  | 
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
 | 
216  | 
in list_comb(const_of thy (c, Ts), us) end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
217  | 
| NONE => (*a variable, not a constant*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
218  | 
let val T = HOLogic.typeT  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
219  | 
val opr = (*a Free variable is typically a Skolem function*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
220  | 
case strip_prefix ResClause.fixed_var_prefix a of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
221  | 
SOME b => Free(b,T)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
222  | 
| NONE =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
223  | 
case strip_prefix ResClause.schematic_var_prefix a of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
224  | 
SOME b => make_var (b,T)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
225  | 
| NONE => make_var (a,T) (*Variable from the ATP, say X1*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
226  | 
in list_comb (opr, List.map (term_of_stree thy) ts) end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
227  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
228  | 
(*Type class literal applied to a type. Returns triple of polarity, class, type.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
229  | 
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
 | 
230  | 
| constraint_of_stree pol t = case t of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
231  | 
Int _ => raise STREE t  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
232  | 
| Br (a,ts) =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
233  | 
(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
 | 
234  | 
(SOME b, [T]) => (pol, b, T)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
235  | 
| _ => raise STREE t);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
236  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
237  | 
(** Accumulate type constraints in a clause: negative type literals **)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
239  | 
fun addix (key,z) = Vartab.map_default (key,[]) (cons z);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
240  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
241  | 
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
 | 
242  | 
| 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
 | 
243  | 
| add_constraint (_, vt) = vt;  | 
| 
 
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  | 
(*False literals (which E includes in its proofs) are deleted*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
246  | 
val nofalses = filter (not o equal HOLogic.false_const);  | 
| 
 
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  | 
(*Accumulate sort constraints in vt, with "real" literals in lits.*)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
249  | 
fun lits_of_strees ctxt (vt, lits) [] = (vt, rev (nofalses lits))  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
250  | 
| lits_of_strees ctxt (vt, lits) (t::ts) =  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
251  | 
lits_of_strees ctxt (add_constraint (constraint_of_stree true t, vt), lits) ts  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
252  | 
handle STREE _ =>  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
253  | 
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
 | 
254  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
255  | 
(*Update TVars/TFrees with detected sort constraints.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
256  | 
fun fix_sorts vt =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
257  | 
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
 | 
258  | 
| 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
 | 
259  | 
| 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
 | 
260  | 
fun tmsubst (Const (a, T)) = Const (a, tysubst T)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
261  | 
| tmsubst (Free (a, T)) = Free (a, tysubst T)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
262  | 
| tmsubst (Var (xi, T)) = Var (xi, tysubst T)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
263  | 
| tmsubst (t as Bound _) = t  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
264  | 
| 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
 | 
265  | 
| tmsubst (t $ u) = tmsubst t $ tmsubst u;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
266  | 
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
 | 
267  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
268  | 
(*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
 | 
269  | 
vt0 holds the initial sort constraints, from the conjecture clauses.*)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
270  | 
fun clause_of_strees_aux ctxt vt0 ts =  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
271  | 
case lits_of_strees ctxt (vt0,[]) ts of  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
272  | 
(_, []) => HOLogic.false_const  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
273  | 
| (vt, lits) =>  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
274  | 
let val dt = fix_sorts vt (foldr1 HOLogic.mk_disj lits)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
275  | 
val infer = Sign.infer_types (ProofContext.pp ctxt) (ProofContext.theory_of ctxt)  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
276  | 
(ProofContext.consts_of ctxt) (Variable.def_type ctxt false)  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
277  | 
(Variable.def_sort ctxt) (Variable.names_of ctxt) true  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
278  | 
in  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
279  | 
#1(infer ([dt], HOLogic.boolT))  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
280  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
281  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
282  | 
(*Quantification over a list of Vars. FUXNE: for term.ML??*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
283  | 
fun list_all_var ([], t: term) = t  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
284  | 
| list_all_var ((v as Var(ix,T)) :: vars, t) =  | 
| 
21979
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
285  | 
(all T) $ Abs(string_of_indexname ix, T, abstract_over (v, list_all_var (vars,t)));  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
286  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
287  | 
fun gen_all_vars t = list_all_var (term_vars t, t);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
288  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
289  | 
fun clause_of_strees thy vt0 ts =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
290  | 
gen_all_vars (HOLogic.mk_Trueprop (clause_of_strees_aux thy vt0 ts));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
291  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
292  | 
fun ints_of_stree_aux (Int n, ns) = n::ns  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
293  | 
| ints_of_stree_aux (Br(_,ts), ns) = foldl ints_of_stree_aux ns ts;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
294  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
295  | 
fun ints_of_stree t = ints_of_stree_aux (t, []);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
296  | 
|
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
297  | 
fun decode_tstp ctxt vt0 (name, role, ts, annots) =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
298  | 
let val deps = case annots of NONE => [] | SOME (source,_) => ints_of_stree source  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
299  | 
in (name, role, clause_of_strees ctxt vt0 ts, deps) end;  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
300  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
301  | 
fun dest_tstp ((((name, role), ts), annots), chs) =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
302  | 
case chs of  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
303  | 
"."::_ => (name, role, ts, annots)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
304  | 
        | _ => error ("TSTP line not terminated by \".\": " ^ implode chs);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
305  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
306  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
307  | 
(** 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
 | 
308  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
309  | 
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
 | 
310  | 
| add_tfree_constraint (_, vt) = vt;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
312  | 
fun tfree_constraints_of_clauses vt [] = vt  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
313  | 
| tfree_constraints_of_clauses vt ([lit]::tss) =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
314  | 
(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
 | 
315  | 
handle STREE _ => (*not a positive type constraint: ignore*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
316  | 
tfree_constraints_of_clauses vt tss)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
317  | 
| 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
 | 
318  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
319  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
320  | 
(**** Translation of TSTP files to Isar Proofs ****)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
321  | 
|
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
322  | 
fun decode_tstp_list ctxt tuples =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
323  | 
let val vt0 = tfree_constraints_of_clauses Vartab.empty (map #3 tuples)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
324  | 
in map (decode_tstp ctxt vt0) tuples end;  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
325  | 
|
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
326  | 
(*FIXME: simmilar function in res_atp. Move to HOLogic?*)  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
327  | 
fun dest_disj_aux (Const ("op |", _) $ t $ t') disjs = dest_disj_aux t (dest_disj_aux t' disjs)
 | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
328  | 
| dest_disj_aux t disjs = t::disjs;  | 
| 
 
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 dest_disj t = dest_disj_aux t [];  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
331  | 
|
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
332  | 
(*Remove types from a term, to eliminate the randomness of type inference*)  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
333  | 
fun smash_types (Const(a,_)) = Const(a,dummyT)  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
334  | 
| smash_types (Free(a,_)) = Free(a,dummyT)  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
335  | 
| smash_types (Var(a,_)) = Var(a,dummyT)  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
336  | 
| smash_types (f$t) = smash_types f $ smash_types t  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
337  | 
| smash_types t = t;  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
338  | 
|
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
339  | 
val sort_lits = sort Term.fast_term_ord o dest_disj o  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
340  | 
smash_types o HOLogic.dest_Trueprop o strip_all_body;  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
341  | 
|
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
342  | 
fun permuted_clause t =  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
343  | 
let val lits = sort_lits t  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
344  | 
fun perm [] = NONE  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
345  | 
| perm (ctm::ctms) =  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
346  | 
if forall (op aconv) (ListPair.zip (lits, sort_lits ctm)) then SOME ctm  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
347  | 
else perm ctms  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
348  | 
in perm end;  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
349  | 
|
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
350  | 
(*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
 | 
351  | 
ATP may have their literals reordered.*)  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
352  | 
fun isar_lines ctxt ctms =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
353  | 
let val string_of = ProofContext.string_of_term ctxt  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
354  | 
fun doline hs (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
 | 
355  | 
(case permuted_clause t ctms of  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
356  | 
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
 | 
357  | 
| NONE => "assume? " ^ lname ^ ": \"" ^ string_of t ^ "\"\n") (*no match!!*)  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
358  | 
| doline hs (lname, t, deps) =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
359  | 
hs ^ lname ^ ": \"" ^ string_of t ^  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
360  | 
"\"\n by (meson " ^ space_implode " " deps ^ ")\n"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
361  | 
fun dolines [(lname, t, deps)] = [doline "show " (lname, t, deps)]  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
362  | 
| dolines ((lname, t, deps)::lines) = doline "have " (lname, t, deps) :: dolines lines  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
363  | 
in setmp show_sorts true dolines end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
364  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
365  | 
fun notequal t (_,t',_) = not (t aconv t');  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
366  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
367  | 
fun eq_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
368  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
369  | 
fun replace_dep (old, new) dep = if dep=old then new else [dep];  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
370  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
371  | 
fun replace_deps (old, new) (lno, t, deps) =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
372  | 
(lno, t, List.concat (map (replace_dep (old, new)) deps));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
373  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
374  | 
(*Discard axioms and also False conjecture clauses (which can only contain type information).  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
375  | 
Consolidate adjacent lines that prove the same clause, since they differ only in type  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
376  | 
information.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
377  | 
fun add_prfline ((lno, "axiom", t, []), lines) = (*axioms are not proof lines*)  | 
| 
22044
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
378  | 
if eq_false t (*must be clsrel/clsarity: type information, so delete refs to it*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
379  | 
then map (replace_deps (lno, [])) lines  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
380  | 
else (case take_prefix (notequal t) lines of  | 
| 
22044
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
381  | 
(_,[]) => lines (*no repetition of proof line*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
382  | 
| (pre, (lno',t',deps')::post) => (*repetition: replace later line by earlier one*)  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
383  | 
pre @ map (replace_deps (lno', [lno])) post)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
384  | 
| add_prfline ((lno, role, t, []), lines) = (*no deps: conjecture clause*)  | 
| 
22044
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
385  | 
if eq_false t (*must be tfree_tcs: type information, so delete refs to it*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
386  | 
then map (replace_deps (lno, [])) lines  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
387  | 
else (lno, t, []) :: lines  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
388  | 
| add_prfline ((lno, role, t, deps), lines) =  | 
| 
22044
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
389  | 
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
 | 
390  | 
(_,[]) => (lno, t, deps) :: lines (*no repetition of proof line*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
391  | 
| (pre, (lno',t',deps')::post) =>  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
392  | 
(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
 | 
393  | 
(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
 | 
394  | 
|
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
395  | 
(*TVars are forbidden in goals. Also, we don't want lines with too few dependencies.  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
396  | 
Deleted lines are replaced by their own dependencies. Note that the previous "add_prfline"  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
397  | 
phase may delete some dependencies, hence this phase comes later.*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
398  | 
fun add_wanted_prfline ((lno, t, []), lines) =  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
399  | 
(lno, t, []) :: lines (*conjecture clauses must be kept*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
400  | 
| add_wanted_prfline ((lno, t, deps), lines) =  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
401  | 
if not (null (term_tvars t)) orelse length deps < !min_deps  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
402  | 
then map (replace_deps (lno, deps)) lines (*Delete line*)  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
403  | 
else (lno, t, deps) :: lines;  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
404  | 
|
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
405  | 
(*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
 | 
406  | 
fun stringify_deps thm_names deps_map [] = []  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
407  | 
| stringify_deps thm_names deps_map ((lno, t, deps) :: lines) =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
408  | 
if lno <= Vector.length thm_names (*axiom*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
409  | 
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
 | 
410  | 
else let val lname = Int.toString (length deps_map)  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
411  | 
fun fix lno = if lno <= Vector.length thm_names  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
412  | 
then SOME(Vector.sub(thm_names,lno-1))  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
413  | 
else AList.lookup op= deps_map lno;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
414  | 
in (lname, t, List.mapPartial fix deps) ::  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
415  | 
stringify_deps thm_names ((lno,lname)::deps_map) lines  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
416  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
417  | 
|
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
418  | 
val proofstart = "\nproof (neg_clausify)\n";  | 
| 
21979
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
419  | 
|
| 
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
420  | 
fun isar_header [] = proofstart  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
421  | 
| isar_header ts = proofstart ^ "fix " ^ space_implode " " ts ^ "\n";  | 
| 
21979
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
422  | 
|
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
423  | 
fun decode_tstp_file cnfs ctxt th sgno thm_names =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
424  | 
let val tuples = map (dest_tstp o tstp_line o explode) cnfs  | 
| 
22044
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
425  | 
val lines = foldr add_wanted_prfline []  | 
| 
 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 
paulson 
parents: 
22012 
diff
changeset
 | 
426  | 
(foldr add_prfline [] (decode_tstp_list ctxt tuples))  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
427  | 
val (ccls,fixes) = ResAxioms.neg_conjecture_clauses th sgno  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
428  | 
val ccls = map forall_intr_vars ccls  | 
| 
21979
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
429  | 
in  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
430  | 
if !Output.show_debug_msgs  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
431  | 
then app (Output.debug o string_of_thm) ccls  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
432  | 
else ();  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21979 
diff
changeset
 | 
433  | 
isar_header (map #1 fixes) ^  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
434  | 
String.concat (isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines))  | 
| 
21979
 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 
paulson 
parents: 
21978 
diff
changeset
 | 
435  | 
end;  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
436  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
437  | 
(*Could use split_lines, but it can return blank lines...*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
438  | 
val lines = String.tokens (equal #"\n");  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
439  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
440  | 
val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
441  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
442  | 
(*The output to the watcher must be a SINGLE line...clearly \t must not be used.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
443  | 
val encode_newlines = String.translate (fn c => if c = #"\n" then "\t" else str c);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
444  | 
val restore_newlines = String.translate (fn c => if c = #"\t" then "\n" else str c);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
445  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
446  | 
fun signal_success probfile toParent ppid msg =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
447  | 
  (trace ("\nReporting Success for" ^ probfile ^ "\n" ^ msg);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
448  | 
TextIO.output (toParent, "Success. " ^ encode_newlines msg ^ "\n");  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
449  | 
TextIO.output (toParent, probfile ^ "\n");  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
450  | 
TextIO.flushOut toParent;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
451  | 
Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
452  | 
|
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
453  | 
fun tstp_extract proofextract probfile toParent ppid ctxt th sgno thm_names =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
454  | 
  let val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
455  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
456  | 
signal_success probfile toParent ppid  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
457  | 
(decode_tstp_file cnfs ctxt th sgno thm_names)  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
458  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
459  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
460  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
461  | 
(**** retrieve the axioms that were used in the proof ****)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
462  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
463  | 
(*Get names of axioms used. Axioms are indexed from 1, while the vector is indexed from 0*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
464  | 
fun get_axiom_names (thm_names: string vector) step_nums =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
465  | 
let fun is_axiom n = n <= Vector.length thm_names  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
466  | 
fun index i = Vector.sub(thm_names, i-1)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
467  | 
val axnums = List.filter is_axiom step_nums  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
468  | 
val axnames = sort_distinct string_ord (map index axnums)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
469  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
470  | 
if length axnums = length step_nums then "UNSOUND!!" :: axnames  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
471  | 
else axnames  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
472  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
473  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
474  | 
(*String contains multiple lines. We want those of the form  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
475  | 
"253[0:Inp] et cetera..."  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
476  | 
A list consisting of the first number in each line is returned. *)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
477  | 
fun get_spass_linenums proofstr =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
478  | 
let val toks = String.tokens (not o Char.isAlphaNum)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
479  | 
fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
480  | 
| inputno _ = NONE  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
481  | 
val lines = String.tokens (fn c => c = #"\n") proofstr  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
482  | 
in List.mapPartial (inputno o toks) lines end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
483  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
484  | 
fun get_axiom_names_spass proofstr thm_names =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
485  | 
get_axiom_names thm_names (get_spass_linenums proofstr);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
486  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
487  | 
fun not_comma c = c <> #",";  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
488  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
489  | 
(*A valid TSTP axiom line has the form cnf(NNN,axiom,...) where NNN is a positive integer.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
490  | 
fun parse_tstp_line s =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
491  | 
  let val ss = Substring.full (unprefix "cnf(" (nospaces s))
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
492  | 
val (intf,rest) = Substring.splitl not_comma ss  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
493  | 
val (rolef,rest) = Substring.splitl not_comma (Substring.triml 1 rest)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
494  | 
(*We only allow negated_conjecture because the line number will be removed in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
495  | 
get_axiom_names above, while suppressing the UNSOUND warning*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
496  | 
val ints = if Substring.string rolef mem_string ["axiom","negated_conjecture"]  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
497  | 
then Substring.string intf  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
498  | 
else "error"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
499  | 
in Int.fromString ints end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
500  | 
handle Fail _ => NONE;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
501  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
502  | 
fun get_axiom_names_tstp proofstr thm_names =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
503  | 
get_axiom_names thm_names (List.mapPartial parse_tstp_line (split_lines proofstr));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
504  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
505  | 
(*String contains multiple lines. We want those of the form  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
506  | 
"*********** [448, input] ***********".  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
507  | 
A list consisting of the first number in each line is returned. *)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
508  | 
fun get_vamp_linenums proofstr =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
509  | 
let val toks = String.tokens (not o Char.isAlphaNum)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
510  | 
fun inputno [ntok,"input"] = Int.fromString ntok  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
511  | 
| inputno _ = NONE  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
512  | 
val lines = String.tokens (fn c => c = #"\n") proofstr  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
513  | 
in List.mapPartial (inputno o toks) lines end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
514  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
515  | 
fun get_axiom_names_vamp proofstr thm_names =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
516  | 
get_axiom_names thm_names (get_vamp_linenums proofstr);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
517  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
518  | 
fun rules_to_string [] = "NONE"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
519  | 
| rules_to_string xs = space_implode " " xs  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
520  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
521  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
522  | 
(*The signal handler in watcher.ML must be able to read the output of this.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
523  | 
fun prover_lemma_list_aux getax proofstr probfile toParent ppid thm_names =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
524  | 
 (trace ("\n\nGetting lemma names. proofstr is " ^ proofstr ^
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
525  | 
" num of clauses is " ^ string_of_int (Vector.length thm_names));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
526  | 
signal_success probfile toParent ppid  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
527  | 
    ("Lemmas used in automatic proof: " ^ rules_to_string (getax proofstr thm_names)))
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
528  | 
handle e => (*FIXME: exn handler is too general!*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
529  | 
  (trace ("\nprover_lemma_list_aux: In exception handler: " ^ Toplevel.exn_message e);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
530  | 
TextIO.output (toParent, "Translation failed for the proof: " ^  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
531  | 
String.toString proofstr ^ "\n");  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
532  | 
TextIO.output (toParent, probfile);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
533  | 
TextIO.flushOut toParent;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
534  | 
Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
535  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
536  | 
val e_lemma_list = prover_lemma_list_aux get_axiom_names_tstp;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
537  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
538  | 
val vamp_lemma_list = prover_lemma_list_aux get_axiom_names_vamp;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
539  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
540  | 
val spass_lemma_list = prover_lemma_list_aux get_axiom_names_spass;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
541  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
542  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
543  | 
(**** Extracting proofs from an ATP's output ****)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
544  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
545  | 
(*Return everything in s that comes before the string t*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
546  | 
fun cut_before t s =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
547  | 
let val (s1,s2) = Substring.position t (Substring.full s)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
548  | 
in if Substring.size s2 = 0 then error "cut_before: string not found"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
549  | 
else Substring.string s2  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
550  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
551  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
552  | 
val start_E = "# Proof object starts here."  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
553  | 
val end_E = "# Proof object ends here."  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
554  | 
val start_V6 = "%================== Proof: ======================"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
555  | 
val end_V6 = "%============== End of proof. =================="  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
556  | 
val start_V8 = "=========== Refutation =========="  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
557  | 
val end_V8 = "======= End of refutation ======="  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
558  | 
val end_SPASS = "Formulae used in the proof"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
559  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
560  | 
(*********************************************************************************)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
561  | 
(* Inspect the output of an ATP process to see if it has found a proof, *)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
562  | 
(* and if so, transfer output to the input pipe of the main Isabelle process *)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
563  | 
(*********************************************************************************)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
564  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
565  | 
(*Returns "true" if it successfully returns a lemma list, otherwise "false", but this  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
566  | 
return value is currently never used!*)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
567  | 
fun startTransfer endS (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names) =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
568  | 
let fun transferInput currentString =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
569  | 
let val thisLine = TextIO.inputLine fromChild  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
570  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
571  | 
if thisLine = "" (*end of file?*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
572  | 
	then (trace ("\n extraction_failed.  End bracket: " ^ endS ^
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
573  | 
"\naccumulated text: " ^ currentString);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
574  | 
false)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
575  | 
else if String.isPrefix endS thisLine  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
576  | 
then let val proofextract = currentString ^ cut_before endS thisLine  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
577  | 
val lemma_list = if endS = end_V8 then vamp_lemma_list  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
578  | 
else if endS = end_SPASS then spass_lemma_list  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
579  | 
else e_lemma_list  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
580  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
581  | 
	       trace ("\nExtracted proof:\n" ^ proofextract); 
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
582  | 
	       if !full andalso String.isPrefix "cnf(" proofextract
 | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
583  | 
then tstp_extract proofextract probfile toParent ppid ctxt th sgno thm_names  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
584  | 
else lemma_list proofextract probfile toParent ppid thm_names;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
585  | 
true  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
586  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
587  | 
else transferInput (currentString^thisLine)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
588  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
589  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
590  | 
transferInput ""  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
591  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
592  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
593  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
594  | 
(*The signal handler in watcher.ML must be able to read the output of this.*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
595  | 
fun signal_parent (toParent, ppid, msg, probfile) =  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
596  | 
(TextIO.output (toParent, msg);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
597  | 
TextIO.output (toParent, probfile ^ "\n");  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
598  | 
TextIO.flushOut toParent;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
599  | 
  trace ("\nSignalled parent: " ^ msg ^ probfile);
 | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
600  | 
Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
601  | 
(*Give the parent time to respond before possibly sending another signal*)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
602  | 
OS.Process.sleep (Time.fromMilliseconds 600));  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
603  | 
|
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
604  | 
(*FIXME: once TSTP output is produced by all ATPs, these three functions can be combined.*)  | 
| 
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
605  | 
|
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
606  | 
(*Called from watcher. Returns true if the Vampire process has returned a verdict.*)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
607  | 
fun checkVampProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
608  | 
let val thisLine = TextIO.inputLine fromChild  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
609  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
610  | 
trace thisLine;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
611  | 
if thisLine = ""  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
612  | 
then (trace "\nNo proof output seen"; false)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
613  | 
else if String.isPrefix start_V8 thisLine  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
614  | 
then startTransfer end_V8 arg  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
615  | 
else if (String.isPrefix "Satisfiability detected" thisLine) orelse  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
616  | 
(String.isPrefix "Refutation not found" thisLine)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
617  | 
then (signal_parent (toParent, ppid, "Failure\n", probfile);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
618  | 
true)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
619  | 
else checkVampProofFound arg  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
620  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
621  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
622  | 
(*Called from watcher. Returns true if the E process has returned a verdict.*)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
623  | 
fun checkEProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
624  | 
let val thisLine = TextIO.inputLine fromChild  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
625  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
626  | 
trace thisLine;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
627  | 
if thisLine = "" then (trace "\nNo proof output seen"; false)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
628  | 
else if String.isPrefix start_E thisLine  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
629  | 
then startTransfer end_E arg  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
630  | 
else if String.isPrefix "# Problem is satisfiable" thisLine  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
631  | 
then (signal_parent (toParent, ppid, "Invalid\n", probfile);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
632  | 
true)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
633  | 
else if String.isPrefix "# Cannot determine problem status within resource limit" thisLine  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
634  | 
then (signal_parent (toParent, ppid, "Failure\n", probfile);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
635  | 
true)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
636  | 
else checkEProofFound arg  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
637  | 
end;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
638  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
639  | 
(*Called from watcher. Returns true if the SPASS process has returned a verdict.*)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
640  | 
fun checkSpassProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
641  | 
let val thisLine = TextIO.inputLine fromChild  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
642  | 
in  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
643  | 
trace thisLine;  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
644  | 
if thisLine = "" then (trace "\nNo proof output seen"; false)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
645  | 
else if String.isPrefix "Here is a proof" thisLine  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
646  | 
then startTransfer end_SPASS arg  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
647  | 
else if thisLine = "SPASS beiseite: Completion found.\n"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
648  | 
then (signal_parent (toParent, ppid, "Invalid\n", probfile);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
649  | 
true)  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
650  | 
else if thisLine = "SPASS beiseite: Ran out of time.\n" orelse  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
651  | 
thisLine = "SPASS beiseite: Maximal number of loops exceeded.\n"  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
652  | 
then (signal_parent (toParent, ppid, "Failure\n", probfile);  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
653  | 
true)  | 
| 
22012
 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 
paulson 
parents: 
21999 
diff
changeset
 | 
654  | 
else checkSpassProofFound arg  | 
| 
21978
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
655  | 
end  | 
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
656  | 
|
| 
 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 
paulson 
parents:  
diff
changeset
 | 
657  | 
end;  |