src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
author blanchet
Sun, 15 Dec 2013 18:01:38 +0100
changeset 54753 688da3388b2d
parent 54751 9b93f9117f8b
child 54754 6b0ca7f79e93
permissions -rw-r--r--
use simplifier for rewrite

(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
    Author:     Jasmin Blanchette, TU Muenchen
    Author:     Steffen Juilf Smolka, TU Muenchen

Isar proof reconstruction from ATP proofs.
*)

signature SLEDGEHAMMER_RECONSTRUCT =
sig
  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
  type 'a atp_proof = 'a ATP_Proof.atp_proof
  type stature = ATP_Problem_Generate.stature
  type one_line_params = Sledgehammer_Reconstructor.one_line_params

  type isar_params =
    bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list *
    thm

  val isar_proof_text :
    Proof.context -> bool option -> isar_params -> one_line_params -> string
  val proof_text :
    Proof.context -> bool option -> (unit -> isar_params) -> int -> one_line_params -> string
end;

structure Sledgehammer_Reconstruct : SLEDGEHAMMER_RECONSTRUCT =
struct

open ATP_Util
open ATP_Problem
open ATP_Proof
open ATP_Problem_Generate
open ATP_Proof_Reconstruct
open Sledgehammer_Util
open Sledgehammer_Reconstructor
open Sledgehammer_Proof
open Sledgehammer_Annotate
open Sledgehammer_Print
open Sledgehammer_Preplay
open Sledgehammer_Compress
open Sledgehammer_Try0
open Sledgehammer_Minimize_Isar

structure String_Redirect = ATP_Proof_Redirect(
  type key = atp_step_name
  val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s')
  val string_of = fst)

open String_Redirect

fun maybe_mk_Trueprop t = t |> fastype_of t = HOLogic.boolT ? HOLogic.mk_Trueprop
val maybe_dest_Trueprop = perhaps (try HOLogic.dest_Trueprop)

val e_skolemize_rule = "skolemize"
val vampire_skolemisation_rule = "skolemisation"
(* TODO: Use "Z3_Proof.string_of_rule" once it is moved to Isabelle *)
val z3_apply_def_rule = "apply-def"
val z3_hypothesis_rule = "hypothesis"
val z3_intro_def_rule = "intro-def"
val z3_lemma_rule = "lemma"
val z3_skolemize_rule = "sk"
val z3_th_lemma_rule = "th-lemma"

val is_skolemize_rule =
  member (op =) [e_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule]

fun raw_label_of_num num = (num, 0)

fun label_of_clause [(num, _)] = raw_label_of_num num
  | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0)

fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss)
  | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names))

fun replace_one_dependency (old, new) dep =
  if is_same_atp_step dep old then new else [dep]
fun replace_dependencies_in_line p (name, role, t, rule, deps) =
  (name, role, t, rule, fold (union (op =) o replace_one_dependency p) deps [])

fun inline_z3_defs _ [] = []
  | inline_z3_defs defs ((line as (name, role, t, rule, deps)) :: lines) =
    if rule = z3_intro_def_rule then
      let val def = t |> maybe_dest_Trueprop |> HOLogic.dest_eq |> swap in
        inline_z3_defs (insert (op =) def defs)
          (map (replace_dependencies_in_line (name, [])) lines)
      end
    else if rule = z3_apply_def_rule then
      inline_z3_defs defs (map (replace_dependencies_in_line (name, [])) lines)
    else
      (name, role, Term.subst_atomic defs t, rule, deps) :: inline_z3_defs defs lines

fun alist_cons_list eq (k, v) = AList.map_default eq (k, []) (cons v)

(* FIXME: use "prop_of_clause" defined below *)
fun add_z3_hypotheses [] = I
  | add_z3_hypotheses hyps =
    HOLogic.dest_Trueprop
    #> curry HOLogic.mk_imp (Library.foldr1 HOLogic.mk_conj (map HOLogic.dest_Trueprop hyps))
    #> HOLogic.mk_Trueprop

