src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
author blanchet
Mon, 16 Dec 2013 20:24:13 +0100
changeset 54770 0e354ef1b167
parent 54769 3d6ac2f68bf3
child 54771 85879aa61334
permissions -rw-r--r--
reverse Skolem function arguments
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     3
    Author:     Steffen Juilf Smolka, TU Muenchen
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     4
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
     5
Isar proof reconstruction from ATP proofs.
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     6
*)
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     7
52374
blanchet
parents: 52369
diff changeset
     8
signature SLEDGEHAMMER_RECONSTRUCT =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     9
sig
54495
237d5be57277 refactored
blanchet
parents: 54093
diff changeset
    10
  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53052
diff changeset
    11
  type 'a atp_proof = 'a ATP_Proof.atp_proof
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    12
  type stature = ATP_Problem_Generate.stature
54495
237d5be57277 refactored
blanchet
parents: 54093
diff changeset
    13
  type one_line_params = Sledgehammer_Reconstructor.one_line_params
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    14
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    15
  type isar_params =
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
    16
    bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list *
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
    17
    thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    18
54760
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
    19
  val isar_proof_text : Proof.context -> bool option -> isar_params -> one_line_params -> string
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
    20
  val proof_text : Proof.context -> bool option -> (unit -> isar_params) -> int ->
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
    21
    one_line_params -> string
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    22
end;
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    23
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
    24
structure Sledgehammer_Reconstruct : SLEDGEHAMMER_RECONSTRUCT =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    25
struct
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    26
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    27
open ATP_Util
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    28
open ATP_Problem
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    29
open ATP_Proof
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    30
open ATP_Problem_Generate
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    31
open ATP_Proof_Reconstruct
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
    32
open Sledgehammer_Util
52555
6811291d1869 moved code -> easier debugging
smolkas
parents: 52454
diff changeset
    33
open Sledgehammer_Reconstructor
50264
a9ec48b98734 renamed sledgehammer_isar_reconstruct to sledgehammer_proof
smolkas
parents: 50262
diff changeset
    34
open Sledgehammer_Proof
50258
1c708d7728c7 put annotate in own structure
smolkas
parents: 50257
diff changeset
    35
open Sledgehammer_Annotate
52555
6811291d1869 moved code -> easier debugging
smolkas
parents: 52454
diff changeset
    36
open Sledgehammer_Print
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
    37
open Sledgehammer_Preplay
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51129
diff changeset
    38
open Sledgehammer_Compress
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
    39
open Sledgehammer_Try0
52611
831f7479c74f minimize dependencies (used facts) of Isar proof steps; remove unreferenced steps
smolkas
parents: 52592
diff changeset
    40
open Sledgehammer_Minimize_Isar
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    41
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    42
structure String_Redirect = ATP_Proof_Redirect(
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53052
diff changeset
    43
  type key = atp_step_name
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    44
  val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s')
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    45
  val string_of = fst)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    46
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    47
open String_Redirect
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    48
54769
3d6ac2f68bf3 correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
blanchet
parents: 54768
diff changeset
    49
val e_skolemize_rules = ["skolemize", "shift_quantors"]
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    50
val vampire_skolemisation_rule = "skolemisation"
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    51
(* TODO: Use "Z3_Proof.string_of_rule" once it is moved to Isabelle *)
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
    52
val z3_apply_def_rule = "apply-def"
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    53
val z3_hypothesis_rule = "hypothesis"
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
    54
val z3_intro_def_rule = "intro-def"
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    55
val z3_lemma_rule = "lemma"
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
    56
val z3_skolemize_rule = "sk"
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
    57
val z3_th_lemma_rule = "th-lemma"
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    58
54769
3d6ac2f68bf3 correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
blanchet
parents: 54768
diff changeset
    59
val skolemize_rules = e_skolemize_rules @ [vampire_skolemisation_rule, z3_skolemize_rule]
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    60
54769
3d6ac2f68bf3 correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
blanchet
parents: 54768
diff changeset
    61
val is_skolemize_rule = member (op =) skolemize_rules
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
    62
val is_arith_rule = String.isPrefix z3_th_lemma_rule
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
    63
54501
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
    64
fun raw_label_of_num num = (num, 0)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    65
54501
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
    66
