src/HOL/Tools/res_reconstruct.ML
author paulson
Wed, 28 Nov 2007 16:26:03 +0100
changeset 25492 4cc7976948ac
parent 25414 3326bd7ecd48
child 25710 4cdf7de81e1b
permissions -rwxr-xr-x
Chained theorems are no longer mentioned in metis calls and (if used) they prevent the output of a structured proof.
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 =
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    10
sig
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    11
  datatype atp = E | SPASS | Vampire
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    12
  val modulus:     int ref
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    13
  val recon_sorts: bool ref
25492
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
    14
  val chained_hint: string
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    15
  val checkEProofFound:
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    16
	TextIO.instream * TextIO.outstream * Posix.Process.pid *
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    17
	string * Proof.context * thm * int * string Vector.vector -> bool
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    18
  val checkVampProofFound:
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    19
	TextIO.instream * TextIO.outstream * Posix.Process.pid *
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    20
	string * Proof.context * thm * int * string Vector.vector -> bool
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    21
  val checkSpassProofFound:
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    22
	TextIO.instream * TextIO.outstream * Posix.Process.pid *
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    23
	string * Proof.context * thm * int * string Vector.vector -> bool
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    24
  val signal_parent: TextIO.outstream * Posix.Process.pid * string * string -> unit
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    25
  val txt_path: string -> Path.T
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    26
  val fix_sorts: sort Vartab.table -> term -> term
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    27
  val invert_const: string -> string
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    28
  val invert_type_const: string -> string
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    29
  val num_typargs: Context.theory -> string -> int
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    30
  val make_tvar: string -> typ
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    31
  val strip_prefix: string -> string -> string option
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    32
end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    33
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    34
structure ResReconstruct : RES_RECONSTRUCT =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    35
struct
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
val trace_path = Path.basic "atp_trace";
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    38
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
    39
fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    40
              else ();
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    41
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
    42
datatype atp = E | SPASS | Vampire;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    43
23833
3fe991a1f805 Full sort information by default.
paulson
parents: 23519
diff changeset
    44
val recon_sorts = ref true;
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
    45
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
    46
