src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
author blanchet
Mon, 03 Feb 2014 15:33:18 +0100
changeset 55286 7bbbd9393ce0
parent 55285 e88ad20035f4
child 55287 ffa306239316
permissions -rw-r--r--
tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_isar.ML
49883
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
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
     8
signature SLEDGEHAMMER_ISAR =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
     9
sig
54771
85879aa61334 move some Z3 specifics out (and into private repository with the rest of the Z3-specific code)
blanchet
parents: 54770
diff changeset
    10
  type atp_step_name = ATP_Proof.atp_step_name
54495
237d5be57277 refactored
blanchet
parents: 54093
diff changeset
    11
  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
    12
  type 'a atp_proof = 'a ATP_Proof.atp_proof
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    13
  type stature = ATP_Problem_Generate.stature
54495
237d5be57277 refactored
blanchet
parents: 54093
diff changeset
    14
  type one_line_params = Sledgehammer_Reconstructor.one_line_params
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    15
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
    16
  val trace : bool Config.T
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
    17
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    18
  type isar_params =
55267
e68fd012bbf3 got rid of 'try0' step that is now redundant
blanchet
parents: 55266
diff changeset
    19
    bool * (string option * string option) * Time.time * real * bool * bool
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
    20
      * (term, string) atp_step list * thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    21
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
    22
  val proof_text : Proof.context -> bool -> bool option -> (unit -> isar_params) -> int ->
54760
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
    23
    one_line_params -> string
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    24
end;
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    25
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    26
structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    27
struct
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    28
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    29
open ATP_Util
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    30
open ATP_Problem
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    31
open ATP_Proof
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
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    35
open Sledgehammer_Isar_Proof
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    36
open Sledgehammer_Isar_Preplay
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    37
open Sledgehammer_Isar_Compress
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    38
open Sledgehammer_Isar_Minimize
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    39
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    40
structure String_Redirect = ATP_Proof_Redirect(
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53052
diff changeset
    41
  type key = atp_step_name
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    42
  val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s')
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    43
  val string_of = fst)
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    44
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    45
open String_Redirect
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
    46
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
    47
val trace = Attrib.setup_config_bool @{binding sledgehammer_isar_trace} (K false)
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
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"]
54836
47857a79bdad recognize datatype reasoning in SPASS-Pirate
blanchet
parents: 54831
diff changeset
    50
val spass_pirate_datatype_rule = "DT"
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    51
val vampire_skolemisation_rule = "skolemisation"
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    52
(* TODO: Use "Z3_Proof.string_of_rule" once it is moved to Isabelle *)
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
    53
val z3_skolemize_rule = "sk"
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
    54
val z3_th_lemma_rule = "th-lemma"
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    55
54772
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54771
diff changeset
    56
val skolemize_rules =
f5fd4a34b0e8 handle Skolems gracefully for SPASS as well
blanchet
parents: 54771
diff changeset
    57
  e_skolemize_rules @ [spass_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule]
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 is_skolemize_rule = member (op =) skolemize_rules
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
    60
val is_arith_rule = String.isPrefix z3_th_lemma_rule
54836
47857a79bdad recognize datatype reasoning in SPASS-Pirate
blanchet
parents: 54831
diff changeset
    61
val is_datatype_rule = String.isPrefix spass_pirate_datatype_rule
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
    62
54501
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
    63
fun raw_label_of_num num = (num, 0)
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    64
54501
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
    65
fun label_of_clause [(num, _)] = raw_label_of_num num
77c9460e01b0 simplified old code
blanchet
parents: 54500
diff changeset
    66
  | 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
    67
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
    68
fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss)
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
    69
  | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    70
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
    71
(* No "real" literals means only type information (tfree_tcs, clsrel, or clsarity). *)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
    72
fun is_only_type_information t = t aconv @{prop True}
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
    73
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
    74
(* Discard facts; consolidate adjacent lines that prove the same formula, since they differ only in
54759
blanchet
parents: 54758
diff changeset
    75
   type information. *)