fun label_of_clause [(num, _)] = raw_label_of_num num
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
    67
  | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0)
50005
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
    68
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
    69
fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss)
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
    70
  | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    71
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    72
fun replace_one_dependency (old, new) dep =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    73
  if is_same_atp_step dep old then new else [dep]
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    74
fun replace_dependencies_in_line p (name, role, t, rule, deps) =
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    75
  (name, role, t, rule, fold (union (op =) o replace_one_dependency p) deps [])
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    76
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    77
fun inline_z3_defs _ [] = []
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    78
  | inline_z3_defs defs ((line as (name, role, t, rule, deps)) :: lines) =
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    79
    if rule = z3_intro_def_rule then
54757
4960647932ec use 'prop' rather than 'bool' systematically in Isar reconstruction code
blanchet
parents: 54756
diff changeset
    80
      let val def = t |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> swap in
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    81
        inline_z3_defs (insert (op =) def defs)
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    82
          (map (replace_dependencies_in_line (name, [])) lines)
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    83
      end
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    84
    else if rule = z3_apply_def_rule then
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    85
      inline_z3_defs defs (map (replace_dependencies_in_line (name, [])) lines)
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    86
    else
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    87
      (name, role, Term.subst_atomic defs t, rule, deps) :: inline_z3_defs defs lines
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    88
54750
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    89
fun alist_cons_list eq (k, v) = AList.map_default eq (k, []) (cons v)
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    90
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    91
fun add_z3_hypotheses [] = I
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    92
  | add_z3_hypotheses hyps =
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    93
    HOLogic.dest_Trueprop
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
    94
    #> curry s_imp (Library.foldr1 s_conj (map HOLogic.dest_Trueprop hyps))
