src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
author blanchet
Fri, 14 Mar 2014 11:05:37 +0100
changeset 57467 e03c0f6ad1b6
parent 57439 8e7a9ad44e14
child 57468 fc937e7ef4c6
permissions -rw-r--r--
tuning
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_isar.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Author:     Steffen Juilf Smolka, TU Muenchen
     4 
     5 Isar proof reconstruction from ATP proofs.
     6 *)
     7 
     8 signature SLEDGEHAMMER_ISAR =
     9 sig
    10   type atp_step_name = ATP_Proof.atp_step_name
    11   type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
    12   type 'a atp_proof = 'a ATP_Proof.atp_proof
    13   type stature = ATP_Problem_Generate.stature
    14   type one_line_params = Sledgehammer_Proof_Methods.one_line_params
    15 
    16   val trace : bool Config.T
    17 
    18   type isar_params =
    19     bool * (string option * string option) * Time.time * real * bool * bool
    20     * (term, string) atp_step list * thm
    21 
    22   val proof_text : Proof.context -> bool -> bool option -> bool option -> (unit -> isar_params) ->
    23     int -> one_line_params -> string
    24 end;
    25 
    26 structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =
    27 struct
    28 
    29 open ATP_Util
    30 open ATP_Problem
    31 open ATP_Proof
    32 open ATP_Proof_Reconstruct
    33 open Sledgehammer_Util
    34 open Sledgehammer_Proof_Methods
    35 open Sledgehammer_Isar_Proof
    36 open Sledgehammer_Isar_Preplay
    37 open Sledgehammer_Isar_Compress
    38 open Sledgehammer_Isar_Minimize
    39 
    40 structure String_Redirect = ATP_Proof_Redirect(
    41   type key = atp_step_name
    42   val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s')
    43   val string_of = fst)
    44 
    45 open String_Redirect
    46 
    47 val trace = Attrib.setup_config_bool @{binding sledgehammer_isar_trace} (K false)
    48 
    49 val e_skolemize_rules = ["skolemize", "shift_quantors"]
    50 val spass_pirate_datatype_rule = "DT"
    51 val vampire_skolemisation_rule = "skolemisation"
    52 (* TODO: Use "Z3_Proof.string_of_rule" once it is moved to Isabelle *)
    53 val z3_skolemize_rule = "sk"
    54 val z3_th_lemma_rule = "th-lemma"
    55 
    56 val skolemize_rules =
    57   e_skolemize_rules @ [spass_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule]
    58 
    59 val is_skolemize_rule = member (op =) skolemize_rules
    60 val is_arith_rule = String.isPrefix z3_th_lemma_rule
    61 val is_datatype_rule = String.isPrefix spass_pirate_datatype_rule
    62 
    63 fun raw_label_of_num num = (num, 0)
    64 
    65 fun label_of_clause [(num, _)] = raw_label_of_num num
    66   | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0)
    67 
    68 fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss)
    69   | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names))
    70 
    71 (* No "real" literals means only type information (tfree_tcs, clsrel, or clsarity). *)
    72 fun is_only_type_information t = t aconv @{prop True}
    73 
    74 (* Discard facts; consolidate adjacent lines that prove the same formula, since they differ only in
    75    type information. *)
    76 fun add_line_pass1 (line as (name, role, t, rule, [])) lines =
    77     (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or
    78        definitions. *)
    79     if role = Lemma orelse role = Conjecture orelse role = Negated_Conjecture orelse
    80        role = Hypothesis orelse is_arith_rule rule then
    81       line :: lines
    82     else if role = Axiom then
    83       (* Facts are not proof lines. *)
    84       lines |> is_only_type_information t ? map (replace_dependencies_in_line (name, []))
    85     else
    86       map (replace_dependencies_in_line (name, [])) lines
    87   | add_line_pass1 line lines = line :: lines
    88 
    89 fun add_lines_pass2 res [] = rev res
    90   | add_lines_pass2 res ((line as (name, role, t, rule, deps)) :: lines) =
    91     let
    92       val is_last_line = null lines
    93 
    94       fun looks_interesting () =
    95         not (is_only_type_information t) andalso null (Term.add_tvars t [])
    96         andalso length deps >= 2 andalso not (can the_single lines)
    97 
    98       fun is_skolemizing_line (_, _, _, rule', deps') =
    99         is_skolemize_rule rule' andalso member (op =) deps' name
   100       fun is_before_skolemize_rule () = exists is_skolemizing_line lines
   101     in
   102       if role <> Plain orelse is_skolemize_rule rule orelse is_arith_rule rule orelse
   103          is_datatype_rule rule orelse is_last_line orelse looks_interesting () orelse
   104          is_before_skolemize_rule () then
   105         add_lines_pass2 (line :: res) lines
   106       else
   107         add_lines_pass2 res (map (replace_dependencies_in_line (name, deps)) lines)
   108     end
   109 
   110 type isar_params =
   111   bool * (string option * string option) * Time.time * real * bool * bool
   112   * (term, string) atp_step list * thm
   113 
   114 val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method]
   115 val simp_based_methods = [Auto_Method, Simp_Method, Fastforce_Method, Force_Method]
   116 val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method]
   117 
   118 val arith_methods = basic_arith_methods @ simp_based_methods @ basic_systematic_methods
   119 val datatype_methods = [Simp_Method, Simp_Size_Method]
   120 val systematic_methods0 = basic_systematic_methods @ basic_arith_methods @ simp_based_methods @
   121   [Metis_Method (SOME no_typesN, NONE)]
   122 val rewrite_methods = simp_based_methods @ basic_systematic_methods @ basic_arith_methods
   123 val skolem_methods = basic_systematic_methods
   124 
   125 fun isar_proof_text ctxt debug isar_proofs smt_proofs isar_params
   126     (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
   127   let
   128     val _ = if debug then Output.urgent_message "Constructing Isar proof..." else ()
   129 
   130     fun isar_proof_of () =
   131       let
   132         val SOME (verbose, alt_metis_args, preplay_timeout, compress_isar, try0_isar, minimize,
   133           atp_proof, goal) = try isar_params ()
   134 
   135         val systematic_methods = insert (op =) (Metis_Method alt_metis_args) systematic_methods0
   136 
   137         fun massage_methods (meths as meth :: _) =
   138           if not try0_isar then [meth]
   139           else if smt_proofs = SOME true then SMT2_Method :: meths
   140           else meths
   141 
   142         val (params, _, concl_t) = strip_subgoal goal subgoal ctxt
   143         val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params
   144         val ctxt = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd
   145 
   146         val do_preplay = preplay_timeout <> Time.zeroTime
   147         val compress_isar = if isar_proofs = NONE andalso do_preplay then 1000.0 else compress_isar
   148 
   149         val is_fixed = Variable.is_declared ctxt orf Name.is_skolem
   150         fun skolems_of t = Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev
   151 
   152         fun get_role keep_role ((num, _), role, t, rule, _) =
   153           if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE
   154 
   155         val atp_proof =
   156           atp_proof
   157           |> rpair [] |-> fold_rev add_line_pass1
   158           |> add_lines_pass2 []
   159 
   160         val conjs =
   161           map_filter (fn (name, role, _, _, _) =>
   162               if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE)
   163             atp_proof
   164         val assms = map_filter (Option.map fst o get_role (curry (op =) Hypothesis)) atp_proof
   165         val lems =
   166           map_filter (get_role (curry (op =) Lemma)) atp_proof
   167           |> map (fn ((l, t), rule) =>
   168             let
   169               val (skos, meths) =
   170                 (if is_skolemize_rule rule then (skolems_of t, skolem_methods)
   171                  else if is_arith_rule rule then ([], arith_methods)
   172                  else ([], rewrite_methods))
   173                 ||> massage_methods
   174             in
   175               Prove ([], skos, l, t, [], ([], []), meths, "")
   176             end)
   177 
   178         val bot = atp_proof |> List.last |> #1
   179 
   180         val refute_graph =
   181           atp_proof
   182           |> map (fn (name, _, _, _, from) => (from, name))
   183           |> make_refute_graph bot
   184           |> fold (Atom_Graph.default_node o rpair ()) conjs
   185 
   186         val axioms = axioms_of_refute_graph refute_graph conjs
   187 
   188         val tainted = tainted_atoms_of_refute_graph refute_graph conjs
   189         val is_clause_tainted = exists (member (op =) tainted)
   190         val steps =
   191           Symtab.empty
   192           |> fold (fn (name as (s, _), role, t, rule, _) =>
   193               Symtab.update_new (s, (rule, t
   194                 |> (if is_clause_tainted [name] then
   195                       HOLogic.dest_Trueprop
   196                       #> role <> Conjecture ? s_not
   197                       #> fold exists_of (map Var (Term.add_vars t []))
   198                       #> HOLogic.mk_Trueprop
   199                     else
   200                       I))))
   201             atp_proof
   202 
   203         val rule_of_clause_id = fst o the o Symtab.lookup steps o fst
   204 
   205         fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> close_form
   206           | prop_of_clause names =
   207             let
   208               val lits = map (HOLogic.dest_Trueprop o snd)
   209                 (map_filter (Symtab.lookup steps o fst) names)
   210             in
   211               (case List.partition (can HOLogic.dest_not) lits of
   212                 (negs as _ :: _, pos as _ :: _) =>
   213                 s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos)
   214               | _ => fold (curry s_disj) lits @{term False})
   215             end
   216             |> HOLogic.mk_Trueprop |> close_form
   217 
   218         fun maybe_show outer c = (outer andalso eq_set (op =) (c, conjs)) ? cons Show
   219 
   220         fun isar_steps outer predecessor accum [] =
   221             accum
   222             |> (if tainted = [] then
   223                   cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [],
   224                     (the_list predecessor, []), massage_methods systematic_methods, ""))
   225                 else
   226                   I)
   227             |> rev
   228           | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) =
   229             let
   230               val l = label_of_clause c
   231               val t = prop_of_clause c
   232               val rule = rule_of_clause_id id
   233               val skolem = is_skolemize_rule rule
   234 
   235               val deps = fold add_fact_of_dependencies gamma ([], [])
   236               val meths =
   237                 (if skolem then skolem_methods
   238                  else if is_arith_rule rule then arith_methods
   239                  else if is_datatype_rule rule then datatype_methods
   240                  else systematic_methods)
   241                 |> massage_methods
   242 
   243               fun prove sub facts = Prove (maybe_show outer c [], [], l, t, sub, facts, meths, "")
   244               fun steps_of_rest step = isar_steps outer (SOME l) (step :: accum) infs
   245             in
   246               if is_clause_tainted c then
   247                 (case gamma of
   248                   [g] =>
   249                   if skolem andalso is_clause_tainted g then
   250                     let val subproof = Proof (skolems_of (prop_of_clause g), [], rev accum) in
   251                       isar_steps outer (SOME l) [prove [subproof] ([], [])] infs
   252                     end
   253                   else
   254                     steps_of_rest (prove [] deps)
   255                 | _ => steps_of_rest (prove [] deps))
   256               else
   257                 steps_of_rest (if skolem then Prove ([], skolems_of t, l, t, [], deps, meths, "")
   258                   else prove [] deps)
   259             end
   260           | isar_steps outer predecessor accum (Cases cases :: infs) =
   261             let
   262               fun isar_case (c, subinfs) =
   263                 isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] subinfs
   264               val c = succedent_of_cases cases
   265               val l = label_of_clause c
   266               val t = prop_of_clause c
   267               val step =
   268                 Prove (maybe_show outer c [], [], l, t,
   269                   map isar_case (filter_out (null o snd) cases),
   270                   (the_list predecessor, []), massage_methods systematic_methods, "")
   271             in
   272               isar_steps outer (SOME l) (step :: accum) infs
   273             end
   274         and isar_proof outer fix assms lems infs =
   275           Proof (fix, assms, lems @ isar_steps outer NONE [] infs)
   276 
   277         val trace = Config.get ctxt trace
   278 
   279         val canonical_isar_proof =
   280           refute_graph
   281           |> trace ? tap (tracing o prefix "Refute graph: " o string_of_refute_graph)
   282           |> redirect_graph axioms tainted bot
   283           |> trace ? tap (tracing o prefix "Direct proof: " o string_of_direct_proof)
   284           |> isar_proof true params assms lems
   285           |> postprocess_isar_proof_remove_unreferenced_steps I
   286           |> relabel_isar_proof_canonically
   287 
   288         val ctxt = ctxt |> enrich_context_with_local_facts canonical_isar_proof
   289 
   290         val preplay_data = Unsynchronized.ref Canonical_Label_Tab.empty
   291 
   292         val _ = fold_isar_steps (fn meth =>
   293             K (set_preplay_outcomes_of_isar_step ctxt debug preplay_timeout preplay_data meth []))
   294           (steps_of_isar_proof canonical_isar_proof) ()
   295 
   296         fun str_of_preplay_outcome outcome =
   297           if Lazy.is_finished outcome then string_of_play_outcome (Lazy.force outcome) else "?"
   298         fun str_of_meth l meth =
   299           string_of_proof_method meth ^ " " ^
   300           str_of_preplay_outcome (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)
   301         fun comment_of l = map (str_of_meth l) #> commas
   302 
   303         fun trace_isar_proof label proof =
   304           if trace then
   305             tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^
   306               string_of_isar_proof ctxt subgoal subgoal_count
   307                 (comment_isar_proof comment_of proof) ^ "\n")
   308           else
   309             ()
   310 
   311         fun comment_of l (meth :: _) =
   312           (case (verbose,
   313               Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)) of
   314             (false, Played _) => ""
   315           | (_, outcome) => string_of_play_outcome outcome)
   316 
   317         val (play_outcome, isar_proof) =
   318           canonical_isar_proof
   319           |> tap (trace_isar_proof "Original")
   320           |> compress_isar_proof ctxt debug compress_isar preplay_timeout preplay_data
   321           |> tap (trace_isar_proof "Compressed")
   322           |> postprocess_isar_proof_remove_unreferenced_steps
   323                (keep_fastest_method_of_isar_step (!preplay_data)
   324                 #> minimize ? minimize_isar_step_dependencies ctxt debug preplay_data)
   325           |> tap (trace_isar_proof "Minimized")
   326           (* It's not clear whether this is worth the trouble (and if so, "isar_compress" has an
   327              unnatural semantics): *)
   328 (*
   329           |> minimize
   330                ? (compress_isar_proof ctxt compress_isar preplay_timeout preplay_data
   331                   #> tap (trace_isar_proof "Compressed again"))
   332 *)
   333           |> `(preplay_outcome_of_isar_proof (!preplay_data))
   334           ||> (comment_isar_proof comment_of
   335                #> chain_isar_proof
   336                #> kill_useless_labels_in_isar_proof
   337                #> relabel_isar_proof_nicely)
   338       in
   339         (case string_of_isar_proof ctxt subgoal subgoal_count isar_proof of
   340           "" =>
   341           if isar_proofs = SOME true then "\nNo structured proof available (proof too simple)."
   342           else ""
   343         | isar_text =>
   344           let
   345             val msg =
   346               (if verbose then
   347                  let val num_steps = add_isar_steps (steps_of_isar_proof isar_proof) 0 in
   348                    [string_of_int num_steps ^ " step" ^ plural_s num_steps]
   349                  end
   350                else
   351                  []) @
   352               (if do_preplay then [string_of_play_outcome play_outcome] else [])
   353           in
   354             "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
   355             Active.sendback_markup [Markup.padding_command] isar_text
   356           end)
   357       end
   358 
   359     val one_line_proof = one_line_proof_text 0 one_line_params
   360     val isar_proof =
   361       if debug then
   362         isar_proof_of ()
   363       else
   364         (case try isar_proof_of () of
   365           SOME s => s
   366         | NONE =>
   367           if isar_proofs = SOME true then "\nWarning: Isar proof construction failed." else "")
   368   in one_line_proof ^ isar_proof end
   369 
   370 fun isar_proof_would_be_a_good_idea smt_proofs (meth, play) =
   371   (case play of
   372     Played _ => meth = SMT2_Method andalso smt_proofs <> SOME true
   373   | Play_Timed_Out time => Time.> (time, Time.zeroTime)
   374   | Play_Failed => true)
   375 
   376 fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained
   377     (one_line_params as (preplay, _, _, _, _, _)) =
   378   (if isar_proofs = SOME true orelse
   379       (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea smt_proofs preplay) then
   380      isar_proof_text ctxt debug isar_proofs smt_proofs isar_params
   381    else
   382      one_line_proof_text num_chained) one_line_params
   383 
   384 end;