src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML
author nipkow
Wed, 29 May 2024 10:43:22 +0200
changeset 80207 ca7e7e41374e
parent 79942 7793e3161d2b
child 80540 f86bcf439837
permissions -rw-r--r--
pretty-printing sledgehammer command: merge indexed theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55287
ffa306239316 renamed ML file
blanchet
parents: 55286
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
     3
    Author:     Steffen Juilf Smolka, TU Muenchen
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
     4
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
     5
Reconstructors.
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
     6
*)
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
     7
55287
ffa306239316 renamed ML file
blanchet
parents: 55286
diff changeset
     8
signature SLEDGEHAMMER_PROOF_METHODS =
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
     9
sig
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    10
  type stature = ATP_Problem_Generate.stature
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    11
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    12
  datatype SMT_backend =
72798
e732c98b02e6 tuned proof preplay to explicitly refer to Z3 backend
desharna
parents: 72518
diff changeset
    13
    SMT_Z3 |
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    14
    SMT_Verit of string
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    15
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    16
  datatype proof_method =
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    17
    Metis_Method of string option * string option |
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    18
    Meson_Method |
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    19
    SMT_Method of SMT_backend |
56852
b38c5b9cf590 added 'satx' to Sledgehammer's portfolio (cf. 'isar_try0')
blanchet
parents: 56093
diff changeset
    20
    SATx_Method |
78695
41273636a82a added argo
blanchet
parents: 75956
diff changeset
    21
    Argo_Method |
55323
253a029335a2 split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents: 55315
diff changeset
    22
    Blast_Method |
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    23
    Simp_Method |
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    24
    Auto_Method |
58092
4ae52c60603a renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
blanchet
parents: 58075
diff changeset
    25
    Fastforce_Method |
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    26
    Force_Method |
58092
4ae52c60603a renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
blanchet
parents: 58075
diff changeset
    27
    Moura_Method |
55323
253a029335a2 split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents: 55315
diff changeset
    28
    Linarith_Method |
253a029335a2 split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents: 55315
diff changeset
    29
    Presburger_Method |
79942
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
    30
    Algebra_Method |
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
    31
    Order_Method
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    32
54824
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
    33
  datatype play_outcome =
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
    34
    Played of Time.time |
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
    35
    Play_Timed_Out of Time.time |
56093
4eeb73a1feec simplified preplaying information
blanchet
parents: 56081
diff changeset
    36
    Play_Failed
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    37
54824
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
    38
  type one_line_params =
57739
blanchet
parents: 57735
diff changeset
    39
    ((string * stature) list * (proof_method * play_outcome)) * string * int * int
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    40
57287
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
    41
  val is_proof_method_direct : proof_method -> bool
72401
2783779b7dd3 removed obsolete unmaintained experimental prover Pirate
blanchet
parents: 70586
diff changeset
    42
  val string_of_proof_method : string list -> proof_method -> string
57054
blanchet
parents: 56985
diff changeset
    43
  val tac_of_proof_method : Proof.context -> thm list * thm list -> proof_method -> int -> tactic
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
    44
  val string_of_play_outcome : play_outcome -> string
70586
57df8a85317a clarified signature;
wenzelm
parents: 63692
diff changeset
    45
  val play_outcome_ord : play_outcome ord
56985
82c83978fbd9 correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
blanchet
parents: 56983
diff changeset
    46
  val one_line_proof_text : Proof.context -> int -> one_line_params -> string
54495
237d5be57277 refactored
blanchet
parents: 52590
diff changeset
    47
end;
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    48
55287
ffa306239316 renamed ML file
blanchet
parents: 55286
diff changeset
    49
structure Sledgehammer_Proof_Methods : SLEDGEHAMMER_PROOF_METHODS =
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    50
struct
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    51
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54824
diff changeset
    52
open ATP_Util
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    53
open ATP_Problem_Generate
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
    54
open ATP_Proof_Reconstruct
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    55
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    56
datatype SMT_backend =
72798
e732c98b02e6 tuned proof preplay to explicitly refer to Z3 backend
desharna
parents: 72518
diff changeset
    57
  SMT_Z3 |
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    58
  SMT_Verit of string
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    59
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    60
datatype proof_method =
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    61
  Metis_Method of string option * string option |
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    62
  Meson_Method |
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    63
  SMT_Method of SMT_backend |
56852
b38c5b9cf590 added 'satx' to Sledgehammer's portfolio (cf. 'isar_try0')
blanchet
parents: 56093
diff changeset
    64
  SATx_Method |
78695
41273636a82a added argo
blanchet
parents: 75956
diff changeset
    65
  Argo_Method |