54750
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    95
    #> HOLogic.mk_Trueprop
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    96
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    97
fun inline_z3_hypotheses _ _ [] = []
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    98
  | inline_z3_hypotheses hyp_names hyps ((name, role, t, rule, deps) :: lines) =
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
    99
    if rule = z3_hypothesis_rule then
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   100
      inline_z3_hypotheses (name :: hyp_names) (alist_cons_list (op =) (t, name) hyps) lines
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   101
    else
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   102
      let val deps' = subtract (op =) hyp_names deps in
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   103
        if rule = z3_lemma_rule then
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   104
          (name, role, t, rule, deps') :: inline_z3_hypotheses hyp_names hyps lines
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   105
        else
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   106
          let
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   107
            val add_hyps = filter_out (null o inter (op =) deps o snd) hyps
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   108
            val t' = add_z3_hypotheses (map fst add_hyps) t
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   109
            val deps' = subtract (op =) hyp_names deps
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   110
            val hyps' = fold (AList.update (op =) o apsnd (insert (op =) name)) add_hyps hyps
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   111
          in
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   112
            (name, role, t', rule, deps') :: inline_z3_hypotheses hyp_names hyps' lines
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   113
          end
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   114
      end
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
   115
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   116
fun simplify_prop (@{const Not} $ t) = s_not (simplify_prop t)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   117
  | simplify_prop (@{const conj} $ t $ u) = s_conj (simplify_prop t, simplify_prop u)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   118
  | simplify_prop (@{const disj} $ t $ u) = s_disj (simplify_prop t, simplify_prop u)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   119
  | simplify_prop (@{const implies} $ t $ u) = s_imp (simplify_prop t, simplify_prop u)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   120
  | simplify_prop (@{const HOL.eq (bool)} $ t $ u) = s_iff (simplify_prop t, simplify_prop u)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   121
  | simplify_prop (t $ u) = simplify_prop t $ simplify_prop u
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   122
  | simplify_prop (Abs (s, T, t)) = Abs (s, T, simplify_prop t)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   123
  | simplify_prop t = t
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   124
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   125
fun simplify_line (name, role, t, rule, deps) = (name, role, simplify_prop t, rule, deps)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   126
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   127
(* No "real" literals means only type information (tfree_tcs, clsrel, or clsarity). *)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   128
fun is_only_type_information t = t aconv @{prop True}
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   129
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   130
(* Discard facts; consolidate adjacent lines that prove the same formula, since they differ only in
54759
blanchet
parents: 54758
diff changeset
   131
   type information. *)
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   132
fun add_line_pass1 (line as (name as (_, ss), role, t, rule, [])) lines =
54770
0e354ef1b167 reverse Skolem function arguments
blanchet
parents: 54769
diff changeset
   133
    (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or
0e354ef1b167 reverse Skolem function arguments
blanchet
parents: 54769
diff changeset
   134
       definitions. *)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   135
    if role = Lemma orelse role = Conjecture orelse role = Negated_Conjecture orelse
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   136
       role = Hypothesis orelse is_arith_rule rule then
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
   137
      line :: lines
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
   138
    else if role = Axiom then
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   139
      (* Facts are not proof lines. *)
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   140
      lines |> is_only_type_information t ? map (replace_dependencies_in_line (name, []))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   141
    else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   142
      map (replace_dependencies_in_line (name, [])) lines
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   143
  | add_line_pass1 line lines = line :: lines
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   144
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   145
(* Recursively delete empty lines (type information) from the proof.
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   146
   (FIXME: needed? And why is "delete_dependency" so complicated?) *)
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   147
fun add_line_pass2 (line as (name, _, t, _, [])) lines =
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   148
    if is_only_type_information t then delete_dependency name lines else line :: lines
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   149
  | add_line_pass2 line lines = line :: lines
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   150
and delete_dependency name lines =
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   151
  fold_rev add_line_pass2 (map (replace_dependencies_in_line (name, [])) lines) []
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   152
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   153
fun add_line_pass3 res [] = rev res
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   154
  | add_line_pass3 res ((name as (_, ss), role, t, rule, deps) :: lines) =
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   155
    if role <> Plain orelse is_skolemize_rule rule orelse is_arith_rule rule orelse
54716
55ed20a29a8c reverse order in which lines are selected, to ensure that the number of dependencies is accurate
blanchet
parents: 54715
diff changeset
   156
       (* the last line must be kept *)
55ed20a29a8c reverse order in which lines are selected, to ensure that the number of dependencies is accurate
blanchet
parents: 54715
diff changeset
   157
       null lines orelse
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
   158
       (not (is_only_type_information t) andalso null (Term.add_tvars t [])
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
   159
        andalso length deps >= 2 andalso
54716
55ed20a29a8c reverse order in which lines are selected, to ensure that the number of dependencies is accurate
blanchet
parents: 54715
diff changeset
   160
        (* don't keep next to last line, which usually results in a trivial step *)
55ed20a29a8c reverse order in which lines are selected, to ensure that the number of dependencies is accurate
blanchet
parents: 54715
diff changeset
   161
        not (can the_single lines)) then
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   162
      add_line_pass3 ((name, role, simplify_prop t, rule, deps) :: res) lines
54716
55ed20a29a8c reverse order in which lines are selected, to ensure that the number of dependencies is accurate
blanchet
parents: 54715
diff changeset
   163
    else
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   164
      add_line_pass3 res (map (replace_dependencies_in_line (name, deps)) lines)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   165
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   166
val add_labels_of_proof =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   167
  steps_of_proof
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   168
  #> fold_isar_steps (byline_of_step #> (fn SOME ((ls, _), _) => union (op =) ls | _ => I))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   169
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   170
fun kill_useless_labels_in_proof proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   171
  let
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   172
    val used_ls = add_labels_of_proof proof []
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   173
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   174
    fun kill_label l = if member (op =) used_ls l then l else no_label
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   175
    fun kill_assms assms = map (apfst kill_label) assms
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   176
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   177
    fun kill_step (Prove (qs, xs, l, t, subproofs, by)) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   178
        Prove (qs, xs, kill_label l, t, map kill_proof subproofs, by)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   179
      | kill_step step = step
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   180
    and kill_proof (Proof (fix, assms, steps)) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   181
      Proof (fix, kill_assms assms, map kill_step steps)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   182
  in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   183
    kill_proof proof
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   184
  end
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   185
54501
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
   186
val assume_prefix = "a"
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
   187
val have_prefix = "f"
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
   188
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   189
val relabel_proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   190
  let
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   191
    fun fresh_label depth prefix (accum as (l, subst, next)) =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   192
      if l = no_label then
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   193
        accum
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   194
      else
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   195
        let val l' = (replicate_string (depth + 1) prefix, next) in
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   196
          (l', (l, l') :: subst, next + 1)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   197
        end
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   198
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   199
    fun relabel_facts subst = apfst (maps (the_list o AList.lookup (op =) subst))
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   200
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   201
    fun relabel_assm depth (l, t) (subst, next) =
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   202
      let val (l, subst, next) = (l, subst, next) |> fresh_label depth assume_prefix in
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   203
        ((l, t), (subst, next))
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   204
      end
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   205
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   206
    fun relabel_assms subst depth assms = fold_map (relabel_assm depth) assms (subst, 1) ||> fst
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   207
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   208
    fun relabel_steps _ _ _ [] = []
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   209
      | relabel_steps subst depth next (Prove (qs, xs, l, t, sub, by) :: steps) =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   210
        let
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   211
          val (l, subst, next) = (l, subst, next) |> fresh_label depth have_prefix
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   212
          val sub = relabel_proofs subst depth sub
54759
blanchet
parents: 54758
diff changeset
   213
          val by = apfst (relabel_facts subst) by
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   214
        in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   215
          Prove (qs, xs, l, t, sub, by) :: relabel_steps subst depth next steps
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   216
        end
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   217
      | relabel_steps subst depth next (step :: steps) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   218
        step :: relabel_steps subst depth next steps
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   219
    and relabel_proof subst depth (Proof (fix, assms, steps)) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   220
      let val (assms, subst) = relabel_assms subst depth assms in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   221
        Proof (fix, assms, relabel_steps subst depth 1 steps)
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   222
      end
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   223
    and relabel_proofs subst depth = map (relabel_proof subst (depth + 1))
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   224
  in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   225
    relabel_proof [] 0
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   226
  end
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   227
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   228
val chain_direct_proof =
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   229
  let
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   230
    fun chain_qs_lfs NONE lfs = ([], lfs)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   231
      | chain_qs_lfs (SOME l0) lfs =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   232
        if member (op =) lfs l0 then ([Then], lfs |> remove (op =) l0) else ([], lfs)
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   233
    fun chain_step lbl (Prove (qs, xs, l, t, subproofs, ((lfs, gfs), methss))) =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   234
        let val (qs', lfs) = chain_qs_lfs lbl lfs in
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   235
          Prove (qs' @ qs, xs, l, t, chain_proofs subproofs, ((lfs, gfs), methss))
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   236
        end
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   237
      | chain_step _ step = step
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   238
    and chain_steps _ [] = []
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   239
      | chain_steps (prev as SOME _) (i :: is) =
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   240
        chain_step prev i :: chain_steps (label_of_step i) is
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   241
      | chain_steps _ (i :: is) = i :: chain_steps (label_of_step i) is
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   242
    and chain_proof (Proof (fix, assms, steps)) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   243
      Proof (fix, assms, chain_steps (try (List.last #> fst) assms) steps)
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   244
    and chain_proofs proofs = map (chain_proof) proofs
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   245
  in
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   246
    chain_proof
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   247
  end
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   248
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   249
type isar_params =
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
   250
  bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list *
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
   251
  thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   252
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   253
val arith_methodss =
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   254
  [[Arith_Method, Simp_Method, Auto_Method, Fastforce_Method, Blast_Method, Force_Method,
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   255
    Metis_Method], [Meson_Method]]
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   256
val metislike_methodss =
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   257
  [[Metis_Method, Simp_Method, Auto_Method, Arith_Method, Blast_Method, Fastforce_Method,
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   258
    Force_Method], [Meson_Method]]
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   259
val rewrite_methodss =
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   260
  [[Simp_Method, Auto_Method, Fastforce_Method, Force_Method, Metis_Method], [Meson_Method]]
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   261
val skolem_methodss = [[Metis_Method, Blast_Method], [Metis_New_Skolem_Method], [Meson_Method]]
54766
6ac273f176cd store alternative proof methods in Isar data structure
blanchet
parents: 54765
diff changeset
   262
53052
a0db255af8c5 sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
wenzelm
parents: 53047
diff changeset
   263
fun isar_proof_text ctxt isar_proofs
54756
blanchet
parents: 54755
diff changeset
   264
    (debug, verbose, metis_type_enc, metis_lam_trans, preplay_timeout, isar_compress, isar_try0,
blanchet
parents: 54755
diff changeset
   265
     atp_proof, goal)
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   266
    (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   267
  let
52196
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52150
diff changeset
   268
    val (params, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52150
diff changeset
   269
    val (_, ctxt) =
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52150
diff changeset
   270
      params
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52150
diff changeset
   271
      |> map (fn (s, T) => (Binding.name s, SOME T, NoSyn))
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   272
      |> (fn fixes => ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes)
53052
a0db255af8c5 sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
wenzelm
parents: 53047
diff changeset
   273
    val one_line_proof = one_line_proof_text 0 one_line_params
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   274
    val do_preplay = preplay_timeout <> SOME Time.zeroTime
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   275
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   276
    val is_fixed = Variable.is_declared ctxt orf can Name.dest_skolem
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   277
    fun skolems_of t = Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   278
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   279
    fun get_role keep_role ((num, _), role, t, rule, _) =
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   280
      if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   281
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   282
    fun isar_proof_of () =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   283
      let
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   284
        val atp_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   285
          atp_proof
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
   286
          |> inline_z3_defs []
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   287
          |> map simplify_line
54750
f50693bab67c inline Z3 hypotheses
blanchet
parents: 54746
diff changeset
   288
          |> inline_z3_hypotheses [] []
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   289
          |> rpair [] |-> fold_rev add_line_pass1
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   290
          |> rpair [] |-> fold_rev add_line_pass2
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   291
          |> add_line_pass3 []
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   292
54535
59737a43e044 support Negated_Conjecture as a TPTP role as well (e.g. for SMT proofs)
blanchet
parents: 54507
diff changeset
   293
        val conjs =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   294
          map_filter (fn (name, role, _, _, _) =>
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   295
              if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   296
            atp_proof
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   297
        val assms = map_filter (Option.map fst o get_role (curry (op =) Hypothesis)) atp_proof
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   298
        val lems =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   299
          map_filter (get_role (curry (op =) Lemma)) atp_proof
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   300
          |> map (fn ((l, t), rule) =>
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
   301
            let
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   302
              val (skos, methss) =
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   303
                if is_skolemize_rule rule then (skolems_of t, skolem_methodss)
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   304
                else if is_arith_rule rule then ([], arith_methodss)
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   305
                else ([], rewrite_methodss)
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
   306
            in
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   307
              Prove ([], skos, l, t, [], (([], []), methss))
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
   308
            end)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   309
51212
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   310
        val bot = atp_proof |> List.last |> #1
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   311
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   312
        val refute_graph =
51212
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   313
          atp_proof
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   314
          |> map (fn (name, _, _, _, from) => (from, name))
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   315
          |> make_refute_graph bot
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   316
          |> fold (Atom_Graph.default_node o rpair ()) conjs
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   317
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   318
        val axioms = axioms_of_refute_graph refute_graph conjs
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   319
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   320
        val tainted = tainted_atoms_of_refute_graph refute_graph conjs
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   321
        val is_clause_tainted = exists (member (op =) tainted)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   322
        val steps =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   323
          Symtab.empty
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   324
          |> fold (fn (name as (s, _), role, t, rule, _) =>
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   325
              Symtab.update_new (s, (rule, t
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   326
                |> (if is_clause_tainted [name] then
54768
ee0881a54c72 fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
blanchet
parents: 54767
diff changeset
   327
                      HOLogic.dest_Trueprop
ee0881a54c72 fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
blanchet
parents: 54767
diff changeset
   328
                      #> role <> Conjecture ? s_not
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   329
                      #> fold exists_of (map Var (Term.add_vars t []))
54768
ee0881a54c72 fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
blanchet
parents: 54767
diff changeset
   330
                      #> HOLogic.mk_Trueprop
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   331
                    else
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   332
                      I))))
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   333
            atp_proof
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   334
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   335
        val rule_of_clause_id = fst o the o Symtab.lookup steps o fst
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   336
54757
4960647932ec use 'prop' rather than 'bool' systematically in Isar reconstruction code
blanchet
parents: 54756
diff changeset
   337
        fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> close_form
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   338
          | prop_of_clause names =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   339
            let
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   340
              val lits = map (HOLogic.dest_Trueprop o snd)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   341
                (map_filter (Symtab.lookup steps o fst) names)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   342
            in
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   343
              (case List.partition (can HOLogic.dest_not) lits of
50018
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   344
                (negs as _ :: _, pos as _ :: _) =>
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   345
                s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos)
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   346
              | _ => fold (curry s_disj) lits @{term False})
50018
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   347
            end
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   348
            |> HOLogic.mk_Trueprop |> close_form
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   349
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   350
        fun maybe_show outer c =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   351
          (outer andalso length c = 1 andalso subset (op =) (c, conjs)) ? cons Show
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   352
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   353
        fun isar_steps outer predecessor accum [] =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   354
            accum
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   355
            |> (if tainted = [] then
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   356
                  cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [],
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   357
                               ((the_list predecessor, []), metislike_methodss)))
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   358
                else
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   359
                  I)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   360
            |> rev
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   361
          | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   362
            let
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   363
              val l = label_of_clause c
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   364
              val t = prop_of_clause c
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   365
              val rule = rule_of_clause_id id
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   366
              val skolem = is_skolemize_rule rule
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   367
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   368
              fun prove sub by = Prove (maybe_show outer c [], [], l, t, sub, by)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   369
              fun do_rest l step = isar_steps outer (SOME l) (step :: accum) infs
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   370
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   371
              val deps = fold add_fact_of_dependencies gamma no_facts
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   372
              val methss = if is_arith_rule rule then arith_methodss else metislike_methodss
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   373
              val by = (deps, methss)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   374
            in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   375
              if is_clause_tainted c then
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   376
                (case gamma of
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   377
                  [g] =>
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   378
                  if skolem andalso is_clause_tainted g then
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   379
                    let val subproof = Proof (skolems_of (prop_of_clause g), [], rev accum) in
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   380
                      isar_steps outer (SOME l) [prove [subproof] (no_facts, skolem_methodss)] []
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   381
                    end
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   382
                  else
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   383
                    do_rest l (prove [] by)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   384
                | _ => do_rest l (prove [] by))
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   385
              else
54765
blanchet
parents: 54763
diff changeset
   386
                do_rest l (if skolem then Prove ([], skolems_of t, l, t, [], by) else prove [] by)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   387
            end
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   388
          | isar_steps outer predecessor accum (Cases cases :: infs) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   389
            let
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   390
              fun isar_case (c, infs) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   391
                isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] infs
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   392
              val c = succedent_of_cases cases
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   393
              val l = label_of_clause c
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   394
              val t = prop_of_clause c
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   395
              val step =
54760
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
   396
                Prove (maybe_show outer c [], [], l, t,
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
   397
                  map isar_case (filter_out (null o snd) cases),
54767
81290fe85890 generalize method list further to list of list (clustering preferred methods together)
blanchet
parents: 54766
diff changeset
   398
                  ((the_list predecessor, []), metislike_methodss))
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   399
            in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   400
              isar_steps outer (SOME l) (step :: accum) infs
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   401
            end
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   402
        and isar_proof outer fix assms lems infs =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   403
          Proof (fix, assms, lems @ isar_steps outer NONE [] infs)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   404
54765
blanchet
parents: 54763
diff changeset
   405
        (* 60 seconds seems like a reasonable interpreation of "no timeout" *)
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   406
        val preplay_timeout = preplay_timeout |> the_default (seconds 60.0)
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   407
53761
4d34e267fba9 use configuration mechanism for low-level tracing
blanchet
parents: 53586
diff changeset
   408
        val (preplay_interface as {overall_preplay_stats, ...}, isar_proof) =
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   409
          refute_graph
54699
65c4fd04b5b2 useful debugging info
blanchet
parents: 54535
diff changeset
   410
(*
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   411
          |> tap (tracing o prefix "Refute graph: " o string_of_refute_graph)
54699
65c4fd04b5b2 useful debugging info
blanchet
parents: 54535
diff changeset
   412
*)
51031
63d71b247323 more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
blanchet
parents: 51026
diff changeset
   413
          |> redirect_graph axioms tainted bot
54699
65c4fd04b5b2 useful debugging info
blanchet
parents: 54535
diff changeset
   414
(*
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   415
          |> tap (tracing o prefix "Direct proof: " o string_of_direct_proof)
54699
65c4fd04b5b2 useful debugging info
blanchet
parents: 54535
diff changeset
   416
*)
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   417
          |> isar_proof true params assms lems
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   418
          |> postprocess_remove_unreferenced_steps I
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   419
          |> relabel_proof_canonically
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   420
          |> `(proof_preplay_interface debug ctxt metis_type_enc metis_lam_trans do_preplay
53761
4d34e267fba9 use configuration mechanism for low-level tracing
blanchet
parents: 53586
diff changeset
   421
               preplay_timeout)
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   422
52626
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52611
diff changeset
   423
        val ((preplay_time, preplay_fail), isar_proof) =
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   424
          isar_proof
54501
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
   425
          |> compress_proof (if isar_proofs = SOME true then isar_compress else 1000.0)
53762
06510d01a07b hard-coded an obscure option
blanchet
parents: 53761
diff changeset
   426
               preplay_interface
52632
23393c31c7fe added sledgehammer parameters isar_try0 and isar_minimize and their negative aliases
smolkas
parents: 52626
diff changeset
   427
          |> isar_try0 ? try0 preplay_timeout preplay_interface
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   428
          |> postprocess_remove_unreferenced_steps (isar_try0 ? min_deps_of_step preplay_interface)
52626
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52611
diff changeset
   429
          |> `overall_preplay_stats
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   430
          ||> (chain_direct_proof #> kill_useless_labels_in_proof #> relabel_proof)
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   431
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   432
        val isar_text =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   433
          string_of_proof ctxt metis_type_enc metis_lam_trans subgoal subgoal_count isar_proof
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   434
      in
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   435
        (case isar_text of
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   436
          "" =>
51190
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   437
          if isar_proofs = SOME true then
50671
blanchet
parents: 50670
diff changeset
   438
            "\nNo structured proof available (proof too simple)."
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   439
          else
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   440
            ""
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   441
        | _ =>
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   442
          let
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   443
            val msg =
51203
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   444
              (if verbose then
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   445
                let
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   446
                  val num_steps = add_proof_steps (steps_of_proof isar_proof) 0
51203
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   447
                in [string_of_int num_steps ^ " step" ^ plural_s num_steps] end
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   448
               else
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   449
                 []) @
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   450
              (if do_preplay then
54763
blanchet
parents: 54760
diff changeset
   451
                [(if preplay_fail then "may fail, " else "") ^ string_of_ext_time preplay_time]
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   452
               else
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   453
                 [])
50277
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   454
          in
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   455
            "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   456
            Active.sendback_markup [Markup.padding_command] isar_text
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   457
          end)
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   458
      end
54760
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
   459
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   460
    val isar_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   461
      if debug then
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   462
        isar_proof_of ()
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   463
      else
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   464
        (case try isar_proof_of () of
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   465
          SOME s => s
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   466
        | NONE =>
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   467
          if isar_proofs = SOME true then "\nWarning: The Isar proof construction failed." else "")
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   468
  in one_line_proof ^ isar_proof end
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   469
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51179
diff changeset
   470
fun isar_proof_would_be_a_good_idea preplay =
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   471
  (case preplay of
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51179
diff changeset
   472
    Played (reconstr, _) => reconstr = SMT
54093
4e299e2c762d no isar proofs if preplay was not attempted
blanchet
parents: 53764
diff changeset
   473
  | Trust_Playable _ => false
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   474
  | Failed_to_Play _ => true)
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51179
diff changeset
   475
53052
a0db255af8c5 sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
wenzelm
parents: 53047
diff changeset
   476
fun proof_text ctxt isar_proofs isar_params num_chained
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   477
               (one_line_params as (preplay, _, _, _, _, _)) =
51190
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   478
  (if isar_proofs = SOME true orelse
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   479
      (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   480
     isar_proof_text ctxt isar_proofs (isar_params ())
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   481
   else
53052
a0db255af8c5 sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
wenzelm
parents: 53047
diff changeset
   482
     one_line_proof_text num_chained) one_line_params
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   483
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   484
end;