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