| author | krauss | 
| Thu, 22 Mar 2007 13:36:57 +0100 | |
| changeset 22498 | 62cdd4b3e96b | 
| parent 22491 | 535fbed859da | 
| child 22545 | bd72c625c930 | 
| 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: 
21978diff
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: 
21978diff
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: 
21999diff
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: 
21999diff
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: 
21999diff
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 | |
| 22130 | 30 | fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 31 | else (); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 32 | |
| 
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 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 36 | val recon_sorts = ref false; | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 37 | |
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 38 | val modulus = ref 1; (*keep every nth proof line*) | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 39 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 40 | (**** PARSING OF TSTP FORMAT ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 41 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 42 | (*Syntax trees, either termlist or formulae*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 43 | datatype stree = Int of int | Br of string * stree list; | 
| 
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 atom x = Br(x,[]); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 46 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 47 | fun scons (x,y) = Br("cons", [x,y]);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 48 | val listof = foldl scons (atom "nil"); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 49 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 50 | (*Strings enclosed in single quotes, e.g. filenames*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 51 | val quoted = $$"'" |-- Scan.repeat (~$$"'") --| $$"'" >> implode; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 52 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 53 | (*Intended for $true and $false*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 54 | 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 | 55 | val truefalse = $$"$" |-- Symbol.scan_id >> (atom o tf); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 56 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 57 | (*Integer constants, typically proof line numbers*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 58 | fun is_digit s = Char.isDigit (String.sub(s,0)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 59 | 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 | 60 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 61 | (*Generalized FO terms, which include filenames, numbers, etc.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 62 | fun termlist x = (term -- Scan.repeat ($$"," |-- term) >> op::) x | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 63 | and term x = (quoted >> atom || integer>>Int || truefalse || | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 64 |               Symbol.scan_id -- Scan.optional ($$"(" |-- termlist --| $$")") [] >> Br ||
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 65 |               $$"(" |-- term --| $$")" ||
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 66 | $$"[" |-- termlist --| $$"]" >> listof) x; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 67 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 68 | fun negate t = Br("c_Not", [t]);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 69 | fun equate (t1,t2) = Br("c_equal", [t1,t2]);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 70 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 71 | (*Apply equal or not-equal to a term*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 72 | fun syn_equal (t, NONE) = t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 73 | | syn_equal (t1, SOME (NONE, t2)) = equate (t1,t2) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 74 | | syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 75 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 76 | (*Literals can involve negation, = and !=.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 77 | val literal = $$"~" |-- term >> negate || | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 78 | (term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal) ; | 
| 
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 | val literals = literal -- Scan.repeat ($$"|" |-- literal) >> op:: ; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 81 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 82 | (*Clause: a list of literals separated by the disjunction sign*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 83 | val clause = $$"(" |-- literals --| $$")";
 | 
| 
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 | val annotations = $$"," |-- term -- Scan.option ($$"," |-- termlist); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 86 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 87 | (*<cnf_annotated> ::=Êcnf(<name>,<formula_role>,<cnf_formula><annotations>). | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 88 | The <name> could be an identifier, but we assume integers.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 89 | val tstp_line = (Scan.this_string "cnf" -- $$"(") |-- 
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 90 | integer --| $$"," -- Symbol.scan_id --| $$"," -- | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 91 | clause -- Scan.option annotations --| $$ ")"; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 92 | |
| 
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 | (**** 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 | 95 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 96 | (*original file: Isabelle_ext.sml*) | 
| 
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 | val A_min_spc = Char.ord #"A" - Char.ord #" "; | 
| 
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 | 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 | 101 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 102 | (*why such a tiny range?*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 103 | fun check_valid_int x = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 104 | let val val_x = cList2int x | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 105 | 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 | 106 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 107 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 108 | fun normalise_s s [] st_ sti = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 109 | String.implode(rev( | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 110 | if st_ | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 111 | then if null sti | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 112 | then (#"_" :: s) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 113 | else if check_valid_int sti | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 114 | then (Char.chr (cList2int sti) :: s) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 115 | else (sti @ (#"_" :: s)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 116 | else s)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 117 | | normalise_s s (#"_"::cs) st_ sti = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 118 | if st_ | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 119 | then let val s' = if null sti | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 120 | then (#"_"::s) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 121 | else if check_valid_int sti | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 122 | then (Char.chr (cList2int sti) :: s) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 123 | else (sti @ (#"_" :: s)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 124 | in normalise_s s' cs false [] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 125 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 126 | else normalise_s s cs true [] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 127 | | normalise_s s (c::cs) true sti = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 128 | if (Char.isDigit c) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 129 | then normalise_s s cs true (c::sti) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 130 | else let val s' = if null sti | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 131 | then if ((c >= #"A") andalso (c<= #"P")) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 132 | then ((Char.chr(Char.ord c - A_min_spc))::s) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 133 | else (c :: (#"_" :: s)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 134 | else if check_valid_int sti | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 135 | then (Char.chr (cList2int sti) :: s) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 136 | else (sti @ (#"_" :: s)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 137 | in normalise_s s' cs false [] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 138 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 139 | | 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 | 140 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 141 | (*This version does not look for standard prefixes first.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 142 | fun normalise_string s = normalise_s [] (String.explode s) false []; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 143 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 144 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 145 | (**** INTERPRETATION OF TSTP SYNTAX TREES ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 146 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 147 | exception STREE of stree; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 148 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 149 | (*If string s has the prefix s1, return the result of deleting it.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 150 | fun strip_prefix s1 s = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 151 | 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 | 152 | else NONE; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 153 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 154 | (*Invert the table of translations between Isabelle and ATPs*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 155 | val type_const_trans_table_inv = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 156 | 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 | 157 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 158 | fun invert_type_const c = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 159 | case Symtab.lookup type_const_trans_table_inv c of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 160 | SOME c' => c' | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 161 | | NONE => c; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 162 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 163 | fun make_tvar b = TVar(("'" ^ b, 0), HOLogic.typeS);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 164 | fun make_var (b,T) = Var((b,0),T); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 165 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 166 | (*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 | 167 | by information from type literals, or by type inference.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 168 | fun type_of_stree t = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 169 | case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 170 | Int _ => raise STREE t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 171 | | Br (a,ts) => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 172 | let val Ts = map type_of_stree ts | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 173 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 174 | case strip_prefix ResClause.tconst_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 175 | SOME b => Type(invert_type_const b, Ts) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 176 | | NONE => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 177 | 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 | 178 | else | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 179 | case strip_prefix ResClause.tfree_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 180 |                     SOME b => TFree("'" ^ b, HOLogic.typeS)
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 181 | | NONE => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 182 | case strip_prefix ResClause.tvar_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 183 | SOME b => make_tvar b | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 184 | | NONE => make_tvar a (*Variable from the ATP, say X1*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 185 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 186 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 187 | (*Invert the table of translations between Isabelle and ATPs*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 188 | val const_trans_table_inv = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 189 | Symtab.make (map swap (Symtab.dest ResClause.const_trans_table)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 190 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 191 | fun invert_const c = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 192 | case Symtab.lookup const_trans_table_inv c of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 193 | SOME c' => c' | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 194 | | NONE => c; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 195 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 196 | (*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 | 197 | 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 | 198 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 199 | (*Generates a constant, given its type arguments*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 200 | 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 | 201 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 202 | (*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 | 203 | them to be inferred.*) | 
| 22428 | 204 | fun term_of_stree args thy t = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 205 | case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 206 | Int _ => raise STREE t | 
| 22428 | 207 |     | Br ("hBOOL",[t]) => term_of_stree [] thy t  (*ignore hBOOL*)
 | 
| 208 |     | Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t
 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 209 | | Br (a,ts) => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 210 | case strip_prefix ResClause.const_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 211 | SOME "equal" => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 212 | if length ts = 2 then | 
| 22428 | 213 |                 list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree [] thy) ts)
 | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 214 | else raise STREE t (*equality needs two arguments*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 215 | | SOME b => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 216 | let val c = invert_const b | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 217 | val nterms = length ts - num_typargs thy c | 
| 22428 | 218 | val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args) | 
| 219 | (*Extra args from hAPP come AFTER any arguments given directly to the | |
| 220 | constant.*) | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 221 | 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 | 222 | in list_comb(const_of thy (c, Ts), us) end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 223 | | NONE => (*a variable, not a constant*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 224 | let val T = HOLogic.typeT | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 225 | val opr = (*a Free variable is typically a Skolem function*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 226 | case strip_prefix ResClause.fixed_var_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 227 | SOME b => Free(b,T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 228 | | NONE => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 229 | case strip_prefix ResClause.schematic_var_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 230 | SOME b => make_var (b,T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 231 | | NONE => make_var (a,T) (*Variable from the ATP, say X1*) | 
| 22428 | 232 | in list_comb (opr, List.map (term_of_stree [] thy) (args@ts)) end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 233 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 234 | (*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 | 235 | 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 | 236 | | constraint_of_stree pol t = case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 237 | Int _ => raise STREE t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 238 | | Br (a,ts) => | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 239 | (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 | 240 | (SOME b, [T]) => (pol, b, T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 241 | | _ => raise STREE t); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 242 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 243 | (** Accumulate type constraints in a clause: negative type literals **) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 244 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 245 | fun addix (key,z) = Vartab.map_default (key,[]) (cons z); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 246 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 247 | fun 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 | 248 | | 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 | 249 | | add_constraint (_, vt) = vt; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 250 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 251 | (*False literals (which E includes in its proofs) are deleted*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 252 | val nofalses = filter (not o equal HOLogic.false_const); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 253 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 254 | (*Final treatment of the list of "real" literals from a clause.*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 255 | fun finish [] = HOLogic.true_const (*No "real" literals means only type information*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 256 | | finish lits = | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 257 | case nofalses lits of | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 258 | [] => HOLogic.false_const (*The empty clause, since we started with real literals*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 259 | | xs => foldr1 HOLogic.mk_disj (rev xs); | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 260 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 261 | (*Accumulate sort constraints in vt, with "real" literals in lits.*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 262 | fun lits_of_strees ctxt (vt, lits) [] = (vt, finish lits) | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 263 | | 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: 
21999diff
changeset | 264 | 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 | 265 | handle STREE _ => | 
| 22428 | 266 | 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 | 267 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 268 | (*Update TVars/TFrees with detected sort constraints.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 269 | fun fix_sorts vt = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 270 | 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 | 271 | | 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 | 272 | | 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 | 273 | fun tmsubst (Const (a, T)) = Const (a, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 274 | | tmsubst (Free (a, T)) = Free (a, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 275 | | tmsubst (Var (xi, T)) = Var (xi, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 276 | | tmsubst (t as Bound _) = t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 277 | | 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 | 278 | | tmsubst (t $ u) = tmsubst t $ tmsubst u; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 279 | 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 | 280 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 281 | (*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 | 282 | 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: 
21999diff
changeset | 283 | fun clause_of_strees_aux ctxt vt0 ts = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 284 | let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 285 | val infer = Sign.infer_types (ProofContext.pp ctxt) (ProofContext.theory_of ctxt) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 286 | (ProofContext.consts_of ctxt) (Variable.def_type ctxt false) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 287 | (Variable.def_sort ctxt) (Variable.names_of ctxt) true | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 288 | in | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 289 | #1(infer ([fix_sorts vt dt], HOLogic.boolT)) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 290 | end; | 
| 21978 
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 | (*Quantification over a list of Vars. FUXNE: for term.ML??*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 293 | fun list_all_var ([], t: term) = t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 294 | | list_all_var ((v as Var(ix,T)) :: vars, t) = | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 295 | (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 | 296 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 297 | 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 | 298 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 299 | fun clause_of_strees thy vt0 ts = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 300 | 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 | 301 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 302 | fun ints_of_stree_aux (Int n, ns) = n::ns | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 303 | | 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 | 304 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 305 | fun ints_of_stree t = ints_of_stree_aux (t, []); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 306 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 307 | fun decode_tstp ctxt vt0 (name, role, ts, annots) = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 308 | 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: 
21999diff
changeset | 309 | 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 | 310 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 311 | fun dest_tstp ((((name, role), ts), annots), chs) = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 312 | case chs of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 313 | "."::_ => (name, role, ts, annots) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 314 |         | _ => error ("TSTP line not terminated by \".\": " ^ implode chs);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 315 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 316 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 317 | (** 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 | 318 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 319 | 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 | 320 | | add_tfree_constraint (_, vt) = vt; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 321 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 322 | fun tfree_constraints_of_clauses vt [] = vt | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 323 | | tfree_constraints_of_clauses vt ([lit]::tss) = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 324 | (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 | 325 | handle STREE _ => (*not a positive type constraint: ignore*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 326 | tfree_constraints_of_clauses vt tss) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 327 | | 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 | 328 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 329 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 330 | (**** Translation of TSTP files to Isar Proofs ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 331 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 332 | fun decode_tstp_list ctxt tuples = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 333 | 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: 
21999diff
changeset | 334 | in map (decode_tstp ctxt vt0) tuples end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 335 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 336 | (*FIXME: simmilar function in res_atp. Move to HOLogic?*) | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 337 | 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: 
21979diff
changeset | 338 | | dest_disj_aux t disjs = t::disjs; | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 339 | |
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 340 | fun dest_disj t = dest_disj_aux t []; | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 341 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 342 | (*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: 
21999diff
changeset | 343 | 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: 
21999diff
changeset | 344 | | 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: 
21999diff
changeset | 345 | | 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: 
21999diff
changeset | 346 | | 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: 
21999diff
changeset | 347 | | smash_types t = t; | 
| 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 348 | |
| 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 349 | 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: 
21999diff
changeset | 350 | 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: 
21979diff
changeset | 351 | |
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 352 | fun permuted_clause t = | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 353 | let val lits = sort_lits t | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 354 | fun perm [] = NONE | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 355 | | perm (ctm::ctms) = | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 356 | 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: 
21979diff
changeset | 357 | else perm ctms | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 358 | in perm end; | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 359 | |
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 360 | fun have_or_show "show " lname = "show \"" | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 361 | | have_or_show have lname = have ^ lname ^ ": \"" | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 362 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 363 | (*ctms is a list of conjecture clauses as yielded by Isabelle. Those returned by the | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 364 | ATP may have their literals reordered.*) | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 365 | fun isar_lines ctxt ctms = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 366 | let val string_of = ProofContext.string_of_term ctxt | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 367 | fun doline have (lname, t, []) = (*No deps: it's a conjecture clause, with no proof.*) | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 368 | (case permuted_clause t ctms of | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 369 | SOME u => "assume " ^ lname ^ ": \"" ^ string_of u ^ "\"\n" | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 370 | | NONE => "assume? " ^ lname ^ ": \"" ^ string_of t ^ "\"\n") (*no match!!*) | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 371 | | doline have (lname, t, deps) = | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 372 | have_or_show have lname ^ string_of t ^ | 
| 22372 | 373 | "\"\n by (metis " ^ space_implode " " deps ^ ")\n" | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 374 | fun dolines [(lname, t, deps)] = [doline "show " (lname, t, deps)] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 375 | | dolines ((lname, t, deps)::lines) = doline "have " (lname, t, deps) :: dolines lines | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 376 | in setmp show_sorts (!recon_sorts) dolines end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 377 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 378 | fun notequal t (_,t',_) = not (t aconv t'); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 379 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 380 | (*No "real" literals means only type information*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 381 | fun eq_types t = t aconv (HOLogic.mk_Trueprop HOLogic.true_const); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 382 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 383 | 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 | 384 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 385 | fun replace_deps (old, new) (lno, t, deps) = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 386 | (lno, t, List.concat (map (replace_dep (old, new)) deps)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 387 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 388 | (*Discard axioms; consolidate adjacent lines that prove the same clause, since they differ | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 389 | only in type information.*) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 390 | fun add_prfline ((lno, "axiom", t, []), lines) = (*axioms are not proof lines*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 391 | if eq_types t (*must be clsrel/clsarity: type information, so delete refs to it*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 392 | then map (replace_deps (lno, [])) lines | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 393 | else | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 394 | (case take_prefix (notequal t) lines of | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 395 | (_,[]) => lines (*no repetition of proof line*) | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 396 | | (pre, (lno',t',deps')::post) => (*repetition: replace later line by earlier one*) | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 397 | pre @ map (replace_deps (lno', [lno])) post) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 398 | | add_prfline ((lno, role, t, []), lines) = (*no deps: conjecture clause*) | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 399 | (lno, t, []) :: lines | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 400 | | add_prfline ((lno, role, t, deps), lines) = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 401 | if eq_types t then (lno, t, deps) :: lines | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 402 | (*Type information will be deleted later; skip repetition test.*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 403 | else (*FIXME: Doesn't this code risk conflating proofs involving different types??*) | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 404 | case take_prefix (notequal t) lines of | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 405 | (_,[]) => (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: 
22012diff
changeset | 406 | | (pre, (lno',t',deps')::post) => | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 407 | (lno, t', deps) :: (*repetition: replace later line by earlier one*) | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 408 | (pre @ map (replace_deps (lno', [lno])) post); | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 409 | |
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 410 | (*Recursively delete empty lines (type information) from the proof.*) | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 411 | fun add_nonnull_prfline ((lno, t, []), lines) = (*no dependencies, so a conjecture clause*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 412 | if eq_types t (*must be type information, tfree_tcs, clsrel, clsarity: delete refs to it*) | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 413 | then delete_dep lno lines | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 414 | else (lno, t, []) :: lines | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 415 | | add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 416 | and delete_dep lno lines = foldr add_nonnull_prfline [] (map (replace_deps (lno, [])) lines); | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 417 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 418 | (*TVars are forbidden in goals. Also, we don't want lines with <2 dependencies. | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 419 | To further compress proofs, setting modulus:=n deletes every nth line, and nlines | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 420 | counts the number of proof lines processed so far. | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 421 | Deleted lines are replaced by their own dependencies. Note that the "add_nonnull_prfline" | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 422 | phase may delete some dependencies, hence this phase comes later.*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 423 | fun add_wanted_prfline ((lno, t, []), (nlines, lines)) = | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 424 | (nlines, (lno, t, []) :: lines) (*conjecture clauses must be kept*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 425 | | add_wanted_prfline (line, (nlines, [])) = (nlines, [line]) (*final line must be kept*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 426 | | add_wanted_prfline ((lno, t, deps), (nlines, lines)) = | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 427 | if not (null (term_tvars t)) orelse length deps < 2 orelse nlines mod !modulus <> 0 | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 428 | then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 429 | else (nlines+1, (lno, t, deps) :: lines); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 430 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 431 | (*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 | 432 | fun stringify_deps thm_names deps_map [] = [] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 433 | | stringify_deps thm_names deps_map ((lno, t, deps) :: lines) = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 434 | if lno <= Vector.length thm_names (*axiom*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 435 | 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: 
21978diff
changeset | 436 | else let val lname = Int.toString (length deps_map) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 437 | fun fix lno = if lno <= Vector.length thm_names | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 438 | then SOME(Vector.sub(thm_names,lno-1)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 439 | else AList.lookup op= deps_map lno; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 440 | in (lname, t, List.mapPartial fix deps) :: | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 441 | stringify_deps thm_names ((lno,lname)::deps_map) lines | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 442 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 443 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 444 | val proofstart = "\nproof (neg_clausify)\n"; | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 445 | |
| 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 446 | fun isar_header [] = proofstart | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 447 | | isar_header ts = proofstart ^ "fix " ^ space_implode " " ts ^ "\n"; | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 448 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 449 | 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 | 450 | let val tuples = map (dest_tstp o tstp_line o explode) cnfs | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 451 | val nonnull_lines = | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 452 | foldr add_nonnull_prfline [] | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 453 | (foldr add_prfline [] (decode_tstp_list ctxt tuples)) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 454 | val (_,lines) = foldr add_wanted_prfline (0,[]) nonnull_lines | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 455 | val (ccls,fixes) = ResAxioms.neg_conjecture_clauses th sgno | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 456 | val ccls = map forall_intr_vars ccls | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 457 | in | 
| 22130 | 458 | app (fn th => Output.debug (fn () => string_of_thm th)) ccls; | 
| 459 | isar_header (map #1 fixes) ^ | |
| 460 | 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: 
21978diff
changeset | 461 | end; | 
| 21978 
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 | (*Could use split_lines, but it can return blank lines...*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 464 | val lines = String.tokens (equal #"\n"); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 465 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 466 | 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 | 467 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 468 | (*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 | 469 | 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 | 470 | 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 | 471 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 472 | fun signal_success probfile toParent ppid msg = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 473 |   (trace ("\nReporting Success for" ^ probfile ^ "\n" ^ msg);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 474 | TextIO.output (toParent, "Success. " ^ encode_newlines msg ^ "\n"); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 475 | TextIO.output (toParent, probfile ^ "\n"); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 476 | TextIO.flushOut toParent; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 477 | 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 | 478 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 479 | 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 | 480 |   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 | 481 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 482 | signal_success probfile toParent ppid | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 483 | (decode_tstp_file cnfs ctxt th sgno thm_names) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 484 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 485 | |
| 
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 | (**** retrieve the axioms that were used in the proof ****) | 
| 
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 | (*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 | 490 | fun get_axiom_names (thm_names: string vector) step_nums = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 491 | let fun is_axiom n = n <= Vector.length thm_names | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 492 | fun index i = Vector.sub(thm_names, i-1) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 493 | val axnums = List.filter is_axiom step_nums | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 494 | val axnames = sort_distinct string_ord (map index axnums) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 495 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 496 | if length axnums = length step_nums then "UNSOUND!!" :: axnames | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 497 | else axnames | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 498 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 499 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 500 | (*String contains multiple lines. We want those of the form | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 501 | "253[0:Inp] et cetera..." | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 502 | 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 | 503 | fun get_spass_linenums proofstr = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 504 | let val toks = String.tokens (not o Char.isAlphaNum) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 505 | fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 506 | | inputno _ = NONE | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 507 | val lines = String.tokens (fn c => c = #"\n") proofstr | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 508 | in List.mapPartial (inputno o toks) lines end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 509 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 510 | fun get_axiom_names_spass proofstr thm_names = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 511 | get_axiom_names thm_names (get_spass_linenums proofstr); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 512 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 513 | fun not_comma c = c <> #","; | 
| 
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 | (*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 | 516 | fun parse_tstp_line s = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 517 |   let val ss = Substring.full (unprefix "cnf(" (nospaces s))
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 518 | val (intf,rest) = Substring.splitl not_comma ss | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 519 | 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 | 520 | (*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 | 521 | get_axiom_names above, while suppressing the UNSOUND warning*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 522 | 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 | 523 | then Substring.string intf | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 524 | else "error" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 525 | in Int.fromString ints end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 526 | handle Fail _ => NONE; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 527 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 528 | fun get_axiom_names_tstp proofstr thm_names = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 529 | 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 | 530 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 531 | (*String contains multiple lines. We want those of the form | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 532 | "*********** [448, input] ***********". | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 533 | 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 | 534 | fun get_vamp_linenums proofstr = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 535 | let val toks = String.tokens (not o Char.isAlphaNum) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 536 | fun inputno [ntok,"input"] = Int.fromString ntok | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 537 | | inputno _ = NONE | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 538 | val lines = String.tokens (fn c => c = #"\n") proofstr | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 539 | in List.mapPartial (inputno o toks) lines end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 540 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 541 | fun get_axiom_names_vamp proofstr thm_names = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 542 | get_axiom_names thm_names (get_vamp_linenums proofstr); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 543 | |
| 22372 | 544 | fun rules_to_metis [] = "metis" | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 545 | | rules_to_metis xs = "(metis " ^ space_implode " " xs ^ ")" | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 546 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 547 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 548 | (*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 | 549 | 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 | 550 |  (trace ("\n\nGetting lemma names. proofstr is " ^ proofstr ^
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 551 | " 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 | 552 | signal_success probfile toParent ppid | 
| 22372 | 553 |     ("Try this proof method: \n" ^ rules_to_metis (getax proofstr thm_names))
 | 
| 554 | ) | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 555 | handle e => (*FIXME: exn handler is too general!*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 556 |   (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 | 557 | TextIO.output (toParent, "Translation failed for the proof: " ^ | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 558 | String.toString proofstr ^ "\n"); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 559 | TextIO.output (toParent, probfile); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 560 | TextIO.flushOut toParent; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 561 | 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 | 562 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 563 | 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 | 564 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 565 | 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 | 566 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 567 | 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 | 568 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 569 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 570 | (**** Extracting proofs from an ATP's output ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 571 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 572 | (*Return everything in s that comes before the string t*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 573 | fun cut_before t s = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 574 | let val (s1,s2) = Substring.position t (Substring.full s) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 575 | 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 | 576 | else Substring.string s2 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 577 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 578 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 579 | val start_E = "# Proof object starts here." | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 580 | val end_E = "# Proof object ends here." | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 581 | val start_V6 = "%================== Proof: ======================" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 582 | val end_V6 = "%============== End of proof. ==================" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 583 | val start_V8 = "=========== Refutation ==========" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 584 | val end_V8 = "======= End of refutation =======" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 585 | val end_SPASS = "Formulae used in the proof" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 586 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 587 | (*********************************************************************************) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 588 | (* 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 | 589 | (* 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 | 590 | (*********************************************************************************) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 591 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 592 | (*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 | 593 | 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: 
21999diff
changeset | 594 | 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 | 595 | let fun transferInput currentString = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 596 | let val thisLine = TextIO.inputLine fromChild | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 597 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 598 | if thisLine = "" (*end of file?*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 599 | 	then (trace ("\n extraction_failed.  End bracket: " ^ endS ^
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 600 | "\naccumulated text: " ^ currentString); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 601 | false) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 602 | else if String.isPrefix endS thisLine | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 603 | then let val proofextract = currentString ^ cut_before endS thisLine | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 604 | 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 | 605 | else if endS = end_SPASS then spass_lemma_list | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 606 | else e_lemma_list | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 607 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 608 | 	       trace ("\nExtracted proof:\n" ^ proofextract); 
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 609 | 	       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: 
21999diff
changeset | 610 | 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 | 611 | else lemma_list proofextract probfile toParent ppid thm_names; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 612 | true | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 613 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 614 | else transferInput (currentString^thisLine) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 615 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 616 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 617 | transferInput "" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 618 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 619 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 620 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 621 | (*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 | 622 | fun signal_parent (toParent, ppid, msg, probfile) = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 623 | (TextIO.output (toParent, msg); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 624 | TextIO.output (toParent, probfile ^ "\n"); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 625 | TextIO.flushOut toParent; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 626 |   trace ("\nSignalled parent: " ^ msg ^ probfile);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 627 | 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 | 628 | (*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 | 629 | OS.Process.sleep (Time.fromMilliseconds 600)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 630 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 631 | (*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: 
21999diff
changeset | 632 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 633 | (*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: 
21999diff
changeset | 634 | 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 | 635 | let val thisLine = TextIO.inputLine fromChild | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 636 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 637 | trace thisLine; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 638 | if thisLine = "" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 639 | then (trace "\nNo proof output seen"; false) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 640 | 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: 
21999diff
changeset | 641 | then startTransfer end_V8 arg | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 642 | else if (String.isPrefix "Satisfiability detected" thisLine) orelse | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 643 | (String.isPrefix "Refutation not found" thisLine) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 644 | then (signal_parent (toParent, ppid, "Failure\n", probfile); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 645 | true) | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 646 | else checkVampProofFound arg | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 647 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 648 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 649 | (*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: 
21999diff
changeset | 650 | 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 | 651 | let val thisLine = TextIO.inputLine fromChild | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 652 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 653 | trace thisLine; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 654 | if thisLine = "" then (trace "\nNo proof output seen"; false) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 655 | 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: 
21999diff
changeset | 656 | then startTransfer end_E arg | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 657 | else if String.isPrefix "# Problem is satisfiable" thisLine | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 658 | then (signal_parent (toParent, ppid, "Invalid\n", probfile); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 659 | true) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 660 | 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 | 661 | then (signal_parent (toParent, ppid, "Failure\n", probfile); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 662 | true) | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 663 | else checkEProofFound arg | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 664 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 665 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 666 | (*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: 
21999diff
changeset | 667 | 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 | 668 | let val thisLine = TextIO.inputLine fromChild | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 669 | in | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 670 | trace thisLine; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 671 | if thisLine = "" then (trace "\nNo proof output seen"; false) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 672 | 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: 
21999diff
changeset | 673 | then startTransfer end_SPASS arg | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 674 | else if thisLine = "SPASS beiseite: Completion found.\n" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 675 | then (signal_parent (toParent, ppid, "Invalid\n", probfile); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 676 | true) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 677 | 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 | 678 | thisLine = "SPASS beiseite: Maximal number of loops exceeded.\n" | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 679 | then (signal_parent (toParent, ppid, "Failure\n", probfile); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 680 | true) | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 681 | else checkSpassProofFound arg | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 682 | end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 683 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 684 | end; |