fun inline_z3_hypotheses _ _ [] = []
  | inline_z3_hypotheses hyp_names hyps ((name, role, t, rule, deps) :: lines) =
    if rule = z3_hypothesis_rule then
      inline_z3_hypotheses (name :: hyp_names) (alist_cons_list (op =) (t, name) hyps) lines
    else
      let val deps' = subtract (op =) hyp_names deps in
        if rule = z3_lemma_rule then
          (name, role, t, rule, deps') :: inline_z3_hypotheses hyp_names hyps lines
        else
          let
            val add_hyps = filter_out (null o inter (op =) deps o snd) hyps
            val t' = add_z3_hypotheses (map fst add_hyps) t
            val deps' = subtract (op =) hyp_names deps
            val hyps' = fold (AList.update (op =) o apsnd (insert (op =) name)) add_hyps hyps
          in
            (name, role, t', rule, deps') :: inline_z3_hypotheses hyp_names hyps' lines
          end
      end

(* No "real" literals means only type information (tfree_tcs, clsrel, or
   clsarity). *)
fun is_only_type_information t = t aconv @{term True}

(* Discard facts; consolidate adjacent lines that prove the same formula, since
   they differ only in type information.*)
fun add_line (line as (name as (_, ss), role, t, rule, [])) lines =
    (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire)
       internal facts or definitions. *)
    if role = Lemma orelse role = Conjecture orelse role = Negated_Conjecture orelse
       role = Hypothesis then
      line :: lines
    else if role = Axiom then
      (* Facts are not proof lines. *)
      lines |> is_only_type_information t ? map (replace_dependencies_in_line (name, []))
    else
      map (replace_dependencies_in_line (name, [])) lines
  | add_line line lines = line :: lines

(* Recursively delete empty lines (type information) from the proof. *)
fun add_nontrivial_line (line as (name, _, t, _, [])) lines =
    if is_only_type_information t then delete_dependency name lines else line :: lines
  | add_nontrivial_line line lines = line :: lines
and delete_dependency name lines =
  fold_rev add_nontrivial_line (map (replace_dependencies_in_line (name, [])) lines) []

fun add_desired_lines res [] = rev res
  | add_desired_lines res ((name as (_, ss), role, t, rule, deps) :: lines) =
    if role <> Plain orelse is_skolemize_rule rule orelse
       (* the last line must be kept *)
       null lines orelse
       (not (is_only_type_information t) andalso null (Term.add_tvars t [])
        andalso length deps >= 2 andalso
        (* don't keep next to last line, which usually results in a trivial step *)
        not (can the_single lines)) then
      add_desired_lines ((name, role, t, rule, deps) :: res) lines
    else
      add_desired_lines res (map (replace_dependencies_in_line (name, deps)) lines)

val add_labels_of_proof =
  steps_of_proof
  #> fold_isar_steps (byline_of_step #> (fn SOME ((ls, _), _) => union (op =) ls | _ => I))

fun kill_useless_labels_in_proof proof =
  let
    val used_ls = add_labels_of_proof proof []
    fun kill_label l = if member (op =) used_ls l then l else no_label
    fun kill_assms assms = map (apfst kill_label) assms
    fun kill_step (Prove (qs, xs, l, t, subproofs, by)) =
        Prove (qs, xs, kill_label l, t, map kill_proof subproofs, by)
      | kill_step step = step
    and kill_proof (Proof (fix, assms, steps)) =
      Proof (fix, kill_assms assms, map kill_step steps)
  in
    kill_proof proof
  end

fun prefix_of_depth n = replicate_string (n + 1)

val assume_prefix = "a"
val have_prefix = "f"

val relabel_proof =
  let
    fun fresh_label depth prefix (old as (l, subst, next)) =
      if l = no_label then
        old
      else
        let val l' = (prefix_of_depth depth prefix, next) in
          (l', (l, l') :: subst, next + 1)
        end

    fun relabel_facts subst = apfst (maps (the_list o AList.lookup (op =) subst))

    fun relabel_assm depth (l, t) (subst, next) =
      let val (l, subst, next) = (l, subst, next) |> fresh_label depth assume_prefix in
        ((l, t), (subst, next))
      end

    fun relabel_assms subst depth assms = fold_map (relabel_assm depth) assms (subst, 1) ||> fst

    fun relabel_steps _ _ _ [] = []
      | relabel_steps subst depth next (Prove (qs, xs, l, t, sub, by) :: steps) =
        let
          val (l, subst, next) = (l, subst, next) |> fresh_label depth have_prefix
          val sub = relabel_proofs subst depth sub
          val by = by |> relabel_byline subst
        in
          Prove (qs, xs, l, t, sub, by) :: relabel_steps subst depth next steps
        end
      | relabel_steps subst depth next (step :: steps) =
        step :: relabel_steps subst depth next steps
    and relabel_proof subst depth (Proof (fix, assms, steps)) =
      let val (assms, subst) = relabel_assms subst depth assms in
        Proof (fix, assms, relabel_steps subst depth 1 steps)
      end
    and relabel_byline subst byline = apfst (relabel_facts subst) byline
    and relabel_proofs subst depth = map (relabel_proof subst (depth + 1))
  in
    relabel_proof [] 0
  end

