src/HOL/Tools/res_reconstruct.ML
changeset 35825 a6aad5a70ed4
parent 35814 234eaa508359
child 35826 1590abc3d42a
equal deleted inserted replaced
35814:234eaa508359 35825:a6aad5a70ed4
     1 (*  Title:      HOL/Tools/res_reconstruct.ML
       
     2     Author:     Lawrence C Paulson and Claire Quigley, Cambridge University Computer Laboratory
       
     3 
       
     4 Transfer of proofs from external provers.
       
     5 *)
       
     6 
       
     7 signature RES_RECONSTRUCT =
       
     8 sig
       
     9   val chained_hint: string
       
    10 
       
    11   val fix_sorts: sort Vartab.table -> term -> term
       
    12   val invert_const: string -> string
       
    13   val invert_type_const: string -> string
       
    14   val num_typargs: theory -> string -> int
       
    15   val make_tvar: string -> typ
       
    16   val strip_prefix: string -> string -> string option
       
    17   val setup: theory -> theory
       
    18   (* extracting lemma list*)
       
    19   val find_failure: string -> string option
       
    20   val lemma_list: bool -> string ->
       
    21     string * string vector * (int * int) * Proof.context * thm * int -> string * string list
       
    22   (* structured proofs *)
       
    23   val structured_proof: string ->
       
    24     string * string vector * (int * int) * Proof.context * thm * int -> string * string list
       
    25 end;
       
    26 
       
    27 structure Res_Reconstruct : RES_RECONSTRUCT =
       
    28 struct
       
    29 
       
    30 val trace_path = Path.basic "atp_trace";
       
    31 
       
    32 fun trace s =
       
    33   if ! Res_Axioms.trace then File.append (File.tmp_path trace_path) s
       
    34   else ();
       
    35 
       
    36 fun string_of_thm ctxt = PrintMode.setmp [] (Display.string_of_thm ctxt);
       
    37 
       
    38 (*For generating structured proofs: keep every nth proof line*)
       
    39 val (modulus, modulus_setup) = Attrib.config_int "sledgehammer_modulus" 1;
       
    40 
       
    41 (*Indicates whether to include sort information in generated proofs*)
       
    42 val (recon_sorts, recon_sorts_setup) = Attrib.config_bool "sledgehammer_sorts" true;
       
    43 
       
    44 (*Indicated whether to generate full proofs or just lemma lists - now via setup of atps*)
       
    45 (* val (full_proofs, full_proofs_setup) = Attrib.config_bool "sledgehammer_full" false; *)
       
    46 
       
    47 val setup = modulus_setup #> recon_sorts_setup;
       
    48 
       
    49 (**** PARSING OF TSTP FORMAT ****)
       
    50 
       
    51 (*Syntax trees, either termlist or formulae*)
       
    52 datatype stree = Int of int | Br of string * stree list;
       
    53 
       
    54 fun atom x = Br(x,[]);
       
    55 
       
    56 fun scons (x,y) = Br("cons", [x,y]);
       
    57 val listof = List.foldl scons (atom "nil");
       
    58 
       
    59 (*Strings enclosed in single quotes, e.g. filenames*)
       
    60 val quoted = $$"'" |-- Scan.repeat (~$$"'") --| $$"'" >> implode;
       
    61 
       
    62 (*Intended for $true and $false*)
       
    63 fun tf s = "c_" ^ str (Char.toUpper (String.sub(s,0))) ^ String.extract(s,1,NONE);
       
    64 val truefalse = $$"$" |-- Symbol.scan_id >> (atom o tf);
       
    65 
       
    66 (*Integer constants, typically proof line numbers*)
       
    67 fun is_digit s = Char.isDigit (String.sub(s,0));
       
    68 val integer = Scan.many1 is_digit >> (the o Int.fromString o implode);
       
    69 
       
    70 (*Generalized FO terms, which include filenames, numbers, etc.*)
       
    71 fun termlist x = (term ::: Scan.repeat ($$"," |-- term)) x
       
    72 and term x = (quoted >> atom || integer>>Int || truefalse ||
       
    73               Symbol.scan_id -- Scan.optional ($$"(" |-- termlist --| $$")") [] >> Br ||
       
    74               $$"(" |-- term --| $$")" ||
       
    75               $$"[" |-- Scan.optional termlist [] --| $$"]" >> listof) x;
       
    76 
       
    77 fun negate t = Br("c_Not", [t]);
       
    78 fun equate (t1,t2) = Br("c_equal", [t1,t2]);
       
    79 
       
    80 (*Apply equal or not-equal to a term*)
       
    81 fun syn_equal (t, NONE) = t
       
    82   | syn_equal (t1, SOME (NONE, t2)) = equate (t1,t2)
       
    83   | syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2));
       
    84 
       
    85 (*Literals can involve negation, = and !=.*)
       
    86 fun literal x = ($$"~" |-- literal >> negate ||
       
    87                  (term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal)) x;
       
    88 
       
    89 val literals = literal ::: Scan.repeat ($$"|" |-- literal);
       
    90 
       
    91 (*Clause: a list of literals separated by the disjunction sign*)
       
    92 val clause = $$"(" |-- literals --| $$")" || Scan.single literal;
       
    93 
       
    94 val annotations = $$"," |-- term -- Scan.option ($$"," |-- termlist);
       
    95 
       
    96 (*<cnf_annotated> ::= cnf(<name>,<formula_role>,<cnf_formula><annotations>).
       
    97   The <name> could be an identifier, but we assume integers.*)
       
    98 val tstp_line = (Scan.this_string "cnf" -- $$"(") |--
       
    99                 integer --| $$"," -- Symbol.scan_id --| $$"," --
       
   100                 clause -- Scan.option annotations --| $$ ")";
       
   101 
       
   102 
       
   103 (**** INTERPRETATION OF TSTP SYNTAX TREES ****)
       
   104 
       
   105 exception STREE of stree;
       
   106 
       
   107 (*If string s has the prefix s1, return the result of deleting it.*)
       
   108 fun strip_prefix s1 s =
       
   109   if String.isPrefix s1 s
       
   110   then SOME (Res_Clause.undo_ascii_of (String.extract (s, size s1, NONE)))
       
   111   else NONE;
       
   112 
       
   113 (*Invert the table of translations between Isabelle and ATPs*)
       
   114 val type_const_trans_table_inv =
       
   115       Symtab.make (map swap (Symtab.dest Res_Clause.type_const_trans_table));
       
   116 
       
   117 fun invert_type_const c =
       
   118     case Symtab.lookup type_const_trans_table_inv c of
       
   119         SOME c' => c'
       
   120       | NONE => c;
       
   121 
       
   122 fun make_tvar b = TVar(("'" ^ b, 0), HOLogic.typeS);
       
   123 fun make_var (b,T) = Var((b,0),T);
       
   124 
       
   125 (*Type variables are given the basic sort, HOL.type. Some will later be constrained
       
   126   by information from type literals, or by type inference.*)
       
   127 fun type_of_stree t =
       
   128   case t of
       
   129       Int _ => raise STREE t
       
   130     | Br (a,ts) =>
       
   131         let val Ts = map type_of_stree ts
       
   132         in
       
   133           case strip_prefix Res_Clause.tconst_prefix a of
       
   134               SOME b => Type(invert_type_const b, Ts)
       
   135             | NONE =>
       
   136                 if not (null ts) then raise STREE t  (*only tconsts have type arguments*)
       
   137                 else
       
   138                 case strip_prefix Res_Clause.tfree_prefix a of
       
   139                     SOME b => TFree("'" ^ b, HOLogic.typeS)
       
   140                   | NONE =>
       
   141                 case strip_prefix Res_Clause.tvar_prefix a of
       
   142                     SOME b => make_tvar b
       
   143                   | NONE => make_tvar a   (*Variable from the ATP, say X1*)
       
   144         end;
       
   145 
       
   146 (*Invert the table of translations between Isabelle and ATPs*)
       
   147 val const_trans_table_inv =
       
   148       Symtab.update ("fequal", "op =")
       
   149         (Symtab.make (map swap (Symtab.dest Res_Clause.const_trans_table)));
       
   150 
       
   151 fun invert_const c =
       
   152     case Symtab.lookup const_trans_table_inv c of
       
   153         SOME c' => c'
       
   154       | NONE => c;
       
   155 
       
   156 (*The number of type arguments of a constant, zero if it's monomorphic*)
       
   157 fun num_typargs thy s = length (Sign.const_typargs thy (s, Sign.the_const_type thy s));
       
   158 
       
   159 (*Generates a constant, given its type arguments*)
       
   160 fun const_of thy (a,Ts) = Const(a, Sign.const_instance thy (a,Ts));
       
   161 
       
   162 (*First-order translation. No types are known for variables. HOLogic.typeT should allow
       
   163   them to be inferred.*)
       
   164 fun term_of_stree args thy t =
       
   165   case t of
       
   166       Int _ => raise STREE t
       
   167     | Br ("hBOOL",[t]) => term_of_stree [] thy t  (*ignore hBOOL*)
       
   168     | Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t
       
   169     | Br (a,ts) =>
       
   170         case strip_prefix Res_Clause.const_prefix a of
       
   171             SOME "equal" =>
       
   172               list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree [] thy) ts)
       
   173           | SOME b =>
       
   174               let val c = invert_const b
       
   175                   val nterms = length ts - num_typargs thy c
       
   176                   val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args)
       
   177                   (*Extra args from hAPP come AFTER any arguments given directly to the
       
   178                     constant.*)
       
   179                   val Ts = List.map type_of_stree (List.drop(ts,nterms))
       
   180               in  list_comb(const_of thy (c, Ts), us)  end
       
   181           | NONE => (*a variable, not a constant*)
       
   182               let val T = HOLogic.typeT
       
   183                   val opr = (*a Free variable is typically a Skolem function*)
       
   184                     case strip_prefix Res_Clause.fixed_var_prefix a of
       
   185                         SOME b => Free(b,T)
       
   186                       | NONE =>
       
   187                     case strip_prefix Res_Clause.schematic_var_prefix a of
       
   188                         SOME b => make_var (b,T)
       
   189                       | NONE => make_var (a,T)    (*Variable from the ATP, say X1*)
       
   190               in  list_comb (opr, List.map (term_of_stree [] thy) (ts@args))  end;
       
   191 
       
   192 (*Type class literal applied to a type. Returns triple of polarity, class, type.*)
       
   193 fun constraint_of_stree pol (Br("c_Not",[t])) = constraint_of_stree (not pol) t
       
   194   | constraint_of_stree pol t = case t of
       
   195         Int _ => raise STREE t
       
   196       | Br (a,ts) =>
       
   197             (case (strip_prefix Res_Clause.class_prefix a, map type_of_stree ts) of
       
   198                  (SOME b, [T]) => (pol, b, T)
       
   199                | _ => raise STREE t);
       
   200 
       
   201 (** Accumulate type constraints in a clause: negative type literals **)
       
   202 
       
   203 fun addix (key,z)  = Vartab.map_default (key,[]) (cons z);
       
   204 
       
   205 fun add_constraint ((false, cl, TFree(a,_)), vt) = addix ((a,~1),cl) vt
       
   206   | add_constraint ((false, cl, TVar(ix,_)), vt) = addix (ix,cl) vt
       
   207   | add_constraint (_, vt) = vt;
       
   208 
       
   209 (*False literals (which E includes in its proofs) are deleted*)
       
   210 val nofalses = filter (not o equal HOLogic.false_const);
       
   211 
       
   212 (*Final treatment of the list of "real" literals from a clause.*)
       
   213 fun finish [] = HOLogic.true_const  (*No "real" literals means only type information*)
       
   214   | finish lits =
       
   215       case nofalses lits of
       
   216           [] => HOLogic.false_const  (*The empty clause, since we started with real literals*)
       
   217         | xs => foldr1 HOLogic.mk_disj (rev xs);
       
   218 
       
   219 (*Accumulate sort constraints in vt, with "real" literals in lits.*)
       
   220 fun lits_of_strees _ (vt, lits) [] = (vt, finish lits)
       
   221   | lits_of_strees ctxt (vt, lits) (t::ts) =
       
   222       lits_of_strees ctxt (add_constraint (constraint_of_stree true t, vt), lits) ts
       
   223       handle STREE _ =>
       
   224       lits_of_strees ctxt (vt, term_of_stree [] (ProofContext.theory_of ctxt) t :: lits) ts;
       
   225 
       
   226 (*Update TVars/TFrees with detected sort constraints.*)
       
   227 fun fix_sorts vt =
       
   228   let fun tysubst (Type (a, Ts)) = Type (a, map tysubst Ts)
       
   229         | tysubst (TVar (xi, s)) = TVar (xi, the_default s (Vartab.lookup vt xi))
       
   230         | tysubst (TFree (x, s)) = TFree (x, the_default s (Vartab.lookup vt (x, ~1)))
       
   231       fun tmsubst (Const (a, T)) = Const (a, tysubst T)
       
   232         | tmsubst (Free (a, T)) = Free (a, tysubst T)
       
   233         | tmsubst (Var (xi, T)) = Var (xi, tysubst T)
       
   234         | tmsubst (t as Bound _) = t
       
   235         | tmsubst (Abs (a, T, t)) = Abs (a, tysubst T, tmsubst t)
       
   236         | tmsubst (t $ u) = tmsubst t $ tmsubst u;
       
   237   in fn t => if Vartab.is_empty vt then t else tmsubst t end;
       
   238 
       
   239 (*Interpret a list of syntax trees as a clause, given by "real" literals and sort constraints.
       
   240   vt0 holds the initial sort constraints, from the conjecture clauses.*)
       
   241 fun clause_of_strees ctxt vt0 ts =
       
   242   let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts in
       
   243     singleton (Syntax.check_terms ctxt) (TypeInfer.constrain HOLogic.boolT (fix_sorts vt dt))
       
   244   end;
       
   245 
       
   246 fun gen_all_vars t = fold_rev Logic.all (OldTerm.term_vars t) t;
       
   247 
       
   248 fun ints_of_stree_aux (Int n, ns) = n::ns
       
   249   | ints_of_stree_aux (Br(_,ts), ns) = List.foldl ints_of_stree_aux ns ts;
       
   250 
       
   251 fun ints_of_stree t = ints_of_stree_aux (t, []);
       
   252 
       
   253 fun decode_tstp vt0 (name, role, ts, annots) ctxt =
       
   254   let val deps = case annots of NONE => [] | SOME (source,_) => ints_of_stree source
       
   255       val cl = clause_of_strees ctxt vt0 ts
       
   256   in  ((name, role, cl, deps), fold Variable.declare_term (OldTerm.term_frees cl) ctxt)  end;
       
   257 
       
   258 fun dest_tstp ((((name, role), ts), annots), chs) =
       
   259   case chs of
       
   260           "."::_ => (name, role, ts, annots)
       
   261         | _ => error ("TSTP line not terminated by \".\": " ^ implode chs);
       
   262 
       
   263 
       
   264 (** Global sort constraints on TFrees (from tfree_tcs) are positive unit clauses. **)
       
   265 
       
   266 fun add_tfree_constraint ((true, cl, TFree(a,_)), vt) = addix ((a,~1),cl) vt
       
   267   | add_tfree_constraint (_, vt) = vt;
       
   268 
       
   269 fun tfree_constraints_of_clauses vt [] = vt
       
   270   | tfree_constraints_of_clauses vt ([lit]::tss) =
       
   271       (tfree_constraints_of_clauses (add_tfree_constraint (constraint_of_stree true lit, vt)) tss
       
   272        handle STREE _ => (*not a positive type constraint: ignore*)
       
   273        tfree_constraints_of_clauses vt tss)
       
   274   | tfree_constraints_of_clauses vt (_::tss) = tfree_constraints_of_clauses vt tss;
       
   275 
       
   276 
       
   277 (**** Translation of TSTP files to Isar Proofs ****)
       
   278 
       
   279 fun decode_tstp_list ctxt tuples =
       
   280   let val vt0 = tfree_constraints_of_clauses Vartab.empty (map #3 tuples)
       
   281   in  #1 (fold_map (decode_tstp vt0) tuples ctxt) end;
       
   282 
       
   283 (** Finding a matching assumption. The literals may be permuted, and variable names
       
   284     may disagree. We have to try all combinations of literals (quadratic!) and
       
   285     match up the variable names consistently. **)
       
   286 
       
   287 fun strip_alls_aux n (Const("all",_)$Abs(a,T,t))  =
       
   288       strip_alls_aux (n+1) (subst_bound (Var ((a,n), T), t))
       
   289   | strip_alls_aux _ t  =  t;
       
   290 
       
   291 val strip_alls = strip_alls_aux 0;
       
   292 
       
   293 exception MATCH_LITERAL;
       
   294 
       
   295 (*Ignore types: they are not to be trusted...*)
       
   296 fun match_literal (t1$u1) (t2$u2) env =
       
   297       match_literal t1 t2 (match_literal u1 u2 env)
       
   298   | match_literal (Abs (_,_,t1)) (Abs (_,_,t2)) env =
       
   299       match_literal t1 t2 env
       
   300   | match_literal (Bound i1) (Bound i2) env =
       
   301       if i1=i2 then env else raise MATCH_LITERAL
       
   302   | match_literal (Const(a1,_)) (Const(a2,_)) env =
       
   303       if a1=a2 then env else raise MATCH_LITERAL
       
   304   | match_literal (Free(a1,_)) (Free(a2,_)) env =
       
   305       if a1=a2 then env else raise MATCH_LITERAL
       
   306   | match_literal (Var(ix1,_)) (Var(ix2,_)) env = insert (op =) (ix1,ix2) env
       
   307   | match_literal _ _ _ = raise MATCH_LITERAL;
       
   308 
       
   309 (*Checking that all variable associations are unique. The list env contains no
       
   310   repetitions, but does it contain say (x,y) and (y,y)? *)
       
   311 fun good env =
       
   312   let val (xs,ys) = ListPair.unzip env
       
   313   in  not (has_duplicates (op=) xs orelse has_duplicates (op=) ys)  end;
       
   314 
       
   315 (*Match one list of literals against another, ignoring types and the order of
       
   316   literals. Sorting is unreliable because we don't have types or variable names.*)
       
   317 fun matches_aux _ [] [] = true
       
   318   | matches_aux env (lit::lits) ts =
       
   319       let fun match1 us [] = false
       
   320             | match1 us (t::ts) =
       
   321                 let val env' = match_literal lit t env
       
   322                 in  (good env' andalso matches_aux env' lits (us@ts)) orelse
       
   323                     match1 (t::us) ts
       
   324                 end
       
   325                 handle MATCH_LITERAL => match1 (t::us) ts
       
   326       in  match1 [] ts  end;
       
   327 
       
   328 (*Is this length test useful?*)
       
   329 fun matches (lits1,lits2) =
       
   330   length lits1 = length lits2  andalso
       
   331   matches_aux [] (map Envir.eta_contract lits1) (map Envir.eta_contract lits2);
       
   332 
       
   333 fun permuted_clause t =
       
   334   let val lits = HOLogic.disjuncts t
       
   335       fun perm [] = NONE
       
   336         | perm (ctm::ctms) =
       
   337             if matches (lits, HOLogic.disjuncts (HOLogic.dest_Trueprop (strip_alls ctm)))
       
   338             then SOME ctm else perm ctms
       
   339   in perm end;
       
   340 
       
   341 fun have_or_show "show " _ = "show \""
       
   342   | have_or_show have lname = have ^ lname ^ ": \""
       
   343 
       
   344 (*ctms is a list of conjecture clauses as yielded by Isabelle. Those returned by the
       
   345   ATP may have their literals reordered.*)
       
   346 fun isar_lines ctxt ctms =
       
   347   let val string_of = PrintMode.setmp [] (fn term => Syntax.string_of_term ctxt term)
       
   348       val _ = trace ("\n\nisar_lines: start\n")
       
   349       fun doline _ (lname, t, []) =  (*No deps: it's a conjecture clause, with no proof.*)
       
   350            (case permuted_clause t ctms of
       
   351                 SOME u => "assume " ^ lname ^ ": \"" ^ string_of u ^ "\"\n"
       
   352               | NONE => "assume? " ^ lname ^ ": \"" ^ string_of t ^ "\"\n")  (*no match!!*)
       
   353         | doline have (lname, t, deps) =
       
   354             have_or_show have lname ^ string_of (gen_all_vars (HOLogic.mk_Trueprop t)) ^
       
   355             "\"\n  by (metis " ^ space_implode " " deps ^ ")\n"
       
   356       fun dolines [(lname, t, deps)] = [doline "show " (lname, t, deps)]
       
   357         | dolines ((lname, t, deps)::lines) = doline "have " (lname, t, deps) :: dolines lines
       
   358   in setmp_CRITICAL show_sorts (Config.get ctxt recon_sorts) dolines end;
       
   359 
       
   360 fun notequal t (_,t',_) = not (t aconv t');
       
   361 
       
   362 (*No "real" literals means only type information*)
       
   363 fun eq_types t = t aconv HOLogic.true_const;
       
   364 
       
   365 fun replace_dep (old:int, new) dep = if dep=old then new else [dep];
       
   366 
       
   367 fun replace_deps (old:int, new) (lno, t, deps) =
       
   368       (lno, t, List.foldl (uncurry (union (op =))) [] (map (replace_dep (old, new)) deps));
       
   369 
       
   370 (*Discard axioms; consolidate adjacent lines that prove the same clause, since they differ
       
   371   only in type information.*)
       
   372 fun add_prfline ((lno, "axiom", t, []), lines) =  (*axioms are not proof lines*)
       
   373       if eq_types t (*must be clsrel/clsarity: type information, so delete refs to it*)
       
   374       then map (replace_deps (lno, [])) lines
       
   375       else
       
   376        (case take_prefix (notequal t) lines of
       
   377            (_,[]) => lines                  (*no repetition of proof line*)
       
   378          | (pre, (lno', _, _) :: post) =>   (*repetition: replace later line by earlier one*)
       
   379              pre @ map (replace_deps (lno', [lno])) post)
       
   380   | add_prfline ((lno, _, t, []), lines) =  (*no deps: conjecture clause*)
       
   381       (lno, t, []) :: lines
       
   382   | add_prfline ((lno, _, t, deps), lines) =
       
   383       if eq_types t then (lno, t, deps) :: lines
       
   384       (*Type information will be deleted later; skip repetition test.*)
       
   385       else (*FIXME: Doesn't this code risk conflating proofs involving different types??*)
       
   386       case take_prefix (notequal t) lines of
       
   387          (_,[]) => (lno, t, deps) :: lines  (*no repetition of proof line*)
       
   388        | (pre, (lno', t', _) :: post) =>
       
   389            (lno, t', deps) ::               (*repetition: replace later line by earlier one*)
       
   390            (pre @ map (replace_deps (lno', [lno])) post);
       
   391 
       
   392 (*Recursively delete empty lines (type information) from the proof.*)
       
   393 fun add_nonnull_prfline ((lno, t, []), lines) = (*no dependencies, so a conjecture clause*)
       
   394      if eq_types t (*must be type information, tfree_tcs, clsrel, clsarity: delete refs to it*)
       
   395      then delete_dep lno lines
       
   396      else (lno, t, []) :: lines
       
   397   | add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines
       
   398 and delete_dep lno lines = List.foldr add_nonnull_prfline [] (map (replace_deps (lno, [])) lines);
       
   399 
       
   400 fun bad_free (Free (a,_)) = String.isPrefix "sko_" a
       
   401   | bad_free _ = false;
       
   402 
       
   403 (*TVars are forbidden in goals. Also, we don't want lines with <2 dependencies.
       
   404   To further compress proofs, setting modulus:=n deletes every nth line, and nlines
       
   405   counts the number of proof lines processed so far.
       
   406   Deleted lines are replaced by their own dependencies. Note that the "add_nonnull_prfline"
       
   407   phase may delete some dependencies, hence this phase comes later.*)
       
   408 fun add_wanted_prfline ctxt ((lno, t, []), (nlines, lines)) =
       
   409       (nlines, (lno, t, []) :: lines)   (*conjecture clauses must be kept*)
       
   410   | add_wanted_prfline ctxt ((lno, t, deps), (nlines, lines)) =
       
   411       if eq_types t orelse not (null (Term.add_tvars t [])) orelse
       
   412          exists_subterm bad_free t orelse
       
   413          (not (null lines) andalso   (*final line can't be deleted for these reasons*)
       
   414           (length deps < 2 orelse nlines mod (Config.get ctxt modulus) <> 0))
       
   415       then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*)
       
   416       else (nlines+1, (lno, t, deps) :: lines);
       
   417 
       
   418 (*Replace numeric proof lines by strings, either from thm_names or sequential line numbers*)
       
   419 fun stringify_deps thm_names deps_map [] = []
       
   420   | stringify_deps thm_names deps_map ((lno, t, deps) :: lines) =
       
   421       if lno <= Vector.length thm_names  (*axiom*)
       
   422       then (Vector.sub(thm_names,lno-1), t, []) :: stringify_deps thm_names deps_map lines
       
   423       else let val lname = Int.toString (length deps_map)
       
   424                fun fix lno = if lno <= Vector.length thm_names
       
   425                              then SOME(Vector.sub(thm_names,lno-1))
       
   426                              else AList.lookup op= deps_map lno;
       
   427            in  (lname, t, map_filter fix (distinct (op=) deps)) ::
       
   428                stringify_deps thm_names ((lno,lname)::deps_map) lines
       
   429            end;
       
   430 
       
   431 val proofstart = "proof (neg_clausify)\n";
       
   432 
       
   433 fun isar_header [] = proofstart
       
   434   | isar_header ts = proofstart ^ "fix " ^ space_implode " " ts ^ "\n";
       
   435 
       
   436 fun decode_tstp_file cnfs ctxt th sgno thm_names =
       
   437   let val _ = trace "\ndecode_tstp_file: start\n"
       
   438       val tuples = map (dest_tstp o tstp_line o explode) cnfs
       
   439       val _ = trace (Int.toString (length tuples) ^ " tuples extracted\n")
       
   440       val ctxt = ProofContext.set_mode ProofContext.mode_schematic ctxt
       
   441       val raw_lines = List.foldr add_prfline [] (decode_tstp_list ctxt tuples)
       
   442       val _ = trace (Int.toString (length raw_lines) ^ " raw_lines extracted\n")
       
   443       val nonnull_lines = List.foldr add_nonnull_prfline [] raw_lines
       
   444       val _ = trace (Int.toString (length nonnull_lines) ^ " nonnull_lines extracted\n")
       
   445       val (_,lines) = List.foldr (add_wanted_prfline ctxt) (0,[]) nonnull_lines
       
   446       val _ = trace (Int.toString (length lines) ^ " lines extracted\n")
       
   447       val (ccls,fixes) = Res_Axioms.neg_conjecture_clauses ctxt th sgno
       
   448       val _ = trace (Int.toString (length ccls) ^ " conjecture clauses\n")
       
   449       val ccls = map forall_intr_vars ccls
       
   450       val _ =
       
   451         if ! Res_Axioms.trace then app (fn th => trace ("\nccl: " ^ string_of_thm ctxt th)) ccls
       
   452         else ()
       
   453       val ilines = isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines)
       
   454       val _ = trace "\ndecode_tstp_file: finishing\n"
       
   455   in
       
   456     isar_header (map #1 fixes) ^ implode ilines ^ "qed\n"
       
   457   end handle STREE _ => error "Could not extract proof (ATP output malformed?)";
       
   458 
       
   459 
       
   460 (*=== EXTRACTING PROOF-TEXT === *)
       
   461 
       
   462 val begin_proof_strings = ["# SZS output start CNFRefutation.",
       
   463   "=========== Refutation ==========",
       
   464   "Here is a proof"];
       
   465 
       
   466 val end_proof_strings = ["# SZS output end CNFRefutation",
       
   467   "======= End of refutation =======",
       
   468   "Formulae used in the proof"];
       
   469 
       
   470 fun get_proof_extract proof =
       
   471   let
       
   472     (*splits to_split by the first possible of a list of splitters*)
       
   473     val (begin_string, end_string) =
       
   474       (find_first (fn s => String.isSubstring s proof) begin_proof_strings,
       
   475       find_first (fn s => String.isSubstring s proof) end_proof_strings)
       
   476   in
       
   477     if is_none begin_string orelse is_none end_string
       
   478     then error "Could not extract proof (no substring indicating a proof)"
       
   479     else proof |> first_field (the begin_string) |> the |> snd
       
   480                |> first_field (the end_string) |> the |> fst
       
   481   end;
       
   482 
       
   483 (* ==== CHECK IF PROOF OF E OR VAMPIRE WAS SUCCESSFUL === *)
       
   484 
       
   485 val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable",
       
   486   "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"];
       
   487 val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"];
       
   488 val failure_strings_SPASS = ["SPASS beiseite: Completion found.",
       
   489   "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."];
       
   490 val failure_strings_remote = ["Remote-script could not extract proof"];
       
   491 fun find_failure proof =
       
   492   let val failures =
       
   493     map_filter (fn s => if String.isSubstring s proof then SOME s else NONE)
       
   494       (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote)
       
   495   val correct = null failures andalso
       
   496     exists (fn s => String.isSubstring s proof) begin_proof_strings andalso
       
   497     exists (fn s => String.isSubstring s proof) end_proof_strings
       
   498   in
       
   499     if correct then NONE
       
   500     else if null failures then SOME "Output of ATP not in proper format"
       
   501     else SOME (hd failures) end;
       
   502 
       
   503 (* === EXTRACTING LEMMAS === *)
       
   504 (* lines have the form "cnf(108, axiom, ...",
       
   505 the number (108) has to be extracted)*)
       
   506 fun get_step_nums false proofextract =
       
   507   let val toks = String.tokens (not o Char.isAlphaNum)
       
   508   fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok
       
   509     | inputno ("cnf"::ntok::"negated"::"conjecture"::_) = Int.fromString ntok
       
   510     | inputno _ = NONE
       
   511   val lines = split_lines proofextract
       
   512   in  map_filter (inputno o toks) lines  end
       
   513 (*String contains multiple lines. We want those of the form
       
   514   "253[0:Inp] et cetera..."
       
   515   A list consisting of the first number in each line is returned. *)
       
   516 |  get_step_nums true proofextract =
       
   517   let val toks = String.tokens (not o Char.isAlphaNum)
       
   518   fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok
       
   519     | inputno _ = NONE
       
   520   val lines = split_lines proofextract
       
   521   in  map_filter (inputno o toks) lines  end
       
   522   
       
   523 (*extracting lemmas from tstp-output between the lines from above*)
       
   524 fun extract_lemmas get_step_nums (proof, thm_names, conj_count, _, _, _) =
       
   525   let
       
   526   (* get the names of axioms from their numbers*)
       
   527   fun get_axiom_names thm_names step_nums =
       
   528     let
       
   529     val last_axiom = Vector.length thm_names
       
   530     fun is_axiom n = n <= last_axiom
       
   531     fun is_conj n = n >= fst conj_count andalso n < fst conj_count + snd conj_count
       
   532     fun getname i = Vector.sub(thm_names, i-1)
       
   533     in
       
   534       (sort_distinct string_ord (filter (fn x => x <> "??.unknown")
       
   535         (map getname (filter is_axiom step_nums))),
       
   536       exists is_conj step_nums)
       
   537     end
       
   538   val proofextract = get_proof_extract proof
       
   539   in
       
   540     get_axiom_names thm_names (get_step_nums proofextract)
       
   541   end;
       
   542 
       
   543 (*Used to label theorems chained into the sledgehammer call*)
       
   544 val chained_hint = "CHAINED";
       
   545 val nochained = filter_out (fn y => y = chained_hint)
       
   546   
       
   547 (* metis-command *)
       
   548 fun metis_line [] = "apply metis"
       
   549   | metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")"
       
   550 
       
   551 (* atp_minimize [atp=<prover>] <lemmas> *)
       
   552 fun minimize_line _ [] = ""
       
   553   | minimize_line name lemmas = "For minimizing the number of lemmas try this command:\n" ^
       
   554         (Markup.markup Markup.sendback) ("atp_minimize [atp=" ^ name ^ "] " ^
       
   555                                          space_implode " " (nochained lemmas))
       
   556 
       
   557 fun sendback_metis_nochained lemmas =
       
   558   (Markup.markup Markup.sendback o metis_line) (nochained lemmas)
       
   559 
       
   560 fun lemma_list dfg name result =
       
   561   let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result
       
   562   in (sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^
       
   563     (if used_conj then ""
       
   564      else "\nWarning: Goal is provable because context is inconsistent."),
       
   565      nochained lemmas)
       
   566   end;
       
   567 
       
   568 (* === Extracting structured Isar-proof === *)
       
   569 fun structured_proof name (result as (proof, thm_names, conj_count, ctxt, goal, subgoalno)) =
       
   570   let
       
   571   (*Could use split_lines, but it can return blank lines...*)
       
   572   val lines = String.tokens (equal #"\n");
       
   573   val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c)
       
   574   val proofextract = get_proof_extract proof
       
   575   val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
       
   576   val (one_line_proof, lemma_names) = lemma_list false name result
       
   577   val structured =
       
   578     if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then ""
       
   579     else decode_tstp_file cnfs ctxt goal subgoalno thm_names
       
   580   in
       
   581   (one_line_proof ^ "\n\n" ^ Markup.markup Markup.sendback structured, lemma_names)
       
   582 end
       
   583 
       
   584 end;