src/HOL/Tools/meson.ML
author wenzelm
Fri, 24 Sep 2010 14:14:21 +0200
changeset 39637 cc3452317b5f
parent 39328 268cd501bdc1
child 39886 8a9f0c97d550
permissions -rw-r--r--
slightly more robust EditBus plumbing wrt. Document_View/Document_Model;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
     1
(*  Title:      HOL/Tools/meson.ML
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
     3
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
     4
The MESON resolution proof procedure for HOL.
29267
8615b4f54047 use exists_subterm directly;
wenzelm
parents: 28397
diff changeset
     5
When making clauses, avoids using the rewriter -- instead uses RS recursively.
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
     6
*)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
     7
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
     8
signature MESON =
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
     9
sig
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    10
  val trace: bool Unsynchronized.ref
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    11
  val term_pair_of: indexname * (typ * 'a) -> term * 'a
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    12
  val flexflex_first_order: thm -> thm
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    13
  val size_of_subgoals: thm -> int
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: 39159
diff changeset
    14
  val has_too_many_clauses: Proof.context -> term -> bool
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
    15
  val make_cnf: thm list -> thm -> Proof.context -> thm list * Proof.context
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    16
  val finish_cnf: thm list -> thm list
38089
ed65a0777e10 perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents: 37926
diff changeset
    17
  val presimplify: thm -> thm
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    18
  val make_nnf: Proof.context -> thm -> thm
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    19
  val skolemize: Proof.context -> thm -> thm
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    20
  val is_fol_term: theory -> term -> bool
35869
cac366550624 start work on direct proof reconstruction for Sledgehammer
blanchet
parents: 35625
diff changeset
    21
  val make_clauses_unsorted: thm list -> thm list
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    22
  val make_clauses: thm list -> thm list
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    23
  val make_horns: thm list -> thm list
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    24
  val best_prolog_tac: (thm -> int) -> thm list -> tactic
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    25
  val depth_prolog_tac: thm list -> tactic
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    26
  val gocls: thm list -> thm list
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    27
  val skolemize_prems_tac: Proof.context -> thm list -> int -> tactic
39037
5146d640aa4a Let MESON take an additional "preskolemization tactic", which can be used to put the goal in definitional CNF
blanchet
parents: 38864
diff changeset
    28
  val MESON:
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: 39159
diff changeset
    29
    tactic -> (thm list -> thm list) -> (thm list -> tactic) -> Proof.context
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39159
diff changeset
    30
    -> int -> tactic
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    31
  val best_meson_tac: (thm -> int) -> Proof.context -> int -> tactic
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    32
  val safe_best_meson_tac: Proof.context -> int -> tactic
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    33
  val depth_meson_tac: Proof.context -> int -> tactic
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    34
  val prolog_step_tac': thm list -> int -> tactic
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    35
  val iter_deepen_prolog_tac: thm list -> tactic
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    36
  val iter_deepen_meson_tac: Proof.context -> thm list -> int -> tactic
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    37
  val make_meta_clause: thm -> thm
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    38
  val make_meta_clauses: thm list -> thm list
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    39
  val meson_tac: Proof.context -> thm list -> int -> tactic
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    40
  val negate_head: thm -> thm
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    41
  val select_literal: int -> thm -> thm
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    42
  val skolemize_tac: Proof.context -> int -> tactic
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
    43
  val setup: theory -> theory
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
    44
end
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
    45
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
    46
structure Meson: MESON =
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
    47
struct
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
    48
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    49
val trace = Unsynchronized.ref false;
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    50
fun trace_msg msg = if ! trace then tracing (msg ()) else ();
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    51
26562
9d25ef112cf6 * Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
paulson
parents: 26424
diff changeset
    52
val max_clauses_default = 60;
38806
0aafc7e81056 more appropriate name for configuration option "meson_max_clauses" (cf. output of 'pront_configs');
wenzelm
parents: 38802
diff changeset
    53
val (max_clauses, setup) = Attrib.config_int "meson_max_clauses" (K max_clauses_default);
26562
9d25ef112cf6 * Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
paulson
parents: 26424
diff changeset
    54
38802
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
    55
(*No known example (on 1-5-2007) needs even thirty*)
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
    56
val iter_deepen_limit = 50;
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
    57
31454
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    58
val disj_forward = @{thm disj_forward};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    59
val disj_forward2 = @{thm disj_forward2};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    60
val make_pos_rule = @{thm make_pos_rule};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    61
val make_pos_rule' = @{thm make_pos_rule'};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    62
val make_pos_goal = @{thm make_pos_goal};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    63
val make_neg_rule = @{thm make_neg_rule};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    64
val make_neg_rule' = @{thm make_neg_rule'};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    65
val make_neg_goal = @{thm make_neg_goal};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    66
val conj_forward = @{thm conj_forward};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    67
val all_forward = @{thm all_forward};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    68
val ex_forward = @{thm ex_forward};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    69
val choice = @{thm choice};
2c0959ab073f dropped legacy ML bindings; tuned
haftmann
parents: 30607
diff changeset
    70
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    71
val not_conjD = @{thm meson_not_conjD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    72
val not_disjD = @{thm meson_not_disjD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    73
val not_notD = @{thm meson_not_notD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    74
val not_allD = @{thm meson_not_allD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    75
val not_exD = @{thm meson_not_exD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    76
val imp_to_disjD = @{thm meson_imp_to_disjD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    77
val not_impD = @{thm meson_not_impD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    78
val iff_to_disjD = @{thm meson_iff_to_disjD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    79
val not_iffD = @{thm meson_not_iffD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    80
val conj_exD1 = @{thm meson_conj_exD1};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    81
val conj_exD2 = @{thm meson_conj_exD2};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    82
val disj_exD = @{thm meson_disj_exD};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    83
val disj_exD1 = @{thm meson_disj_exD1};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    84
val disj_exD2 = @{thm meson_disj_exD2};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    85
val disj_assoc = @{thm meson_disj_assoc};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    86
val disj_comm = @{thm meson_disj_comm};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    87
val disj_FalseD1 = @{thm meson_disj_FalseD1};
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
    88
val disj_FalseD2 = @{thm meson_disj_FalseD2};
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
    89
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
    90
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
    91
(**** Operators for forward proof ****)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
    92
20417
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
    93
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
    94
(** First-order Resolution **)
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
    95
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
    96
fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t);
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
    97
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
    98
(*FIXME: currently does not "rename variables apart"*)
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
    99
fun first_order_resolve thA thB =
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   100
  (case
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   101
    try (fn () =>
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   102
      let val thy = theory_of_thm thA
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   103
          val tmA = concl_of thA
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   104
          val Const("==>",_) $ tmB $ _ = prop_of thB
37398
e194213451c9 beta-eta-contract, to respect "first_order_match"'s specification;
blanchet
parents: 37388
diff changeset
   105
          val tenv =
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   106
            Pattern.first_order_match thy (tmB, tmA)
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   107
                                          (Vartab.empty, Vartab.empty) |> snd
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   108
          val ct_pairs = map (pairself (cterm_of thy) o term_pair_of) (Vartab.dest tenv)
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   109
      in  thA RS (cterm_instantiate ct_pairs thB)  end) () of
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   110
    SOME th => th
37398
e194213451c9 beta-eta-contract, to respect "first_order_match"'s specification;
blanchet
parents: 37388
diff changeset
   111
  | NONE => raise THM ("first_order_resolve", 0, [thA, thB]))
18175
7858b777569a new version of "tryres" allowing multiple unifiers (apparently needed for
paulson
parents: 18141
diff changeset
   112
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   113
fun flexflex_first_order th =
38709
04414091f3b5 structure Thm: less pervasive names;
wenzelm
parents: 38623
diff changeset
   114
  case Thm.tpairs_of th of
23440
37860871f241 Added flexflex_first_order and tidied first_order_resolution
paulson
parents: 23262
diff changeset
   115
      [] => th
37860871f241 Added flexflex_first_order and tidied first_order_resolution
paulson
parents: 23262
diff changeset
   116
    | pairs =>
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   117
        let val thy = theory_of_thm th
39328
blanchet
parents: 39277
diff changeset
   118
            val (_, tenv) =
32032
a6a6e8031c14 tuned/modernized Envir operations;
wenzelm
parents: 31945
diff changeset
   119
              fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   120
            val t_pairs = map term_pair_of (Vartab.dest tenv)
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   121
            val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   122
        in  th'  end
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   123
        handle THM _ => th;
23440
37860871f241 Added flexflex_first_order and tidied first_order_resolution
paulson
parents: 23262
diff changeset
   124
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   125
(*Forward proof while preserving bound variables names*)
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   126
fun rename_bvs_RS th rl =
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   127
  let val th' = th RS rl
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   128
  in  Thm.rename_boundvars (concl_of th') (concl_of th) th' end;
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   129
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   130
(*raises exception if no rules apply*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   131
fun tryres (th, rls) =
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18024
diff changeset
   132
  let fun tryall [] = raise THM("tryres", 0, th::rls)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   133
        | tryall (rl::rls) = (rename_bvs_RS th rl handle THM _ => tryall rls)
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18024
diff changeset
   134
  in  tryall rls  end;
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   135
21050
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   136
(*Permits forward proof from rules that discharge assumptions. The supplied proof state st,
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   137
  e.g. from conj_forward, should have the form
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   138
    "[| P' ==> ?P; Q' ==> ?Q |] ==> ?P & ?Q"
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   139
  and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*)
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   140
fun forward_res ctxt nf st =
21050
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   141
  let fun forward_tacf [prem] = rtac (nf prem) 1
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   142
        | forward_tacf prems =
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   143
            error (cat_lines
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   144
              ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" ::
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   145
                Display.string_of_thm ctxt st ::
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   146
                "Premises:" :: map (Display.string_of_thm ctxt) prems))
21050
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   147
  in
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37539
diff changeset
   148
    case Seq.pull (ALLGOALS (Misc_Legacy.METAHYPS forward_tacf) st)
21050
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   149
    of SOME(th,_) => th
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   150
     | NONE => raise THM("forward_res", 0, [st])
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   151
  end;
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   152
20134
73cb53843190 has_consts renamed to has_conn, now actually parses the first-order formula
paulson
parents: 20119
diff changeset
   153
(*Are any of the logical connectives in "bs" present in the term?*)
73cb53843190 has_consts renamed to has_conn, now actually parses the first-order formula
paulson
parents: 20119
diff changeset
   154
fun has_conns bs =
39328
blanchet
parents: 39277
diff changeset
   155
  let fun has (Const _) = false
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   156
        | has (Const(@{const_name Trueprop},_) $ p) = has p
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   157
        | has (Const(@{const_name Not},_) $ p) = has p
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   158
        | has (Const(@{const_name HOL.disj},_) $ p $ q) = member (op =) bs @{const_name HOL.disj} orelse has p orelse has q
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   159
        | has (Const(@{const_name HOL.conj},_) $ p $ q) = member (op =) bs @{const_name HOL.conj} orelse has p orelse has q
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   160
        | has (Const(@{const_name All},_) $ Abs(_,_,p)) = member (op =) bs @{const_name All} orelse has p
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   161
        | has (Const(@{const_name Ex},_) $ Abs(_,_,p)) = member (op =) bs @{const_name Ex} orelse has p
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   162
        | has _ = false
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   163
  in  has  end;
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   164
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   165
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   166
(**** Clause handling ****)
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   167
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   168
fun literals (Const(@{const_name Trueprop},_) $ P) = literals P
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   169
  | literals (Const(@{const_name HOL.disj},_) $ P $ Q) = literals P @ literals Q
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   170
  | literals (Const(@{const_name Not},_) $ P) = [(false,P)]
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   171
  | literals P = [(true,P)];
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   172
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   173
(*number of literals in a term*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   174
val nliterals = length o literals;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   175
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   176
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   177
(*** Tautology Checking ***)
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   178
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   179
fun signed_lits_aux (Const (@{const_name HOL.disj}, _) $ P $ Q) (poslits, neglits) =
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   180
      signed_lits_aux Q (signed_lits_aux P (poslits, neglits))
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   181
  | signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits)
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   182
  | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits);
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   183
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   184
fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]);
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   185
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   186
(*Literals like X=X are tautologous*)
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38806
diff changeset
   187
fun taut_poslit (Const(@{const_name HOL.eq},_) $ t $ u) = t aconv u
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   188
  | taut_poslit (Const(@{const_name True},_)) = true
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   189
  | taut_poslit _ = false;
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   190
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   191
fun is_taut th =
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   192
  let val (poslits,neglits) = signed_lits th
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   193
  in  exists taut_poslit poslits
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   194
      orelse
20073
da82b545d2de removed obsolete mem_term;
wenzelm
parents: 20018
diff changeset
   195
      exists (member (op aconv) neglits) (HOLogic.false_const :: poslits)
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19875
diff changeset
   196
  end
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   197
  handle TERM _ => false;       (*probably dest_Trueprop on a weird theorem*)
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   198
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   199
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   200
(*** To remove trivial negated equality literals from clauses ***)
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   201
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   202
(*They are typically functional reflexivity axioms and are the converses of
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   203
  injectivity equivalences*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   204
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 39037
diff changeset
   205
val not_refl_disj_D = @{thm meson_not_refl_disj_D};
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   206
20119
7923aacc10c6 fix to refl_clause_aux: added occurs check
paulson
parents: 20073
diff changeset
   207
(*Is either term a Var that does not properly occur in the other term?*)
7923aacc10c6 fix to refl_clause_aux: added occurs check
paulson
parents: 20073
diff changeset
   208
fun eliminable (t as Var _, u) = t aconv u orelse not (Logic.occs(t,u))
7923aacc10c6 fix to refl_clause_aux: added occurs check
paulson
parents: 20073
diff changeset
   209
  | eliminable (u, t as Var _) = t aconv u orelse not (Logic.occs(t,u))
7923aacc10c6 fix to refl_clause_aux: added occurs check
paulson
parents: 20073
diff changeset
   210
  | eliminable _ = false;
7923aacc10c6 fix to refl_clause_aux: added occurs check
paulson
parents: 20073
diff changeset
   211
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   212
fun refl_clause_aux 0 th = th
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   213
  | refl_clause_aux n th =
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   214
       case HOLogic.dest_Trueprop (concl_of th) of
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   215
          (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _) =>
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   216
            refl_clause_aux n (th RS disj_assoc)    (*isolate an atom as first disjunct*)
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38806
diff changeset
   217
        | (Const (@{const_name HOL.disj}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ t $ u)) $ _) =>
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   218
            if eliminable(t,u)
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   219
            then refl_clause_aux (n-1) (th RS not_refl_disj_D)  (*Var inequation: delete*)
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   220
            else refl_clause_aux (n-1) (th RS disj_comm)  (*not between Vars: ignore*)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   221
        | (Const (@{const_name HOL.disj}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   222
        | _ => (*not a disjunction*) th;
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   223
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   224
fun notequal_lits_count (Const (@{const_name HOL.disj}, _) $ P $ Q) =
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   225
      notequal_lits_count P + notequal_lits_count Q
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38806
diff changeset
   226
  | notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _)) = 1
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   227
  | notequal_lits_count _ = 0;
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   228
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   229
(*Simplify a clause by applying reflexivity to its negated equality literals*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   230
fun refl_clause th =
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   231
  let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (concl_of th))
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19875
diff changeset
   232
  in  zero_var_indexes (refl_clause_aux neqs th)  end
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   233
  handle TERM _ => th;  (*probably dest_Trueprop on a weird theorem*)
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   234
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   235
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   236
(*** Removal of duplicate literals ***)
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   237
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   238
(*Forward proof, passing extra assumptions as theorems to the tactic*)
39328
blanchet
parents: 39277
diff changeset
   239
fun forward_res2 nf hyps st =
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   240
  case Seq.pull
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   241
        (REPEAT
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37539
diff changeset
   242
         (Misc_Legacy.METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   243
         st)
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   244
  of SOME(th,_) => th
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   245
   | NONE => raise THM("forward_res2", 0, [st]);
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   246
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   247
(*Remove duplicates in P|Q by assuming ~P in Q
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   248
  rls (initially []) accumulates assumptions of the form P==>False*)
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   249
fun nodups_aux ctxt rls th = nodups_aux ctxt rls (th RS disj_assoc)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   250
    handle THM _ => tryres(th,rls)
39328
blanchet
parents: 39277
diff changeset
   251
    handle THM _ => tryres(forward_res2 (nodups_aux ctxt) rls (th RS disj_forward2),
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   252
                           [disj_FalseD1, disj_FalseD2, asm_rl])
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   253
    handle THM _ => th;
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   254
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   255
(*Remove duplicate literals, if there are any*)
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   256
fun nodups ctxt th =
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   257
  if has_duplicates (op =) (literals (prop_of th))
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   258
    then nodups_aux ctxt [] th
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   259
    else th;
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   260
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   261
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   262
(*** The basic CNF transformation ***)
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   263
39328
blanchet
parents: 39277
diff changeset
   264
fun estimated_num_clauses bound t =
26562
9d25ef112cf6 * Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
paulson
parents: 26424
diff changeset
   265
 let
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: 39159
diff changeset
   266
  fun sum x y = if x < bound andalso y < bound then x+y else bound
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39159
diff changeset
   267
  fun prod x y = if x < bound andalso y < bound then x*y else bound
26562
9d25ef112cf6 * Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
paulson
parents: 26424
diff changeset
   268
  
9d25ef112cf6 * Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
paulson
parents: 26424
diff changeset
   269
  (*Estimate the number of clauses in order to detect infeasible theorems*)
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   270
  fun signed_nclauses b (Const(@{const_name Trueprop},_) $ t) = signed_nclauses b t
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   271
    | signed_nclauses b (Const(@{const_name Not},_) $ t) = signed_nclauses (not b) t
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   272
    | signed_nclauses b (Const(@{const_name HOL.conj},_) $ t $ u) =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   273
        if b then sum (signed_nclauses b t) (signed_nclauses b u)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   274
             else prod (signed_nclauses b t) (signed_nclauses b u)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   275
    | signed_nclauses b (Const(@{const_name HOL.disj},_) $ t $ u) =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   276
        if b then prod (signed_nclauses b t) (signed_nclauses b u)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   277
             else sum (signed_nclauses b t) (signed_nclauses b u)
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38709
diff changeset
   278
    | signed_nclauses b (Const(@{const_name HOL.implies},_) $ t $ u) =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   279
        if b then prod (signed_nclauses (not b) t) (signed_nclauses b u)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   280
             else sum (signed_nclauses (not b) t) (signed_nclauses b u)
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38806
diff changeset
   281
    | signed_nclauses b (Const(@{const_name HOL.eq}, Type ("fun", [T, _])) $ t $ u) =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   282
        if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   283
            if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u))
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   284
                          (prod (signed_nclauses (not b) u) (signed_nclauses b t))
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   285
                 else sum (prod (signed_nclauses b t) (signed_nclauses b u))
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   286
                          (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u))
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   287
        else 1
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   288
    | signed_nclauses b (Const(@{const_name Ex}, _) $ Abs (_,_,t)) = signed_nclauses b t
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   289
    | signed_nclauses b (Const(@{const_name All},_) $ Abs (_,_,t)) = signed_nclauses b t
26562
9d25ef112cf6 * Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
paulson
parents: 26424
diff changeset
   290
    | signed_nclauses _ _ = 1; (* literal *)
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: 39159
diff changeset
   291
 in signed_nclauses true t end
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39159
diff changeset
   292
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39159
diff changeset
   293
fun has_too_many_clauses ctxt t =
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39159
diff changeset
   294
  let val max_cl = Config.get ctxt max_clauses in
39328
blanchet
parents: 39277
diff changeset
   295
    estimated_num_clauses (max_cl + 1) t > max_cl
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: 39159
diff changeset
   296
  end
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19875
diff changeset
   297
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   298
(*Replaces universally quantified variables by FREE variables -- because
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   299
  assumptions may not contain scheme variables.  Later, generalize using Variable.export. *)
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   300
local  
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   301
  val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec))));
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   302
  val spec_varT = #T (Thm.rep_cterm spec_var);
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   303
  fun name_of (Const (@{const_name All}, _) $ Abs(x,_,_)) = x | name_of _ = Name.uu;
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   304
in  
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   305
  fun freeze_spec th ctxt =
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   306
    let
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   307
      val cert = Thm.cterm_of (ProofContext.theory_of ctxt);
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   308
      val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (concl_of th))] ctxt;
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   309
      val spec' = Thm.instantiate ([], [(spec_var, cert (Free (x, spec_varT)))]) spec;
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   310
    in (th RS spec', ctxt') end
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   311
end;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   312
15998
bc916036cf84 new cnf function taking Skolemization theorems as an extra argument
paulson
parents: 15965
diff changeset
   313
