src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
author blanchet
Tue, 19 Nov 2013 17:37:35 +0100
changeset 54500 f625e0e79dd1
parent 54499 319f8659267d
child 54501 77c9460e01b0
permissions -rw-r--r--
refactoring
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 =
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    16
    bool * bool * string * string * Time.time option * real * bool * (string * stature) list vector
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    17
    * (term, string) atp_step list * thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    18
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    19
  val isar_proof_text :
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
    20
    Proof.context -> bool option -> isar_params -> one_line_params -> string
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    21
  val proof_text :
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
    22
    Proof.context -> bool option -> (unit -> isar_params) -> int -> one_line_params -> string
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    23
end;
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    24
52374
blanchet
parents: 52369
diff changeset
    25
structure Sledgehammer_Reconstruct : SLEDGEHAMMER_RECONSTRUCT =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    26
struct
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    27
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    28
open ATP_Util
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    29
open ATP_Problem
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    30
open ATP_Proof
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    31
open ATP_Problem_Generate
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    32
open ATP_Proof_Reconstruct
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
    33
open Sledgehammer_Util
52555
6811291d1869 moved code -> easier debugging
smolkas
parents: 52454
diff changeset
    34
open Sledgehammer_Reconstructor
50264
a9ec48b98734 renamed sledgehammer_isar_reconstruct to sledgehammer_proof
smolkas
parents: 50262
diff changeset
    35
open Sledgehammer_Proof
50258
1c708d7728c7 put annotate in own structure
smolkas
parents: 50257
diff changeset
    36
open Sledgehammer_Annotate
52555
6811291d1869 moved code -> easier debugging
smolkas
parents: 52454
diff changeset
    37
open Sledgehammer_Print
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
    38
open Sledgehammer_Preplay
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51129
diff changeset
    39
open Sledgehammer_Compress
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
    40
open Sledgehammer_Try0
52611
831f7479c74f minimize dependencies (used facts) of Isar proof steps; remove unreferenced steps
smolkas
parents: 52592
diff changeset
    41
open Sledgehammer_Minimize_Isar
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    42
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    43
structure String_Redirect = ATP_Proof_Redirect(
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53052
diff changeset
    44
  type key = atp_step_name
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    45
  val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s')
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    46
  val string_of = fst)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    47
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    48
open String_Redirect
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    49
50017
d9c1b11a78d2 avoid name clashes
blanchet
parents: 50016
diff changeset
    50
val assume_prefix = "a"
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    51
val have_prefix = "f"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    52
val raw_prefix = "x"
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    53
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51976
diff changeset
    54
fun raw_label_of_name (num, ss) =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    55
  case resolve_conjecture ss of
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    56
    [j] => (conjecture_prefix, j)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    57
  | _ => (raw_prefix ^ ascii_of num, 0)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    58
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51976
diff changeset
    59
fun label_of_clause [name] = raw_label_of_name name
51976
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
    60
  | label_of_clause c =
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51976
diff changeset
    61
    (space_implode "___" (map (fst o raw_label_of_name) c), 0)
50005
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
    62
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
    63
fun add_fact_of_dependencies fact_names (names as [(_, ss)]) =
50005
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
    64
    if is_fact fact_names ss then
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
    65
      apsnd (union (op =) (map fst (resolve_fact fact_names ss)))
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
    66
    else
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
    67
      apfst (insert (op =) (label_of_clause names))
52034
11b48e7a4e7e correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents: 52031
diff changeset
    68
  | add_fact_of_dependencies _ names =
