| 39948 |      1 | (*  Title:      HOL/Tools/Meson/meson_tactic.ML
 | 
|  |      2 |     Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
 | 
|  |      3 |     Author:     Jasmin Blanchette, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | The "meson" proof method for HOL.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | signature MESON_TACTIC =
 | 
|  |      9 | sig
 | 
|  |     10 |   val meson_general_tac : Proof.context -> thm list -> int -> tactic
 | 
|  |     11 |   val setup: theory -> theory
 | 
|  |     12 | end;
 | 
|  |     13 | 
 | 
|  |     14 | structure Meson_Tactic : MESON_TACTIC =
 | 
|  |     15 | struct
 | 
|  |     16 | 
 | 
|  |     17 | open Meson_Clausify
 | 
|  |     18 | 
 | 
|  |     19 | fun meson_general_tac ctxt ths =
 | 
| 42793 |     20 |   let val ctxt' = put_claset HOL_cs ctxt
 | 
|  |     21 |   in Meson.meson_tac ctxt' (maps (snd o cnf_axiom ctxt' false 0) ths) end
 | 
| 39948 |     22 | 
 | 
|  |     23 | val setup =
 | 
|  |     24 |   Method.setup @{binding meson} (Attrib.thms >> (fn ths => fn ctxt =>
 | 
|  |     25 |      SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ctxt ths)))
 | 
|  |     26 |      "MESON resolution proof procedure"
 | 
|  |     27 | 
 | 
|  |     28 | end;
 |