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