val chain_direct_proof =
  let
    fun chain_qs_lfs NONE lfs = ([], lfs)
      | chain_qs_lfs (SOME l0) lfs =
        if member (op =) lfs l0 then ([Then], lfs |> remove (op =) l0) else ([], lfs)
    fun chain_step lbl (Prove (qs, xs, l, t, subproofs, ((lfs, gfs), method))) =
        let val (qs', lfs) = chain_qs_lfs lbl lfs in
          Prove (qs' @ qs, xs, l, t, chain_proofs subproofs, ((lfs, gfs), method))
        end
      | chain_step _ step = step
    and chain_steps _ [] = []
      | chain_steps (prev as SOME _) (i :: is) =
        chain_step prev i :: chain_steps (label_of_step i) is
      | chain_steps _ (i :: is) = i :: chain_steps (label_of_step i) is
    and chain_proof (Proof (fix, assms, steps)) =
      Proof (fix, assms, chain_steps (try (List.last #> fst) assms) steps)
    and chain_proofs proofs = map (chain_proof) proofs
  in chain_proof end

type isar_params =
  bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list *
  thm

fun isar_proof_text ctxt isar_proofs
    (debug, verbose, metis_type_enc, metis_lam_trans, preplay_timeout, isar_compress,
     isar_try0, atp_proof, goal)
    (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
  let
    val (params, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt
    val (_, ctxt) =
      params
      |> map (fn (s, T) => (Binding.name s, SOME T, NoSyn))
      |> (fn fixes => ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes)
    val one_line_proof = one_line_proof_text 0 one_line_params
    val do_preplay = preplay_timeout <> SOME Time.zeroTime

    val is_fixed = Variable.is_declared ctxt orf can Name.dest_skolem
    fun skolems_of t = Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev

    fun get_role keep_role ((num, _), role, t, rule, _) =
      if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE

    fun isar_proof_of () =
      let
        val atp_proof =
          atp_proof
          |> inline_z3_defs []
          |> inline_z3_hypotheses [] []
          |> rpair [] |-> fold_rev add_line
          |> rpair [] |-> fold_rev add_nontrivial_line
          |> add_desired_lines []

        val conjs =
          map_filter (fn (name, role, _, _, _) =>
              if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE)
            atp_proof
        val assms = map_filter (Option.map fst o get_role (curry (op =) Hypothesis)) atp_proof
        val lems =
          map_filter (get_role (curry (op =) Lemma)) atp_proof
          |> map (fn ((l, t), rule) =>
            let
              val (skos, meth) =
                if is_skolemize_rule rule then (skolems_of t, MetisM)
                else if rule = z3_th_lemma_rule then ([], ArithM)
                else ([], SimpM)
            in
              Prove ([], skos, l, maybe_mk_Trueprop t, [], (([], []), meth))
            end)

        val bot = atp_proof |> List.last |> #1

        val refute_graph =
          atp_proof
          |> map (fn (name, _, _, _, from) => (from, name))
          |> make_refute_graph bot
          |> fold (Atom_Graph.default_node o rpair ()) conjs

        val axioms = axioms_of_refute_graph refute_graph conjs

        val tainted = tainted_atoms_of_refute_graph refute_graph conjs
        val is_clause_tainted = exists (member (op =) tainted)
        val steps =
          Symtab.empty
          |> fold (fn (name as (s, _), role, t, rule, _) =>
                      Symtab.update_new (s, (rule,
                        t |> (if is_clause_tainted [name] then
                                role <> Conjecture ? (maybe_dest_Trueprop #> s_not)
                                #> fold exists_of (map Var (Term.add_vars t []))
                              else
                                I))))
                  atp_proof

        fun is_clause_skolemize_rule [(s, _)] =
            Option.map (is_skolemize_rule o fst) (Symtab.lookup steps s) = SOME true
          | is_clause_skolemize_rule _ = false

        (* The assumptions and conjecture are "prop"s; the other formulas are "bool"s (for ATPs) or
           "prop"s (for Z3). TODO: Always use "prop". *)
        fun prop_of_clause [(num, _)] =
            Symtab.lookup steps num |> the |> snd |> maybe_mk_Trueprop |> close_form
          | prop_of_clause names =
            let
              val lits =
                map (maybe_dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names)
            in
              case List.partition (can HOLogic.dest_not) lits of
                (negs as _ :: _, pos as _ :: _) =>
                s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos)
              | _ => fold (curry s_disj) lits @{term False}
            end
            |> HOLogic.mk_Trueprop |> close_form

        fun maybe_show outer c =
          (outer andalso length c = 1 andalso subset (op =) (c, conjs)) ? cons Show

        fun isar_steps outer predecessor accum [] =
            accum
            |> (if tainted = [] then
                  cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [],
                               (([the predecessor], []), MetisM)))
                else
                  I)
            |> rev
          | isar_steps outer _ accum (Have (gamma, c) :: infs) =
            let
              val l = label_of_clause c
              val t = prop_of_clause c
              val by = (fold add_fact_of_dependencies gamma no_facts, MetisM)
              fun prove sub by = Prove (maybe_show outer c [], [], l, t, sub, by)
              fun do_rest l step = isar_steps outer (SOME l) (step :: accum) infs
            in
              if is_clause_tainted c then
                (case gamma of
                  [g] =>
                  if is_clause_skolemize_rule g andalso is_clause_tainted g then
                    let val subproof = Proof (skolems_of (prop_of_clause g), [], rev accum) in
                      isar_steps outer (SOME l) [prove [subproof] (no_facts, MetisM)] []
                    end
                  else
                    do_rest l (prove [] by)
                | _ => do_rest l (prove [] by))
              else
                (if is_clause_skolemize_rule c then Prove ([], skolems_of t, l, t, [], by)
                 else prove [] by)
                |> do_rest l
            end
          | isar_steps outer predecessor accum (Cases cases :: infs) =
            let
              fun isar_case (c, infs) =
                isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] infs
              val c = succedent_of_cases cases
              val l = label_of_clause c
              val t = prop_of_clause c
              val step =
                Prove (maybe_show outer c [], [], l, t, map isar_case cases,
                  ((the_list predecessor, []), MetisM))
            in
              isar_steps outer (SOME l) (step :: accum) infs
            end
        and isar_proof outer fix assms lems infs =
          Proof (fix, assms, lems @ isar_steps outer NONE [] infs)

        val isar_proof_of_direct_proof = isar_proof true params assms lems

        (* 60 seconds seems like a good interpreation of "no timeout" *)
        val preplay_timeout = preplay_timeout |> the_default (seconds 60.0)

        val clean_up_labels_in_proof =
          chain_direct_proof
          #> kill_useless_labels_in_proof
          #> relabel_proof
        val (preplay_interface as {overall_preplay_stats, ...}, isar_proof) =
          refute_graph
(*
          |> tap (tracing o prefix "Refute graph: " o string_of_refute_graph)
*)
          |> redirect_graph axioms tainted bot
(*
          |> tap (tracing o prefix "Direct proof: " o string_of_direct_proof)
*)
          |> isar_proof_of_direct_proof
          |> postprocess_remove_unreferenced_steps I
          |> relabel_proof_canonically
          |> `(proof_preplay_interface debug ctxt metis_type_enc metis_lam_trans do_preplay
               preplay_timeout)
        val ((preplay_time, preplay_fail), isar_proof) =
          isar_proof
          |> compress_proof (if isar_proofs = SOME true then isar_compress else 1000.0)
               preplay_interface
          |> isar_try0 ? try0 preplay_timeout preplay_interface
          |> postprocess_remove_unreferenced_steps (isar_try0 ? min_deps_of_step preplay_interface)
          |> `overall_preplay_stats
          ||> clean_up_labels_in_proof
        val isar_text =
          string_of_proof ctxt metis_type_enc metis_lam_trans subgoal subgoal_count isar_proof
      in
        case isar_text of
          "" =>
          if isar_proofs = SOME true then
            "\nNo structured proof available (proof too simple)."
          else
            ""
        | _ =>
          let
            val msg =
              (if verbose then
                let
                  val num_steps = add_proof_steps (steps_of_proof isar_proof) 0
                in [string_of_int num_steps ^ " step" ^ plural_s num_steps] end
               else
                 []) @
              (if do_preplay then
                [(if preplay_fail then "may fail, " else "") ^
                   string_of_preplay_time preplay_time]
               else
                 [])
          in
            "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
            Active.sendback_markup [Markup.padding_command] isar_text
          end
      end
    val isar_proof =
      if debug then
        isar_proof_of ()
      else case try isar_proof_of () of
        SOME s => s
      | NONE => if isar_proofs = SOME true then
                  "\nWarning: The Isar proof construction failed."
                else
                  ""
  in one_line_proof ^ isar_proof end

fun isar_proof_would_be_a_good_idea preplay =
  case preplay of
    Played (reconstr, _) => reconstr = SMT
  | Trust_Playable _ => false
  | Failed_to_Play _ => true

fun proof_text ctxt isar_proofs isar_params num_chained
               (one_line_params as (preplay, _, _, _, _, _)) =
  (if isar_proofs = SOME true orelse
      (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then
     isar_proof_text ctxt isar_proofs (isar_params ())
   else
     one_line_proof_text num_chained) one_line_params

end;