(*Used with METAHYPS below. There is one assumption, which gets bound to prem
bc916036cf84 new cnf function taking Skolemization theorems as an extra argument
paulson
parents: 15965
diff changeset
   314
  and then normalized via function nf. The normal form is given to resolve_tac,
22515
f4061faa5fda "generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
paulson
parents: 22381
diff changeset
   315
  instantiate a Boolean variable created by resolution with disj_forward. Since
f4061faa5fda "generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
paulson
parents: 22381
diff changeset
   316
  (nf prem) returns a LIST of theorems, we can backtrack to get all combinations.*)
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   317
fun resop nf [prem] = resolve_tac (nf prem) 1;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   318
39037
5146d640aa4a Let MESON take an additional "preskolemization tactic", which can be used to put the goal in definitional CNF
blanchet
parents: 38864
diff changeset
   319
(* Any need to extend this list with "HOL.type_class", "HOL.eq_class",
5146d640aa4a Let MESON take an additional "preskolemization tactic", which can be used to put the goal in definitional CNF
blanchet
parents: 38864
diff changeset
   320
   and "Pure.term"? *)
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   321
val has_meta_conn = exists_Const (member (op =) ["==", "==>", "=simp=>", "all", "prop"] o #1);
20417
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
   322
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   323
fun apply_skolem_theorem (th, rls) =
37398
e194213451c9 beta-eta-contract, to respect "first_order_match"'s specification;
blanchet
parents: 37388
diff changeset
   324
  let
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   325
    fun tryall [] = raise THM ("apply_skolem_theorem", 0, th::rls)
37398
e194213451c9 beta-eta-contract, to respect "first_order_match"'s specification;
blanchet
parents: 37388
diff changeset
   326
      | tryall (rl :: rls) =
e194213451c9 beta-eta-contract, to respect "first_order_match"'s specification;
blanchet
parents: 37388
diff changeset
   327
        first_order_resolve th rl handle THM _ => tryall rls
e194213451c9 beta-eta-contract, to respect "first_order_match"'s specification;
blanchet
parents: 37388
diff changeset
   328
  in tryall rls end
22515
f4061faa5fda "generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
paulson
parents: 22381
diff changeset
   329
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   330
(* Conjunctive normal form, adding clauses from th in front of ths (for foldr).
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   331
   Strips universal quantifiers and breaks up conjunctions.
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   332
   Eliminates existential quantifiers using Skolemization theorems. *)
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   333
fun cnf skolem_ths ctxt (th, ths) =
33222
89ced80833ac critical comments;
wenzelm
parents: 32960
diff changeset
   334
  let val ctxtr = Unsynchronized.ref ctxt   (* FIXME ??? *)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   335
      fun cnf_aux (th,ths) =
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   336
        if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   337
        else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (prop_of th))
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   338
        then nodups ctxt th :: ths (*no work to do, terminate*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   339
        else case head_of (HOLogic.dest_Trueprop (concl_of th)) of
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   340
            Const (@{const_name HOL.conj}, _) => (*conjunction*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   341
                cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths))
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   342
          | Const (@{const_name All}, _) => (*universal quantifier*)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   343
                let val (th',ctxt') = freeze_spec th (!ctxtr)
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   344
                in  ctxtr := ctxt'; cnf_aux (th', ths) end
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   345
          | Const (@{const_name Ex}, _) =>
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   346
              (*existential quantifier: Insert Skolem functions*)
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   347
              cnf_aux (apply_skolem_theorem (th, skolem_ths), ths)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   348
          | Const (@{const_name HOL.disj}, _) =>
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   349
              (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   350
                all combinations of converting P, Q to CNF.*)
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   351
              let val tac =
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37539
diff changeset
   352
                  Misc_Legacy.METAHYPS (resop cnf_nil) 1 THEN
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37539
diff changeset
   353
                   (fn st' => st' |> Misc_Legacy.METAHYPS (resop cnf_nil) 1)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   354
              in  Seq.list_of (tac (th RS disj_forward)) @ ths  end
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   355
          | _ => nodups ctxt th :: ths  (*no work to do*)
