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