src/HOL/Tools/meson.ML
author paulson
Thu, 27 Sep 2007 17:55:28 +0200
changeset 24742 73b8b42a36b6
parent 24300 e170cee91c66
child 24827 646bdc51eb7d
permissions -rw-r--r--
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
theorems of Nat.thy are hidden by the Ordering.thy versions
     1 (*  Title:      HOL/Tools/meson.ML
     2     ID:         $Id$
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     4     Copyright   1992  University of Cambridge
     5 
     6 The MESON resolution proof procedure for HOL.
     7 
     8 When making clauses, avoids using the rewriter -- instead uses RS recursively
     9 
    10 NEED TO SORT LITERALS BY # OF VARS, USING ==>I/E.  ELIMINATES NEED FOR
    11 FUNCTION nodups -- if done to goal clauses too!
    12 *)
    13 
    14 signature MESON =
    15 sig
    16   val term_pair_of: indexname * (typ * 'a) -> term * 'a
    17   val first_order_resolve: thm -> thm -> thm
    18   val flexflex_first_order: thm -> thm
    19   val size_of_subgoals: thm -> int
    20   val make_cnf: thm list -> thm -> thm list
    21   val finish_cnf: thm list -> thm list
    22   val generalize: thm -> thm
    23   val make_nnf: thm -> thm
    24   val make_nnf1: thm -> thm
    25   val skolemize: thm -> thm
    26   val is_fol_term: theory -> term -> bool
    27   val make_clauses: thm list -> thm list
    28   val make_horns: thm list -> thm list
    29   val best_prolog_tac: (thm -> int) -> thm list -> tactic
    30   val depth_prolog_tac: thm list -> tactic
    31   val gocls: thm list -> thm list
    32   val skolemize_prems_tac: thm list -> int -> tactic
    33   val MESON: (thm list -> thm list) -> (thm list -> tactic) -> int -> tactic
    34   val best_meson_tac: (thm -> int) -> int -> tactic
    35   val safe_best_meson_tac: int -> tactic
    36   val depth_meson_tac: int -> tactic
    37   val prolog_step_tac': thm list -> int -> tactic
    38   val iter_deepen_prolog_tac: thm list -> tactic
    39   val iter_deepen_meson_tac: thm list -> int -> tactic
    40   val make_meta_clause: thm -> thm
    41   val make_meta_clauses: thm list -> thm list
    42   val meson_claset_tac: thm list -> claset -> int -> tactic
    43   val meson_tac: int -> tactic
    44   val negate_head: thm -> thm
    45   val select_literal: int -> thm -> thm
    46   val skolemize_tac: int -> tactic
    47 end
    48 
    49 structure Meson: MESON =
    50 struct
    51 
    52 val not_conjD = thm "meson_not_conjD";
    53 val not_disjD = thm "meson_not_disjD";
    54 val not_notD = thm "meson_not_notD";
    55 val not_allD = thm "meson_not_allD";
    56 val not_exD = thm "meson_not_exD";
    57 val imp_to_disjD = thm "meson_imp_to_disjD";
    58 val not_impD = thm "meson_not_impD";
    59 val iff_to_disjD = thm "meson_iff_to_disjD";
    60 val not_iffD = thm "meson_not_iffD";
    61 val conj_exD1 = thm "meson_conj_exD1";
    62 val conj_exD2 = thm "meson_conj_exD2";
    63 val disj_exD = thm "meson_disj_exD";
    64 val disj_exD1 = thm "meson_disj_exD1";
    65 val disj_exD2 = thm "meson_disj_exD2";
    66 val disj_assoc = thm "meson_disj_assoc";
    67 val disj_comm = thm "meson_disj_comm";
    68 val disj_FalseD1 = thm "meson_disj_FalseD1";
    69 val disj_FalseD2 = thm "meson_disj_FalseD2";
    70 
    71 
    72 (**** Operators for forward proof ****)
    73 
    74 
    75 (** First-order Resolution **)
    76 
    77 fun typ_pair_of (ix, (sort,ty)) = (TVar (ix,sort), ty);
    78 fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t);
    79 
    80 val Envir.Envir {asol = tenv0, iTs = tyenv0, ...} = Envir.empty 0
    81 
    82 (*FIXME: currently does not "rename variables apart"*)
    83 fun first_order_resolve thA thB =
    84   let val thy = theory_of_thm thA
    85       val tmA = concl_of thA
    86       val Const("==>",_) $ tmB $ _ = prop_of thB
    87       val (tyenv,tenv) = Pattern.first_order_match thy (tmB,tmA) (tyenv0,tenv0)
    88       val ct_pairs = map (pairself (cterm_of thy) o term_pair_of) (Vartab.dest tenv)
    89   in  thA RS (cterm_instantiate ct_pairs thB)  end
    90   handle _ => raise THM ("first_order_resolve", 0, [thA,thB]);
    91 
    92 fun flexflex_first_order th =
    93   case (tpairs_of th) of
    94       [] => th
    95     | pairs =>
    96         let val thy = theory_of_thm th
    97             val (tyenv,tenv) =
    98                   foldl (uncurry (Pattern.first_order_match thy)) (tyenv0,tenv0) pairs
    99             val t_pairs = map term_pair_of (Vartab.dest tenv)
   100             val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th
   101         in  th'  end
   102         handle THM _ => th;
   103 
   104 (*raises exception if no rules apply -- unlike RL*)
   105 fun tryres (th, rls) =
   106   let fun tryall [] = raise THM("tryres", 0, th::rls)
   107         | tryall (rl::rls) = (th RS rl handle THM _ => tryall rls)
   108   in  tryall rls  end;
   109 
   110 (*Permits forward proof from rules that discharge assumptions. The supplied proof state st,
   111   e.g. from conj_forward, should have the form
   112     "[| P' ==> ?P; Q' ==> ?Q |] ==> ?P & ?Q"
   113   and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*)
   114 fun forward_res nf st =
   115   let fun forward_tacf [prem] = rtac (nf prem) 1
   116         | forward_tacf prems =
   117             error ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:\n" ^
   118                    string_of_thm st ^
   119                    "\nPremises:\n" ^
   120                    cat_lines (map string_of_thm prems))
   121   in
   122     case Seq.pull (ALLGOALS (METAHYPS forward_tacf) st)
   123     of SOME(th,_) => th
   124      | NONE => raise THM("forward_res", 0, [st])
   125   end;
   126 
   127 (*Are any of the logical connectives in "bs" present in the term?*)
   128 fun has_conns bs =
   129   let fun has (Const(a,_)) = false
   130         | has (Const("Trueprop",_) $ p) = has p
   131         | has (Const("Not",_) $ p) = has p
   132         | has (Const("op |",_) $ p $ q) = member (op =) bs "op |" orelse has p orelse has q
   133         | has (Const("op &",_) $ p $ q) = member (op =) bs "op &" orelse has p orelse has q
   134         | has (Const("All",_) $ Abs(_,_,p)) = member (op =) bs "All" orelse has p
   135         | has (Const("Ex",_) $ Abs(_,_,p)) = member (op =) bs "Ex" orelse has p
   136         | has _ = false
   137   in  has  end;
   138 
   139 
   140 (**** Clause handling ****)
   141 
   142 fun literals (Const("Trueprop",_) $ P) = literals P
   143   | literals (Const("op |",_) $ P $ Q) = literals P @ literals Q
   144   | literals (Const("Not",_) $ P) = [(false,P)]
   145   | literals P = [(true,P)];
   146 
   147 (*number of literals in a term*)
   148 val nliterals = length o literals;
   149 
   150 
   151 (*** Tautology Checking ***)
   152 
   153 fun signed_lits_aux (Const ("op |", _) $ P $ Q) (poslits, neglits) =
   154       signed_lits_aux Q (signed_lits_aux P (poslits, neglits))
   155   | signed_lits_aux (Const("Not",_) $ P) (poslits, neglits) = (poslits, P::neglits)
   156   | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits);
   157 
   158 fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]);
   159 
   160 (*Literals like X=X are tautologous*)
   161 fun taut_poslit (Const("op =",_) $ t $ u) = t aconv u
   162   | taut_poslit (Const("True",_)) = true
   163   | taut_poslit _ = false;
   164 
   165 fun is_taut th =
   166   let val (poslits,neglits) = signed_lits th
   167   in  exists taut_poslit poslits
   168       orelse
   169       exists (member (op aconv) neglits) (HOLogic.false_const :: poslits)
   170   end
   171   handle TERM _ => false;       (*probably dest_Trueprop on a weird theorem*)
   172 
   173 
   174 (*** To remove trivial negated equality literals from clauses ***)
   175 
   176 (*They are typically functional reflexivity axioms and are the converses of
   177   injectivity equivalences*)
   178 
   179 val not_refl_disj_D = thm"meson_not_refl_disj_D";
   180 
   181 (*Is either term a Var that does not properly occur in the other term?*)
   182 fun eliminable (t as Var _, u) = t aconv u orelse not (Logic.occs(t,u))
   183   | eliminable (u, t as Var _) = t aconv u orelse not (Logic.occs(t,u))
   184   | eliminable _ = false;
   185 
   186 fun refl_clause_aux 0 th = th
   187   | refl_clause_aux n th =
   188        case HOLogic.dest_Trueprop (concl_of th) of
   189           (Const ("op |", _) $ (Const ("op |", _) $ _ $ _) $ _) =>
   190             refl_clause_aux n (th RS disj_assoc)    (*isolate an atom as first disjunct*)
   191         | (Const ("op |", _) $ (Const("Not",_) $ (Const("op =",_) $ t $ u)) $ _) =>
   192             if eliminable(t,u)
   193             then refl_clause_aux (n-1) (th RS not_refl_disj_D)  (*Var inequation: delete*)
   194             else refl_clause_aux (n-1) (th RS disj_comm)  (*not between Vars: ignore*)
   195         | (Const ("op |", _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
   196         | _ => (*not a disjunction*) th;
   197 
   198 fun notequal_lits_count (Const ("op |", _) $ P $ Q) =
   199       notequal_lits_count P + notequal_lits_count Q
   200   | notequal_lits_count (Const("Not",_) $ (Const("op =",_) $ _ $ _)) = 1
   201   | notequal_lits_count _ = 0;
   202 
   203 (*Simplify a clause by applying reflexivity to its negated equality literals*)
   204 fun refl_clause th =
   205   let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (concl_of th))
   206   in  zero_var_indexes (refl_clause_aux neqs th)  end
   207   handle TERM _ => th;  (*probably dest_Trueprop on a weird theorem*)
   208 
   209 
   210 (*** The basic CNF transformation ***)
   211 
   212 val max_clauses = 40;
   213 
   214 fun sum x y = if x < max_clauses andalso y < max_clauses then x+y else max_clauses;
   215 fun prod x y = if x < max_clauses andalso y < max_clauses then x*y else max_clauses;
   216 
   217 (*Estimate the number of clauses in order to detect infeasible theorems*)
   218 fun signed_nclauses b (Const("Trueprop",_) $ t) = signed_nclauses b t
   219   | signed_nclauses b (Const("Not",_) $ t) = signed_nclauses (not b) t
   220   | signed_nclauses b (Const("op &",_) $ t $ u) =
   221       if b then sum (signed_nclauses b t) (signed_nclauses b u)
   222            else prod (signed_nclauses b t) (signed_nclauses b u)
   223   | signed_nclauses b (Const("op |",_) $ t $ u) =
   224       if b then prod (signed_nclauses b t) (signed_nclauses b u)
   225            else sum (signed_nclauses b t) (signed_nclauses b u)
   226   | signed_nclauses b (Const("op -->",_) $ t $ u) =
   227       if b then prod (signed_nclauses (not b) t) (signed_nclauses b u)
   228            else sum (signed_nclauses (not b) t) (signed_nclauses b u)
   229   | signed_nclauses b (Const("op =", Type ("fun", [T, _])) $ t $ u) =
   230       if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*)
   231           if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u))
   232                         (prod (signed_nclauses (not b) u) (signed_nclauses b t))
   233                else sum (prod (signed_nclauses b t) (signed_nclauses b u))
   234                         (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u))
   235       else 1
   236   | signed_nclauses b (Const("Ex", _) $ Abs (_,_,t)) = signed_nclauses b t
   237   | signed_nclauses b (Const("All",_) $ Abs (_,_,t)) = signed_nclauses b t
   238   | signed_nclauses _ _ = 1; (* literal *)
   239 
   240 val nclauses = signed_nclauses true;
   241 
   242 fun too_many_clauses t = nclauses t >= max_clauses;
   243 
   244 (*Replaces universally quantified variables by FREE variables -- because
   245   assumptions may not contain scheme variables.  Later, we call "generalize". *)
   246 fun freeze_spec th =
   247   let val newname = gensym "mes_"
   248       val spec' = read_instantiate [("x", newname)] spec
   249   in  th RS spec'  end;
   250 
   251 (*Used with METAHYPS below. There is one assumption, which gets bound to prem
   252   and then normalized via function nf. The normal form is given to resolve_tac,
   253   instantiate a Boolean variable created by resolution with disj_forward. Since
   254   (nf prem) returns a LIST of theorems, we can backtrack to get all combinations.*)
   255 fun resop nf [prem] = resolve_tac (nf prem) 1;
   256 
   257 (*Any need to extend this list with
   258   "HOL.type_class","HOL.eq_class","ProtoPure.term"?*)
   259 val has_meta_conn =
   260     exists_Const (member (op =) ["==", "==>", "all", "prop"] o #1);
   261 
   262 fun apply_skolem_ths (th, rls) =
   263   let fun tryall [] = raise THM("apply_skolem_ths", 0, th::rls)
   264         | tryall (rl::rls) = (first_order_resolve th rl handle THM _ => tryall rls)
   265   in  tryall rls  end;
   266 
   267 (*Conjunctive normal form, adding clauses from th in front of ths (for foldr).
   268   Strips universal quantifiers and breaks up conjunctions.
   269   Eliminates existential quantifiers using skoths: Skolemization theorems.*)
   270 fun cnf skoths (th,ths) =
   271   let fun cnf_aux (th,ths) =
   272         if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*)
   273         else if not (has_conns ["All","Ex","op &"] (prop_of th))
   274         then th::ths (*no work to do, terminate*)
   275         else case head_of (HOLogic.dest_Trueprop (concl_of th)) of
   276             Const ("op &", _) => (*conjunction*)
   277                 cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths))
   278           | Const ("All", _) => (*universal quantifier*)
   279                 cnf_aux (freeze_spec th, ths)
   280           | Const ("Ex", _) =>
   281               (*existential quantifier: Insert Skolem functions*)
   282               cnf_aux (apply_skolem_ths (th,skoths), ths)
   283           | Const ("op |", _) =>
   284               (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using
   285                 all combinations of converting P, Q to CNF.*)
   286               let val tac =
   287                   (METAHYPS (resop cnf_nil) 1) THEN
   288                    (fn st' => st' |> METAHYPS (resop cnf_nil) 1)
   289               in  Seq.list_of (tac (th RS disj_forward)) @ ths  end
   290           | _ => (*no work to do*) th::ths
   291       and cnf_nil th = cnf_aux (th,[])
   292   in
   293     if too_many_clauses (concl_of th)
   294     then (Output.debug (fn () => ("cnf is ignoring: " ^ string_of_thm th)); ths)
   295     else cnf_aux (th,ths)
   296   end;
   297 
   298 fun all_names (Const ("all", _) $ Abs(x,_,P)) = x :: all_names P
   299   | all_names _ = [];
   300 
   301 fun new_names n [] = []
   302   | new_names n (x::xs) =
   303       if String.isPrefix "mes_" x then (x, radixstring(26,"A",n)) :: new_names (n+1) xs
   304       else new_names n xs;
   305 
   306 (*The gensym names are ugly, so don't let the user see them. When forall_elim_vars
   307   is called, it will ensure that no name clauses ensue.*)
   308 fun nice_names th =
   309   let val old_names = all_names (prop_of th)
   310   in  Drule.rename_bvars (new_names 0 old_names) th  end;
   311 
   312 (*Convert all suitable free variables to schematic variables,
   313   but don't discharge assumptions.*)
   314 fun generalize th = Thm.varifyT (forall_elim_vars 0 (nice_names (forall_intr_frees th)));
   315 
   316 fun make_cnf skoths th = cnf skoths (th, []);
   317 
   318 (*Generalization, removal of redundant equalities, removal of tautologies.*)
   319 fun finish_cnf ths = filter (not o is_taut) (map (refl_clause o generalize) ths);
   320 
   321 
   322 (**** Removal of duplicate literals ****)
   323 
   324 (*Forward proof, passing extra assumptions as theorems to the tactic*)
   325 fun forward_res2 nf hyps st =
   326   case Seq.pull
   327         (REPEAT
   328          (METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)
   329          st)
   330   of SOME(th,_) => th
   331    | NONE => raise THM("forward_res2", 0, [st]);
   332 
   333 (*Remove duplicates in P|Q by assuming ~P in Q
   334   rls (initially []) accumulates assumptions of the form P==>False*)
   335 fun nodups_aux rls th = nodups_aux rls (th RS disj_assoc)
   336     handle THM _ => tryres(th,rls)
   337     handle THM _ => tryres(forward_res2 nodups_aux rls (th RS disj_forward2),
   338                            [disj_FalseD1, disj_FalseD2, asm_rl])
   339     handle THM _ => th;
   340 
   341 (*Remove duplicate literals, if there are any*)
   342 fun nodups th =
   343   if has_duplicates (op =) (literals (prop_of th))
   344     then nodups_aux [] th
   345     else th;
   346 
   347 
   348 (**** Generation of contrapositives ****)
   349 
   350 fun is_left (Const ("Trueprop", _) $
   351                (Const ("op |", _) $ (Const ("op |", _) $ _ $ _) $ _)) = true
   352   | is_left _ = false;
   353 
   354 (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)
   355 fun assoc_right th =
   356   if is_left (prop_of th) then assoc_right (th RS disj_assoc)
   357   else th;
   358 
   359 (*Must check for negative literal first!*)
   360 val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule];
   361 
   362 (*For ordinary resolution. *)
   363 val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule'];
   364 
   365 (*Create a goal or support clause, conclusing False*)
   366 fun make_goal th =   (*Must check for negative literal first!*)
   367     make_goal (tryres(th, clause_rules))
   368   handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]);
   369 
   370 (*Sort clauses by number of literals*)
   371 fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2);
   372 
   373 fun sort_clauses ths = sort (make_ord fewerlits) ths;
   374 
   375 (*True if the given type contains bool anywhere*)
   376 fun has_bool (Type("bool",_)) = true
   377   | has_bool (Type(_, Ts)) = exists has_bool Ts
   378   | has_bool _ = false;
   379 
   380 (*Is the string the name of a connective? Really only | and Not can remain,
   381   since this code expects to be called on a clause form.*)
   382 val is_conn = member (op =)
   383     ["Trueprop", "op &", "op |", "op -->", "Not",
   384      "All", "Ex", "Ball", "Bex"];
   385 
   386 (*True if the term contains a function--not a logical connective--where the type
   387   of any argument contains bool.*)
   388 val has_bool_arg_const =
   389     exists_Const
   390       (fn (c,T) => not(is_conn c) andalso exists (has_bool) (binder_types T));
   391 
   392 (*A higher-order instance of a first-order constant? Example is the definition of
   393   HOL.one, 1, at a function type in theory SetsAndFunctions.*)
   394 fun higher_inst_const thy (c,T) =
   395   case binder_types T of
   396       [] => false (*not a function type, OK*)
   397     | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts;
   398 
   399 (*Returns false if any Vars in the theorem mention type bool.
   400   Also rejects functions whose arguments are Booleans or other functions.*)
   401 fun is_fol_term thy t =
   402     Term.is_first_order ["all","All","Ex"] t andalso
   403     not (exists (has_bool o fastype_of) (term_vars t)  orelse
   404          has_bool_arg_const t  orelse
   405          exists_Const (higher_inst_const thy) t orelse
   406          has_meta_conn t);
   407 
   408 fun rigid t = not (is_Var (head_of t));
   409 
   410 fun ok4horn (Const ("Trueprop",_) $ (Const ("op |", _) $ t $ _)) = rigid t
   411   | ok4horn (Const ("Trueprop",_) $ t) = rigid t
   412   | ok4horn _ = false;
   413 
   414 (*Create a meta-level Horn clause*)
   415 fun make_horn crules th =
   416   if ok4horn (concl_of th)
   417   then make_horn crules (tryres(th,crules)) handle THM _ => th
   418   else th;
   419 
   420 (*Generate Horn clauses for all contrapositives of a clause. The input, th,
   421   is a HOL disjunction.*)
   422 fun add_contras crules (th,hcs) =
   423   let fun rots (0,th) = hcs
   424         | rots (k,th) = zero_var_indexes (make_horn crules th) ::
   425                         rots(k-1, assoc_right (th RS disj_comm))
   426   in case nliterals(prop_of th) of
   427         1 => th::hcs
   428       | n => rots(n, assoc_right th)
   429   end;
   430 
   431 (*Use "theorem naming" to label the clauses*)
   432 fun name_thms label =
   433     let fun name1 (th, (k,ths)) =
   434           (k-1, PureThy.put_name_hint (label ^ string_of_int k) th :: ths)
   435     in  fn ths => #2 (foldr name1 (length ths, []) ths)  end;
   436 
   437 (*Is the given disjunction an all-negative support clause?*)
   438 fun is_negative th = forall (not o #1) (literals (prop_of th));
   439 
   440 val neg_clauses = List.filter is_negative;
   441 
   442 
   443 (***** MESON PROOF PROCEDURE *****)
   444 
   445 fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi,
   446            As) = rhyps(phi, A::As)
   447   | rhyps (_, As) = As;
   448 
   449 (** Detecting repeated assumptions in a subgoal **)
   450 
   451 (*The stringtree detects repeated assumptions.*)
   452 fun ins_term (net,t) = Net.insert_term (op aconv) (t,t) net;
   453 
   454 (*detects repetitions in a list of terms*)
   455 fun has_reps [] = false
   456   | has_reps [_] = false
   457   | has_reps [t,u] = (t aconv u)
   458   | has_reps ts = (Library.foldl ins_term (Net.empty, ts);  false)
   459                   handle Net.INSERT => true;
   460 
   461 (*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*)
   462 fun TRYING_eq_assume_tac 0 st = Seq.single st
   463   | TRYING_eq_assume_tac i st =
   464        TRYING_eq_assume_tac (i-1) (eq_assumption i st)
   465        handle THM _ => TRYING_eq_assume_tac (i-1) st;
   466 
   467 fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (nprems_of st) st;
   468 
   469 (*Loop checking: FAIL if trying to prove the same thing twice
   470   -- if *ANY* subgoal has repeated literals*)
   471 fun check_tac st =
   472   if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st)
   473   then  Seq.empty  else  Seq.single st;
   474 
   475 
   476 (* net_resolve_tac actually made it slower... *)
   477 fun prolog_step_tac horns i =
   478     (assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN
   479     TRYALL_eq_assume_tac;
   480 
   481 (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*)
   482 fun addconcl(prem,sz) = size_of_term(Logic.strip_assums_concl prem) + sz
   483 
   484 fun size_of_subgoals st = foldr addconcl 0 (prems_of st);
   485 
   486 
   487 (*Negation Normal Form*)
   488 val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD,
   489                not_impD, not_iffD, not_allD, not_exD, not_notD];
   490 
   491 fun ok4nnf (Const ("Trueprop",_) $ (Const ("Not", _) $ t)) = rigid t
   492   | ok4nnf (Const ("Trueprop",_) $ t) = rigid t
   493   | ok4nnf _ = false;
   494 
   495 fun make_nnf1 th =
   496   if ok4nnf (concl_of th)
   497   then make_nnf1 (tryres(th, nnf_rls))
   498     handle THM _ =>
   499         forward_res make_nnf1
   500            (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward]))
   501     handle THM _ => th
   502   else th;
   503 
   504 (*The simplification removes defined quantifiers and occurrences of True and False.
   505   nnf_ss also includes the one-point simprocs,
   506   which are needed to avoid the various one-point theorems from generating junk clauses.*)
   507 val nnf_simps =
   508      [simp_implies_def, Ex1_def, Ball_def, Bex_def, if_True,
   509       if_False, if_cancel, if_eq_cancel, cases_simp];
   510 val nnf_extra_simps =
   511       thms"split_ifs" @ ex_simps @ all_simps @ simp_thms;
   512 
   513 val nnf_ss =
   514   HOL_basic_ss addsimps nnf_extra_simps
   515     addsimprocs [defALL_regroup,defEX_regroup, @{simproc neq}, @{simproc let_simp}];
   516 
   517 fun make_nnf th = case prems_of th of
   518     [] => th |> rewrite_rule (map safe_mk_meta_eq nnf_simps)
   519              |> simplify nnf_ss
   520              |> make_nnf1
   521   | _ => raise THM ("make_nnf: premises in argument", 0, [th]);
   522 
   523 (*Pull existential quantifiers to front. This accomplishes Skolemization for
   524   clauses that arise from a subgoal.*)
   525 fun skolemize th =
   526   if not (has_conns ["Ex"] (prop_of th)) then th
   527   else (skolemize (tryres(th, [choice, conj_exD1, conj_exD2,
   528                               disj_exD, disj_exD1, disj_exD2])))
   529     handle THM _ =>
   530         skolemize (forward_res skolemize
   531                    (tryres (th, [conj_forward, disj_forward, all_forward])))
   532     handle THM _ => forward_res skolemize (th RS ex_forward);
   533 
   534 
   535 (*Make clauses from a list of theorems, previously Skolemized and put into nnf.
   536   The resulting clauses are HOL disjunctions.*)
   537 fun make_clauses ths =
   538     (sort_clauses (map (generalize o nodups) (foldr (cnf[]) [] ths)));
   539 
   540 (*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*)
   541 fun make_horns ths =
   542     name_thms "Horn#"
   543       (distinct Thm.eq_thm_prop (foldr (add_contras clause_rules) [] ths));
   544 
   545 (*Could simply use nprems_of, which would count remaining subgoals -- no
   546   discrimination as to their size!  With BEST_FIRST, fails for problem 41.*)
   547 
   548 fun best_prolog_tac sizef horns =
   549     BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1);
   550 
   551 fun depth_prolog_tac horns =
   552     DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1);
   553 
   554 (*Return all negative clauses, as possible goal clauses*)
   555 fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls));
   556 
   557 fun skolemize_prems_tac prems =
   558     cut_facts_tac (map (skolemize o make_nnf) prems)  THEN'
   559     REPEAT o (etac exE);
   560 
   561 (*Basis of all meson-tactics.  Supplies cltac with clauses: HOL disjunctions.
   562   Function mkcl converts theorems to clauses.*)
   563 fun MESON mkcl cltac i st =
   564   SELECT_GOAL
   565     (EVERY [ObjectLogic.atomize_prems_tac 1,
   566             rtac ccontr 1,
   567             METAHYPS (fn negs =>
   568                       EVERY1 [skolemize_prems_tac negs,
   569                               METAHYPS (cltac o mkcl)]) 1]) i st
   570   handle THM _ => no_tac st;    (*probably from make_meta_clause, not first-order*)
   571 
   572 (** Best-first search versions **)
   573 
   574 (*ths is a list of additional clauses (HOL disjunctions) to use.*)
   575 fun best_meson_tac sizef =
   576   MESON make_clauses
   577     (fn cls =>
   578          THEN_BEST_FIRST (resolve_tac (gocls cls) 1)
   579                          (has_fewer_prems 1, sizef)
   580                          (prolog_step_tac (make_horns cls) 1));
   581 
   582 (*First, breaks the goal into independent units*)
   583 val safe_best_meson_tac =
   584      SELECT_GOAL (TRY (CLASET safe_tac) THEN
   585                   TRYALL (best_meson_tac size_of_subgoals));
   586 
   587 (** Depth-first search version **)
   588 
   589 val depth_meson_tac =
   590   MESON make_clauses
   591     (fn cls => EVERY [resolve_tac (gocls cls) 1, depth_prolog_tac (make_horns cls)]);
   592 
   593 
   594 (** Iterative deepening version **)
   595 
   596 (*This version does only one inference per call;
   597   having only one eq_assume_tac speeds it up!*)
   598 fun prolog_step_tac' horns =
   599     let val (horn0s, hornps) = (*0 subgoals vs 1 or more*)
   600             take_prefix Thm.no_prems horns
   601         val nrtac = net_resolve_tac horns
   602     in  fn i => eq_assume_tac i ORELSE
   603                 match_tac horn0s i ORELSE  (*no backtracking if unit MATCHES*)
   604                 ((assume_tac i APPEND nrtac i) THEN check_tac)
   605     end;
   606 
   607 fun iter_deepen_prolog_tac horns =
   608     ITER_DEEPEN (has_fewer_prems 1) (prolog_step_tac' horns);
   609 
   610 fun iter_deepen_meson_tac ths = MESON make_clauses
   611  (fn cls =>
   612       case (gocls (cls@ths)) of
   613            [] => no_tac  (*no goal clauses*)
   614          | goes =>
   615              let val horns = make_horns (cls@ths)
   616                  val _ =
   617                      Output.debug (fn () => ("meson method called:\n" ^
   618                                   space_implode "\n" (map string_of_thm (cls@ths)) ^
   619                                   "\nclauses:\n" ^
   620                                   space_implode "\n" (map string_of_thm horns)))
   621              in THEN_ITER_DEEPEN (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns)
   622              end
   623  );
   624 
   625 fun meson_claset_tac ths cs =
   626   SELECT_GOAL (TRY (safe_tac cs) THEN TRYALL (iter_deepen_meson_tac ths));
   627 
   628 val meson_tac = CLASET' (meson_claset_tac []);
   629 
   630 
   631 (**** Code to support ordinary resolution, rather than Model Elimination ****)
   632 
   633 (*Convert a list of clauses (disjunctions) to meta-level clauses (==>),
   634   with no contrapositives, for ordinary resolution.*)
   635 
   636 (*Rules to convert the head literal into a negated assumption. If the head
   637   literal is already negated, then using notEfalse instead of notEfalse'
   638   prevents a double negation.*)
   639 val notEfalse = read_instantiate [("R","False")] notE;
   640 val notEfalse' = rotate_prems 1 notEfalse;
   641 
   642 fun negated_asm_of_head th =
   643     th RS notEfalse handle THM _ => th RS notEfalse';
   644 
   645 (*Converting one clause*)
   646 val make_meta_clause =
   647   zero_var_indexes o negated_asm_of_head o make_horn resolution_clause_rules;
   648 
   649 fun make_meta_clauses ths =
   650     name_thms "MClause#"
   651       (distinct Thm.eq_thm_prop (map make_meta_clause ths));
   652 
   653 (*Permute a rule's premises to move the i-th premise to the last position.*)
   654 fun make_last i th =
   655   let val n = nprems_of th
   656   in  if 1 <= i andalso i <= n
   657       then Thm.permute_prems (i-1) 1 th
   658       else raise THM("select_literal", i, [th])
   659   end;
   660 
   661 (*Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing
   662   double-negations.*)
   663 val negate_head = rewrite_rule [atomize_not, not_not RS eq_reflection];
   664 
   665 (*Maps the clause  [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P*)
   666 fun select_literal i cl = negate_head (make_last i cl);
   667 
   668 
   669 (*Top-level Skolemization. Allows part of the conversion to clauses to be
   670   expressed as a tactic (or Isar method).  Each assumption of the selected
   671   goal is converted to NNF and then its existential quantifiers are pulled
   672   to the front. Finally, all existential quantifiers are eliminated,
   673   leaving !!-quantified variables. Perhaps Safe_tac should follow, but it
   674   might generate many subgoals.*)
   675 
   676 fun skolemize_tac i st =
   677   let val ts = Logic.strip_assums_hyp (List.nth (prems_of st, i-1))
   678   in
   679      EVERY' [METAHYPS
   680             (fn hyps => (cut_facts_tac (map (skolemize o make_nnf) hyps) 1
   681                          THEN REPEAT (etac exE 1))),
   682             REPEAT_DETERM_N (length ts) o (etac thin_rl)] i st
   683   end
   684   handle Subscript => Seq.empty;
   685 
   686 end;
   687