19154
f48e36b7d8d4 fixed but in freeze_spec
paulson
parents: 19112
diff changeset
   356
      and cnf_nil th = cnf_aux (th,[])
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: 39159
diff changeset
   357
      val cls =
c2795d8a2461 use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
blanchet
parents: 39159
diff changeset
   358
            if has_too_many_clauses ctxt (concl_of th)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   359
            then (trace_msg (fn () => "cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32955
diff changeset
   360
            else cnf_aux (th,ths)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   361
  in  (cls, !ctxtr)  end;
22515
f4061faa5fda "generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
paulson
parents: 22381
diff changeset
   362
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37398
diff changeset
   363
fun make_cnf skolem_ths th ctxt = cnf skolem_ths ctxt (th, []);
20417
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
   364
c611b1412056 better skolemization, using first-order resolution rather than hoping for the right result
paulson
parents: 20361
diff changeset
   365
(*Generalization, removal of redundant equalities, removal of tautologies.*)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   366
fun finish_cnf ths = filter (not o is_taut) (map refl_clause ths);
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   367
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   368
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   369
(**** Generation of contrapositives ****)
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   370
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   371
fun is_left (Const (@{const_name Trueprop}, _) $
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   372
               (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _)) = true
21102
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   373
  | is_left _ = false;
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   374
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   375
(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   376
fun assoc_right th =
21102
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   377
  if is_left (prop_of th) then assoc_right (th RS disj_assoc)
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   378
  else th;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   379
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   380
(*Must check for negative literal first!*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   381
val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule];
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   382
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   383
(*For ordinary resolution. *)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   384
val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule'];
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   385
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   386
(*Create a goal or support clause, conclusing False*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   387
fun make_goal th =   (*Must check for negative literal first!*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   388
    make_goal (tryres(th, clause_rules))
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   389
  handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]);
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   390
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   391
(*Sort clauses by number of literals*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   392
fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2);
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   393
18389
8352b1d3b639 removal of functional reflexivity axioms
paulson
parents: 18194
diff changeset
   394
fun sort_clauses ths = sort (make_ord fewerlits) ths;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   395
38099
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   396
fun has_bool @{typ bool} = true
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   397
  | has_bool (Type (_, Ts)) = exists has_bool Ts
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   398
  | has_bool _ = false
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   399
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   400
fun has_fun (Type (@{type_name fun}, _)) = true
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   401
  | has_fun (Type (_, Ts)) = exists has_fun Ts
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   402
  | has_fun _ = false
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   403
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   404
(*Is the string the name of a connective? Really only | and Not can remain,
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   405
  since this code expects to be called on a clause form.*)
19875
7405ce9d4f25 avoid unqualified exception names;
wenzelm
parents: 19728
diff changeset
   406
val is_conn = member (op =)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   407
    [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38709
diff changeset
   408
     @{const_name HOL.implies}, @{const_name Not},
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   409
     @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex}];
15613
ab90e95ae02e meson now checks that problems are first-order
paulson
parents: 15581
diff changeset
   410
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   411
(*True if the term contains a function--not a logical connective--where the type
20524
1493053fc111 Tweaks to is_fol_term, the first-order test. We don't count "=" as a connective
paulson
parents: 20417
diff changeset
   412
  of any argument contains bool.*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   413
val has_bool_arg_const =
15613
ab90e95ae02e meson now checks that problems are first-order
paulson
parents: 15581
diff changeset
   414
    exists_Const
38099
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   415
      (fn (c,T) => not(is_conn c) andalso exists has_bool (binder_types T));
22381
cb145d434284 The first-order test now tests for the obscure case of a polymorphic constant like 1 being
paulson
parents: 22360
diff changeset
   416
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   417
(*A higher-order instance of a first-order constant? Example is the definition of
38622
86fc906dcd86 split and enriched theory SetsAndFunctions
haftmann
parents: 38557
diff changeset
   418
  one, 1, at a function type in theory Function_Algebras.*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   419
fun higher_inst_const thy (c,T) =
22381
cb145d434284 The first-order test now tests for the obscure case of a polymorphic constant like 1 being
paulson
parents: 22360
diff changeset
   420
  case binder_types T of
cb145d434284 The first-order test now tests for the obscure case of a polymorphic constant like 1 being
paulson
parents: 22360
diff changeset
   421
      [] => false (*not a function type, OK*)
cb145d434284 The first-order test now tests for the obscure case of a polymorphic constant like 1 being
paulson
parents: 22360
diff changeset
   422
    | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts;
cb145d434284 The first-order test now tests for the obscure case of a polymorphic constant like 1 being
paulson
parents: 22360
diff changeset
   423
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24300
diff changeset
   424
(*Returns false if any Vars in the theorem mention type bool.
21102
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   425
  Also rejects functions whose arguments are Booleans or other functions.*)
22381
cb145d434284 The first-order test now tests for the obscure case of a polymorphic constant like 1 being
paulson
parents: 22360
diff changeset
   426
fun is_fol_term thy t =
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   427
    Term.is_first_order ["all", @{const_name All}, @{const_name Ex}] t andalso
38099
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   428
    not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   429
                           | _ => false) t orelse
e3bb96b83807 fix Meson's definition of first-orderness to prevent errors later on elsewhere (e.g. in Metis)
blanchet
parents: 38089
diff changeset
   430
         has_bool_arg_const t orelse
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   431
         exists_Const (higher_inst_const thy) t orelse
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   432
         has_meta_conn t);
19204
b8f7de7ef629 Tidying, and getting rid of SELECT_GOAL (as it does something different now)
paulson
parents: 19154
diff changeset
   433
21102
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   434
fun rigid t = not (is_Var (head_of t));
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   435
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   436
fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name HOL.disj}, _) $ t $ _)) = rigid t
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   437
  | ok4horn (Const (@{const_name Trueprop},_) $ t) = rigid t