54799
blanchet
parents: 54772
diff changeset
    76
fun add_line_pass1 (line as (name, role, t, rule, [])) lines =
54770
0e354ef1b167 reverse Skolem function arguments
blanchet
parents: 54769
diff changeset
    77
    (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or
0e354ef1b167 reverse Skolem function arguments
blanchet
parents: 54769
diff changeset
    78
       definitions. *)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
    79
    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
    80
       role = Hypothesis orelse is_arith_rule rule then
54746
6db5fbc02436 better handling of Z3 proof blocks
blanchet
parents: 54716
diff changeset
    81
      line :: lines
54505
d023838eb252 more refactoring to accommodate SMT proofs
blanchet
parents: 54503
diff changeset
    82
    else if role = Axiom then
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    83
      (* Facts are not proof lines. *)
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
    84
      lines |> is_only_type_information t ? map (replace_dependencies_in_line (name, []))
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    85
    else
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    86
      map (replace_dependencies_in_line (name, [])) lines
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
    87
  | add_line_pass1 line lines = line :: lines
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
    88
55191
f127ab7368cf killed needless pass
blanchet
parents: 55186
diff changeset
    89
fun add_lines_pass2 res [] = rev res
f127ab7368cf killed needless pass
blanchet
parents: 55186
diff changeset
    90
  | add_lines_pass2 res ((name, role, t, rule, deps) :: lines) =
55184
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    91
    let
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    92
      val is_last_line = null lines
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    93
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    94
      fun looks_interesting () =
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    95
        not (is_only_type_information t) andalso null (Term.add_tvars t [])
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    96
        andalso length deps >= 2 andalso not (can the_single lines)
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    97
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    98
      fun is_skolemizing_line (_, _, _, rule', deps') =
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
    99
        is_skolemize_rule rule' andalso member (op =) deps' name
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
   100
      fun is_before_skolemize_rule () = exists is_skolemizing_line lines
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
   101
    in
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
   102
      if role <> Plain orelse is_skolemize_rule rule orelse is_arith_rule rule orelse
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
   103
         is_datatype_rule rule orelse is_last_line orelse looks_interesting () orelse
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
   104
         is_before_skolemize_rule () then
55191
f127ab7368cf killed needless pass
blanchet
parents: 55186
diff changeset
   105
        add_lines_pass2 ((name, role, t, rule, deps) :: res) lines
55184
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
   106
      else
55191
f127ab7368cf killed needless pass
blanchet
parents: 55186
diff changeset
   107
        add_lines_pass2 res (map (replace_dependencies_in_line (name, deps)) lines)
55184
6e2295db4cf8 keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
blanchet
parents: 55183
diff changeset
   108
    end
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   109
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   110
type isar_params =
55267
e68fd012bbf3 got rid of 'try0' step that is now redundant
blanchet
parents: 55266
diff changeset
   111
  bool * (string option * string option) * Time.time * real * bool * bool
e68fd012bbf3 got rid of 'try0' step that is now redundant
blanchet
parents: 55266
diff changeset
   112
    * (term, string) atp_step list * thm
