src/HOL/Tools/Sledgehammer/metis_tactics.ML
author blanchet
Mon, 04 Oct 2010 14:36:18 +0200
changeset 39933 e764c5cf01fe
parent 39914 2f7b060d0c8d
child 39934 9f116d095e5e
permissions -rw-r--r--
instantiate foralls and release exists in the order described by the topological order
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
39891
8e12f1956fcd "meson_new_skolemizer" -> "metis_new_skolemizer" option (since Meson doesn't support the new skolemizer (yet))
blanchet
parents: 39890
diff changeset
    14
  val new_skolemizer : bool Config.T
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    15
  val metis_tac : Proof.context -> thm list -> int -> tactic
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    16
  val metisF_tac : Proof.context -> thm list -> int -> tactic
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    17
  val metisFT_tac : Proof.context -> thm list -> int -> tactic
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    18
  val setup : theory -> theory
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    19
end
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    20
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    21
structure Metis_Tactics : METIS_TACTICS =
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    22
struct
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    23
39494
bf7dd4902321 rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
blanchet
parents: 39450
diff changeset
    24
open Metis_Translate
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    25
open Metis_Reconstruct
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    26
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
    27
structure Int_Pair_Graph =
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
    28
  Graph(type key = int * int val ord = prod_ord int_ord int_ord)
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
    29
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    30
fun trace_msg msg = if !trace then tracing (msg ()) else ()
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32952
diff changeset
    31
39891
8e12f1956fcd "meson_new_skolemizer" -> "metis_new_skolemizer" option (since Meson doesn't support the new skolemizer (yet))
blanchet
parents: 39890
diff changeset
    32
val (type_lits, type_lits_setup) = Attrib.config_bool "metis_type_lits" (K true)
8e12f1956fcd "meson_new_skolemizer" -> "metis_new_skolemizer" option (since Meson doesn't support the new skolemizer (yet))
blanchet
parents: 39890
diff changeset
    33
val (new_skolemizer, new_skolemizer_setup) =
8e12f1956fcd "meson_new_skolemizer" -> "metis_new_skolemizer" option (since Meson doesn't support the new skolemizer (yet))
blanchet
parents: 39890
diff changeset
    34
  Attrib.config_bool "metis_new_skolemizer" (K false)
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
    35
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    36
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
    37
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    38
fun have_common_thm ths1 ths2 =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    39
  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
    40
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
    41
(*Determining which axiom clauses are actually used*)
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
    42
fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th)
32994
ccc07fbbfefd removed some unreferenced material;
wenzelm
parents: 32956
diff changeset
    43
  | used_axioms _ _ = NONE;
24855
161eb8381b49 metis method: used theorems
paulson
parents: 24827
diff changeset
    44
39450
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    45
val clause_params =
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    46
  {ordering = Metis_KnuthBendixOrder.default,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    47
   orderLiterals = Metis_Clause.UnsignedLiteralOrder,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    48
   orderTerms = true}
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    49
val active_params =
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    50
  {clause = clause_params,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    51
   prefactor = #prefactor Metis_Active.default,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    52
   postfactor = #postfactor Metis_Active.default}
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    53
val waiting_params =
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    54
  {symbolsWeight = 1.0,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    55
   variablesWeight = 0.0,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    56
   literalsWeight = 0.0,
7e9879fbb7c5 supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents: 39419
diff changeset
    57
   models = []}
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
    58
val resolution_params = {active = active_params, waiting = waiting_params}
37573
7f987e8582a7 fewer dependencies
blanchet
parents: 37572
diff changeset
    59
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    60
fun shuffle_key (((axiom_no, (_, index_no)), _), _) = (index_no, axiom_no)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    61
fun shuffle_ord p =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    62
  rev_order (prod_ord int_ord int_ord (pairself shuffle_key p))
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    63
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    64
val copy_prem = @{lemma "P ==> (P ==> P ==> Q) ==> Q" by fast}
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    65
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    66
fun copy_prems_tac [] ns i =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    67
    if forall (curry (op =) 1) ns then all_tac else copy_prems_tac (rev ns) [] i
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    68
  | copy_prems_tac (1 :: ms) ns i =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    69
    rotate_tac 1 i THEN copy_prems_tac ms (1 :: ns) i
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    70
  | copy_prems_tac (m :: ms) ns i =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    71
    etac copy_prem i THEN copy_prems_tac ms (m div 2 :: (m + 1) div 2 :: ns) i
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    72
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    73
fun instantiate_forall_tac thy params t i =
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
    74
  let
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    75
    fun repair (t as (Var ((s, _), _))) =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    76
        (case find_index (fn ((s', _), _) => s' = s) params of
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    77
           ~1 => t
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    78
         | j => Bound j)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    79
      | repair (t $ u) = repair t $ repair u
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    80
      | repair t = t
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    81
    val t' = t |> repair |> fold (curry absdummy) (map snd params)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    82
    fun do_instantiate th =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    83
      let val var = Term.add_vars (prop_of th) [] |> the_single in
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    84
        th |> Thm.instantiate ([], [(cterm_of thy (Var var), cterm_of thy t')])
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    85
      end
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    86
  in
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    87
    etac @{thm allE} i
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    88
    THEN rotate_tac ~1 i
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    89
    THEN PRIMITIVE do_instantiate
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    90
  end
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
    91
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    92
(*### TODO: fix confusion between ax_no and prem_no *)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    93
fun release_clusters_tac _ _ _ _ [] = K all_tac
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    94
  | release_clusters_tac thy ax_counts substs params
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    95
                         ((ax_no, cluster_no) :: clusters) =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    96
    let
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    97
      val n = AList.lookup (op =) ax_counts ax_no |> the
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    98
      fun in_right_cluster s =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
    99
        (s |> Meson_Clausify.cluster_of_zapped_var_name |> fst |> snd |> fst)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   100
        = cluster_no
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   101
      val alls =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   102
        substs
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   103
        |> maps (fn (ax_no', (_, (_, tsubst))) =>
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   104
                    if ax_no' = ax_no then
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   105
                      tsubst |> filter (in_right_cluster
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   106
                                        o fst o fst o dest_Var o fst)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   107
                             |> map snd
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   108
                    else
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   109
                      [])
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   110
      val params' = params
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   111
    in
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   112
      rotate_tac ax_no
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   113
      THEN' EVERY' (map (instantiate_forall_tac thy params) alls)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   114
(*      THEN' ### *)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   115
      THEN' rotate_tac (~ ax_no)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   116
      THEN' release_clusters_tac thy ax_counts substs params' clusters
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   117
   end
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   118
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   119
val cluster_ord =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   120
  prod_ord (prod_ord int_ord (prod_ord int_ord int_ord)) bool_ord
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   121
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   122
(* Attempts to derive the theorem "False" from a theorem of the form
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   123
   "P1 ==> ... ==> Pn ==> False", where the "Pi"s are to be discharged using the
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   124
   specified axioms. The axioms have leading "All" and "Ex" quantifiers, which
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   125
   must be eliminated first. *)
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   126
fun discharge_skolem_premises ctxt axioms prems_imp_false =
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   127
  case prop_of prems_imp_false of
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   128
    @{prop False} => prems_imp_false
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   129
  | prems_imp_false_prop =>
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   130
    let
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   131
      val thy = ProofContext.theory_of ctxt
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   132
      fun match_term p =
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   133
        let
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   134
          val (tyenv, tenv) =
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   135
            Pattern.first_order_match thy p (Vartab.empty, Vartab.empty)
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   136
          val tsubst =
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   137
            tenv |> Vartab.dest
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   138
                 |> sort (cluster_ord
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   139
                          o pairself (Meson_Clausify.cluster_of_zapped_var_name
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   140
                                      o fst o fst))
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   141
                 |> map (Meson.term_pair_of
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   142
                         #> pairself (Envir.subst_term_types tyenv))
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   143
        in (tyenv, tsubst) end
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   144
      fun subst_info_for_prem prem_no prem =
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   145
        case prem of
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   146
          _ $ (Const (@{const_name skolem}, _) $ (_ $ t $ num)) =>
39902
bb43fe4fac93 avoid dependency on "int"
blanchet
parents: 39901
diff changeset
   147
          let val ax_no = HOLogic.dest_nat num in
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   148
            (ax_no, (prem_no, match_term (nth axioms ax_no |> snd, t)))
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   149
          end
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   150
        | _ => raise TERM ("discharge_skolem_premises: Malformed premise",
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   151
                           [prem])
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   152
      fun cluster_of_var_name skolem s =
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   153
        let
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   154
          val ((ax_no, (cluster_no, _)), skolem') =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   155
            Meson_Clausify.cluster_of_zapped_var_name s
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   156
        in
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   157
          if skolem' = skolem andalso cluster_no > 0 then
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   158
            SOME (ax_no, cluster_no)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   159
          else
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   160
            NONE
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   161
        end
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   162
      fun clusters_in_term skolem t =
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   163
        Term.add_var_names t [] |> map_filter (cluster_of_var_name skolem o fst)
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   164
      fun deps_for_term_subst (var, t) =
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   165
        case clusters_in_term false var of
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   166
          [] => NONE
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   167
        | [(ax_no, cluster_no)] =>
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   168
          SOME ((ax_no, cluster_no),
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   169
                clusters_in_term true t
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   170
                |> cluster_no > 1 ? cons (ax_no, cluster_no - 1))
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   171
        | _ => raise TERM ("discharge_skolem_premises: Expected Var", [var])
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   172
      val prems = Logic.strip_imp_prems prems_imp_false_prop
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   173
      val substs = prems |> map2 subst_info_for_prem (0 upto length prems - 1)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   174
                         |> sort (int_ord o pairself fst)
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   175
      val depss = maps (map_filter deps_for_term_subst o snd o snd o snd) substs
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   176
      val clusters = maps (op ::) depss
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   177
      val ordered_clusters =
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   178
        Int_Pair_Graph.empty
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   179
        |> fold Int_Pair_Graph.default_node (map (rpair ()) clusters)
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   180
        |> fold Int_Pair_Graph.add_deps_acyclic depss
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   181
        |> Int_Pair_Graph.topological_order
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   182
        handle Int_Pair_Graph.CYCLES _ =>
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   183
               error "Cannot replay Metis proof in Isabelle without axiom of \
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   184
                     \choice."
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   185
      val params0 =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   186
        [] |> fold Term.add_vars (map snd axioms)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   187
           |> map (`(Meson_Clausify.cluster_of_zapped_var_name o fst o fst))
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   188
           |> filter (fn (((_, (cluster_no, _)), skolem), _) =>
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   189
                         cluster_no = 0 andalso skolem)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   190
           |> sort shuffle_ord |> map snd
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   191
      val ax_counts =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   192
        Inttab.empty
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   193
        |> fold (fn (ax_no, _) => Inttab.map_default (ax_no, 0) (Integer.add 1))
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   194
                substs
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   195
        |> Inttab.dest
39914
2f7b060d0c8d some Poly/ML-specific debugging code escaped in the wild -- comment it out
blanchet
parents: 39902
diff changeset
   196
(* for debugging:
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   197
      fun string_for_subst_info (ax_no, (prem_no, (tyenv, tsubst))) =
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   198
        "ax: " ^ string_of_int ax_no ^ "; asm: " ^ string_of_int prem_no ^
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   199
        "; tyenv: " ^ PolyML.makestring tyenv ^ "; tsubst: {" ^
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   200
        commas (map ((fn (s, t) => s ^ " |-> " ^ t)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   201
                     o pairself (Syntax.string_of_term ctxt)) tsubst) ^ "}"
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   202
      val _ = tracing ("SUBSTS:\n" ^ cat_lines (map string_for_subst_info substs))
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   203
      val _ = tracing ("OUTERMOST SKOLEMS: " ^ PolyML.makestring params0)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   204
      val _ = tracing ("ORDERED CLUSTERS: " ^ PolyML.makestring ordered_clusters)
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   205
      val _ = tracing ("AXIOM COUNT: " ^ PolyML.makestring ax_counts)
39914
2f7b060d0c8d some Poly/ML-specific debugging code escaped in the wild -- comment it out
blanchet
parents: 39902
diff changeset
   206
*)
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   207
    in
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   208
      Goal.prove ctxt [] [] @{prop False}
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   209
          (K (cut_rules_tac (map (fst o nth axioms o fst) ax_counts) 1
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   210
              THEN TRY (REPEAT_ALL_NEW (etac @{thm exE}) 1)
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   211
              THEN copy_prems_tac (map snd ax_counts) [] 1
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   212
              THEN print_tac "copied axioms:"
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   213
              THEN release_clusters_tac thy ax_counts substs params0
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   214
                                        ordered_clusters 1
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   215
              THEN print_tac "released axioms:"
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39894
diff changeset
   216
              THEN match_tac [prems_imp_false] 1
39933
e764c5cf01fe instantiate foralls and release exists in the order described by the topological order
blanchet
parents: 39914
diff changeset
   217
              THEN print_tac "applied rule:"
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   218
              THEN DETERM_UNTIL_SOLVED
39894
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39892
diff changeset
   219
                       (rtac @{thm skolem_COMBK_I} 1
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   220
                        THEN assume_tac 1)))
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   221
    end
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   222
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   223
(* Main function to start Metis proof and reconstruction *)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   224
fun FOL_SOLVE mode ctxt cls ths0 =
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   225
  let val thy = ProofContext.theory_of ctxt
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   226
      val type_lits = Config.get ctxt type_lits
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39899
diff changeset
   227
      val new_skolemizer =
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39899
diff changeset
   228
        Config.get ctxt new_skolemizer orelse null (Meson_Choices.get ctxt)
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
   229
      val th_cls_pairs =
39894
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39892
diff changeset
   230
        map2 (fn j => fn th =>
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39892
diff changeset
   231
                (Thm.get_name_hint th,
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39899
diff changeset
   232
                 Meson_Clausify.cnf_axiom ctxt new_skolemizer j th))
39894
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39892
diff changeset
   233
             (0 upto length ths0 - 1) ths0
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   234
      val thss = map (snd o snd) th_cls_pairs
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   235
      val dischargers = map_filter (fst o snd) th_cls_pairs
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   236
      val _ = trace_msg (fn () => "FOL_SOLVE: CONJECTURE CLAUSES")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   237
      val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   238
      val _ = trace_msg (fn () => "THEOREM CLAUSES")
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   239
      val _ = app (app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th))) thss
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   240
      val (mode, {axioms, tfrees, old_skolems}) =
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   241
        build_logic_map mode ctxt type_lits cls thss
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   242
      val _ = if null tfrees then ()
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   243
              else (trace_msg (fn () => "TFREE CLAUSES");
37643
f576af716aa6 rewrote the TPTP problem generation code more or less from scratch;
blanchet
parents: 37632
diff changeset
   244
                    app (fn TyLitFree ((s, _), (s', _)) =>
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   245
                            trace_msg (fn () => s ^ "(" ^ s' ^ ")")) tfrees)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   246
      val _ = trace_msg (fn () => "CLAUSES GIVEN TO METIS")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   247
      val thms = map #1 axioms
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
   248
      val _ = app (fn th => trace_msg (fn () => Metis_Thm.toString th)) thms
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   249
      val _ = trace_msg (fn () => "mode = " ^ string_of_mode mode)
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   250
      val _ = trace_msg (fn () => "START METIS PROVE PROCESS")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   251
  in
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33316
diff changeset
   252
      case filter (is_false o prop_of) cls of
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   253
          false_th::_ => [false_th RS @{thm FalseE}]
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   254
        | [] =>
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   255
      case Metis_Resolution.new resolution_params {axioms = thms, conjecture = []}
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   256
           |> Metis_Resolution.loop of
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
   257
          Metis_Resolution.Contradiction mth =>
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   258
            let val _ = trace_msg (fn () => "METIS RECONSTRUCTION START: " ^
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
   259
                          Metis_Thm.toString mth)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   260
                val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   261
                             (*add constraints arising from converting goal to clause form*)
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
   262
                val proof = Metis_Proof.proof mth
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   263
                val result =
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   264
                  fold (replay_one_inference ctxt' mode old_skolems) proof axioms
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   265
                and used = map_filter (used_axioms axioms) proof
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   266
                val _ = trace_msg (fn () => "METIS COMPLETED...clauses actually used:")
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   267
                val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) used
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   268
                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
   269
                  if have_common_thm used cls then NONE else SOME name)
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   270
            in
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   271
                if not (null cls) andalso not (have_common_thm used cls) then
36383
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
   272
                  warning "Metis: The assumptions are inconsistent."
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
   273
                else
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
   274
                  ();
6adf1068ac0f better error reporting;
blanchet
parents: 36230
diff changeset
   275
                if not (null unused) then
36230
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   276
                  warning ("Metis: Unused theorems: " ^ commas_quote unused
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   277
                           ^ ".")
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   278
                else
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   279
                  ();
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   280
                case result of
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   281
                    (_,ith)::_ =>
36230
43d10a494c91 added warning about inconsistent context to Metis;
blanchet
parents: 36170
diff changeset
   282
                        (trace_msg (fn () => "Success: " ^ Display.string_of_thm ctxt ith);
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   283
                         [discharge_skolem_premises ctxt dischargers ith])
38097
5e4ad2df09f3 revert exception throwing in FOL_SOLVE, since they're not caught anyway
blanchet
parents: 38028
diff changeset
   284
                  | _ => (trace_msg (fn () => "Metis: No result"); [])
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   285
            end
39419
c9accfd621a5 "Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents: 39376
diff changeset
   286
        | Metis_Resolution.Satisfiable _ =>
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   287
            (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
   288
             [])
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   289
  end;
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   290
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   291
(* 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
   292
   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
   293
   compatibility. *)
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   294
fun also_extensionalize_theorem th =
39890
a1695e2169d0 finished renaming file and module
blanchet
parents: 39887
diff changeset
   295
  let val th' = Meson_Clausify.extensionalize_theorem th in
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   296
    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
   297
    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
   298
  end
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   299
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   300
val neg_clausify =
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   301
  single
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   302
  #> 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
   303
  #> maps also_extensionalize_theorem
39890
a1695e2169d0 finished renaming file and module
blanchet
parents: 39887
diff changeset
   304
  #> map Meson_Clausify.introduce_combinators_in_theorem
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   305
  #> Meson.finish_cnf
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
   306
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
   307
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
   308
  (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
   309
             (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
   310
     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
   311
   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
   312
     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
   313
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38632
diff changeset
   314
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
   315
  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
   316
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   317
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
   318
  let
e6ff246c0cdb renamings + only need second component of name pool to reconstruct proofs
blanchet
parents: 37925
diff changeset
   319
    val _ = trace_msg (fn () =>
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   320
        "Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths))
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   321
  in
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37625
diff changeset
   322
    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
   323
      (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
   324
    else
39594
624d6c0e220d revert b96941dddd04 and c13b4589fddf, which dramatically inflate proof terms
blanchet
parents: 39560
diff changeset
   325
      Meson.MESON (preskolem_tac ctxt) (maps neg_clausify)
624d6c0e220d revert b96941dddd04 and c13b4589fddf, which dramatically inflate proof terms
blanchet
parents: 39560
diff changeset
   326
                  (fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1)
624d6c0e220d revert b96941dddd04 and c13b4589fddf, which dramatically inflate proof terms
blanchet
parents: 39560
diff changeset
   327
                  ctxt i st0
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   328
  end
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   329
37516
c81c86bfc18a have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents: 37509
diff changeset
   330
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
   331
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
   332
val metisFT_tac = generic_metis_tac FT
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   333
38632
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   334
(* 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
   335
   "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
   336
   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
   337
   (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
   338
   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
   339
val has_tvar =
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   340
  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
   341
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
   342
  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
   343
    METHOD (fn facts =>
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   344
               let
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   345
                 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
   346
                   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
   347
               in
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   348
                 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
   349
                           CHANGED_PROP
9cde57cdd0e3 treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents: 38614
diff changeset
   350
                           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
   351
               end)))
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   352
32956
c39860141415 tuned white space;
wenzelm
parents: 32955
diff changeset
   353
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
   354
  type_lits_setup
39891
8e12f1956fcd "meson_new_skolemizer" -> "metis_new_skolemizer" option (since Meson doesn't support the new skolemizer (yet))
blanchet
parents: 39890
diff changeset
   355
  #> new_skolemizer_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
   356
  #> 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
   357
  #> 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
   358
  #> 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
   359
            "Metis for FOL/HOL problems with fully-typed translation"
23442
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   360
028e39e5e8f3 The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff changeset
   361
end;