21102
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   438
  | ok4horn _ = false;
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   439
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   440
(*Create a meta-level Horn clause*)
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   441
fun make_horn crules th =
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   442
  if ok4horn (concl_of th)
21102
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   443
  then make_horn crules (tryres(th,crules)) handle THM _ => th
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   444
  else th;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   445
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   446
(*Generate Horn clauses for all contrapositives of a clause. The input, th,
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   447
  is a HOL disjunction.*)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   448
fun add_contras crules th hcs =
39328
blanchet
parents: 39277
diff changeset
   449
  let fun rots (0,_) = hcs
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   450
        | rots (k,th) = zero_var_indexes (make_horn crules th) ::
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   451
                        rots(k-1, assoc_right (th RS disj_comm))
15862
67574c1b15a0 removed unnecessary (?) check
paulson
parents: 15779
diff changeset
   452
  in case nliterals(prop_of th) of
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   453
        1 => th::hcs
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   454
      | n => rots(n, assoc_right th)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   455
  end;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   456
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   457
(*Use "theorem naming" to label the clauses*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   458
fun name_thms label =
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   459
    let fun name1 th (k, ths) =
27865
27a8ad9612a3 moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
wenzelm
parents: 27239
diff changeset
   460
          (k-1, Thm.put_name_hint (label ^ string_of_int k) th :: ths)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   461
    in  fn ths => #2 (fold_rev name1 ths (length ths, []))  end;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   462
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   463
(*Is the given disjunction an all-negative support clause?*)
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   464
fun is_negative th = forall (not o #1) (literals (prop_of th));
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   465
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33245
diff changeset
   466
val neg_clauses = filter is_negative;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   467
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   468
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   469
(***** MESON PROOF PROCEDURE *****)
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   470
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   471
fun rhyps (Const("==>",_) $ (Const(@{const_name Trueprop},_) $ A) $ phi,
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   472
           As) = rhyps(phi, A::As)
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   473
  | rhyps (_, As) = As;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   474
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   475
(** Detecting repeated assumptions in a subgoal **)
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   476
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   477
(*The stringtree detects repeated assumptions.*)
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33222
diff changeset
   478
fun ins_term t net = Net.insert_term (op aconv) (t, t) net;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   479
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   480
(*detects repetitions in a list of terms*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   481
fun has_reps [] = false
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   482
  | has_reps [_] = false
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   483
  | has_reps [t,u] = (t aconv u)
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33222
diff changeset
   484
  | has_reps ts = (fold ins_term ts Net.empty; false) handle Net.INSERT => true;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   485
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   486
(*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*)
18508
paulson
parents: 18405
diff changeset
   487
fun TRYING_eq_assume_tac 0 st = Seq.single st
paulson
parents: 18405
diff changeset
   488
  | TRYING_eq_assume_tac i st =
31945
d5f186aa0bed structure Thm: less pervasive names;
wenzelm
parents: 31454
diff changeset
   489
       TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st)
18508
paulson
parents: 18405
diff changeset
   490
       handle THM _ => TRYING_eq_assume_tac (i-1) st;
paulson
parents: 18405
diff changeset
   491
paulson
parents: 18405
diff changeset
   492
fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (nprems_of st) st;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   493
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   494
(*Loop checking: FAIL if trying to prove the same thing twice
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   495
  -- if *ANY* subgoal has repeated literals*)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   496
fun check_tac st =
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   497
  if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   498
  then  Seq.empty  else  Seq.single st;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   499
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   500
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   501
(* net_resolve_tac actually made it slower... *)
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   502
fun prolog_step_tac horns i =
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   503
    (assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN
18508
paulson
parents: 18405
diff changeset
   504
    TRYALL_eq_assume_tac;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   505
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   506
(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   507
fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz;
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   508
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   509
fun size_of_subgoals st = fold_rev addconcl (prems_of st) 0;
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15574
diff changeset
   510
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   511
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   512
(*Negation Normal Form*)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   513
val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD,
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   514
               not_impD, not_iffD, not_allD, not_exD, not_notD];
15581
f07e865d9d40 now checks for higher-order vars
paulson
parents: 15579
diff changeset
   515
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   516
fun ok4nnf (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ t)) = rigid t
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   517
  | ok4nnf (Const (@{const_name Trueprop},_) $ t) = rigid t
21102
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   518
  | ok4nnf _ = false;
7f2ebe5c5b72 Conversion to clause form now tolerates Boolean variables without looping.
paulson
parents: 21095
diff changeset
   519
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   520
fun make_nnf1 ctxt th =
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   521
  if ok4nnf (concl_of th)
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   522
  then make_nnf1 ctxt (tryres(th, nnf_rls))
28174
626f0a79a4b9 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
paulson
parents: 27865
diff changeset
   523
    handle THM ("tryres", _, _) =>
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   524
        forward_res ctxt (make_nnf1 ctxt)
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   525
           (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward]))