55323
253a029335a2 split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents: 55315
diff changeset
    66
  Blast_Method |
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    67
  Simp_Method |
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    68
  Auto_Method |
58092
4ae52c60603a renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
blanchet
parents: 58075
diff changeset
    69
  Fastforce_Method |
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
    70
  Force_Method |
58092
4ae52c60603a renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
blanchet
parents: 58075
diff changeset
    71
  Moura_Method |
55323
253a029335a2 split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents: 55315
diff changeset
    72
  Linarith_Method |
253a029335a2 split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents: 55315
diff changeset
    73
  Presburger_Method |
79942
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
    74
  Algebra_Method |
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
    75
  Order_Method
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    76
54824
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
    77
datatype play_outcome =
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
    78
  Played of Time.time |
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
    79
  Play_Timed_Out of Time.time |
56093
4eeb73a1feec simplified preplaying information
blanchet
parents: 56081
diff changeset
    80
  Play_Failed
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
    81
57739
blanchet
parents: 57735
diff changeset
    82
type one_line_params =
blanchet
parents: 57735
diff changeset
    83
  ((string * stature) list * (proof_method * play_outcome)) * string * int * int
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
    84
57287
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
    85
fun is_proof_method_direct (Metis_Method _) = true
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
    86
  | is_proof_method_direct Meson_Method = true
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
    87
  | is_proof_method_direct (SMT_Method _) = true
57287
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
    88
  | is_proof_method_direct Simp_Method = true
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
    89
  | is_proof_method_direct _ = false
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
    90
58499
094efe6ac459 don't affect other subgoals with 'auto' in one-liner proofs
blanchet
parents: 58092
diff changeset
    91
fun is_proof_method_multi_goal Auto_Method = true
094efe6ac459 don't affect other subgoals with 'auto' in one-liner proofs
blanchet
parents: 58092
diff changeset
    92
  | is_proof_method_multi_goal _ = false
094efe6ac459 don't affect other subgoals with 'auto' in one-liner proofs
blanchet
parents: 58092
diff changeset
    93
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
    94
fun maybe_paren s = s |> not (Symbol_Pos.is_identifier s) ? enclose "(" ")"
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
    95
80207
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
    96
(*
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
    97
Combine indexed fact names for pretty-printing.
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
    98
Example: [... "assms(1)" "assms(3)" ...] becomes [... "assms(1,3)" ...]
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
    99
Combines only adjacent same names.
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   100
Input should not have same name with and without index.
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   101
*)
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   102
fun merge_indexed_facts (ss: string list) :string list =
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   103
  let
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   104
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   105
    fun split (s: string) : string * string =
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   106
      case first_field "(" s of
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   107
        NONE => (s,"")
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   108
      | SOME (name,isp) => (name, String.substring (isp, 0, size isp - 1))
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   109
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   110
    fun merge ((name1,is1) :: (name2,is2) :: zs) =
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   111
        if name1 = name2
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   112
        then merge ((name1,is1 ^ "," ^ is2) :: zs)
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   113
        else (name1,is1) :: merge ((name2,is2) :: zs)
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   114
      | merge xs = xs;
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   115
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   116
    fun parents is = if is = "" then "" else "(" ^ is ^ ")"
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   117
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   118
  in
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   119
    map (fn (name,is) => name ^ parents is) (merge (map split ss))
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   120
  end
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   121
72401
2783779b7dd3 removed obsolete unmaintained experimental prover Pirate
blanchet
parents: 70586
diff changeset
   122
fun string_of_proof_method ss meth =
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   123
  let
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   124
    val meth_s =
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   125
      (case meth of
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   126
        Metis_Method (NONE, NONE) => "metis"
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   127
      | Metis_Method (type_enc_opt, lam_trans_opt) =>
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   128
        "metis (" ^ commas (map_filter I [type_enc_opt, lam_trans_opt]) ^ ")"
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   129
      | Meson_Method => "meson"
72798
e732c98b02e6 tuned proof preplay to explicitly refer to Z3 backend
desharna
parents: 72518
diff changeset
   130
      | SMT_Method SMT_Z3 => "smt (z3)"
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   131
      | SMT_Method (SMT_Verit strategy) =>
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   132
        "smt (" ^ commas ("verit" :: (if strategy = "default" then [] else [strategy])) ^ ")"
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   133
      | SATx_Method => "satx"
78695
41273636a82a added argo
blanchet
parents: 75956
diff changeset
   134
      | Argo_Method => "argo"
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   135
      | Blast_Method => "blast"
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   136
      | Simp_Method => if null ss then "simp" else "simp add:"
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   137
      | Auto_Method => "auto"
58092
4ae52c60603a renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
blanchet
parents: 58075
diff changeset
   138
      | Fastforce_Method => "fastforce"
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   139
      | Force_Method => "force"
58092
4ae52c60603a renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
blanchet
parents: 58075
diff changeset
   140
      | Moura_Method => "moura"
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   141
      | Linarith_Method => "linarith"
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   142
      | Presburger_Method => "presburger"
79942
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
   143
      | Algebra_Method => "algebra"
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
   144
      | Order_Method => "order")
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   145
  in
