src/HOL/Tools/meson.ML
author paulson
Wed, 19 May 2004 11:31:26 +0200
changeset 14763 c1fd297712ba
parent 14744 7ccfc167e451
child 14813 826edbc317e6
permissions -rw-r--r--
has_consts now handles the @-operator
     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 local
    15 
    16  val not_conjD = thm "meson_not_conjD";
    17  val not_disjD = thm "meson_not_disjD";
    18  val not_notD = thm "meson_not_notD";
    19  val not_allD = thm "meson_not_allD";
    20  val not_exD = thm "meson_not_exD";
    21  val imp_to_disjD = thm "meson_imp_to_disjD";
    22  val not_impD = thm "meson_not_impD";
    23  val iff_to_disjD = thm "meson_iff_to_disjD";
    24  val not_iffD = thm "meson_not_iffD";
    25  val conj_exD1 = thm "meson_conj_exD1";
    26  val conj_exD2 = thm "meson_conj_exD2";
    27  val disj_exD = thm "meson_disj_exD";
    28  val disj_exD1 = thm "meson_disj_exD1";
    29  val disj_exD2 = thm "meson_disj_exD2";
    30  val disj_assoc = thm "meson_disj_assoc";
    31  val disj_comm = thm "meson_disj_comm";
    32  val disj_FalseD1 = thm "meson_disj_FalseD1";
    33  val disj_FalseD2 = thm "meson_disj_FalseD2";
    34 
    35 
    36  (**** Operators for forward proof ****)
    37 
    38  (*raises exception if no rules apply -- unlike RL*)
    39  fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))
    40    | tryres (th, []) = raise THM("tryres", 0, [th]);
    41 
    42  val prop_of = #prop o rep_thm;
    43 
    44  (*Permits forward proof from rules that discharge assumptions*)
    45  fun forward_res nf st =
    46    case Seq.pull (ALLGOALS (METAHYPS (fn [prem] => rtac (nf prem) 1)) st)
    47    of Some(th,_) => th
    48     | None => raise THM("forward_res", 0, [st]);
    49 
    50 
    51  (*Are any of the constants in "bs" present in the term?*)
    52  fun has_consts bs =
    53    let fun has (Const(a,_)) = a mem bs
    54 	 | has (Const ("Hilbert_Choice.Eps",_) $ _) = false
    55                       (*ignore constants within @-terms*)
    56          | has (f$u) = has f orelse has u
    57          | has (Abs(_,_,t)) = has t
    58          | has _ = false
    59    in  has  end;
    60 
    61 
    62  (**** Clause handling ****)
    63 
    64  fun literals (Const("Trueprop",_) $ P) = literals P
    65    | literals (Const("op |",_) $ P $ Q) = literals P @ literals Q
    66    | literals (Const("Not",_) $ P) = [(false,P)]
    67    | literals P = [(true,P)];
    68 
    69  (*number of literals in a term*)
    70  val nliterals = length o literals;
    71 
    72  (*to detect, and remove, tautologous clauses*)
    73  fun taut_lits [] = false
    74    | taut_lits ((flg,t)::ts) = (not flg,t) mem ts orelse taut_lits ts;
    75 
    76  (*Include False as a literal: an occurrence of ~False is a tautology*)
    77  fun is_taut th = taut_lits ((true, HOLogic.false_const) ::
    78                              literals (prop_of th));
    79 
    80  (*Generation of unique names -- maxidx cannot be relied upon to increase!
    81    Cannot rely on "variant", since variables might coincide when literals
    82    are joined to make a clause...
    83    19 chooses "U" as the first variable name*)
    84  val name_ref = ref 19;
    85 
    86  (*Replaces universally quantified variables by FREE variables -- because
    87    assumptions may not contain scheme variables.  Later, call "generalize". *)
    88  fun freeze_spec th =
    89    let val sth = th RS spec
    90        val newname = (name_ref := !name_ref + 1;
    91                       radixstring(26, "A", !name_ref))
    92    in  read_instantiate [("x", newname)] sth  end;
    93 
    94  fun resop nf [prem] = resolve_tac (nf prem) 1;
    95 
    96  (*Conjunctive normal form, detecting tautologies early.
    97    Strips universal quantifiers and breaks up conjunctions. *)
    98  fun cnf_aux seen (th,ths) =
    99    if taut_lits (literals(prop_of th) @ seen)  
   100    then ths     (*tautology ignored*)
   101    else if not (has_consts ["All","op &"] (prop_of th))  
   102    then th::ths (*no work to do, terminate*)
   103    else (*conjunction?*)
   104          cnf_aux seen (th RS conjunct1,
   105                        cnf_aux seen (th RS conjunct2, ths))
   106    handle THM _ => (*universal quant?*)
   107          cnf_aux  seen (freeze_spec th,  ths)
   108    handle THM _ => (*disjunction?*)
   109      let val tac =
   110          (METAHYPS (resop (cnf_nil seen)) 1) THEN
   111          (fn st' => st' |>
   112                  METAHYPS (resop (cnf_nil (literals (concl_of st') @ seen))) 1)
   113      in  Seq.list_of (tac (th RS disj_forward)) @ ths  end
   114  and cnf_nil seen th = cnf_aux seen (th,[]);
   115 
   116  (*Top-level call to cnf -- it's safe to reset name_ref*)
   117  fun cnf (th,ths) =
   118     (name_ref := 19;  cnf (th RS conjunct1, cnf (th RS conjunct2, ths))
   119      handle THM _ => (*not a conjunction*) cnf_aux [] (th, ths));
   120 
   121  (**** Removal of duplicate literals ****)
   122 
   123  (*Forward proof, passing extra assumptions as theorems to the tactic*)
   124  fun forward_res2 nf hyps st =
   125    case Seq.pull
   126          (REPEAT
   127           (METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)
   128           st)
   129    of Some(th,_) => th
   130     | None => raise THM("forward_res2", 0, [st]);
   131 
   132  (*Remove duplicates in P|Q by assuming ~P in Q
   133    rls (initially []) accumulates assumptions of the form P==>False*)
   134  fun nodups_aux rls th = nodups_aux rls (th RS disj_assoc)
   135      handle THM _ => tryres(th,rls)
   136      handle THM _ => tryres(forward_res2 nodups_aux rls (th RS disj_forward2),
   137                             [disj_FalseD1, disj_FalseD2, asm_rl])
   138      handle THM _ => th;
   139 
   140  (*Remove duplicate literals, if there are any*)
   141  fun nodups th =
   142      if null(findrep(literals(prop_of th))) then th
   143      else nodups_aux [] th;
   144 
   145 
   146  (**** Generation of contrapositives ****)
   147 
   148  (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)
   149  fun assoc_right th = assoc_right (th RS disj_assoc)
   150          handle THM _ => th;
   151 
   152  (*Must check for negative literal first!*)
   153  val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule];
   154 
   155  (*For ordinary resolution. *)
   156  val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule'];
   157 
   158  (*Create a goal or support clause, conclusing False*)
   159  fun make_goal th =   (*Must check for negative literal first!*)
   160      make_goal (tryres(th, clause_rules))
   161    handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]);
   162 
   163  (*Sort clauses by number of literals*)
   164  fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2);
   165 
   166  (*TAUTOLOGY CHECK SHOULD NOT BE NECESSARY!*)
   167  fun sort_clauses ths = sort (make_ord fewerlits) (filter (not o is_taut) ths);
   168 
   169  (*Convert all suitable free variables to schematic variables*)
   170  fun generalize th = forall_elim_vars 0 (forall_intr_frees th);
   171 
   172  (*Create a meta-level Horn clause*)
   173  fun make_horn crules th = make_horn crules (tryres(th,crules))
   174                            handle THM _ => th;
   175 
   176  (*Generate Horn clauses for all contrapositives of a clause*)
   177  fun add_contras crules (th,hcs) =
   178    let fun rots (0,th) = hcs
   179          | rots (k,th) = zero_var_indexes (make_horn crules th) ::
   180                          rots(k-1, assoc_right (th RS disj_comm))
   181    in case nliterals(prop_of th) of
   182          1 => th::hcs
   183        | n => rots(n, assoc_right th)
   184    end;
   185 
   186  (*Use "theorem naming" to label the clauses*)
   187  fun name_thms label =
   188      let fun name1 (th, (k,ths)) =
   189            (k-1, Thm.name_thm (label ^ string_of_int k, th) :: ths)
   190 
   191      in  fn ths => #2 (foldr name1 (ths, (length ths, [])))  end;
   192 
   193  (*Find an all-negative support clause*)
   194  fun is_negative th = forall (not o #1) (literals (prop_of th));
   195 
   196  val neg_clauses = filter is_negative;
   197 
   198 
   199  (***** MESON PROOF PROCEDURE *****)
   200 
   201  fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi,
   202             As) = rhyps(phi, A::As)
   203    | rhyps (_, As) = As;
   204 
   205  (** Detecting repeated assumptions in a subgoal **)
   206 
   207  (*The stringtree detects repeated assumptions.*)
   208  fun ins_term (net,t) = Net.insert_term((t,t), net, op aconv);
   209 
   210  (*detects repetitions in a list of terms*)
   211  fun has_reps [] = false
   212    | has_reps [_] = false
   213    | has_reps [t,u] = (t aconv u)
   214    | has_reps ts = (foldl ins_term (Net.empty, ts);  false)
   215                    handle INSERT => true;
   216 
   217  (*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*)
   218  fun TRYALL_eq_assume_tac 0 st = Seq.single st
   219    | TRYALL_eq_assume_tac i st =
   220         TRYALL_eq_assume_tac (i-1) (eq_assumption i st)
   221         handle THM _ => TRYALL_eq_assume_tac (i-1) st;
   222 
   223  (*Loop checking: FAIL if trying to prove the same thing twice
   224    -- if *ANY* subgoal has repeated literals*)
   225  fun check_tac st =
   226    if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st)
   227    then  Seq.empty  else  Seq.single st;
   228 
   229 
   230  (* net_resolve_tac actually made it slower... *)
   231  fun prolog_step_tac horns i =
   232      (assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN
   233      TRYALL eq_assume_tac;
   234 
   235 
   236 in
   237 
   238 
   239 (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*)
   240 local fun addconcl(prem,sz) = size_of_term(Logic.strip_assums_concl prem) + sz
   241 in
   242 fun size_of_subgoals st = foldr addconcl (prems_of st, 0)
   243 end;
   244 
   245 (*Negation Normal Form*)
   246 val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD,
   247                not_impD, not_iffD, not_allD, not_exD, not_notD];
   248 fun make_nnf th = make_nnf (tryres(th, nnf_rls))
   249     handle THM _ =>
   250         forward_res make_nnf
   251            (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward]))
   252     handle THM _ => th;
   253 
   254 (*Pull existential quantifiers (Skolemization)*)
   255 fun skolemize th =
   256   if not (has_consts ["Ex"] (prop_of th)) then th
   257   else skolemize (tryres(th, [choice, conj_exD1, conj_exD2,
   258                               disj_exD, disj_exD1, disj_exD2]))
   259     handle THM _ =>
   260         skolemize (forward_res skolemize
   261                    (tryres (th, [conj_forward, disj_forward, all_forward])))
   262     handle THM _ => forward_res skolemize (th RS ex_forward);
   263 
   264 
   265 (*Make clauses from a list of theorems, previously Skolemized and put into nnf.
   266   The resulting clauses are HOL disjunctions.*)
   267 fun make_clauses ths =
   268     sort_clauses (map (generalize o nodups) (foldr cnf (ths,[])));
   269 
   270 (*Convert a list of clauses to (contrapositive) Horn clauses*)
   271 fun make_horns ths =
   272     name_thms "Horn#"
   273       (gen_distinct Drule.eq_thm_prop (foldr (add_contras clause_rules) (ths,[])));
   274 
   275 (*Could simply use nprems_of, which would count remaining subgoals -- no
   276   discrimination as to their size!  With BEST_FIRST, fails for problem 41.*)
   277 
   278 fun best_prolog_tac sizef horns =
   279     BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1);
   280 
   281 fun depth_prolog_tac horns =
   282     DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1);
   283 
   284 (*Return all negative clauses, as possible goal clauses*)
   285 fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls));
   286 
   287 
   288 fun skolemize_tac prems =
   289     cut_facts_tac (map (skolemize o make_nnf) prems)  THEN'
   290     REPEAT o (etac exE);
   291 
   292 (*Shell of all meson-tactics.  Supplies cltac with clauses: HOL disjunctions*)
   293 fun MESON cltac = SELECT_GOAL
   294  (EVERY1 [rtac ccontr,
   295           METAHYPS (fn negs =>
   296                     EVERY1 [skolemize_tac negs,
   297                             METAHYPS (cltac o make_clauses)])]);
   298 
   299 (** Best-first search versions **)
   300 
   301 fun best_meson_tac sizef =
   302   MESON (fn cls =>
   303          THEN_BEST_FIRST (resolve_tac (gocls cls) 1)
   304                          (has_fewer_prems 1, sizef)
   305                          (prolog_step_tac (make_horns cls) 1));
   306 
   307 (*First, breaks the goal into independent units*)
   308 val safe_best_meson_tac =
   309      SELECT_GOAL (TRY Safe_tac THEN
   310                   TRYALL (best_meson_tac size_of_subgoals));
   311 
   312 (** Depth-first search version **)
   313 
   314 val depth_meson_tac =
   315      MESON (fn cls => EVERY [resolve_tac (gocls cls) 1,
   316                              depth_prolog_tac (make_horns cls)]);
   317 
   318 
   319 
   320 (** Iterative deepening version **)
   321 
   322 (*This version does only one inference per call;
   323   having only one eq_assume_tac speeds it up!*)
   324 fun prolog_step_tac' horns =
   325     let val (horn0s, hornps) = (*0 subgoals vs 1 or more*)
   326             take_prefix Thm.no_prems horns
   327         val nrtac = net_resolve_tac horns
   328     in  fn i => eq_assume_tac i ORELSE
   329                 match_tac horn0s i ORELSE  (*no backtracking if unit MATCHES*)
   330                 ((assume_tac i APPEND nrtac i) THEN check_tac)
   331     end;
   332 
   333 fun iter_deepen_prolog_tac horns =
   334     ITER_DEEPEN (has_fewer_prems 1) (prolog_step_tac' horns);
   335 
   336 val iter_deepen_meson_tac =
   337   MESON (fn cls =>
   338          (THEN_ITER_DEEPEN (resolve_tac (gocls cls) 1)
   339                            (has_fewer_prems 1)
   340                            (prolog_step_tac' (make_horns cls))));
   341 
   342 fun meson_claset_tac cs =
   343   SELECT_GOAL (TRY (safe_tac cs) THEN TRYALL iter_deepen_meson_tac);
   344 
   345 val meson_tac = CLASET' meson_claset_tac;
   346 
   347 
   348 (** Code to support ordinary resolution, rather than Model Elimination **)
   349 
   350 (*Convert a list of clauses to meta-level clauses, with no contrapositives,
   351   for ordinary resolution.*)
   352 
   353 (*Rules to convert the head literal into a negated assumption. If the head
   354   literal is already negated, then using notEfalse instead of notEfalse'
   355   prevents a double negation.*)
   356 val notEfalse = read_instantiate [("R","False")] notE;
   357 val notEfalse' = rotate_prems 1 notEfalse;
   358 
   359 fun negate_nead th = 
   360     th RS notEfalse handle THM _ => th RS notEfalse';
   361 
   362 (*Converting one clause*)
   363 fun make_meta_clause th = negate_nead (make_horn resolution_clause_rules th);
   364 
   365 fun make_meta_clauses ths =
   366     name_thms "MClause#"
   367       (gen_distinct Drule.eq_thm_prop (map make_meta_clause ths));
   368 
   369 (*Permute a rule's premises to move the i-th premise to the last position.*)
   370 fun make_last i th =
   371   let val n = nprems_of th 
   372   in  if 1 <= i andalso i <= n 
   373       then Thm.permute_prems (i-1) 1 th
   374       else raise THM("make_last", i, [th])
   375   end;
   376 
   377 (*Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing
   378   double-negations.*)
   379 val negate_head = rewrite_rule [atomize_not, not_not RS eq_reflection];
   380 
   381 (*Maps the clause  [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P*)
   382 fun select_literal i cl = negate_head (make_last i cl);
   383 
   384 
   385 
   386 (** proof method setup **)
   387 
   388 local
   389 
   390 fun meson_meth ctxt =
   391   Method.SIMPLE_METHOD' HEADGOAL
   392     (CHANGED_PROP o meson_claset_tac (Classical.get_local_claset ctxt));
   393 
   394 in
   395 
   396 val meson_setup =
   397  [Method.add_methods
   398   [("meson", Method.ctxt_args meson_meth, "The MESON resolution proof procedure")]];
   399 
   400 end;
   401 
   402 end;