50005
e9a9bff107da handle non-unit clauses gracefully
blanchet
parents: 50004
diff changeset
    69
    apfst (insert (op =) (label_of_clause names))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    70
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    71
fun replace_one_dependency (old, new) dep =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    72
  if is_same_atp_step dep old then new else [dep]
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    73
fun replace_dependencies_in_line p (name, role, t, rule, deps) =
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    74
  (name, role, t, rule, fold (union (op =) o replace_one_dependency p) deps [])
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    75
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    76
(* No "real" literals means only type information (tfree_tcs, clsrel, or
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    77
   clsarity). *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    78
fun is_only_type_information t = t aconv @{term True}
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    79
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
    80
fun s_maybe_not role = role <> Conjecture ? s_not
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
    81
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    82
fun is_same_inference (role, t) (_, role', t', _, _) =
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    83
  s_maybe_not role t aconv s_maybe_not role' t'
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    84
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    85
(* Discard facts; consolidate adjacent lines that prove the same formula, since
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    86
   they differ only in type information.*)
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    87
fun add_line fact_names (name as (_, ss), role, t, rule, []) lines =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    88
    (* No dependencies: fact, conjecture, or (for Vampire) internal facts or
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    89
       definitions. *)
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
    90
    if is_conjecture ss then
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
    91
      (name, role, t, rule, []) :: lines
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
    92
    else if is_fact fact_names ss then
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    93
      (* Facts are not proof lines. *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    94
      if is_only_type_information t then
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    95
        map (replace_dependencies_in_line (name, [])) lines
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
    96
      else
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
    97
        lines
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    98
    else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    99
      map (replace_dependencies_in_line (name, [])) lines
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   100
  | add_line _ (line as (name, role, t, _, _)) lines =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   101
    (* Type information will be deleted later; skip repetition test. *)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   102
    if is_only_type_information t then
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   103
      line :: lines
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   104
    (* Is there a repetition? If so, replace later line by earlier one. *)
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   105
    else case take_prefix (not o is_same_inference (role, t)) lines of
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   106
      (_, []) => line :: lines
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   107
    | (pre, (name', _, _, _, _) :: post) =>
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   108
      line :: pre @ map (replace_dependencies_in_line (name', [name])) post
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   109
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   110
(* Recursively delete empty lines (type information) from the proof. *)
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   111
fun add_nontrivial_line (line as (name, _, t, _, [])) lines =
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   112
    if is_only_type_information t then delete_dependency name lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   113
    else line :: lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   114
  | add_nontrivial_line line lines = line :: lines
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   115
and delete_dependency name lines =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   116
  fold_rev add_nontrivial_line
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   117
           (map (replace_dependencies_in_line (name, [])) lines) []
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   118
50675
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   119
val e_skolemize_rule = "skolemize"
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   120
val vampire_skolemisation_rule = "skolemisation"
e3e707c8ac57 keep E's and Vampire's skolemization steps
blanchet
parents: 50674
diff changeset
   121
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   122
val is_skolemize_rule =
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   123
  member (op =) [e_skolemize_rule, vampire_skolemisation_rule]
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   124
52125
ac7830871177 improved handling of free variables' types in Isar proofs
blanchet
parents: 52077
diff changeset
   125
fun add_desired_line fact_names (name as (_, ss), role, t, rule, deps)
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   126
                     (j, lines) =
51198
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   127
  (j + 1,
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   128
   if is_fact fact_names ss orelse
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   129
      is_conjecture ss orelse
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   130
      is_skolemize_rule rule orelse
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   131
      (* the last line must be kept *)
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   132
      j = 0 orelse
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   133
      (not (is_only_type_information t) andalso
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   134
       null (Term.add_tvars t []) andalso
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   135
       length deps >= 2 andalso
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   136
       (* kill next to last line, which usually results in a trivial step *)
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   137
       j <> 1) then
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   138
     (name, role, t, rule, deps) :: lines  (* keep line *)
51198
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   139
   else
4dd63cf5bba5 got rid of rump support for Vampire definitions
blanchet
parents: 51194
diff changeset
   140
     map (replace_dependencies_in_line (name, deps)) lines)  (* drop line *)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   141
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   142
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   143
val add_labels_of_proof =
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   144
  steps_of_proof #> fold_isar_steps
52756
1ac8a0d0ddb1 parse nonnumeric identifiers in E proofs correctly
blanchet
parents: 52697
diff changeset
   145
    (byline_of_step #> (fn SOME (By ((ls, _), _)) => union (op =) ls | _ => I))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   146
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   147
fun kill_useless_labels_in_proof proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   148
  let
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   149
    val used_ls = add_labels_of_proof proof []
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   150
    fun do_label l = if member (op =) used_ls l then l else no_label
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
   151
    fun do_assms (Assume assms) = Assume (map (apfst do_label) assms)
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   152
    fun do_step (Prove (qs, xs, l, t, subproofs, by)) =
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   153
          Prove (qs, xs, do_label l, t, map do_proof subproofs, by)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   154
      | do_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
   155
    and do_proof (Proof (fix, assms, 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
   156
          Proof (fix, do_assms assms, map do_step steps)
51128
0021ea861129 introduced subblock in isar_step datatype for conjecture herbrandization
smolkas
parents: 51031
diff changeset
   157
  in do_proof proof end
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   158
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51976
diff changeset
   159
fun prefix_of_depth n = replicate_string (n + 1)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   160
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   161
val relabel_proof =
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   162
  let
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
   163
    fun fresh_label depth prefix (old as (l, subst, next)) =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   164
      if l = no_label then
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   165
        old
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   166
      else
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51976
diff changeset
   167
        let val l' = (prefix_of_depth depth 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
   168
          (l', (l, l') :: subst, next + 1)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   169
        end
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   170
    fun do_facts subst =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   171
      apfst (maps (the_list o AList.lookup (op =) subst))
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
   172
    fun do_assm depth (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
   173
      let
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
   174
        val (l, 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
   175
          (l, subst, next) |> fresh_label depth assume_prefix
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
   176
      in
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
   177
        ((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
   178
      end
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
   179
    fun do_assms subst depth (Assume assms) =
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
   180
      fold_map (do_assm depth) assms (subst, 1)
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
   181
      |> apfst Assume
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
   182
      |> apsnd fst
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
   183
    fun do_steps _ _ _ [] = []
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   184
      | do_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
   185
        let
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
   186
          val (l, 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
   187
            (l, subst, next) |> fresh_label depth have_prefix
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   188
          val sub = do_proofs subst depth sub
52626
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52611
diff changeset
   189
          val by = by |> do_byline subst
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   190
        in Prove (qs, xs, l, t, sub, by) :: do_steps subst depth next steps end
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
   191
      | do_steps subst depth next (step :: 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
   192
        step :: do_steps subst depth next 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
   193
    and do_proof subst depth (Proof (fix, assms, 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
   194
      let val (assms, subst) = do_assms subst depth assms in
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
   195
        Proof (fix, assms, do_steps subst depth 1 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
   196
      end
52626
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52611
diff changeset
   197
    and do_byline subst byline =
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   198
      map_facts_of_byline (do_facts subst) byline
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
   199
    and do_proofs subst depth = map (do_proof subst (depth + 1))
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
   200
  in do_proof [] 0 end
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   201
50004
c96e8e40d789 several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
blanchet
parents: 49994
diff changeset
   202
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
   203
  let
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   204
    fun do_qs_lfs NONE lfs = ([], lfs)
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   205
      | do_qs_lfs (SOME l0) lfs =
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   206
        if member (op =) lfs l0 then ([Then], lfs |> remove (op =) l0)
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   207
        else ([], lfs)
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   208
    fun chain_step lbl
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   209
      (Prove (qs, xs, l, t, subproofs, By ((lfs, gfs), method))) =
52453
2cba5906d836 simplified data structure
smolkas
parents: 52374
diff changeset
   210
          let val (qs', lfs) = do_qs_lfs lbl lfs in
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   211
            Prove (qs' @ qs, xs, l, t, chain_proofs subproofs,
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   212
                   By ((lfs, gfs), method))
52453
2cba5906d836 simplified data structure
smolkas
parents: 52374
diff changeset
   213
          end
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50671
diff changeset
   214
      | 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
   215
    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
   216
      | 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
   217
        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
   218
      | chain_steps _ (i :: is) = 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
   219
    and chain_proof (Proof (fix, Assume assms, 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
   220
      Proof (fix, Assume assms,
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
   221
             chain_steps (try (List.last #> fst) assms) 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
   222
    and chain_proofs proofs = map (chain_proof) proofs
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
   223
  in chain_proof end
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   224
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   225
type isar_params =
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   226
  bool * bool * string * string * Time.time option * real * bool * (string * stature) list vector
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   227
  * (term, string) atp_step list * thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   228
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
   229
fun isar_proof_text ctxt isar_proofs
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   230
    (debug, verbose, metis_type_enc, metis_lam_trans, preplay_timeout, isar_compress,
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   231
     isar_try0, fact_names, atp_proof, goal)
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   232
    (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   233
  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
   234
    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
   235
    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
   236
      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
   237
      |> map (fn (s, T) => (Binding.name s, SOME T, NoSyn))
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
   238
      |> (fn fixes =>
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
   239
             ctxt |> Variable.set_body false
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
   240
                  |> 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
   241
    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
   242
    val do_preplay = preplay_timeout <> SOME Time.zeroTime
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   243
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   244
    fun isar_proof_of () =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   245
      let
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   246
        val atp_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   247
          atp_proof
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   248
          |> rpair [] |-> fold_rev (add_line fact_names)
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   249
          |> rpair [] |-> fold_rev add_nontrivial_line
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   250
          |> rpair (0, [])
52125
ac7830871177 improved handling of free variables' types in Isar proofs
blanchet
parents: 52077
diff changeset
   251
          |-> fold_rev (add_desired_line fact_names)
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   252
          |> snd
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   253
        val conj_name = conjecture_prefix ^ string_of_int (length hyp_ts)
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   254
        val conjs =
50010
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   255
          atp_proof |> map_filter
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   256
            (fn (name as (_, ss), _, _, _, []) =>
50010
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   257
                if member (op =) ss conj_name then SOME name else NONE
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   258
              | _ => NONE)
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   259
        val assms =
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   260
          atp_proof |> map_filter
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   261
            (fn (name as (_, ss), _, _, _, []) =>
50013
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   262
                (case resolve_conjecture ss of
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   263
                   [j] =>
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   264
                   if j = length hyp_ts then NONE
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51976
diff changeset
   265
                   else SOME (raw_label_of_name name, nth hyp_ts j)
50013
cceec179bdca use original formulas for hypotheses and conclusion to avoid mismatches
blanchet
parents: 50012
diff changeset
   266
                 | _ => NONE)
50010
17488e45eb5a proper handling of assumptions arising from the goal's being expressed in rule format, for Isar proof construction
blanchet
parents: 50005
diff changeset
   267
              | _ => NONE)
51212
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   268
        val bot = atp_proof |> List.last |> #1
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   269
        val refute_graph =
51212
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   270
          atp_proof
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   271
          |> map (fn (name, _, _, _, from) => (from, name))
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   272
          |> make_refute_graph bot
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   273
          |> fold (Atom_Graph.default_node o rpair ()) conjs
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   274
        val axioms = axioms_of_refute_graph refute_graph conjs
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   275
        val tainted = tainted_atoms_of_refute_graph refute_graph conjs
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   276
        val is_clause_tainted = exists (member (op =) tainted)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   277
        val steps =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   278
          Symtab.empty
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   279
          |> fold (fn (name as (s, _), role, t, rule, _) =>
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   280
                      Symtab.update_new (s, (rule,
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   281
                        t |> (if is_clause_tainted [name] then
50905
db99fcf69761 more improvements to Isar proof reconstructions
blanchet
parents: 50705
diff changeset
   282
                                s_maybe_not role
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   283
                                #> fold exists_of (map Var (Term.add_vars t []))
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   284
                              else
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   285
                                I))))
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   286
                  atp_proof
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   287
        fun is_clause_skolemize_rule [(s, _)] =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   288
            Option.map (is_skolemize_rule o fst) (Symtab.lookup steps s) =
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   289
            SOME true
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   290
          | is_clause_skolemize_rule _ = false
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   291
        (* The assumptions and conjecture are "prop"s; the other formulas are
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   292
           "bool"s. *)
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
   293
        fun prop_of_clause [(s, ss)] =
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   294
            (case resolve_conjecture ss of
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   295
               [j] => if j = length hyp_ts then concl_t else nth hyp_ts j
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   296
             | _ => the_default ("", @{term False}) (Symtab.lookup steps s)
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   297
                    |> snd |> HOLogic.mk_Trueprop |> close_form)
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   298
          | prop_of_clause names =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   299
            let
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   300
              val lits = map snd (map_filter (Symtab.lookup steps o fst) names)
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   301
            in
50018
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   302
              case List.partition (can HOLogic.dest_not) lits of
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   303
                (negs as _ :: _, pos as _ :: _) =>
51212
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   304
                s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs),
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   305
                       Library.foldr1 s_disj pos)
50018
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   306
              | _ => fold (curry s_disj) lits @{term False}
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   307
            end
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   308
            |> HOLogic.mk_Trueprop |> close_form
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
   309
        fun isar_proof_of_direct_proof infs =
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
   310
          let
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
   311
            fun maybe_show outer c =
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
   312
              (outer andalso length c = 1 andalso subset (op =) (c, conjs))
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
   313
              ? cons Show
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
   314
            val is_fixed = Variable.is_declared ctxt orf can Name.dest_skolem
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
   315
            fun skolems_of t =
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
   316
              Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev
51976
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   317
            fun do_steps outer predecessor accum [] =
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   318
                accum
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   319
                |> (if tainted = [] then
52453
2cba5906d836 simplified data structure
smolkas
parents: 52374
diff changeset
   320
                      cons (Prove (if outer then [Show] else [],
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   321
                                   Fix [], no_label, concl_t, [],
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   322
                                   By_Metis ([the predecessor], [])))
51976
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   323
                    else
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   324
                      I)
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   325
                |> rev
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
   326
              | do_steps outer _ accum (Have (gamma, c) :: infs) =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   327
                let
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   328
                  val l = label_of_clause c
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   329
                  val t = prop_of_clause c
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   330
                  val by =
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   331
                    By_Metis
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   332
                      (fold (add_fact_of_dependencies fact_names) gamma no_facts)
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   333
                  fun prove sub by =
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   334
                    Prove (maybe_show outer c [], Fix [], l, t, sub, by)
51976
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   335
                  fun do_rest l step =
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   336
                    do_steps outer (SOME l) (step :: accum) infs
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   337
                in
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   338
                  if is_clause_tainted c then
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   339
                    case gamma of
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   340
                      [g] =>
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   341
                      if is_clause_skolemize_rule g andalso
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   342
                         is_clause_tainted g then
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   343
                        let
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   344
                          val subproof =
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
   345
                            Proof (Fix (skolems_of (prop_of_clause g)),
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
   346
                                   Assume [], rev accum)
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   347
                        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
   348
                          do_steps outer (SOME l)
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   349
                              [prove [subproof] (By_Metis no_facts)] []
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   350
                        end
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   351
                      else
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   352
                        do_rest l (prove [] by)
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   353
                    | _ => do_rest l (prove [] by)
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   354
                  else
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   355
                    if is_clause_skolemize_rule c then
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   356
                      do_rest l (Prove ([], Fix (skolems_of t), l, t, [], by))
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   357
                    else
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   358
                      do_rest l (prove [] by)
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   359
                end
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
   360
              | do_steps outer predecessor accum (Cases cases :: infs) =
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   361
                let
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   362
                  fun do_case (c, infs) =
51976
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   363
                    do_proof false [] [(label_of_clause c, prop_of_clause c)]
e5303bd748f2 generate valid direct Isar proof also if the facts are contradictory
blanchet
parents: 51879
diff changeset
   364
                             infs
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   365
                  val c = succedent_of_cases cases
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51165
diff changeset
   366
                  val l = label_of_clause c
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
   367
                  val t = prop_of_clause c
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
   368
                  val step =
52453
2cba5906d836 simplified data structure
smolkas
parents: 52374
diff changeset
   369
                    Prove (maybe_show outer c [], Fix [], l, t,
52454
b528a975b256 tuned: cleaned up data structure
smolkas
parents: 52453
diff changeset
   370
                      map do_case cases, By_Metis (the_list predecessor, []))
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   371
                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
   372
                  do_steps outer (SOME l) (step :: accum) infs
51149
4f0147ed8bcb skolemize conjecture properly in Isar proof
blanchet
parents: 51148
diff changeset
   373
                end
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
   374
            and do_proof outer fix assms infs =
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
   375
              Proof (Fix fix, Assume assms, do_steps outer NONE [] infs)
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
   376
          in
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
   377
            do_proof true params assms infs
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
   378
          end
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
   379
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   380
        (* 60 seconds seems like a good interpreation of "no timeout" *)
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   381
        val preplay_timeout = preplay_timeout |> the_default (seconds 60.0)
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   382
51741
blanchet
parents: 51258
diff changeset
   383
        val clean_up_labels_in_proof =
51165
blanchet
parents: 51164
diff changeset
   384
          chain_direct_proof
blanchet
parents: 51164
diff changeset
   385
          #> kill_useless_labels_in_proof
blanchet
parents: 51164
diff changeset
   386
          #> relabel_proof
53761
4d34e267fba9 use configuration mechanism for low-level tracing
blanchet
parents: 53586
diff changeset
   387
        val (preplay_interface as {overall_preplay_stats, ...}, isar_proof) =
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   388
          refute_graph
51031
63d71b247323 more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
blanchet
parents: 51026
diff changeset
   389
          |> redirect_graph axioms tainted bot
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
   390
          |> isar_proof_of_direct_proof
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   391
          |> relabel_proof_canonically
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   392
          |> `(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
   393
               preplay_timeout)
52626
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52611
diff changeset
   394
        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
   395
          isar_proof
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   396
          |> compress_proof
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   397
               (if isar_proofs = SOME true then isar_compress else 1000.0)
53762
06510d01a07b hard-coded an obscure option
blanchet
parents: 53761
diff changeset
   398
               preplay_interface
52632
23393c31c7fe added sledgehammer parameters isar_try0 and isar_minimize and their negative aliases
smolkas
parents: 52626
diff changeset
   399
          |> isar_try0 ? try0 preplay_timeout preplay_interface
53764
eba0d1c069b8 merged "isar_try0" and "isar_minimize" options
blanchet
parents: 53763
diff changeset
   400
          |> minimize_dependencies_and_remove_unrefed_steps isar_try0
52632
23393c31c7fe added sledgehammer parameters isar_try0 and isar_minimize and their negative aliases
smolkas
parents: 52626
diff changeset
   401
               preplay_interface
52626
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52611
diff changeset
   402
          |> `overall_preplay_stats
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52611
diff changeset
   403
          ||> clean_up_labels_in_proof
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   404
        val isar_text =
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   405
          string_of_proof ctxt metis_type_enc metis_lam_trans subgoal subgoal_count isar_proof
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   406
      in
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   407
        case isar_text of
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   408
          "" =>
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
   409
          if isar_proofs = SOME true then
50671
blanchet
parents: 50670
diff changeset
   410
            "\nNo structured proof available (proof too simple)."
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   411
          else
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   412
            ""
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   413
        | _ =>
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   414
          let
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   415
            val msg =
51203
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   416
              (if verbose then
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   417
                let
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52556
diff changeset
   418
                  val num_steps = add_proof_steps (steps_of_proof isar_proof) 0
51203
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   419
                in [string_of_int num_steps ^ " step" ^ plural_s num_steps] end
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   420
               else
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   421
                 []) @
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   422
              (if do_preplay then
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50905
diff changeset
   423
                [(if preplay_fail then "may fail, " else "") ^
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52555
diff changeset
   424
                   string_of_preplay_time preplay_time]
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   425
               else
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   426
                 [])
50277
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   427
          in
51741
blanchet
parents: 51258
diff changeset
   428
            "\n\nStructured proof"
blanchet
parents: 51258
diff changeset
   429
              ^ (commas msg |> not (null msg) ? enclose " (" ")")
52697
6fb98a20c349 explicit padding on command boundary for "auto" generated sendback -- do not replace the corresponding goal command, but append to it;
wenzelm
parents: 52632
diff changeset
   430
              ^ ":\n" ^
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
   431
              Active.sendback_markup [Markup.padding_command] isar_text
50277
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   432
          end
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   433
      end
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   434
    val isar_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   435
      if debug then
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   436
        isar_proof_of ()
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   437
      else case try isar_proof_of () of
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   438
        SOME s => s
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
   439
      | NONE => if isar_proofs = SOME true then
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   440
                  "\nWarning: The Isar proof construction failed."
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   441
                else
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   442
                  ""
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   443
  in one_line_proof ^ isar_proof end
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   444
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
   445
fun isar_proof_would_be_a_good_idea preplay =
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
   446
  case preplay of
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
   447
    Played (reconstr, _) => reconstr = SMT
54093
4e299e2c762d no isar proofs if preplay was not attempted
blanchet
parents: 53764
diff changeset
   448
  | Trust_Playable _ => false
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
   449
  | Failed_to_Play _ => true
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
   450
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
   451
fun proof_text ctxt isar_proofs isar_params num_chained
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   452
               (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
   453
  (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
   454
      (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then
54500
f625e0e79dd1 refactoring
blanchet
parents: 54499
diff changeset
   455
     isar_proof_text ctxt isar_proofs (isar_params ())
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   456
   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
   457
     one_line_proof_text num_chained) one_line_params
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   458
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   459
end;