src/HOL/Tools/Sledgehammer/metis_tactics.ML
author blanchet
Mon, 20 Sep 2010 16:29:55 +0200
changeset 39560 c13b4589fddf
parent 39550 f97fa74c388e
child 39594 624d6c0e220d
permissions -rw-r--r--
preprocess "Ex" before doing clausification in Metis; this is dual to "All" in b96941dddd04
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/metis_tactics.ML
38027
505657ddb047 standardize "Author" tags
blanchet
parents: 38016
diff changeset
     2
    Author:     Kong W. Susanto, Cambridge University Computer Laboratory
505657ddb047 standardize "Author" tags
blanchet
parents: 38016
diff changeset
     3
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
505657ddb047 standardize "Author" tags
blanchet
parents: 38016
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
     5
    Copyright   Cambridge University 2007
23447
1f16190e3836 tuned comments;
wenzelm
parents: 23442
diff changeset
     6
29266
4a478f9d2847 use regular Term.add_vars, Term.add_frees etc.;
wenzelm
parents: 28700
diff changeset
     7
HOL setup for the Metis prover.
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
     8
*)
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
     9
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    10
signature METIS_TACTICS =
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    11
sig
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    12
  val trace : bool Unsynchronized.ref
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    13
  val type_lits : bool Config.T
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    14
  val metis_tac : Proof.context -> thm list -> int -> tactic
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    15
  val metisF_tac : Proof.context -> thm list -> int -> tactic
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    16
  val metisFT_tac : Proof.context -> thm list -> int -> tactic
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    17
  val setup : theory -> theory
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    18
end
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    19
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    20
structure Metis_Tactics : METIS_TACTICS =
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    21
struct
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    22
39494
bf7dd4902321 rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
blanchet
parents: 39450
diff changeset
    23
open Metis_Translate
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    24
open Metis_Reconstruct
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    25
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    26
fun trace_msg msg = if !trace then tracing (msg ()) else ()
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32952
diff changeset
    27
36001
992839c4be90 static defaults for configuration options;
wenzelm
parents: 35865
diff changeset
    28
val (type_lits, type_lits_setup) = Attrib.config_bool "metis_type_lits" (K true);
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    29
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    30
fun is_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    31
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    32
fun have_common_thm ths1 ths2 =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    33
  exists (member Thm.eq_thm ths1) (map Meson.make_meta_clause ths2)
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    34
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    35
(*Determining which axiom clauses are actually used*)
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
    36
fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th)
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32956
diff changeset
    37
  | used_axioms _ _ = NONE;
24855
161eb8381b49 metis method: used theorems
paulson
parents: 24827
diff changeset
    38