80207
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   146
    maybe_paren (space_implode " " (meth_s :: merge_indexed_facts ss))
56983
132142089ea6 use 'simp add:' syntax in Sledgehammer rather than 'using'
blanchet
parents: 56965
diff changeset
   147
  end
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
   148
57054
blanchet
parents: 56985
diff changeset
   149
fun tac_of_proof_method ctxt (local_facts, global_facts) meth =
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   150
  let
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   151
    fun tac_of_metis (type_enc_opt, lam_trans_opt) =
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   152
      let
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   153
        val ctxt = ctxt
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   154
          |> Config.put Metis_Tactic.verbose false
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   155
          |> Config.put Metis_Tactic.trace false
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   156
      in
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   157
        SELECT_GOAL (Metis_Tactic.metis_method ((Option.map single type_enc_opt, lam_trans_opt),
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   158
          global_facts) ctxt local_facts)
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   159
      end
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   160
72798
e732c98b02e6 tuned proof preplay to explicitly refer to Z3 backend
desharna
parents: 72518
diff changeset
   161
    fun tac_of_smt SMT_Z3 = SMT_Solver.smt_tac
75956
1e2a9d2251b0 remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 75040
diff changeset
   162
      | tac_of_smt (SMT_Verit strategy) = Verit_Strategies.verit_tac_stgy strategy
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   163
  in
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
   164
    (case meth of
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   165
      Metis_Method options => tac_of_metis options
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   166
    | SMT_Method backend => tac_of_smt backend ctxt (local_facts @ global_facts)
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61841
diff changeset
   167
    | _ =>
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   168
      Method.insert_tac ctxt local_facts THEN'
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61841
diff changeset
   169
      (case meth of
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   170
        Meson_Method => Meson_Tactic.meson_general_tac ctxt global_facts
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   171
      | Simp_Method => Simplifier.asm_full_simp_tac (ctxt addsimps global_facts)
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   172
      | _ =>
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   173
        Method.insert_tac ctxt global_facts THEN'
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   174
        (case meth of
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   175
          SATx_Method => SAT.satx_tac ctxt
78695
41273636a82a added argo
blanchet
parents: 75956
diff changeset
   176
        | Argo_Method => Argo_Tactic.argo_tac ctxt []
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   177
        | Blast_Method => blast_tac ctxt
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   178
        | Auto_Method => SELECT_GOAL (Clasimp.auto_tac ctxt)
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   179
        | Fastforce_Method => Clasimp.fast_force_tac ctxt
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   180
        | Force_Method => Clasimp.force_tac ctxt
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   181
        | Moura_Method => moura_tac ctxt
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   182
        | Linarith_Method => Lin_Arith.tac ctxt
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   183
        | Presburger_Method => Cooper.tac true [] [] ctxt
79942
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
   184
        | Algebra_Method => Groebner.algebra_tac [] [] ctxt
7793e3161d2b try proof method "order" in Sledgehammer's proof reconstruction
desharna
parents: 78736
diff changeset
   185
        | Order_Method => HOL_Order_Tac.tac [] ctxt)))
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   186
  end
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
   187
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54824
diff changeset
   188
fun string_of_play_outcome (Played time) = string_of_ext_time (false, time)
56093
4eeb73a1feec simplified preplaying information
blanchet
parents: 56081
diff changeset
   189
  | string_of_play_outcome (Play_Timed_Out time) =
4eeb73a1feec simplified preplaying information
blanchet
parents: 56081
diff changeset
   190
    if time = Time.zeroTime then "" else string_of_ext_time (true, time) ^ ", timed out"
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54824
diff changeset
   191
  | string_of_play_outcome Play_Failed = "failed"
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54824
diff changeset
   192
78736
45867a453a3f used standard Time.compare in Sledgehammer's preplay
desharna
parents: 78695
diff changeset
   193
fun play_outcome_ord (Played time1, Played time2) = Time.compare (time1, time2)
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55212
diff changeset
   194
  | play_outcome_ord (Played _, _) = LESS
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55212
diff changeset
   195
  | play_outcome_ord (_, Played _) = GREATER