49914
23e36a4d28f1 refactor code
blanchet
parents: 49913
diff changeset
   113
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   114
val arith_methods =
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   115
  [Arith_Method, Simp_Method, Auto_Method, Fastforce_Method, Blast_Method, Force_Method,
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   116
   Algebra_Method, Metis_Method (NONE, NONE), Meson_Method]
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   117
val datatype_methods = [Simp_Method, Simp_Size_Method]
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   118
val metislike_methods0 =
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   119
  [Metis_Method (NONE, NONE), Simp_Method, Auto_Method, Arith_Method, Blast_Method,
55274
b84867d6550b added a new version of 'metis' to the mix
blanchet
parents: 55273
diff changeset
   120
   Fastforce_Method, Force_Method, Algebra_Method, Meson_Method,
b84867d6550b added a new version of 'metis' to the mix
blanchet
parents: 55273
diff changeset
   121
   Metis_Method (SOME no_typesN, NONE)]
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   122
val rewrite_methods =
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   123
  [Auto_Method, Simp_Method, Fastforce_Method, Force_Method, Metis_Method (NONE, NONE),
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   124
   Meson_Method]
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   125
val skolem_methods = [Metis_Method (NONE, NONE), Blast_Method, Meson_Method]
54766
6ac273f176cd store alternative proof methods in Isar data structure
blanchet
parents: 54765
diff changeset
   126
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   127
fun isar_proof_text ctxt debug isar_proofs isar_params
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49917
diff changeset
   128
    (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   129
  let
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   130
    fun isar_proof_of () =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   131
      let
55267
e68fd012bbf3 got rid of 'try0' step that is now redundant
blanchet
parents: 55266
diff changeset
   132
        val SOME (verbose, alt_metis_args, preplay_timeout, compress_isar, try0_isar, minimize,
e68fd012bbf3 got rid of 'try0' step that is now redundant
blanchet
parents: 55266
diff changeset
   133
          atp_proof, goal) = try isar_params ()
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   134
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   135
        val metislike_methods = insert (op =) (Metis_Method alt_metis_args) metislike_methods0
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   136
55273
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   137
        val massage_meths = not try0_isar ? single o hd
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   138
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   139
        val (params, _, concl_t) = strip_subgoal goal subgoal ctxt
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   140
        val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   141
        val ctxt = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   142
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   143
        val do_preplay = preplay_timeout <> Time.zeroTime
55253
cfd276a7dbeb unform treatment of preplay_timeout = 0 and > 0
blanchet
parents: 55245
diff changeset
   144
        val compress_isar = if isar_proofs = NONE andalso do_preplay then 1000.0 else compress_isar
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   145
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   146
        val is_fixed = Variable.is_declared ctxt orf can Name.dest_skolem
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   147
        fun skolems_of t = Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   148
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   149
        fun get_role keep_role ((num, _), role, t, rule, _) =
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   150
          if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   151
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   152
        val atp_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   153
          atp_proof
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   154
          |> rpair [] |-> fold_rev add_line_pass1
55191
f127ab7368cf killed needless pass
blanchet
parents: 55186
diff changeset
   155
          |> add_lines_pass2 []
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   156
54535
59737a43e044 support Negated_Conjecture as a TPTP role as well (e.g. for SMT proofs)
blanchet
parents: 54507
diff changeset
   157
        val conjs =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   158
          map_filter (fn (name, role, _, _, _) =>
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   159
              if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   160
            atp_proof
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   161
        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
   162
        val lems =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   163
          map_filter (get_role (curry (op =) Lemma)) atp_proof
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   164
          |> map (fn ((l, t), rule) =>
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
   165
            let
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   166
              val (skos, meths) =
55273
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   167
                (if is_skolemize_rule rule then (skolems_of t, skolem_methods)
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   168
                 else if is_arith_rule rule then ([], arith_methods)
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   169
                 else ([], rewrite_methods))
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   170
                ||> massage_meths
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
   171
            in
55280
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   172
              Prove ([], skos, l, t, [], ([], []), meths)
54753
688da3388b2d use simplifier for rewrite
blanchet
parents: 54751
diff changeset
   173
            end)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   174
51212
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   175
        val bot = atp_proof |> List.last |> #1
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   176
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   177
        val refute_graph =
51212
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   178
          atp_proof
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   179
          |> map (fn (name, _, _, _, from) => (from, name))
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   180
          |> make_refute_graph bot
2bbcc9cc12b4 ensure all conjecture clauses are in the graph -- to prevent exceptions later
blanchet
parents: 51208
diff changeset
   181
          |> fold (Atom_Graph.default_node o rpair ()) conjs
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   182
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   183
        val axioms = axioms_of_refute_graph refute_graph conjs
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   184
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   185
        val tainted = tainted_atoms_of_refute_graph refute_graph conjs
51156
cbb640c3d203 made check for conjecture skolemization sound
blanchet
parents: 51149
diff changeset
   186
        val is_clause_tainted = exists (member (op =) tainted)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   187
        val steps =
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   188
          Symtab.empty
51201
f176855a1ee2 tuning (removed redundant datatype)
blanchet
parents: 51200
diff changeset
   189
          |> fold (fn (name as (s, _), role, t, rule, _) =>
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   190
              Symtab.update_new (s, (rule, t
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   191
                |> (if is_clause_tainted [name] then
54768
ee0881a54c72 fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
blanchet
parents: 54767
diff changeset
   192
                      HOLogic.dest_Trueprop
ee0881a54c72 fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
blanchet
parents: 54767
diff changeset
   193
                      #> role <> Conjecture ? s_not
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   194
                      #> 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
   195
                      #> HOLogic.mk_Trueprop
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   196
                    else
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   197
                      I))))
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   198
            atp_proof
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   199
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   200
        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
   201
54757
4960647932ec use 'prop' rather than 'bool' systematically in Isar reconstruction code
blanchet
parents: 54756
diff changeset
   202
        fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> close_form
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   203
          | prop_of_clause names =
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   204
            let
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   205
              val lits = map (HOLogic.dest_Trueprop o snd)
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   206
                (map_filter (Symtab.lookup steps o fst) names)
50676
83b8a5a39709 generate "obtain" steps corresponding to skolemization inferences
blanchet
parents: 50675
diff changeset
   207
            in
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   208
              (case List.partition (can HOLogic.dest_not) lits of
50018
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   209
                (negs as _ :: _, pos as _ :: _) =>
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   210
                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
   211
              | _ => fold (curry s_disj) lits @{term False})
50018
4ea26c74d7ea use implications rather than disjunctions to improve readability
blanchet
parents: 50017
diff changeset
   212
            end
50016
0ae5328ded8c fixed more "Trueprop" issues
blanchet
parents: 50015
diff changeset
   213
            |> HOLogic.mk_Trueprop |> close_form
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   214
55169
fda77499eef5 proper 'show' detection
blanchet
parents: 55168
diff changeset
   215
        fun maybe_show outer c = (outer andalso eq_set (op =) (c, conjs)) ? cons Show
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   216
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   217
        fun isar_steps outer predecessor accum [] =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   218
            accum
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   219
            |> (if tainted = [] then
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   220
                  cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [],
55280
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   221
                    (the_list predecessor, []), massage_meths metislike_methods))
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   222
                else
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   223
                  I)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   224
            |> rev
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   225
          | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) =
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   226
            let
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   227
              val l = label_of_clause c
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   228
              val t = prop_of_clause c
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   229
              val rule = rule_of_clause_id id
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   230
              val skolem = is_skolemize_rule rule
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   231
55279
df41d34d1324 tuned data structure
blanchet
parents: 55275
diff changeset
   232
              val deps = fold add_fact_of_dependencies gamma ([], [])
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   233
              val meths =
