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