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