39450
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    39
val clause_params =
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    40
  {ordering = Metis_KnuthBendixOrder.default,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    41
   orderLiterals = Metis_Clause.UnsignedLiteralOrder,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    42
   orderTerms = true}
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    43
val active_params =
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    44
  {clause = clause_params,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    45
   prefactor = #prefactor Metis_Active.default,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    46
   postfactor = #postfactor Metis_Active.default}
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    47
val waiting_params =
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    48
  {symbolsWeight = 1.0,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    49
   variablesWeight = 0.0,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    50
   literalsWeight = 0.0,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    51
   models = []}
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    52
val resolution_params = {active = active_params, waiting = waiting_params}
37573
7f987e8582a7 fewer dependencies
blanchet
parents: 37572
diff changeset
    53
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
    54
(* Main function to start Metis proof and reconstruction *)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    55
fun FOL_SOLVE mode ctxt cls ths0 =
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    56
  let val thy = ProofContext.theory_of ctxt
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    57
      val type_lits = Config.get ctxt type_lits
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    58
      val th_cls_pairs =
38016
135f7d489492 get rid of more dead wood
blanchet
parents: 37994
diff changeset
    59
        map (fn th => (Thm.get_name_hint th, Clausifier.cnf_axiom thy th)) ths0
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    60
      val ths = maps #2 th_cls_pairs
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    61
      val _ = trace_msg (fn () => "FOL_SOLVE: CONJECTURE CLAUSES")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    62
      val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    63
      val _ = trace_msg (fn () => "THEOREM CLAUSES")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    64
      val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) ths
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    65
      val (mode, {axioms, tfrees, skolems}) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    66
        build_logic_map mode ctxt type_lits cls ths
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    67
      val _ = if null tfrees then ()
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    68
              else (trace_msg (fn () => "TFREE CLAUSES");
37643
f576af716aa6 rewrote the TPTP problem generation code more or less from scratch;
blanchet
parents: 37632
diff changeset
    69
                    app (fn TyLitFree ((s, _), (s', _)) =>
37573
7f987e8582a7 fewer dependencies
blanchet
parents: 37572
diff changeset
    70
                            trace_msg (fn _ => s ^ "(" ^ s' ^ ")")) tfrees)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    71
      val _ = trace_msg (fn () => "CLAUSES GIVEN TO METIS")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    72
      val thms = map #1 axioms
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
    73
      val _ = app (fn th => trace_msg (fn () => Metis_Thm.toString th)) thms
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    74
      val _ = trace_msg (fn () => "mode = " ^ string_of_mode mode)
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    75
      val _ = trace_msg (fn () => "START METIS PROVE PROCESS")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    76
  in
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33316
diff changeset
    77
      case filter (is_false o prop_of) cls of
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    78
          false_th::_ => [false_th RS @{thm FalseE}]
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    79
        | [] =>
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    80
      case Metis_Resolution.new resolution_params {axioms = thms, conjecture = []}
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    81
           |> Metis_Resolution.loop of
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
    82
          Metis_Resolution.Contradiction mth =>
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    83
            let val _ = trace_msg (fn () => "METIS RECONSTRUCTION START: " ^
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
    84
                          Metis_Thm.toString mth)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    85
                val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    86
                             (*add constraints arising from converting goal to clause form*)
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
    87
                val proof = Metis_Proof.proof mth
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    88
                val result = fold (replay_one_inference ctxt' mode skolems) proof axioms
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    89
                and used = map_filter (used_axioms axioms) proof
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    90
                val _ = trace_msg (fn () => "METIS COMPLETED...clauses actually used:")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    91
                val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) used
33305
wenzelm
parents: 33243
diff changeset
    92
                val unused = th_cls_pairs |> map_filter (fn (name, cls) =>
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    93
                  if have_common_thm used cls then NONE else SOME name)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    94
            in
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    95
                if not (null cls) andalso not (have_common_thm used cls) then
36383
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
    96
                  warning "Metis: The assumptions are inconsistent."
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
    97
                else
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
    98
                  ();
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
    99
                if not (null unused) then
36230
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   100
                  warning ("Metis: Unused theorems: " ^ commas_quote unused
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   101
                           ^ ".")
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   102
                else
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   103
                  ();
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   104
                case result of
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   105
                    (_,ith)::_ =>
36230
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   106
                        (trace_msg (fn () => "Success: " ^ Display.string_of_thm ctxt ith);
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   107
                         [ith])
38097
5e4ad2df09f3 revert exception throwing in FOL_SOLVE, since they're not caught anyway
blanchet
parents: 38028
diff changeset
   108
                  | _ => (trace_msg (fn () => "Metis: No result"); [])
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   109
            end
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
   110
        | Metis_Resolution.Satisfiable _ =>
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   111
            (trace_msg (fn () => "Metis: No first-order proof with the lemmas supplied");
38097
5e4ad2df09f3 revert exception throwing in FOL_SOLVE, since they're not caught anyway
blanchet
parents: 38028
diff changeset
   112
             [])
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   113
  end;
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   114
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   115
(* Extensionalize "th", because that makes sense and that's what Sledgehammer
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   116
   does, but also keep an unextensionalized version of "th" for backward
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   117
   compatibility. *)
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   118
fun also_extensionalize_theorem th =
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   119
  let val th' = Clausifier.extensionalize_theorem th in
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   120
    if Thm.eq_thm (th, th') then [th]
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   121
    else th :: Meson.make_clauses_unsorted [th']
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   122
  end
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   123
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   124
val neg_clausify =
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   125
  single
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   126
  #> Meson.make_clauses_unsorted
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   127
  #> maps also_extensionalize_theorem
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   128
  #> map Clausifier.introduce_combinators_in_theorem
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   129
  #> Meson.finish_cnf
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   130
39269
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39267
diff changeset
   131
fun preskolem_tac ctxt st0 =
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39267
diff changeset
   132
  (if exists (Meson.has_too_many_clauses ctxt)
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39267
diff changeset
   133
             (Logic.prems_of_goal (prop_of st0) 1) then
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39267
diff changeset
   134
     cnf.cnfx_rewrite_tac ctxt 1
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39267
diff changeset
   135
   else
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39267
diff changeset
   136
     all_tac) st0
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39267
diff changeset
   137
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38632
diff changeset
   138
val type_has_top_sort =
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38632
diff changeset
   139
  exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38632
diff changeset
   140
39560
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   141
val preproc_ss =
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   142
  HOL_basic_ss addsimps @{thms all_simps[symmetric] ex_simps[symmetric]}
39548
b96941dddd04 preprocess "All" before doing clausification in Metis;
blanchet
parents: 39497
diff changeset
   143
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   144
fun generic_metis_tac mode ctxt ths i st0 =
37926
e6ff246c0cdb renamings + only need second component of name pool to reconstruct proofs
blanchet
parents: 37925
diff changeset
   145
  let
e6ff246c0cdb renamings + only need second component of name pool to reconstruct proofs
blanchet
parents: 37925
diff changeset
   146
    val _ = trace_msg (fn () =>
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   147
        "Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths))
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   148
  in
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37625
diff changeset
   149
    if exists_type type_has_top_sort (prop_of st0) then
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   150
      (warning ("Metis: Proof state contains the universal sort {}"); Seq.empty)
35568
8fbbfc39508f renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents: 34087
diff changeset
   151
    else
39548
b96941dddd04 preprocess "All" before doing clausification in Metis;
blanchet
parents: 39497
diff changeset
   152
      Tactical.SOLVED'
39560
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   153
          ((TRY o Simplifier.full_simp_tac preproc_ss)
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   154
           THEN' (REPEAT_DETERM o match_tac @{thms allI})
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   155
           THEN' (REPEAT_DETERM o ematch_tac @{thms exE})
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   156
           THEN' (TRY o Meson.MESON (preskolem_tac ctxt) (maps neg_clausify)
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   157
                         (fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1)
c13b4589fddf preprocess "Ex" before doing clausification in Metis;
blanchet
parents: 39550
diff changeset
   158
                         ctxt)) i st0
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   159
  end
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   160
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   161
val metis_tac = generic_metis_tac HO
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   162
val metisF_tac = generic_metis_tac FO
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   163
val metisFT_tac = generic_metis_tac FT
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   164
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   165
(* Whenever "X" has schematic type variables, we treat "using X by metis" as
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   166
   "by (metis X)", to prevent "Subgoal.FOCUS" from freezing the type variables.
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   167
   We don't do it for nonschematic facts "X" because this breaks a few proofs
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   168
   (in the rare and subtle case where a proof relied on extensionality not being
38994
7c655a491bce fiddled with fudge factor (based on Mirabelle)
blanchet
parents: 38864
diff changeset
   169
   applied) and brings few benefits. *)
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   170
val has_tvar =
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   171
  exists_type (exists_subtype (fn TVar _ => true | _ => false)) o prop_of
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   172
fun method name mode =
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   173
  Method.setup name (Attrib.thms >> (fn ths => fn ctxt =>
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   174
    METHOD (fn facts =>
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   175
               let
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   176
                 val (schem_facts, nonschem_facts) =
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   177
                   List.partition has_tvar facts
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   178
               in
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   179
                 HEADGOAL (Method.insert_tac nonschem_facts THEN'
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   180
                           CHANGED_PROP
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   181
                           o generic_metis_tac mode ctxt (schem_facts @ ths))
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   182
               end)))
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   183
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   184
val setup =
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   185
  type_lits_setup
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   186
  #> method @{binding metis} HO "Metis for FOL/HOL problems"
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   187
  #> method @{binding metisF} FO "Metis for FOL problems"
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   188
  #> method @{binding metisFT} FT
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   189
            "Metis for FOL/HOL problems with fully-typed translation"
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   190
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   191
end;