val modulus = ref 1;    (*keep every nth proof line*)
22044
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
    47
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    48
(**** PARSING OF TSTP FORMAT ****)
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
(*Syntax trees, either termlist or formulae*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    51
datatype stree = Int of int | Br of string * stree list;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    52
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    53
fun atom x = Br(x,[]);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    54
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    55
fun scons (x,y) = Br("cons", [x,y]);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    56
val listof = foldl scons (atom "nil");
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
(*Strings enclosed in single quotes, e.g. filenames*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    59
val quoted = $$"'" |-- Scan.repeat (~$$"'") --| $$"'" >> implode;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    60
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    61
(*Intended for $true and $false*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    62
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
    63
val truefalse = $$"$" |-- Symbol.scan_id >> (atom o tf);
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
(*Integer constants, typically proof line numbers*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    66
fun is_digit s = Char.isDigit (String.sub(s,0));
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    67
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
    68
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    69
(*Generalized FO terms, which include filenames, numbers, etc.*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    70
fun termlist x = (term -- Scan.repeat ($$"," |-- term) >> op::) x
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    71
and term x = (quoted >> atom || integer>>Int || truefalse ||
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    72
              Symbol.scan_id -- Scan.optional ($$"(" |-- termlist --| $$")") [] >> Br ||
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    73
              $$"(" |-- term --| $$")" ||
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
    74
              $$"[" |-- Scan.optional termlist [] --| $$"]" >> listof) x;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    75
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    76
fun negate t = Br("c_Not", [t]);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    77
fun equate (t1,t2) = Br("c_equal", [t1,t2]);
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
(*Apply equal or not-equal to a term*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    80
fun syn_equal (t, NONE) = t
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    81
  | syn_equal (t1, SOME (NONE, t2)) = equate (t1,t2)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    82
  | syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2));
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
(*Literals can involve negation, = and !=.*)
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
    85
fun literal x = ($$"~" |-- literal >> negate ||
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
    86
                 (term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal)) x;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    87
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    88
val literals = literal -- Scan.repeat ($$"|" |-- literal) >> op:: ;
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
(*Clause: a list of literals separated by the disjunction sign*)
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
    91
val clause = $$"(" |-- literals --| $$")" || Scan.single literal;
21978
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
val annotations = $$"," |-- term -- Scan.option ($$"," |-- termlist);
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
(*<cnf_annotated> ::=Êcnf(<name>,<formula_role>,<cnf_formula><annotations>).
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    96
  The <name> could be an identifier, but we assume integers.*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
    97
val tstp_line = (Scan.this_string "cnf" -- $$"(") |--
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
    98
                integer --| $$"," -- Symbol.scan_id --| $$"," --
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
    99
                clause -- Scan.option annotations --| $$ ")";
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   100
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   101
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   102
(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   103
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   104
exception STREE of stree;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   105
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   106
(*If string s has the prefix s1, return the result of deleting it.*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   107
fun strip_prefix s1 s =
24182
a39c5e7de6a7 Fixing the code to undo the function ascii_of
paulson
parents: 23833
diff changeset
   108
  if String.isPrefix s1 s 
a39c5e7de6a7 Fixing the code to undo the function ascii_of
paulson
parents: 23833
diff changeset
   109
  then SOME (ResClause.undo_ascii_of (String.extract (s, size s1, NONE)))
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   110
  else NONE;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   111
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   112
(*Invert the table of translations between Isabelle and ATPs*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   113
val type_const_trans_table_inv =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   114
      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
   115
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   116
fun invert_type_const c =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   117
    case Symtab.lookup type_const_trans_table_inv c of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   118
        SOME c' => c'
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   119
      | NONE => c;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   120
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   121
fun make_tvar b = TVar(("'" ^ b, 0), HOLogic.typeS);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   122
fun make_var (b,T) = Var((b,0),T);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   123
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   124
(*Type variables are given the basic sort, HOL.type. Some will later be constrained
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   125
  by information from type literals, or by type inference.*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   126
fun type_of_stree t =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   127
  case t of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   128
      Int _ => raise STREE t
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   129
    | Br (a,ts) =>
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   130
        let val Ts = map type_of_stree ts
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   131
        in
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   132
          case strip_prefix ResClause.tconst_prefix a of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   133
              SOME b => Type(invert_type_const b, Ts)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   134
            | NONE =>
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   135
                if not (null ts) then raise STREE t  (*only tconsts have type arguments*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   136
                else
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   137
                case strip_prefix ResClause.tfree_prefix a of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   138
                    SOME b => TFree("'" ^ b, HOLogic.typeS)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   139
                  | NONE =>
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   140
                case strip_prefix ResClause.tvar_prefix a of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   141
                    SOME b => make_tvar b
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   142
                  | NONE => make_tvar a   (*Variable from the ATP, say X1*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   143
        end;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   144
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   145
(*Invert the table of translations between Isabelle and ATPs*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   146
val const_trans_table_inv =
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   147
      Symtab.update ("fequal", "op =")
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22728
diff changeset
   148
        (Symtab.make (map swap (Symtab.dest ResClause.const_trans_table)));
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   149
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   150
fun invert_const c =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   151
    case Symtab.lookup const_trans_table_inv c of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   152
        SOME c' => c'
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   153
      | NONE => c;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   154
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   155
(*The number of type arguments of a constant, zero if it's monomorphic*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   156
fun num_typargs thy s = length (Sign.const_typargs thy (s, Sign.the_const_type thy s));
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   157
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   158
(*Generates a constant, given its type arguments*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   159
fun const_of thy (a,Ts) = Const(a, Sign.const_instance thy (a,Ts));
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   160
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   161
(*First-order translation. No types are known for variables. HOLogic.typeT should allow
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   162
  them to be inferred.*)
22428
1755e6381b2c First stab at reconstructing HO problems
paulson
parents: 22372
diff changeset
   163
fun term_of_stree args thy t =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   164
  case t of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   165
      Int _ => raise STREE t
22428
1755e6381b2c First stab at reconstructing HO problems
paulson
parents: 22372
diff changeset
   166
    | Br ("hBOOL",[t]) => term_of_stree [] thy t  (*ignore hBOOL*)
1755e6381b2c First stab at reconstructing HO problems
paulson
parents: 22372
diff changeset
   167
    | Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   168
    | Br (a,ts) =>
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   169
        case strip_prefix ResClause.const_prefix a of
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   170
            SOME "equal" =>
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22728
diff changeset
   171
              list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree [] thy) ts)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   172
          | SOME b =>
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   173
              let val c = invert_const b
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   174
                  val nterms = length ts - num_typargs thy c
22428
1755e6381b2c First stab at reconstructing HO problems
paulson
parents: 22372
diff changeset
   175
                  val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args)
1755e6381b2c First stab at reconstructing HO problems
paulson
parents: 22372
diff changeset
   176
                  (*Extra args from hAPP come AFTER any arguments given directly to the
1755e6381b2c First stab at reconstructing HO problems
paulson
parents: 22372
diff changeset
   177
                    constant.*)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   178
                  val Ts = List.map type_of_stree (List.drop(ts,nterms))
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   179
              in  list_comb(const_of thy (c, Ts), us)  end
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   180
          | NONE => (*a variable, not a constant*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   181
              let val T = HOLogic.typeT
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   182
                  val opr = (*a Free variable is typically a Skolem function*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   183
                    case strip_prefix ResClause.fixed_var_prefix a of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   184
                        SOME b => Free(b,T)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   185
                      | NONE =>
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   186
                    case strip_prefix ResClause.schematic_var_prefix a of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   187
                        SOME b => make_var (b,T)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   188
                      | NONE => make_var (a,T)    (*Variable from the ATP, say X1*)
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   189
              in  list_comb (opr, List.map (term_of_stree [] thy) (ts@args))  end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   190
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   191
(*Type class literal applied to a type. Returns triple of polarity, class, type.*)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   192
fun constraint_of_stree pol (Br("c_Not",[t])) = constraint_of_stree (not pol) t
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   193
  | constraint_of_stree pol t = case t of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   194
        Int _ => raise STREE t
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   195
      | Br (a,ts) =>
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   196
            (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
   197
                 (SOME b, [T]) => (pol, b, T)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   198
               | _ => raise STREE t);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   199
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   200
(** Accumulate type constraints in a clause: negative type literals **)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   201
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   202
fun addix (key,z)  = Vartab.map_default (key,[]) (cons z);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   203
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   204
fun add_constraint ((false, cl, TFree(a,_)), vt) = addix ((a,~1),cl) vt
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   205
  | add_constraint ((false, cl, TVar(ix,_)), vt) = addix (ix,cl) vt
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   206
  | add_constraint (_, vt) = vt;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   207
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   208
(*False literals (which E includes in its proofs) are deleted*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   209
val nofalses = filter (not o equal HOLogic.false_const);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   210
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   211
(*Final treatment of the list of "real" literals from a clause.*)
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   212
fun finish [] = HOLogic.true_const  (*No "real" literals means only type information*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   213
  | finish lits =
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   214
      case nofalses lits of
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   215
          [] => HOLogic.false_const  (*The empty clause, since we started with real literals*)
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   216
        | xs => foldr1 HOLogic.mk_disj (rev xs);
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   217
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   218
(*Accumulate sort constraints in vt, with "real" literals in lits.*)
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   219
fun lits_of_strees ctxt (vt, lits) [] = (vt, finish lits)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   220
  | lits_of_strees ctxt (vt, lits) (t::ts) =
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   221
      lits_of_strees ctxt (add_constraint (constraint_of_stree true t, vt), lits) ts
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   222
      handle STREE _ =>
22428
1755e6381b2c First stab at reconstructing HO problems
paulson
parents: 22372
diff changeset
   223
      lits_of_strees ctxt (vt, term_of_stree [] (ProofContext.theory_of ctxt) t :: lits) ts;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   224
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   225
(*Update TVars/TFrees with detected sort constraints.*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   226
fun fix_sorts vt =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   227
  let fun tysubst (Type (a, Ts)) = Type (a, map tysubst Ts)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   228
        | 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
   229
        | 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
   230
      fun tmsubst (Const (a, T)) = Const (a, tysubst T)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   231
        | tmsubst (Free (a, T)) = Free (a, tysubst T)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   232
        | tmsubst (Var (xi, T)) = Var (xi, tysubst T)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   233
        | tmsubst (t as Bound _) = t
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   234
        | tmsubst (Abs (a, T, t)) = Abs (a, tysubst T, tmsubst t)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   235
        | tmsubst (t $ u) = tmsubst t $ tmsubst u;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   236
  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
   237
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   238
(*Interpret a list of syntax trees as a clause, given by "real" literals and sort constraints.
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   239
  vt0 holds the initial sort constraints, from the conjecture clauses.*)
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   240
fun clause_of_strees ctxt vt0 ts =
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22692
diff changeset
   241
  let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts in
24680
0d355aa59e67 TypeInfer.constrain: canonical argument order;
wenzelm
parents: 24632
diff changeset
   242
    singleton (Syntax.check_terms ctxt) (TypeInfer.constrain HOLogic.boolT (fix_sorts vt dt))
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22692
diff changeset
   243
  end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   244
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   245
(*Quantification over a list of Vars. FIXME: for term.ML??*)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   246
fun list_all_var ([], t: term) = t
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   247
  | 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
   248
      (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
   249
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   250
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
   251
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   252
fun ints_of_stree_aux (Int n, ns) = n::ns
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   253
  | 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
   254
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   255
fun ints_of_stree t = ints_of_stree_aux (t, []);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   256
25086
729f9aad1f50 Improving the propagation of type constraints for Frees
paulson
parents: 24958
diff changeset
   257
fun decode_tstp vt0 (name, role, ts, annots) ctxt =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   258
  let val deps = case annots of NONE => [] | SOME (source,_) => ints_of_stree source
25086
729f9aad1f50 Improving the propagation of type constraints for Frees
paulson
parents: 24958
diff changeset
   259
      val cl = clause_of_strees ctxt vt0 ts
729f9aad1f50 Improving the propagation of type constraints for Frees
paulson
parents: 24958
diff changeset
   260
  in  ((name, role, cl, deps), fold Variable.declare_term (term_frees cl) ctxt)  end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   261
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   262
fun dest_tstp ((((name, role), ts), annots), chs) =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   263
  case chs of
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   264
          "."::_ => (name, role, ts, annots)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   265
        | _ => error ("TSTP line not terminated by \".\": " ^ implode chs);
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
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   268
(** 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
   269
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   270
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
   271
  | add_tfree_constraint (_, vt) = vt;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   272
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   273
fun tfree_constraints_of_clauses vt [] = vt
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   274
  | tfree_constraints_of_clauses vt ([lit]::tss) =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   275
      (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
   276
       handle STREE _ => (*not a positive type constraint: ignore*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   277
       tfree_constraints_of_clauses vt tss)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   278
  | 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
   279
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   280
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   281
(**** Translation of TSTP files to Isar Proofs ****)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   282
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   283
fun decode_tstp_list ctxt tuples =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   284
  let val vt0 = tfree_constraints_of_clauses Vartab.empty (map #3 tuples)
25086
729f9aad1f50 Improving the propagation of type constraints for Frees
paulson
parents: 24958
diff changeset
   285
  in  #1 (fold_map (decode_tstp vt0) tuples ctxt) end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   286
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   287
(** Finding a matching assumption. The literals may be permuted, and variable names
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   288
    may disagree. We have to try all combinations of literals (quadratic!) and 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   289
    match up the variable names consistently. **)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   290
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   291
fun strip_alls_aux n (Const("all",_)$Abs(a,T,t))  =  
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   292
      strip_alls_aux (n+1) (subst_bound (Var ((a,n), T), t))
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   293
  | strip_alls_aux _ t  =  t;
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   294
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   295
val strip_alls = strip_alls_aux 0;
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   296
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   297
exception MATCH_LITERAL;
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   298
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   299
(*Ignore types: they are not to be trusted...*)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   300
fun match_literal (t1$u1) (t2$u2) env =
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   301
      match_literal t1 t2 (match_literal u1 u2 env)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   302
  | match_literal (Abs (_,_,t1)) (Abs (_,_,t2)) env = 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   303
      match_literal t1 t2 env
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   304
  | match_literal (Bound i1) (Bound i2) env = 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   305
      if i1=i2 then env else raise MATCH_LITERAL
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   306
  | match_literal (Const(a1,_)) (Const(a2,_)) env = 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   307
      if a1=a2 then env else raise MATCH_LITERAL
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   308
  | match_literal (Free(a1,_)) (Free(a2,_)) env = 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   309
      if a1=a2 then env else raise MATCH_LITERAL
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   310
  | match_literal (Var(ix1,_)) (Var(ix2,_)) env = insert (op =) (ix1,ix2) env
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   311
  | match_literal _ _ env = raise MATCH_LITERAL;
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   312
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   313
(*Checking that all variable associations are unique. The list env contains no
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   314
  repetitions, but does it contain say (x,y) and (y,y)? *)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   315
fun good env = 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   316
  let val (xs,ys) = ListPair.unzip env
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   317
  in  not (has_duplicates (op=) xs orelse has_duplicates (op=) ys)  end;
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   318
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   319
(*Match one list of literals against another, ignoring types and the order of
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   320
  literals. Sorting is unreliable because we don't have types or variable names.*)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   321
fun matches_aux _ [] [] = true
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   322
  | matches_aux env (lit::lits) ts =
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   323
      let fun match1 us [] = false
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   324
            | match1 us (t::ts) =
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   325
                let val env' = match_literal lit t env
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   326
                in  (good env' andalso matches_aux env' lits (us@ts)) orelse 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   327
                    match1 (t::us) ts  
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   328
                end
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   329
                handle MATCH_LITERAL => match1 (t::us) ts
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   330
      in  match1 [] ts  end; 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   331
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   332
(*Is this length test useful?*)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   333
fun matches (lits1,lits2) = 
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   334
  length lits1 = length lits2  andalso  
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   335
  matches_aux [] (map Envir.eta_contract lits1) (map Envir.eta_contract lits2);
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   336
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   337
fun permuted_clause t =
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24956
diff changeset
   338
  let val lits = HOLogic.disjuncts t
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   339
      fun perm [] = NONE
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   340
        | perm (ctm::ctms) =
24958
ff15f76741bd rationalized redundant code
paulson
parents: 24956
diff changeset
   341
            if matches (lits, HOLogic.disjuncts (HOLogic.dest_Trueprop (strip_alls ctm)))
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   342
            then SOME ctm else perm ctms
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   343
  in perm end;
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   344
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   345
fun have_or_show "show " lname = "show \""
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   346
  | have_or_show have lname = have ^ lname ^ ": \""
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   347
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   348
(*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
   349
  ATP may have their literals reordered.*)
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   350
fun isar_lines ctxt ctms =
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24680
diff changeset
   351
  let val string_of = Syntax.string_of_term ctxt
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   352
      fun doline have (lname, t, []) =  (*No deps: it's a conjecture clause, with no proof.*)
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   353
           (case permuted_clause t ctms of
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   354
                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
   355
              | NONE => "assume? " ^ lname ^ ": \"" ^ string_of t ^ "\"\n")  (*no match!!*)
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   356
        | doline have (lname, t, deps) =
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   357
            have_or_show have lname ^ string_of (gen_all_vars (HOLogic.mk_Trueprop t)) ^
22372
02fc0ceb094a Now outputs metis calls
paulson
parents: 22130
diff changeset
   358
            "\"\n  by (metis " ^ space_implode " " deps ^ ")\n"
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   359
      fun dolines [(lname, t, deps)] = [doline "show " (lname, t, deps)]
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   360
        | dolines ((lname, t, deps)::lines) = doline "have " (lname, t, deps) :: dolines lines
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   361
  in setmp show_sorts (!recon_sorts) dolines end;
21978
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
fun notequal t (_,t',_) = not (t aconv t');
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   364
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   365
(*No "real" literals means only type information*)
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   366
fun eq_types t = t aconv HOLogic.true_const;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   367
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22728
diff changeset
   368
fun replace_dep (old:int, new) dep = if dep=old then new else [dep];
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   369
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   370
fun replace_deps (old:int, new) (lno, t, deps) =
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22728
diff changeset
   371
      (lno, t, foldl (op union_int) [] (map (replace_dep (old, new)) deps));
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   372
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   373
(*Discard axioms; consolidate adjacent lines that prove the same clause, since they differ
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   374
  only in type information.*)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   375
fun add_prfline ((lno, "axiom", t, []), lines) =  (*axioms are not proof lines*)
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   376
      if eq_types t (*must be clsrel/clsarity: type information, so delete refs to it*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   377
      then map (replace_deps (lno, [])) lines
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   378
      else
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   379
       (case take_prefix (notequal t) lines of
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   380
           (_,[]) => lines                  (*no repetition of proof line*)
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   381
         | (pre, (lno',t',deps')::post) =>  (*repetition: replace later line by earlier one*)
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   382
             pre @ map (replace_deps (lno', [lno])) post)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   383
  | add_prfline ((lno, role, t, []), lines) =  (*no deps: conjecture clause*)
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   384
      (lno, t, []) :: lines
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   385
  | add_prfline ((lno, role, t, deps), lines) =
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   386
      if eq_types t then (lno, t, deps) :: lines
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   387
      (*Type information will be deleted later; skip repetition test.*)
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   388
      else (*FIXME: Doesn't this code risk conflating proofs involving different types??*)
22044
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
   389
      case take_prefix (notequal t) lines of
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
   390
         (_,[]) => (lno, t, deps) :: lines  (*no repetition of proof line*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   391
       | (pre, (lno',t',deps')::post) =>
22044
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
   392
           (lno, t', deps) ::               (*repetition: replace later line by earlier one*)
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
   393
           (pre @ map (replace_deps (lno', [lno])) post);
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
   394
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   395
(*Recursively delete empty lines (type information) from the proof.*)
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   396
fun add_nonnull_prfline ((lno, t, []), lines) = (*no dependencies, so a conjecture clause*)
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   397
     if eq_types t (*must be type information, tfree_tcs, clsrel, clsarity: delete refs to it*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   398
     then delete_dep lno lines
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   399
     else (lno, t, []) :: lines
22470
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   400
  | add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   401
and delete_dep lno lines = foldr add_nonnull_prfline [] (map (replace_deps (lno, [])) lines);
0d52e5157124 No label on "show"; tries to remove dependencies more cleanly
paulson
parents: 22428
diff changeset
   402
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24920
diff changeset
   403
fun bad_free (Free (a,_)) = String.isPrefix "sko_" a
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22728
diff changeset
   404
  | bad_free _ = false;
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22728
diff changeset
   405
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   406
(*TVars are forbidden in goals. Also, we don't want lines with <2 dependencies.
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   407
  To further compress proofs, setting modulus:=n deletes every nth line, and nlines
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   408
  counts the number of proof lines processed so far.
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   409
  Deleted lines are replaced by their own dependencies. Note that the "add_nonnull_prfline"
22044
6c0702a96076 More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
paulson
parents: 22012
diff changeset
   410
  phase may delete some dependencies, hence this phase comes later.*)
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   411
fun add_wanted_prfline ((lno, t, []), (nlines, lines)) =
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   412
      (nlines, (lno, t, []) :: lines)   (*conjecture clauses must be kept*)
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   413
  | add_wanted_prfline ((lno, t, deps), (nlines, lines)) =
22545
bd72c625c930 Now checks for types-only clause before outputting.
paulson
parents: 22491
diff changeset
   414
      if eq_types t orelse not (null (term_tvars t)) orelse
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24920
diff changeset
   415
         exists bad_free (term_frees t) orelse
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24920
diff changeset
   416
         (not (null lines) andalso   (*final line can't be deleted for these reasons*)
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24920
diff changeset
   417
          (length deps < 2 orelse nlines mod !modulus <> 0))   
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   418
      then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*)
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   419
      else (nlines+1, (lno, t, deps) :: lines);
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   420
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   421
(*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
   422
fun stringify_deps thm_names deps_map [] = []
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   423
  | stringify_deps thm_names deps_map ((lno, t, deps) :: lines) =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   424
      if lno <= Vector.length thm_names  (*axiom*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   425
      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
   426
      else let val lname = Int.toString (length deps_map)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   427
               fun fix lno = if lno <= Vector.length thm_names
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   428
                             then SOME(Vector.sub(thm_names,lno-1))
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   429
                             else AList.lookup op= deps_map lno;
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22728
diff changeset
   430
           in  (lname, t, List.mapPartial fix (distinct (op=) deps)) ::
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   431
               stringify_deps thm_names ((lno,lname)::deps_map) lines
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   432
           end;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   433
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   434
val proofstart = "proof (neg_clausify)\n";
21979
80e10f181c46 Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents: 21978
diff changeset
   435
80e10f181c46 Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents: 21978
diff changeset
   436
fun isar_header [] = proofstart
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   437
  | isar_header ts = proofstart ^ "fix " ^ space_implode " " ts ^ "\n";
21979
80e10f181c46 Improvements to proof reconstruction. Now "fixes" is inserted
paulson
parents: 21978
diff changeset
   438
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   439
fun decode_tstp_file cnfs ctxt th sgno thm_names =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   440
  let val tuples = map (dest_tstp o tstp_line o explode) cnfs
24552
bb7fdd10de9a allow TVars during type inference
paulson
parents: 24547
diff changeset
   441
      val ctxt = ProofContext.set_mode ProofContext.mode_schematic ctxt
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   442
      val nonnull_lines =
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   443
              foldr add_nonnull_prfline []
22491
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   444
                    (foldr add_prfline [] (decode_tstp_list ctxt tuples))
535fbed859da Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
paulson
parents: 22470
diff changeset
   445
      val (_,lines) = foldr add_wanted_prfline (0,[]) nonnull_lines
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21979
diff changeset
   446
      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
   447
      val ccls = map forall_intr_vars ccls
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   448
  in
22130
0906fd95e0b5 Output.debug: non-strict;
wenzelm
parents: 22044
diff changeset
   449
    app (fn th => Output.debug (fn () => string_of_thm th)) ccls;
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   450
    isar_header (map #1 fixes) ^
25492
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   451
    String.concat (isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines)) ^
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   452
    "qed\n"
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   453
  end
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   454
  handle e => (*FIXME: exn handler is too general!*)
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   455
    "Translation of TSTP raised an exception: " ^ Toplevel.exn_message e;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   456
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   457
(*Could use split_lines, but it can return blank lines...*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   458
val lines = String.tokens (equal #"\n");
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   459
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   460
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
   461
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   462
val txt_path = Path.ext "txt" o Path.explode o nospaces;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   463
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   464
fun signal_success probfile toParent ppid msg =
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   465
  let val _ = trace ("\nReporting Success for " ^ probfile ^ "\n" ^ msg)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   466
  in
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   467
    (*We write the proof to a file because sending very long lines may fail...*)
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   468
    File.write (txt_path probfile) msg;
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   469
    TextIO.output (toParent, "Success.\n");
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   470
    TextIO.output (toParent, probfile ^ "\n");
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   471
    TextIO.flushOut toParent;
25414
3326bd7ecd48 back to sigusr2, after Poly/ML 5.1 has been adapted;
wenzelm
parents: 25413
diff changeset
   472
    Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2);
23833
3fe991a1f805 Full sort information by default.
paulson
parents: 23519
diff changeset
   473
    (*Give the parent time to respond before possibly sending another signal*)
3fe991a1f805 Full sort information by default.
paulson
parents: 23519
diff changeset
   474
    OS.Process.sleep (Time.fromMilliseconds 600)
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23139
diff changeset
   475
  end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   476
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   477
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   478
(**** retrieve the axioms that were used in the proof ****)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   479
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   480
(*PureThy.get_name_hint returns "??.unknown" if no name is available.*)
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   481
fun goodhint x = (x <> "??.unknown");  
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   482
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   483
(*Get names of axioms used. Axioms are indexed from 1, while the vector is indexed from 0*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   484
fun get_axiom_names (thm_names: string vector) step_nums =
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   485
    let fun is_axiom n = n <= Vector.length thm_names
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   486
        fun getname i = Vector.sub(thm_names, i-1)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   487
    in
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   488
	sort_distinct string_ord (filter goodhint (map getname (filter is_axiom step_nums)))
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   489
    end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   490
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   491
 (*String contains multiple lines. We want those of the form
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   492
     "253[0:Inp] et cetera..."
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   493
  A list consisting of the first number in each line is returned. *)
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   494
fun get_spass_linenums proofextract =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   495
  let val toks = String.tokens (not o Char.isAlphaNum)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   496
      fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   497
        | inputno _ = NONE
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   498
      val lines = String.tokens (fn c => c = #"\n") proofextract
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   499
  in  List.mapPartial (inputno o toks) lines  end
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   500
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   501
fun get_axiom_names_spass proofextract thm_names =
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   502
   get_axiom_names thm_names (get_spass_linenums proofextract);
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   503
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   504
fun not_comma c = c <>  #",";
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   505
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   506
(*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
   507
fun parse_tstp_line s =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   508
  let val ss = Substring.full (unprefix "cnf(" (nospaces s))
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   509
      val (intf,rest) = Substring.splitl not_comma ss
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   510
      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
   511
      (*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
   512
        get_axiom_names above, while suppressing the UNSOUND warning*)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   513
      val ints = if Substring.string rolef mem_string ["axiom","negated_conjecture"]
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   514
                 then Substring.string intf
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   515
                 else "error"
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   516
  in  Int.fromString ints  end
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   517
  handle Fail _ => NONE;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   518
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   519
fun get_axiom_names_tstp proofextract thm_names =
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   520
   get_axiom_names thm_names (List.mapPartial parse_tstp_line (split_lines proofextract));
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   521
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   522
 (*String contains multiple lines. We want those of the form
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24552
diff changeset
   523
     "*********** [448, input] ***********"
779fc4fcbf8b metis now available in PreList
paulson
parents: 24552
diff changeset
   524
   or possibly those of the form
779fc4fcbf8b metis now available in PreList
paulson
parents: 24552
diff changeset
   525
     "cnf(108, axiom, ..."
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   526
  A list consisting of the first number in each line is returned. *)
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   527
fun get_vamp_linenums proofextract =
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   528
  let val toks = String.tokens (not o Char.isAlphaNum)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   529
      fun inputno [ntok,"input"] = Int.fromString ntok
24632
779fc4fcbf8b metis now available in PreList
paulson
parents: 24552
diff changeset
   530
        | inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   531
        | inputno _ = NONE
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   532
      val lines = String.tokens (fn c => c = #"\n") proofextract
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   533
  in  List.mapPartial (inputno o toks) lines  end
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   534
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   535
fun get_axiom_names_vamp proofextract thm_names =
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   536
   get_axiom_names thm_names (get_vamp_linenums proofextract);
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   537
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   538
fun get_axiom_names_for E       = get_axiom_names_tstp
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   539
  | get_axiom_names_for SPASS   = get_axiom_names_spass
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   540
  | get_axiom_names_for Vampire = get_axiom_names_vamp;
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   541
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   542
fun metis_line [] = "apply metis"
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   543
  | metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")"
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   544
25492
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   545
(*Used to label theorems chained into the sledgehammer call*)
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   546
val chained_hint = "CHAINED";
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   547
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   548
val nochained = filter_out (fn y => y = chained_hint);
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   549
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   550
(*The signal handler in watcher.ML must be able to read the output of this.*)
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   551
fun lemma_list atp proofextract thm_names probfile toParent ppid =
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   552
  signal_success probfile toParent ppid 
25492
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   553
                 (metis_line (nochained (get_axiom_names_for atp proofextract thm_names)));
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   554
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   555
fun tstp_extract proofextract thm_names probfile toParent ppid ctxt th sgno =
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   556
  let val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
25492
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   557
      val names = get_axiom_names_tstp proofextract thm_names
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   558
      val line1 = metis_line (nochained names)
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   559
      val line2 = if chained_hint mem_string names then ""
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   560
                  else decode_tstp_file cnfs ctxt th sgno thm_names
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   561
  in
25492
4cc7976948ac Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
paulson
parents: 25414
diff changeset
   562
    signal_success probfile toParent ppid (line1 ^ "\n" ^ line2)
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   563
  end;
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   564
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   565
(**** Extracting proofs from an ATP's output ****)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   566
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   567
val start_TSTP = "SZS output start CNFRefutation"
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   568
val end_TSTP   = "SZS output end CNFRefutation"
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   569
val start_E = "# Proof object starts here."
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   570
val end_E   = "# Proof object ends here."
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   571
val start_V8 = "=========== Refutation =========="
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   572
val end_V8 = "======= End of refutation ======="
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   573
val start_SPASS = "Here is a proof"
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   574
val end_SPASS = "Formulae used in the proof"
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   575
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   576
fun any_substring ss ln = exists (fn s => String.isSubstring s ln) ss;
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   577
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   578
(*********************************************************************************)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   579
(*  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
   580
(*  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
   581
(*********************************************************************************)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   582
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   583
(*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
   584
  return value is currently never used!*)
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   585
fun startTransfer endS (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names) =
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   586
 let val atp = if endS = end_V8 then Vampire
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   587
	       else if endS = end_SPASS then SPASS
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   588
	       else E
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   589
     fun transferInput proofextract =
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   590
       case TextIO.inputLine fromChild of
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   591
	   NONE =>  (*end of file?*)
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   592
	     (trace ("\n extraction_failed.  End bracket: " ^ endS ^
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   593
		     "\naccumulated text: " ^ proofextract);
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   594
	      false)
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   595
	 | SOME thisLine =>
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   596
	     if any_substring [endS,end_TSTP] thisLine
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   597
	     then
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   598
	      (trace ("\nExtracted proof:\n" ^ proofextract);
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   599
	       if String.isPrefix "cnf(" proofextract
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   600
	       then tstp_extract proofextract thm_names probfile toParent ppid ctxt th sgno
24425
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   601
	       else lemma_list atp proofextract thm_names probfile toParent ppid;
ca97c6f3d9cd Returning both a "one-line" proof and a structured proof
paulson
parents: 24387
diff changeset
   602
	       true)
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   603
	     else transferInput (proofextract ^ thisLine)
21978
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
     transferInput ""
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   606
 end
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   607
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   608
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   609
(*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
   610
fun signal_parent (toParent, ppid, msg, probfile) =
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   611
 (TextIO.output (toParent, msg);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   612
  TextIO.output (toParent, probfile ^ "\n");
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   613
  TextIO.flushOut toParent;
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   614
  trace ("\nSignalled parent: " ^ msg ^ probfile);
25414
3326bd7ecd48 back to sigusr2, after Poly/ML 5.1 has been adapted;
wenzelm
parents: 25413
diff changeset
   615
  Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2);
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   616
  (*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
   617
  OS.Process.sleep (Time.fromMilliseconds 600));
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   618
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   619
(*FIXME: once TSTP output is produced by all ATPs, these three functions can be combined.*)
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   620
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   621
(*Called from watcher. Returns true if the Vampire process has returned a verdict.*)
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   622
fun checkVampProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   623
  (case TextIO.inputLine fromChild of
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   624
    NONE => (trace "\nNo proof output seen"; false)
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   625
  | SOME thisLine =>
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   626
     if any_substring [start_V8,start_TSTP] thisLine
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   627
     then startTransfer end_V8 arg
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   628
     else if (String.isSubstring "Satisfiability detected" thisLine) orelse
24956
4992239a414e failure messages
paulson
parents: 24937
diff changeset
   629
             (String.isSubstring "Refutation not found" thisLine) orelse
4992239a414e failure messages
paulson
parents: 24937
diff changeset
   630
             (String.isSubstring "CANNOT PROVE" thisLine)
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   631
     then (signal_parent (toParent, ppid, "Failure\n", probfile);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   632
	   true)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   633
     else checkVampProofFound arg);
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   634
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   635
(*Called from watcher. Returns true if the E process has returned a verdict.*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   636
fun checkEProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   637
  (case TextIO.inputLine fromChild of
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   638
    NONE => (trace "\nNo proof output seen"; false)
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   639
  | SOME thisLine =>
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   640
     if any_substring [start_E,start_TSTP] thisLine 
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   641
     then startTransfer end_E arg
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   642
     else if String.isSubstring "SZS status: Satisfiable" thisLine
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   643
     then (signal_parent (toParent, ppid, "Invalid\n", probfile);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   644
	   true)
24956
4992239a414e failure messages
paulson
parents: 24937
diff changeset
   645
     else if String.isSubstring "SZS status: ResourceOut" thisLine orelse
4992239a414e failure messages
paulson
parents: 24937
diff changeset
   646
             String.isSubstring "# Cannot determine problem status" thisLine
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   647
     then (signal_parent (toParent, ppid, "Failure\n", probfile);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   648
	   true)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   649
     else checkEProofFound arg);
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   650
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   651
(*Called from watcher. Returns true if the SPASS process has returned a verdict.*)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   652
fun checkSpassProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   653
  (case TextIO.inputLine fromChild of
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   654
    NONE => (trace "\nNo proof output seen"; false)
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   655
  | SOME thisLine =>
24547
64c20ee76bc1 Vampire structured proofs. Better parsing; some bug fixes.
paulson
parents: 24493
diff changeset
   656
     if any_substring [start_SPASS,start_TSTP] thisLine
22012
adf68479ae1b Proof.context now sent to watcher and used in type inference step of proof reconstruction
paulson
parents: 21999
diff changeset
   657
     then startTransfer end_SPASS arg
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   658
     else if thisLine = "SPASS beiseite: Completion found.\n"
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   659
     then (signal_parent (toParent, ppid, "Invalid\n", probfile);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   660
	   true)
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   661
     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
   662
             thisLine = "SPASS beiseite: Maximal number of loops exceeded.\n"
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   663
     then (signal_parent (toParent, ppid, "Failure\n", probfile);
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   664
	   true)
23139
aa899bce7c3b TextIO.inputLine: use present SML B library version;
wenzelm
parents: 23083
diff changeset
   665
    else checkSpassProofFound arg);
21978
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   666
72c21698a055 Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
paulson
parents:
diff changeset
   667
end;