28174
626f0a79a4b9 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
paulson
parents: 27865
diff changeset
   526
    handle THM ("tryres", _, _) => th
38608
01ed56c46259 beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents: 38606
diff changeset
   527
  else th
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   528
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   529
(*The simplification removes defined quantifiers and occurrences of True and False.
20018
5419a71d0baa removed the "tagging" feature
paulson
parents: 19894
diff changeset
   530
  nnf_ss also includes the one-point simprocs,
18405
afb1a52a7011 removal of some redundancies (e.g. one-point-rules) in clause production
paulson
parents: 18389
diff changeset
   531
  which are needed to avoid the various one-point theorems from generating junk clauses.*)
19894
7c7e15b27145 the "all_theorems" option and some fixes
paulson
parents: 19875
diff changeset
   532
val nnf_simps =
37539
c80e77e8d036 cosmetics
blanchet
parents: 37410
diff changeset
   533
  @{thms simp_implies_def Ex1_def Ball_def Bex_def if_True if_False if_cancel
c80e77e8d036 cosmetics
blanchet
parents: 37410
diff changeset
   534
         if_eq_cancel cases_simp}
c80e77e8d036 cosmetics
blanchet
parents: 37410
diff changeset
   535
val nnf_extra_simps = @{thms split_ifs ex_simps all_simps simp_thms}
18405
afb1a52a7011 removal of some redundancies (e.g. one-point-rules) in clause production
paulson
parents: 18389
diff changeset
   536