55273
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   234
                (if skolem then skolem_methods
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   235
                 else if is_arith_rule rule then arith_methods
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   236
                 else if is_datatype_rule rule then datatype_methods
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   237
                 else metislike_methods)
e887c0743614 implemented new 'try0_isar' semantics
blanchet
parents: 55267
diff changeset
   238
                |> massage_meths
55280
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   239
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   240
              fun prove sub facts = Prove (maybe_show outer c [], [], l, t, sub, facts, meths)
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   241
              fun steps_of_rest step = isar_steps outer (SOME l) (step :: accum) infs
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   242
            in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   243
              if is_clause_tainted c then
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   244
                (case gamma of
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   245
                  [g] =>
54755
2eb43ddde491 use 'arith' when appropriate in Z3 proofs
blanchet
parents: 54754
diff changeset
   246
                  if skolem andalso is_clause_tainted g then
54751
9b93f9117f8b implemented Z3 skolemization
blanchet
parents: 54750
diff changeset
   247
                    let val subproof = Proof (skolems_of (prop_of_clause g), [], rev accum) in
55280
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   248
                      isar_steps outer (SOME l) [prove [subproof] ([], [])] infs
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   249
                    end
51148
2246a2e17f92 tuning -- refactoring in preparation for handling skolemization of conjecture
blanchet
parents: 51147
diff changeset
   250
                  else
55280
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   251
                    steps_of_rest (prove [] deps)
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   252
                | _ => steps_of_rest (prove [] deps))
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   253
              else
