src/HOL/Tools/res_axioms.ML
author wenzelm
Sun, 07 Mar 2010 12:19:47 +0100
changeset 35625 9c818cab0dd0
parent 35624 c4e29a0bb8c1
permissions -rw-r--r--
modernized structure Object_Logic;
     1 (*  Title:      HOL/Tools/res_axioms.ML
     2     Author:     Jia Meng, Cambridge University Computer Laboratory
     3 
     4 Transformation of axiom rules (elim/intro/etc) into CNF forms.
     5 *)
     6 
     7 signature RES_AXIOMS =
     8 sig
     9   val trace: bool Unsynchronized.ref
    10   val trace_msg: (unit -> string) -> unit
    11   val cnf_axiom: theory -> thm -> thm list
    12   val pairname: thm -> string * thm
    13   val multi_base_blacklist: string list
    14   val bad_for_atp: thm -> bool
    15   val type_has_topsort: typ -> bool
    16   val cnf_rules_pairs: theory -> (string * thm) list -> (thm * (string * int)) list
    17   val neg_clausify: thm list -> thm list
    18   val expand_defs_tac: thm -> tactic
    19   val combinators: thm -> thm
    20   val neg_conjecture_clauses: Proof.context -> thm -> int -> thm list * (string * typ) list
    21   val suppress_endtheory: bool Unsynchronized.ref
    22     (*for emergency use where endtheory causes problems*)
    23   val setup: theory -> theory
    24 end;
    25 
    26 structure Res_Axioms: RES_AXIOMS =
    27 struct
    28 
    29 val trace = Unsynchronized.ref false;
    30 fun trace_msg msg = if ! trace then tracing (msg ()) else ();
    31 
    32 fun freeze_thm th = #1 (Drule.legacy_freeze_thaw th);
    33 
    34 val type_has_topsort = Term.exists_subtype
    35   (fn TFree (_, []) => true
    36     | TVar (_, []) => true
    37     | _ => false);
    38 
    39 
    40 (**** Transformation of Elimination Rules into First-Order Formulas****)
    41 
    42 val cfalse = cterm_of @{theory HOL} HOLogic.false_const;
    43 val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop HOLogic.false_const);
    44 
    45 (*Converts an elim-rule into an equivalent theorem that does not have the
    46   predicate variable.  Leaves other theorems unchanged.  We simply instantiate the
    47   conclusion variable to False.*)
    48 fun transform_elim th =
    49   case concl_of th of    (*conclusion variable*)
    50        Const("Trueprop",_) $ (v as Var(_,Type("bool",[]))) =>
    51            Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th
    52     | v as Var(_, Type("prop",[])) =>
    53            Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th
    54     | _ => th;
    55 
    56 (*To enforce single-threading*)
    57 exception Clausify_failure of theory;
    58 
    59 
    60 (**** SKOLEMIZATION BY INFERENCE (lcp) ****)
    61 
    62 fun rhs_extra_types lhsT rhs =
    63   let val lhs_vars = Term.add_tfreesT lhsT []
    64       fun add_new_TFrees (TFree v) =
    65             if member (op =) lhs_vars v then I else insert (op =) (TFree v)
    66         | add_new_TFrees _ = I
    67       val rhs_consts = fold_aterms (fn Const c => insert (op =) c | _ => I) rhs []
    68   in fold (#2 #> Term.fold_atyps add_new_TFrees) rhs_consts [] end;
    69 
    70 (*Traverse a theorem, declaring Skolem function definitions. String s is the suggested
    71   prefix for the Skolem constant.*)
    72 fun declare_skofuns s th =
    73   let
    74     val nref = Unsynchronized.ref 0    (* FIXME ??? *)
    75     fun dec_sko (Const ("Ex",_) $ (xtp as Abs (_, T, p))) (axs, thy) =
    76           (*Existential: declare a Skolem function, then insert into body and continue*)
    77           let
    78             val cname = "sko_" ^ s ^ "_" ^ Int.toString (Unsynchronized.inc nref)
    79             val args0 = OldTerm.term_frees xtp  (*get the formal parameter list*)
    80             val Ts = map type_of args0
    81             val extraTs = rhs_extra_types (Ts ---> T) xtp
    82             val argsx = map (fn T => Free (gensym "vsk", T)) extraTs
    83             val args = argsx @ args0
    84             val cT = extraTs ---> Ts ---> T
    85             val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp)
    86                     (*Forms a lambda-abstraction over the formal parameters*)
    87             val (c, thy') =
    88               Sign.declare_const ((Binding.conceal (Binding.name cname), cT), NoSyn) thy
    89             val cdef = cname ^ "_def"
    90             val thy'' =
    91               Theory.add_defs_i true false [(Binding.name cdef, Logic.mk_equals (c, rhs))] thy'
    92             val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef)
    93           in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end
    94       | dec_sko (Const ("All", _) $ (Abs (a, T, p))) thx =
    95           (*Universal quant: insert a free variable into body and continue*)
    96           let val fname = Name.variant (OldTerm.add_term_names (p, [])) a
    97           in dec_sko (subst_bound (Free (fname, T), p)) thx end
    98       | dec_sko (Const ("op &", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
    99       | dec_sko (Const ("op |", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
   100       | dec_sko (Const ("Trueprop", _) $ p) thx = dec_sko p thx
   101       | dec_sko t thx = thx (*Do nothing otherwise*)
   102   in fn thy => dec_sko (Thm.prop_of th) ([], thy) end;
   103 
   104 (*Traverse a theorem, accumulating Skolem function definitions.*)
   105 fun assume_skofuns s th =
   106   let val sko_count = Unsynchronized.ref 0   (* FIXME ??? *)
   107       fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) defs =
   108             (*Existential: declare a Skolem function, then insert into body and continue*)
   109             let val skos = map (#1 o Logic.dest_equals) defs  (*existing sko fns*)
   110                 val args = subtract (op =) skos (OldTerm.term_frees xtp) (*the formal parameters*)
   111                 val Ts = map type_of args
   112                 val cT = Ts ---> T
   113                 val id = "sko_" ^ s ^ "_" ^ Int.toString (Unsynchronized.inc sko_count)
   114                 val c = Free (id, cT)
   115                 val rhs = list_abs_free (map dest_Free args,
   116                                          HOLogic.choice_const T $ xtp)
   117                       (*Forms a lambda-abstraction over the formal parameters*)
   118                 val def = Logic.mk_equals (c, rhs)
   119             in dec_sko (subst_bound (list_comb(c,args), p))
   120                        (def :: defs)
   121             end
   122         | dec_sko (Const ("All",_) $ Abs (a, T, p)) defs =
   123             (*Universal quant: insert a free variable into body and continue*)
   124             let val fname = Name.variant (OldTerm.add_term_names (p,[])) a
   125             in dec_sko (subst_bound (Free(fname,T), p)) defs end
   126         | dec_sko (Const ("op &", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
   127         | dec_sko (Const ("op |", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
   128         | dec_sko (Const ("Trueprop", _) $ p) defs = dec_sko p defs
   129         | dec_sko t defs = defs (*Do nothing otherwise*)
   130   in  dec_sko (prop_of th) []  end;
   131 
   132 
   133 (**** REPLACING ABSTRACTIONS BY COMBINATORS ****)
   134 
   135 (*Returns the vars of a theorem*)
   136 fun vars_of_thm th =
   137   map (Thm.cterm_of (theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th []);
   138 
   139 (*Make a version of fun_cong with a given variable name*)
   140 local
   141     val fun_cong' = fun_cong RS asm_rl; (*renumber f, g to prevent clashes with (a,0)*)
   142     val cx = hd (vars_of_thm fun_cong');
   143     val ty = typ_of (ctyp_of_term cx);
   144     val thy = theory_of_thm fun_cong;
   145     fun mkvar a = cterm_of thy (Var((a,0),ty));
   146 in
   147 fun xfun_cong x = Thm.instantiate ([], [(cx, mkvar x)]) fun_cong'
   148 end;
   149 
   150 (*Removes the lambdas from an equation of the form t = (%x. u).  A non-negative n,
   151   serves as an upper bound on how many to remove.*)
   152 fun strip_lambdas 0 th = th
   153   | strip_lambdas n th =
   154       case prop_of th of
   155           _ $ (Const ("op =", _) $ _ $ Abs (x,_,_)) =>
   156               strip_lambdas (n-1) (freeze_thm (th RS xfun_cong x))
   157         | _ => th;
   158 
   159 val lambda_free = not o Term.has_abs;
   160 
   161 val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B}));
   162 val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C}));
   163 val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S}));
   164 
   165 (*FIXME: requires more use of cterm constructors*)
   166 fun abstract ct =
   167   let
   168       val thy = theory_of_cterm ct
   169       val Abs(x,_,body) = term_of ct
   170       val Type("fun",[xT,bodyT]) = typ_of (ctyp_of_term ct)
   171       val cxT = ctyp_of thy xT and cbodyT = ctyp_of thy bodyT
   172       fun makeK() = instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] @{thm abs_K}
   173   in
   174       case body of
   175           Const _ => makeK()
   176         | Free _ => makeK()
   177         | Var _ => makeK()  (*though Var isn't expected*)
   178         | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
   179         | rator$rand =>
   180             if loose_bvar1 (rator,0) then (*C or S*)
   181                if loose_bvar1 (rand,0) then (*S*)
   182                  let val crator = cterm_of thy (Abs(x,xT,rator))
   183                      val crand = cterm_of thy (Abs(x,xT,rand))
   184                      val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
   185                      val (_,rhs) = Thm.dest_equals (cprop_of abs_S')
   186                  in
   187                    Thm.transitive abs_S' (Conv.binop_conv abstract rhs)
   188                  end
   189                else (*C*)
   190                  let val crator = cterm_of thy (Abs(x,xT,rator))
   191                      val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
   192                      val (_,rhs) = Thm.dest_equals (cprop_of abs_C')
   193                  in
   194                    Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs)
   195                  end
   196             else if loose_bvar1 (rand,0) then (*B or eta*)
   197                if rand = Bound 0 then eta_conversion ct
   198                else (*B*)
   199                  let val crand = cterm_of thy (Abs(x,xT,rand))
   200                      val crator = cterm_of thy rator
   201                      val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
   202                      val (_,rhs) = Thm.dest_equals (cprop_of abs_B')
   203                  in
   204                    Thm.transitive abs_B' (Conv.arg_conv abstract rhs)
   205                  end
   206             else makeK()
   207         | _ => error "abstract: Bad term"
   208   end;
   209 
   210 (*Traverse a theorem, declaring abstraction function definitions. String s is the suggested
   211   prefix for the constants.*)
   212 fun combinators_aux ct =
   213   if lambda_free (term_of ct) then reflexive ct
   214   else
   215   case term_of ct of
   216       Abs _ =>
   217         let val (cv, cta) = Thm.dest_abs NONE ct
   218             val (v, _) = dest_Free (term_of cv)
   219             val u_th = combinators_aux cta
   220             val cu = Thm.rhs_of u_th
   221             val comb_eq = abstract (Thm.cabs cv cu)
   222         in transitive (abstract_rule v cv u_th) comb_eq end
   223     | _ $ _ =>
   224         let val (ct1, ct2) = Thm.dest_comb ct
   225         in  combination (combinators_aux ct1) (combinators_aux ct2)  end;
   226 
   227 fun combinators th =
   228   if lambda_free (prop_of th) then th
   229   else
   230     let val th = Drule.eta_contraction_rule th
   231         val eqth = combinators_aux (cprop_of th)
   232     in  equal_elim eqth th   end
   233     handle THM (msg,_,_) =>
   234       (warning (cat_lines
   235         ["Error in the combinator translation of " ^ Display.string_of_thm_without_context th,
   236           "  Exception message: " ^ msg]);
   237        TrueI);  (*A type variable of sort {} will cause make abstraction fail.*)
   238 
   239 (*cterms are used throughout for efficiency*)
   240 val cTrueprop = Thm.cterm_of @{theory HOL} HOLogic.Trueprop;
   241 
   242 (*cterm version of mk_cTrueprop*)
   243 fun c_mkTrueprop A = Thm.capply cTrueprop A;
   244 
   245 (*Given an abstraction over n variables, replace the bound variables by free
   246   ones. Return the body, along with the list of free variables.*)
   247 fun c_variant_abs_multi (ct0, vars) =
   248       let val (cv,ct) = Thm.dest_abs NONE ct0
   249       in  c_variant_abs_multi (ct, cv::vars)  end
   250       handle CTERM _ => (ct0, rev vars);
   251 
   252 (*Given the definition of a Skolem function, return a theorem to replace
   253   an existential formula by a use of that function.
   254    Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B"  [.] *)
   255 fun skolem_of_def def =
   256   let val (c,rhs) = Thm.dest_equals (cprop_of (freeze_thm def))
   257       val (ch, frees) = c_variant_abs_multi (rhs, [])
   258       val (chilbert,cabs) = Thm.dest_comb ch
   259       val thy = Thm.theory_of_cterm chilbert
   260       val t = Thm.term_of chilbert
   261       val T = case t of Const ("Hilbert_Choice.Eps", Type("fun",[_,T])) => T
   262                       | _ => raise THM ("skolem_of_def: expected Eps", 0, [def])
   263       val cex = Thm.cterm_of thy (HOLogic.exists_const T)
   264       val ex_tm = c_mkTrueprop (Thm.capply cex cabs)
   265       and conc =  c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees)));
   266       fun tacf [prem] = rewrite_goals_tac [def] THEN rtac (prem RS @{thm someI_ex}) 1
   267   in  Goal.prove_internal [ex_tm] conc tacf
   268        |> forall_intr_list frees
   269        |> Thm.forall_elim_vars 0  (*Introduce Vars, but don't discharge defs.*)
   270        |> Thm.varifyT
   271   end;
   272 
   273 
   274 (*Converts an Isabelle theorem (intro, elim or simp format, even higher-order) into NNF.*)
   275 fun to_nnf th ctxt0 =
   276   let val th1 = th |> transform_elim |> zero_var_indexes
   277       val ((_, [th2]), ctxt) = Variable.import true [th1] ctxt0
   278       val th3 = th2
   279         |> Conv.fconv_rule Object_Logic.atomize
   280         |> Meson.make_nnf ctxt |> strip_lambdas ~1
   281   in  (th3, ctxt)  end;
   282 
   283 (*Generate Skolem functions for a theorem supplied in nnf*)
   284 fun assume_skolem_of_def s th =
   285   map (skolem_of_def o assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th);
   286 
   287 
   288 (*** Blacklisting (duplicated in Res_ATP?) ***)
   289 
   290 val max_lambda_nesting = 3;
   291 
   292 fun excessive_lambdas (f$t, k) = excessive_lambdas (f,k) orelse excessive_lambdas (t,k)
   293   | excessive_lambdas (Abs(_,_,t), k) = k=0 orelse excessive_lambdas (t,k-1)
   294   | excessive_lambdas _ = false;
   295 
   296 fun is_formula_type T = (T = HOLogic.boolT orelse T = propT);
   297 
   298 (*Don't count nested lambdas at the level of formulas, as they are quantifiers*)
   299 fun excessive_lambdas_fm Ts (Abs(_,T,t)) = excessive_lambdas_fm (T::Ts) t
   300   | excessive_lambdas_fm Ts t =
   301       if is_formula_type (fastype_of1 (Ts, t))
   302       then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t))
   303       else excessive_lambdas (t, max_lambda_nesting);
   304 
   305 (*The max apply_depth of any metis call in Metis_Examples (on 31-10-2007) was 11.*)
   306 val max_apply_depth = 15;
   307 
   308 fun apply_depth (f$t) = Int.max (apply_depth f, apply_depth t + 1)
   309   | apply_depth (Abs(_,_,t)) = apply_depth t
   310   | apply_depth _ = 0;
   311 
   312 fun too_complex t =
   313   apply_depth t > max_apply_depth orelse
   314   Meson.too_many_clauses NONE t orelse
   315   excessive_lambdas_fm [] t;
   316 
   317 fun is_strange_thm th =
   318   case head_of (concl_of th) of
   319       Const (a, _) => (a <> "Trueprop" andalso a <> "==")
   320     | _ => false;
   321 
   322 fun bad_for_atp th =
   323   too_complex (prop_of th)
   324   orelse exists_type type_has_topsort (prop_of th)
   325   orelse is_strange_thm th;
   326 
   327 val multi_base_blacklist =
   328   ["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm",
   329    "cases","ext_cases"];  (* FIXME put other record thms here, or declare as "noatp" *)
   330 
   331 (*Keep the full complexity of the original name*)
   332 fun flatten_name s = space_implode "_X" (Long_Name.explode s);
   333 
   334 fun fake_name th =
   335   if Thm.has_name_hint th then flatten_name (Thm.get_name_hint th)
   336   else gensym "unknown_thm_";
   337 
   338 (*Skolemize a named theorem, with Skolem functions as additional premises.*)
   339 fun skolem_thm (s, th) =
   340   if member (op =) multi_base_blacklist (Long_Name.base_name s) orelse bad_for_atp th then []
   341   else
   342     let
   343       val ctxt0 = Variable.thm_context th
   344       val (nnfth, ctxt1) = to_nnf th ctxt0
   345       val (cnfs, ctxt2) = Meson.make_cnf (assume_skolem_of_def s nnfth) nnfth ctxt1
   346     in  cnfs |> map combinators |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf  end
   347     handle THM _ => [];
   348 
   349 (*The cache prevents repeated clausification of a theorem, and also repeated declaration of
   350   Skolem functions.*)
   351 structure ThmCache = Theory_Data
   352 (
   353   type T = thm list Thmtab.table * unit Symtab.table;
   354   val empty = (Thmtab.empty, Symtab.empty);
   355   val extend = I;
   356   fun merge ((cache1, seen1), (cache2, seen2)) : T =
   357     (Thmtab.merge (K true) (cache1, cache2), Symtab.merge (K true) (seen1, seen2));
   358 );
   359 
   360 val lookup_cache = Thmtab.lookup o #1 o ThmCache.get;
   361 val already_seen = Symtab.defined o #2 o ThmCache.get;
   362 
   363 val update_cache = ThmCache.map o apfst o Thmtab.update;
   364 fun mark_seen name = ThmCache.map (apsnd (Symtab.update (name, ())));
   365 
   366 (*Exported function to convert Isabelle theorems into axiom clauses*)
   367 fun cnf_axiom thy th0 =
   368   let val th = Thm.transfer thy th0 in
   369     case lookup_cache thy th of
   370       NONE => map Thm.close_derivation (skolem_thm (fake_name th, th))
   371     | SOME cls => cls
   372   end;
   373 
   374 
   375 (**** Rules from the context ****)
   376 
   377 fun pairname th = (Thm.get_name_hint th, th);
   378 
   379 
   380 (**** Translate a set of theorems into CNF ****)
   381 
   382 fun pair_name_cls k (n, []) = []
   383   | pair_name_cls k (n, cls::clss) = (cls, (n,k)) :: pair_name_cls (k+1) (n, clss)
   384 
   385 fun cnf_rules_pairs_aux _ pairs [] = pairs
   386   | cnf_rules_pairs_aux thy pairs ((name,th)::ths) =
   387       let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs
   388                        handle THM _ => pairs | Res_Clause.CLAUSE _ => pairs
   389       in  cnf_rules_pairs_aux thy pairs' ths  end;
   390 
   391 (*The combination of rev and tail recursion preserves the original order*)
   392 fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l);
   393 
   394 
   395 (**** Convert all facts of the theory into clauses (Res_Clause.clause, or Res_HOL_Clause.clause) ****)
   396 
   397 local
   398 
   399 fun skolem_def (name, th) thy =
   400   let val ctxt0 = Variable.thm_context th in
   401     (case try (to_nnf th) ctxt0 of
   402       NONE => (NONE, thy)
   403     | SOME (nnfth, ctxt1) =>
   404         let val (defs, thy') = declare_skofuns (flatten_name name) nnfth thy
   405         in (SOME (th, ctxt0, ctxt1, nnfth, defs), thy') end)
   406   end;
   407 
   408 fun skolem_cnfs (th, ctxt0, ctxt1, nnfth, defs) =
   409   let
   410     val (cnfs, ctxt2) = Meson.make_cnf (map skolem_of_def defs) nnfth ctxt1;
   411     val cnfs' = cnfs
   412       |> map combinators
   413       |> Variable.export ctxt2 ctxt0
   414       |> Meson.finish_cnf
   415       |> map Thm.close_derivation;
   416     in (th, cnfs') end;
   417 
   418 in
   419 
   420 fun saturate_skolem_cache thy =
   421   let
   422     val facts = PureThy.facts_of thy;
   423     val new_facts = (facts, []) |-> Facts.fold_static (fn (name, ths) =>
   424       if Facts.is_concealed facts name orelse already_seen thy name then I
   425       else cons (name, ths));
   426     val new_thms = (new_facts, []) |-> fold (fn (name, ths) =>
   427       if member (op =) multi_base_blacklist (Long_Name.base_name name) then I
   428       else fold_index (fn (i, th) =>
   429         if bad_for_atp th orelse is_some (lookup_cache thy th) then I
   430         else cons (name ^ "_" ^ string_of_int (i + 1), Thm.transfer thy th)) ths);
   431   in
   432     if null new_facts then NONE
   433     else
   434       let
   435         val (defs, thy') = thy
   436           |> fold (mark_seen o #1) new_facts
   437           |> fold_map skolem_def (sort_distinct (Thm.thm_ord o pairself snd) new_thms)
   438           |>> map_filter I;
   439         val cache_entries = Par_List.map skolem_cnfs defs;
   440       in SOME (fold update_cache cache_entries thy') end
   441   end;
   442 
   443 end;
   444 
   445 val suppress_endtheory = Unsynchronized.ref false;
   446 
   447 fun clause_cache_endtheory thy =
   448   if ! suppress_endtheory then NONE
   449   else saturate_skolem_cache thy;
   450 
   451 
   452 (*The cache can be kept smaller by inspecting the prop of each thm. Can ignore all that are
   453   lambda_free, but then the individual theory caches become much bigger.*)
   454 
   455 
   456 (*** meson proof methods ***)
   457 
   458 (*Expand all new definitions of abstraction or Skolem functions in a proof state.*)
   459 fun is_absko (Const ("==", _) $ Free (a,_) $ u) = String.isPrefix "sko_" a
   460   | is_absko _ = false;
   461 
   462 fun is_okdef xs (Const ("==", _) $ t $ u) =   (*Definition of Free, not in certain terms*)
   463       is_Free t andalso not (member (op aconv) xs t)
   464   | is_okdef _ _ = false
   465 
   466 (*This function tries to cope with open locales, which introduce hypotheses of the form
   467   Free == t, conjecture clauses, which introduce various hypotheses, and also definitions
   468   of sko_ functions. *)
   469 fun expand_defs_tac st0 st =
   470   let val hyps0 = #hyps (rep_thm st0)
   471       val hyps = #hyps (crep_thm st)
   472       val newhyps = filter_out (member (op aconv) hyps0 o Thm.term_of) hyps
   473       val defs = filter (is_absko o Thm.term_of) newhyps
   474       val remaining_hyps = filter_out (member (op aconv) (map Thm.term_of defs))
   475                                       (map Thm.term_of hyps)
   476       val fixed = OldTerm.term_frees (concl_of st) @
   477                   fold (union (op aconv)) (map OldTerm.term_frees remaining_hyps) []
   478   in Seq.of_list [Local_Defs.expand (filter (is_okdef fixed o Thm.term_of) defs) st] end;
   479 
   480 
   481 fun meson_general_tac ctxt ths i st0 =
   482   let
   483     val thy = ProofContext.theory_of ctxt
   484     val ctxt0 = Classical.put_claset HOL_cs ctxt
   485   in (Meson.meson_tac ctxt0 (maps (cnf_axiom thy) ths) i THEN expand_defs_tac st0) st0 end;
   486 
   487 val meson_method_setup =
   488   Method.setup @{binding meson} (Attrib.thms >> (fn ths => fn ctxt =>
   489     SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ctxt ths)))
   490     "MESON resolution proof procedure";
   491 
   492 
   493 (*** Converting a subgoal into negated conjecture clauses. ***)
   494 
   495 fun neg_skolemize_tac ctxt =
   496   EVERY' [rtac ccontr, Object_Logic.atomize_prems_tac, Meson.skolemize_tac ctxt];
   497 
   498 val neg_clausify = Meson.make_clauses #> map combinators #> Meson.finish_cnf;
   499 
   500 fun neg_conjecture_clauses ctxt st0 n =
   501   let
   502     val st = Seq.hd (neg_skolemize_tac ctxt n st0)
   503     val ({params, prems, ...}, _) = Subgoal.focus (Variable.set_body false ctxt) n st
   504   in (neg_clausify prems, map (Term.dest_Free o Thm.term_of o #2) params) end;
   505 
   506 (*Conversion of a subgoal to conjecture clauses. Each clause has
   507   leading !!-bound universal variables, to express generality. *)
   508 fun neg_clausify_tac ctxt =
   509   neg_skolemize_tac ctxt THEN'
   510   SUBGOAL (fn (prop, i) =>
   511     let val ts = Logic.strip_assums_hyp prop in
   512       EVERY'
   513        [Subgoal.FOCUS
   514          (fn {prems, ...} =>
   515            (Method.insert_tac
   516              (map forall_intr_vars (neg_clausify prems)) i)) ctxt,
   517         REPEAT_DETERM_N (length ts) o etac thin_rl] i
   518      end);
   519 
   520 val neg_clausify_setup =
   521   Method.setup @{binding neg_clausify} (Scan.succeed (SIMPLE_METHOD' o neg_clausify_tac))
   522   "conversion of goal to conjecture clauses";
   523 
   524 
   525 (** Attribute for converting a theorem into clauses **)
   526 
   527 val clausify_setup =
   528   Attrib.setup @{binding clausify}
   529     (Scan.lift OuterParse.nat >>
   530       (fn i => Thm.rule_attribute (fn context => fn th =>
   531           Meson.make_meta_clause (nth (cnf_axiom (Context.theory_of context) th) i))))
   532   "conversion of theorem to clauses";
   533 
   534 
   535 
   536 (** setup **)
   537 
   538 val setup =
   539   meson_method_setup #>
   540   neg_clausify_setup #>
   541   clausify_setup #>
   542   perhaps saturate_skolem_cache #>
   543   Theory.at_end clause_cache_endtheory;
   544 
   545 end;