afb1a52a7011 removal of some redundancies (e.g. one-point-rules) in clause production
paulson
parents: 18389
diff changeset
   537
val nnf_ss =
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   538
  HOL_basic_ss addsimps nnf_extra_simps
24040
0d5cf52ebf87 proper simproc_setup for "neq", "let_simp";
wenzelm
parents: 23894
diff changeset
   539
    addsimprocs [defALL_regroup,defEX_regroup, @{simproc neq}, @{simproc let_simp}];
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15862
diff changeset
   540
38089
ed65a0777e10 perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents: 37926
diff changeset
   541
val presimplify =
ed65a0777e10 perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents: 37926
diff changeset
   542
  rewrite_rule (map safe_mk_meta_eq nnf_simps)
ed65a0777e10 perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents: 37926
diff changeset
   543
  #> simplify nnf_ss
ed65a0777e10 perform the presimplification done by Metis.make_nnf in Sledgehammer again, to deal with "If" and similar constructs
blanchet
parents: 37926
diff changeset
   544
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   545
fun make_nnf ctxt th = case prems_of th of
38606
3003ddbd46d9 encode "fequal" reasoning rules in Metis problem, just as is done for Sledgehammer -- otherwise any proof that relies on "fequal" found by Sledgehammer can't be reconstructed
blanchet
parents: 38099
diff changeset
   546
    [] => th |> presimplify |> make_nnf1 ctxt
21050
d0447a511edb More robust error handling in make_nnf and forward_res
paulson
parents: 21046
diff changeset
   547
  | _ => raise THM ("make_nnf: premises in argument", 0, [th]);
15581
f07e865d9d40 now checks for higher-order vars
paulson
parents: 15579
diff changeset
   548
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15946
diff changeset
   549
(*Pull existential quantifiers to front. This accomplishes Skolemization for
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15946
diff changeset
   550
  clauses that arise from a subgoal.*)
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   551
fun skolemize1 ctxt th =
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
   552
  if not (has_conns [@{const_name Ex}] (prop_of th)) then th
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   553
  else (skolemize1 ctxt (tryres(th, [choice, conj_exD1, conj_exD2,
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15653
diff changeset
   554
                              disj_exD, disj_exD1, disj_exD2])))
28174
626f0a79a4b9 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
paulson
parents: 27865
diff changeset
   555
    handle THM ("tryres", _, _) =>
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   556
        skolemize1 ctxt (forward_res ctxt (skolemize1 ctxt)
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   557
                   (tryres (th, [conj_forward, disj_forward, all_forward])))
