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