55280
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   254
                steps_of_rest (if skolem then Prove ([], skolems_of t, l, t, [], deps, meths)
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   255
                  else prove [] deps)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   256
            end
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   257
          | isar_steps outer predecessor accum (Cases cases :: infs) =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   258
            let
55186
fafdf2424c57 don't forget the last inference(s) after conjecture skolemization
blanchet
parents: 55184
diff changeset
   259
              fun isar_case (c, subinfs) =
fafdf2424c57 don't forget the last inference(s) after conjecture skolemization
blanchet
parents: 55184
diff changeset
   260
                isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] subinfs
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   261
              val c = succedent_of_cases cases
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   262
              val l = label_of_clause c
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   263
              val t = prop_of_clause c
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   264
              val step =
54760
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
   265
                Prove (maybe_show outer c [], [], l, t,
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
   266
                  map isar_case (filter_out (null o snd) cases),
55280
f0187a12b8f2 tuned data structure
blanchet
parents: 55279
diff changeset
   267
                  (the_list predecessor, []), massage_meths metislike_methods)
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   268
            in
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   269
              isar_steps outer (SOME l) (step :: accum) infs
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   270
            end
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   271
        and isar_proof outer fix assms lems infs =
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   272
          Proof (fix, assms, lems @ isar_steps outer NONE [] infs)
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54699
diff changeset
   273
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   274
        val string_of_isar_proof = string_of_isar_proof ctxt subgoal subgoal_count
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   275
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   276
        val trace = Config.get ctxt trace
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   277
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55253
diff changeset
   278
        val canonical_isar_proof =
51145
280ece22765b tuned code
blanchet
parents: 51130
diff changeset
   279
          refute_graph
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   280
          |> trace ? tap (tracing o prefix "Refute graph: " o string_of_refute_graph)
51031
63d71b247323 more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
blanchet
parents: 51026
diff changeset
   281
          |> redirect_graph axioms tainted bot
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   282
          |> trace ? tap (tracing o prefix "Direct proof: " o string_of_direct_proof)
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   283
          |> isar_proof true params assms lems
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   284
          |> postprocess_isar_proof_remove_unreferenced_steps I
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   285
          |> relabel_isar_proof_canonically
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   286
55286
blanchet
parents: 55285
diff changeset
   287
        val ctxt = ctxt |> enrich_context_with_local_facts canonical_isar_proof
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55253
diff changeset
   288
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   289
        val preplay_data = Unsynchronized.ref Canonical_Label_Tab.empty
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   290
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   291
        val _ = fold_isar_steps (fn meth =>
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   292
            K (set_preplay_outcomes_of_isar_step ctxt preplay_timeout preplay_data meth []))
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   293
          (steps_of_isar_proof canonical_isar_proof) ()
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   294
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   295
        fun str_of_preplay_outcome outcome =
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   296
          if Lazy.is_finished outcome then string_of_play_outcome (Lazy.force outcome) else "?"
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   297
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55222
diff changeset
   298
        fun str_of_meth l meth =
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   299
          string_of_proof_method meth ^ " " ^
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
   300
          str_of_preplay_outcome (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)