28174
626f0a79a4b9 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
paulson
parents: 27865
diff changeset
   558
    handle THM ("tryres", _, _) => 
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   559
        forward_res ctxt (skolemize1 ctxt) (rename_bvs_RS th ex_forward);
29684
40bf9f4e7a4e Minor reorganisation of the Skolemization code
paulson
parents: 29267
diff changeset
   560
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   561
fun skolemize ctxt th = skolemize1 ctxt (make_nnf ctxt th);
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   562
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   563
fun skolemize_nnf_list _ [] = []
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   564
  | skolemize_nnf_list ctxt (th::ths) =
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   565
      skolemize ctxt th :: skolemize_nnf_list ctxt ths
25710
4cdf7de81e1b Replaced refs by config params; finer critical section in mets method
paulson
parents: 25694
diff changeset
   566
      handle THM _ => (*RS can fail if Unify.search_bound is too small*)
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
   567
       (trace_msg (fn () => "Failed to Skolemize " ^ Display.string_of_thm ctxt th);
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   568
        skolemize_nnf_list ctxt ths);
25694
cbb59ba6bf0c Skolemization now catches exception THM, which may be raised if unification fails.
paulson
parents: 24937
diff changeset
   569
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   570
fun add_clauses th cls =
36603
d5d6111761a6 renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
wenzelm
parents: 36001
diff changeset
   571
  let val ctxt0 = Variable.global_thm_context th
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   572
      val (cnfs, ctxt) = make_cnf [] th ctxt0
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24827
diff changeset
   573
  in Variable.export ctxt ctxt0 cnfs @ cls end;
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   574
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   575
(*Make clauses from a list of theorems, previously Skolemized and put into nnf.
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   576
  The resulting clauses are HOL disjunctions.*)
39235
cda88e68106d Meson.make_clauses_unsorted: removed spurious debug code stemming from 5146d640aa4a -- must not handle arbitrary exceptions in user space;
wenzelm
parents: 39159
diff changeset
   577
fun make_clauses_unsorted ths = fold_rev add_clauses ths [];
35869
cac366550624 start work on direct proof reconstruction for Sledgehammer
blanchet
parents: 35625
diff changeset
   578
val make_clauses = sort_clauses o make_clauses_unsorted;
15773
f14ae2432710 Completed integration of reconstruction code. Now finds and displays proofs when used with modified version
quigley
parents: 15736
diff changeset
   579
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   580
(*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*)
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   581
fun make_horns ths =
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   582
    name_thms "Horn#"
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   583
      (distinct Thm.eq_thm_prop (fold_rev (add_contras clause_rules) ths []));
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   584
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   585
(*Could simply use nprems_of, which would count remaining subgoals -- no
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   586
  discrimination as to their size!  With BEST_FIRST, fails for problem 41.*)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   587
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   588
fun best_prolog_tac sizef horns =
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   589
    BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1);
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   590
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   591
fun depth_prolog_tac horns =
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   592
    DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1);
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   593
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   594
(*Return all negative clauses, as possible goal clauses*)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   595
fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls));
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   596
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   597
fun skolemize_prems_tac ctxt prems =
37926
e6ff246c0cdb renamings + only need second component of name pool to reconstruct proofs
blanchet
parents: 37781
diff changeset
   598
  cut_facts_tac (skolemize_nnf_list ctxt prems) THEN' REPEAT o etac exE
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   599
22546
c40d7ab8cbc5 MESON tactical takes an additional argument: the clausification function.
paulson
parents: 22515
diff changeset
   600
(*Basis of all meson-tactics.  Supplies cltac with clauses: HOL disjunctions.
c40d7ab8cbc5 MESON tactical takes an additional argument: the clausification function.
paulson
parents: 22515
diff changeset
   601
  Function mkcl converts theorems to clauses.*)
39037
5146d640aa4a Let MESON take an additional "preskolemization tactic", which can be used to put the goal in definitional CNF
blanchet
parents: 38864
diff changeset
   602