78736
45867a453a3f used standard Time.compare in Sledgehammer's preplay
desharna
parents: 78695
diff changeset
   196
  | play_outcome_ord (Play_Timed_Out time1, Play_Timed_Out time2) = Time.compare (time1, time2)
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55212
diff changeset
   197
  | play_outcome_ord (Play_Timed_Out _, _) = LESS
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55212
diff changeset
   198
  | play_outcome_ord (_, Play_Timed_Out _) = GREATER
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55212
diff changeset
   199
  | play_outcome_ord (Play_Failed, Play_Failed) = EQUAL
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55212
diff changeset
   200
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
   201
fun apply_on_subgoal _ 1 = "by "
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
   202
  | apply_on_subgoal 1 _ = "apply "
57287
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
   203
  | apply_on_subgoal i n = "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
   204
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55269
diff changeset
   205
(* FIXME *)
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   206
fun proof_method_command meth i n used_chaineds _(*num_chained*) extras =
60252
2c468c062589 improved one-line preplaying (don't rely on 'using x by simp' to mean 'by (simp add: x)' and beware of inaccessible '(local.)this')
blanchet
parents: 59058
diff changeset
   207
  let
2c468c062589 improved one-line preplaying (don't rely on 'using x by simp' to mean 'by (simp add: x)' and beware of inaccessible '(local.)this')
blanchet
parents: 59058
diff changeset
   208
    val (indirect_ss, direct_ss) =
73975
8d93f9ca6518 revisited ac28714b7478: more faithful preplaying with chained facts
blanchet
parents: 72798
diff changeset
   209
      if is_proof_method_direct meth then ([], extras) else (extras, [])
60252
2c468c062589 improved one-line preplaying (don't rely on 'using x by simp' to mean 'by (simp add: x)' and beware of inaccessible '(local.)this')
blanchet
parents: 59058
diff changeset
   210
  in
80207
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   211
    (if null indirect_ss then ""
ca7e7e41374e pretty-printing sledgehammer command: merge indexed theorems
nipkow
parents: 79942
diff changeset
   212
     else "using " ^ space_implode " " (merge_indexed_facts indirect_ss) ^ " ") ^
72401
2783779b7dd3 removed obsolete unmaintained experimental prover Pirate
blanchet
parents: 70586
diff changeset
   213
    apply_on_subgoal i n ^ string_of_proof_method direct_ss meth ^
58499
094efe6ac459 don't affect other subgoals with 'auto' in one-liner proofs
blanchet
parents: 58092
diff changeset
   214
    (if is_proof_method_multi_goal meth andalso n <> 1 then "[1]" else "")
57287
68aa585269ac given two one-liners, only show the best of the two
blanchet
parents: 57054
diff changeset
   215
  end
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
   216
56093
4eeb73a1feec simplified preplaying information
blanchet
parents: 56081
diff changeset
   217
fun try_command_line banner play command =
4eeb73a1feec simplified preplaying information
blanchet
parents: 56081
diff changeset
   218
  let val s = string_of_play_outcome play in
75040
fada390d49dd tweaked Auto Sledgehammer's behavior and output
blanchet
parents: 75034
diff changeset
   219
    banner ^ Active.sendback_markup_command command ^ (s |> s <> "" ? enclose " (" ")")
56093
4eeb73a1feec simplified preplaying information
blanchet
parents: 56081
diff changeset
   220
  end
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
   221
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   222
fun one_line_proof_text _ num_chained
57739
blanchet
parents: 57735
diff changeset
   223
    ((used_facts, (meth, play)), banner, subgoal, subgoal_count) =
57777
563df8185d98 more informative preplay failures
blanchet
parents: 57776
diff changeset
   224
  let val (chained, extra) = List.partition (fn (_, (sc, _)) => sc = Chained) used_facts in
57735
056a55b44ec7 eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
blanchet
parents: 57720
diff changeset
   225
    map fst extra
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 72401
diff changeset
   226
    |> proof_method_command meth subgoal subgoal_count (map fst chained) num_chained
63692
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 63518
diff changeset
   227
    |> (if play = Play_Failed then prefix "One-line proof reconstruction failed: "
57735
056a55b44ec7 eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
blanchet
parents: 57720
diff changeset
   228
        else try_command_line banner play)
55211
5d027af93a08 moved ML code around
blanchet
parents: 55180
diff changeset
   229
  end
52555
6811291d1869 moved code -> easier debugging
smolkas
parents:
diff changeset
   230
54495
237d5be57277 refactored
blanchet
parents: 52590
diff changeset
   231
end;