55244
12e1a5d8ee48 simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
blanchet
parents: 55223
diff changeset
   301
        fun comment_of l = map (str_of_meth l) #> commas
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   302
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   303
        fun trace_isar_proof label proof =
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   304
          if trace then
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   305
            tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^ string_of_isar_proof comment_of proof ^
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   306
              "\n")
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   307
          else
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   308
            ()
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   309
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54827
diff changeset
   310
        val (play_outcome, isar_proof) =
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55253
diff changeset
   311
          canonical_isar_proof
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   312
          |> tap (trace_isar_proof "Original")
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   313
          |> compress_isar_proof ctxt compress_isar preplay_data
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   314
          |> tap (trace_isar_proof "Compressed")
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   315
          |> postprocess_isar_proof_remove_unreferenced_steps
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
   316
               (keep_fastest_method_of_isar_step (!preplay_data)
55267
e68fd012bbf3 got rid of 'try0' step that is now redundant
blanchet
parents: 55266
diff changeset
   317
                #> minimize ? minimize_isar_step_dependencies ctxt preplay_data)
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   318
          |> tap (trace_isar_proof "Minimized")
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   319
          |> `(preplay_outcome_of_isar_proof (!preplay_data))
55220
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   320
          ||> chain_isar_proof
9d833fe96c51 moved code around
blanchet
parents: 55219
diff changeset
   321
          ||> kill_useless_labels_in_isar_proof
55282
blanchet
parents: 55280
diff changeset
   322
          ||> relabel_isar_proof_nicely
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   323
      in
55222
a4ef6eb1fc20 added a 'trace' option
blanchet
parents: 55220
diff changeset
   324
        (case string_of_isar_proof (K (K "")) isar_proof of
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   325
          "" =>
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   326
          if isar_proofs = SOME true then "\nNo structured proof available (proof too simple)."
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   327
          else ""
55214
48a347b40629 better tracing + syntactically correct 'metis' calls
blanchet
parents: 55213
diff changeset
   328
        | isar_text =>
50670
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   329
          let
eaa540986291 properly take the existential closure of skolems
blanchet
parents: 50557
diff changeset
   330
            val msg =
51203
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   331
              (if verbose then
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   332
                 let val num_steps = add_isar_steps (steps_of_isar_proof isar_proof) 0 in
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   333
                   [string_of_int num_steps ^ " step" ^ plural_s num_steps]
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
   334
                 end
51203
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   335
               else
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51202
diff changeset
   336
                 []) @
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54827
diff changeset
   337
              (if do_preplay then [string_of_play_outcome play_outcome] else [])
50277
e0a4d8404c76 tweaked calculation of sledgehammer messages
smolkas
parents: 50276
diff changeset
   338
          in
54507
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   339
            "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
d983a020489d whitespace tuning
blanchet
parents: 54505
diff changeset
   340
            Active.sendback_markup [Markup.padding_command] isar_text
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   341
          end)
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   342
      end
54760
a1ac3eaa0d11 generate proper succedent for cases with trivial branches
blanchet
parents: 54759
diff changeset
   343
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   344
    val one_line_proof = one_line_proof_text 0 one_line_params
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   345
    val isar_proof =
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   346
      if debug then
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   347
        isar_proof_of ()
54754
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   348
      else
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   349
        (case try isar_proof_of () of
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   350
          SOME s => s
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   351
        | NONE =>
6b0ca7f79e93 robustness in degenerate case + tuning
blanchet
parents: 54753
diff changeset
   352
          if isar_proofs = SOME true then "\nWarning: The Isar proof construction failed." else "")
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   353
  in one_line_proof ^ isar_proof end
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   354
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55282
diff changeset
   355
fun isar_proof_would_be_a_good_idea (meth, play) =
54824
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
   356
  (case play of
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55282
diff changeset
   357
    Played _ => meth = SMT_Method
54823
a510fc40559c distinguish not preplayed & timed out
blanchet
parents: 54816
diff changeset
   358
  | Play_Timed_Out _ => true
54824
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
   359
  | Play_Failed => true
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
   360
  | Not_Played => 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
   361
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   362
fun proof_text ctxt debug isar_proofs isar_params num_chained
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55253
diff changeset
   363
    (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
   364
  (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
   365
      (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then
55168
948e8b7ea82f correctly handle exceptions arising from (experimental) Isar proof code
blanchet
parents: 54838
diff changeset
   366
     isar_proof_text ctxt debug isar_proofs isar_params
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   367
   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
   368
     one_line_proof_text num_chained) one_line_params
49883
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   369
a6ebdaf8e267 added missing file
blanchet
parents:
diff changeset
   370
end;