fun MESON preskolem_tac mkcl cltac ctxt i st =
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   603
  SELECT_GOAL
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35410
diff changeset
   604
    (EVERY [Object_Logic.atomize_prems_tac 1,
23552
6403d06abe25 to handle non-atomic assumptions
paulson
parents: 23440
diff changeset
   605
            rtac ccontr 1,
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: 39159
diff changeset
   606
            preskolem_tac,
32283
3bebc195c124 qualified Subgoal.FOCUS;
wenzelm
parents: 32262
diff changeset
   607
            Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} =>
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: 39159
diff changeset
   608
                      EVERY1 [skolemize_prems_tac ctxt negs,
32283
3bebc195c124 qualified Subgoal.FOCUS;
wenzelm
parents: 32262
diff changeset
   609
                              Subgoal.FOCUS (cltac o mkcl o #prems) ctxt']) ctxt 1]) i st
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   610
  handle THM _ => no_tac st;    (*probably from make_meta_clause, not first-order*)
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   611
39037
5146d640aa4a Let MESON take an additional "preskolemization tactic", which can be used to put the goal in definitional CNF
blanchet
parents: 38864
diff changeset
   612
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   613
(** Best-first search versions **)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   614
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   615
(*ths is a list of additional clauses (HOL disjunctions) to use.*)
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   616
fun best_meson_tac sizef =
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: 39159
diff changeset
   617
  MESON all_tac make_clauses
22546
c40d7ab8cbc5 MESON tactical takes an additional argument: the clausification function.
paulson
parents: 22515
diff changeset
   618
    (fn cls =>
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   619
         THEN_BEST_FIRST (resolve_tac (gocls cls) 1)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   620
                         (has_fewer_prems 1, sizef)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   621
                         (prolog_step_tac (make_horns cls) 1));
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   622
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   623
(*First, breaks the goal into independent units*)
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   624
fun safe_best_meson_tac ctxt =
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   625
     SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   626
                  TRYALL (best_meson_tac size_of_subgoals ctxt));
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   627
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   628
(** Depth-first search version **)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   629
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   630
val depth_meson_tac =
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: 39159
diff changeset
   631
  MESON all_tac make_clauses
22546
c40d7ab8cbc5 MESON tactical takes an additional argument: the clausification function.
paulson
parents: 22515
diff changeset
   632
    (fn cls => EVERY [resolve_tac (gocls cls) 1, depth_prolog_tac (make_horns cls)]);
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   633
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   634
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   635
(** Iterative deepening version **)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   636
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   637
(*This version does only one inference per call;
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   638
  having only one eq_assume_tac speeds it up!*)
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   639
fun prolog_step_tac' horns =
39328
blanchet
parents: 39277
diff changeset
   640
    let val (horn0s, _) = (*0 subgoals vs 1 or more*)
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   641
            take_prefix Thm.no_prems horns
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   642
        val nrtac = net_resolve_tac horns
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   643
    in  fn i => eq_assume_tac i ORELSE
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   644
                match_tac horn0s i ORELSE  (*no backtracking if unit MATCHES*)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   645
                ((assume_tac i APPEND nrtac i) THEN check_tac)
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   646
    end;
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   647
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   648
fun iter_deepen_prolog_tac horns =
38802
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
   649
    ITER_DEEPEN iter_deepen_limit (has_fewer_prems 1) (prolog_step_tac' horns);
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   650
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: 39159
diff changeset
   651
fun iter_deepen_meson_tac ctxt ths = ctxt |> MESON all_tac make_clauses
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   652
  (fn cls =>
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   653
    (case (gocls (cls @ ths)) of
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   654
      [] => no_tac  (*no goal clauses*)
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   655
    | goes =>
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   656
        let
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   657
          val horns = make_horns (cls @ ths)
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
   658
          val _ = trace_msg (fn () =>
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 32032
diff changeset
   659
            cat_lines ("meson method called:" ::
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   660
              map (Display.string_of_thm ctxt) (cls @ ths) @
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   661
              ["clauses:"] @ map (Display.string_of_thm ctxt) horns))
38802
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
   662
        in
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
   663
          THEN_ITER_DEEPEN iter_deepen_limit
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
   664
            (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns)
a925c0ee42f7 clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents: 38795
diff changeset
   665
        end));
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   666
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   667
fun meson_tac ctxt ths =
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   668
  SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN TRYALL (iter_deepen_meson_tac ctxt ths));
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   669
95dca9f991f2 improved meson setup;
wenzelm
parents: 9840
diff changeset
   670
14813
826edbc317e6 new skolemize_tac and skolemize method
paulson
parents: 14763
diff changeset
   671
(**** Code to support ordinary resolution, rather than Model Elimination ****)
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   672
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   673
(*Convert a list of clauses (disjunctions) to meta-level clauses (==>),
15008
5abd18710a1f new method for explicit classical resolution
paulson
parents: 14890
diff changeset
   674
  with no contrapositives, for ordinary resolution.*)
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   675
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   676
(*Rules to convert the head literal into a negated assumption. If the head
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   677
  literal is already negated, then using notEfalse instead of notEfalse'
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   678
  prevents a double negation.*)
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27230
diff changeset
   679
val notEfalse = read_instantiate @{context} [(("R", 0), "False")] notE;
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   680
val notEfalse' = rotate_prems 1 notEfalse;
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   681
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   682
fun negated_asm_of_head th =
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   683
    th RS notEfalse handle THM _ => th RS notEfalse';
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   684
26066
19df083a2bbf make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
paulson
parents: 25710
diff changeset
   685
(*Converting one theorem from a disjunction to a meta-level clause*)
19df083a2bbf make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
paulson
parents: 25710
diff changeset
   686
fun make_meta_clause th =
33832
cff42395c246 explicitly mark some legacy freeze operations;
wenzelm
parents: 33339
diff changeset
   687
  let val (fth,thaw) = Drule.legacy_freeze_thaw_robust th
26066
19df083a2bbf make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
paulson
parents: 25710
diff changeset
   688
  in  
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   689
      (zero_var_indexes o Thm.varifyT_global o thaw 0 o 
26066
19df083a2bbf make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
paulson
parents: 25710
diff changeset
   690
       negated_asm_of_head o make_horn resolution_clause_rules) fth
19df083a2bbf make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
paulson
parents: 25710
diff changeset
   691
  end;
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   692
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   693
fun make_meta_clauses ths =
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   694
    name_thms "MClause#"
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 22130
diff changeset
   695
      (distinct Thm.eq_thm_prop (map make_meta_clause ths));
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   696
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   697
(*Permute a rule's premises to move the i-th premise to the last position.*)
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   698
fun make_last i th =
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   699
  let val n = nprems_of th
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   700
  in  if 1 <= i andalso i <= n
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   701
      then Thm.permute_prems (i-1) 1 th
15118
e2bd080c7975 make_clauses now meta
paulson
parents: 15032
diff changeset
   702
      else raise THM("select_literal", i, [th])
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   703
  end;
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   704
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   705
(*Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   706
  double-negations.*)
35410
1ea89d2a1bd4 more antiquotations;
wenzelm
parents: 34974
diff changeset
   707
val negate_head = rewrite_rule [@{thm atomize_not}, not_not RS eq_reflection];
14744
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   708
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   709
(*Maps the clause  [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P*)
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   710
fun select_literal i cl = negate_head (make_last i cl);
7ccfc167e451 clauses for ordinary resolution
paulson
parents: 14733
diff changeset
   711
18508
paulson
parents: 18405
diff changeset
   712
14813
826edbc317e6 new skolemize_tac and skolemize method
paulson
parents: 14763
diff changeset
   713
(*Top-level Skolemization. Allows part of the conversion to clauses to be
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   714
  expressed as a tactic (or Isar method).  Each assumption of the selected
14813
826edbc317e6 new skolemize_tac and skolemize method
paulson
parents: 14763
diff changeset
   715
  goal is converted to NNF and then its existential quantifiers are pulled
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   716
  to the front. Finally, all existential quantifiers are eliminated,
14813
826edbc317e6 new skolemize_tac and skolemize method
paulson
parents: 14763
diff changeset
   717
  leaving !!-quantified variables. Perhaps Safe_tac should follow, but it
826edbc317e6 new skolemize_tac and skolemize method
paulson
parents: 14763
diff changeset
   718
  might generate many subgoals.*)
18194
940515d2fa26 -- removed "check_is_fol" from "make_nnf" so that the NNF procedure doesn't check whether a thm is FOL.
mengj
parents: 18175
diff changeset
   719
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   720
fun skolemize_tac ctxt = SUBGOAL (fn (goal, i) =>
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   721
  let val ts = Logic.strip_assums_hyp goal
24300
e170cee91c66 proper signature for Meson;
wenzelm
parents: 24040
diff changeset
   722
  in
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   723
    EVERY'
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37539
diff changeset
   724
     [Misc_Legacy.METAHYPS (fn hyps =>
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   725
        (cut_facts_tac (skolemize_nnf_list ctxt hyps) 1
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   726
          THEN REPEAT (etac exE 1))),
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   727
      REPEAT_DETERM_N (length ts) o (etac thin_rl)] i
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32231
diff changeset
   728
  end);
18194
940515d2fa26 -- removed "check_is_fol" from "make_nnf" so that the NNF procedure doesn't check whether a thm is FOL.
mengj
parents: 18175
diff changeset
   729
9840
9dfcb0224f8c meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
paulson
parents:
diff changeset
   730
end;