# HG changeset patch # User blanchet # Date 1286221062 -7200 # Node ID 1f01c9b2b76b8f929b13ad59e5435209e82cb12a # Parent 6e9aff5ee9b5d90e979ee1ebee2715ec57498280 move MESON files together diff -r 6e9aff5ee9b5 -r 1f01c9b2b76b src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Oct 04 20:55:55 2010 +0200 +++ b/src/HOL/IsaMakefile Mon Oct 04 21:37:42 2010 +0200 @@ -201,6 +201,8 @@ Tools/inductive_realizer.ML \ Tools/inductive_set.ML \ Tools/lin_arith.ML \ + Tools/Meson/meson.ML \ + Tools/Meson/meson_clausify.ML \ Tools/nat_arith.ML \ Tools/primrec.ML \ Tools/prop_logic.ML \ @@ -275,7 +277,6 @@ Tools/int_arith.ML \ Tools/groebner.ML \ Tools/list_code.ML \ - Tools/meson.ML \ Tools/nat_numeral_simprocs.ML \ Tools/Nitpick/kodkod.ML \ Tools/Nitpick/kodkod_sat.ML \ @@ -315,7 +316,6 @@ Tools/recdef.ML \ Tools/record.ML \ Tools/semiring_normalizer.ML \ - Tools/Sledgehammer/meson_clausify.ML \ Tools/Sledgehammer/metis_reconstruct.ML \ Tools/Sledgehammer/metis_translate.ML \ Tools/Sledgehammer/metis_tactics.ML \ diff -r 6e9aff5ee9b5 -r 1f01c9b2b76b src/HOL/Tools/Meson/meson.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Meson/meson.ML Mon Oct 04 21:37:42 2010 +0200 @@ -0,0 +1,712 @@ +(* Title: HOL/Tools/meson.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + +The MESON resolution proof procedure for HOL. +When making clauses, avoids using the rewriter -- instead uses RS recursively. +*) + +signature MESON = +sig + val trace: bool Unsynchronized.ref + val term_pair_of: indexname * (typ * 'a) -> term * 'a + val size_of_subgoals: thm -> int + val has_too_many_clauses: Proof.context -> term -> bool + val make_cnf: thm list -> thm -> Proof.context -> thm list * Proof.context + val finish_cnf: thm list -> thm list + val presimplify: thm -> thm + val make_nnf: Proof.context -> thm -> thm + val skolemize_with_choice_thms : Proof.context -> thm list -> thm -> thm + val skolemize : Proof.context -> thm -> thm + val is_fol_term: theory -> term -> bool + val make_clauses_unsorted: thm list -> thm list + val make_clauses: thm list -> thm list + val make_horns: thm list -> thm list + val best_prolog_tac: (thm -> int) -> thm list -> tactic + val depth_prolog_tac: thm list -> tactic + val gocls: thm list -> thm list + val skolemize_prems_tac : Proof.context -> thm list -> int -> tactic + val MESON: + tactic -> (thm list -> thm list) -> (thm list -> tactic) -> Proof.context + -> int -> tactic + val best_meson_tac: (thm -> int) -> Proof.context -> int -> tactic + val safe_best_meson_tac: Proof.context -> int -> tactic + val depth_meson_tac: Proof.context -> int -> tactic + val prolog_step_tac': thm list -> int -> tactic + val iter_deepen_prolog_tac: thm list -> tactic + val iter_deepen_meson_tac: Proof.context -> thm list -> int -> tactic + val make_meta_clause: thm -> thm + val make_meta_clauses: thm list -> thm list + val meson_tac: Proof.context -> thm list -> int -> tactic + val setup: theory -> theory +end + +structure Meson : MESON = +struct + +val trace = Unsynchronized.ref false; +fun trace_msg msg = if ! trace then tracing (msg ()) else (); + +val max_clauses_default = 60; +val (max_clauses, setup) = Attrib.config_int "meson_max_clauses" (K max_clauses_default); + +(*No known example (on 1-5-2007) needs even thirty*) +val iter_deepen_limit = 50; + +val disj_forward = @{thm disj_forward}; +val disj_forward2 = @{thm disj_forward2}; +val make_pos_rule = @{thm make_pos_rule}; +val make_pos_rule' = @{thm make_pos_rule'}; +val make_pos_goal = @{thm make_pos_goal}; +val make_neg_rule = @{thm make_neg_rule}; +val make_neg_rule' = @{thm make_neg_rule'}; +val make_neg_goal = @{thm make_neg_goal}; +val conj_forward = @{thm conj_forward}; +val all_forward = @{thm all_forward}; +val ex_forward = @{thm ex_forward}; + +val not_conjD = @{thm meson_not_conjD}; +val not_disjD = @{thm meson_not_disjD}; +val not_notD = @{thm meson_not_notD}; +val not_allD = @{thm meson_not_allD}; +val not_exD = @{thm meson_not_exD}; +val imp_to_disjD = @{thm meson_imp_to_disjD}; +val not_impD = @{thm meson_not_impD}; +val iff_to_disjD = @{thm meson_iff_to_disjD}; +val not_iffD = @{thm meson_not_iffD}; +val conj_exD1 = @{thm meson_conj_exD1}; +val conj_exD2 = @{thm meson_conj_exD2}; +val disj_exD = @{thm meson_disj_exD}; +val disj_exD1 = @{thm meson_disj_exD1}; +val disj_exD2 = @{thm meson_disj_exD2}; +val disj_assoc = @{thm meson_disj_assoc}; +val disj_comm = @{thm meson_disj_comm}; +val disj_FalseD1 = @{thm meson_disj_FalseD1}; +val disj_FalseD2 = @{thm meson_disj_FalseD2}; + + +(**** Operators for forward proof ****) + + +(** First-order Resolution **) + +fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t); + +(*FIXME: currently does not "rename variables apart"*) +fun first_order_resolve thA thB = + (case + try (fn () => + let val thy = theory_of_thm thA + val tmA = concl_of thA + val Const("==>",_) $ tmB $ _ = prop_of thB + val tenv = + Pattern.first_order_match thy (tmB, tmA) + (Vartab.empty, Vartab.empty) |> snd + val ct_pairs = map (pairself (cterm_of thy) o term_pair_of) (Vartab.dest tenv) + in thA RS (cterm_instantiate ct_pairs thB) end) () of + SOME th => th + | NONE => raise THM ("first_order_resolve", 0, [thA, thB])) + +(* Applying "choice" swaps the bound variable names. We tweak + "Thm.rename_boundvars"'s input to get the desired names. *) +fun fix_bounds (_ $ (Const (@{const_name Ex}, _) + $ Abs (_, _, Const (@{const_name All}, _) $ _))) + (t0 $ (Const (@{const_name All}, T1) + $ Abs (a1, T1', Const (@{const_name Ex}, T2) + $ Abs (a2, T2', t')))) = + t0 $ (Const (@{const_name All}, T1) + $ Abs (a2, T1', Const (@{const_name Ex}, T2) $ Abs (a1, T2', t'))) + | fix_bounds _ t = t + +(* Hack to make it less likely that we lose our precious bound variable names in + "rename_bvs_RS" below, because of a clash. *) +val protect_prefix = "_" + +fun protect_bounds (t $ u) = protect_bounds t $ protect_bounds u + | protect_bounds (Abs (s, T, t')) = + Abs (protect_prefix ^ s, T, protect_bounds t') + | protect_bounds t = t + +(* Forward proof while preserving bound variables names*) +fun rename_bvs_RS th rl = + let + val t = concl_of th + val r = concl_of rl + val th' = th RS Thm.rename_boundvars r (protect_bounds r) rl + val t' = concl_of th' + in Thm.rename_boundvars t' (fix_bounds t' t) th' end + +(*raises exception if no rules apply*) +fun tryres (th, rls) = + let fun tryall [] = raise THM("tryres", 0, th::rls) + | tryall (rl::rls) = (rename_bvs_RS th rl handle THM _ => tryall rls) + in tryall rls end; + +(*Permits forward proof from rules that discharge assumptions. The supplied proof state st, + e.g. from conj_forward, should have the form + "[| P' ==> ?P; Q' ==> ?Q |] ==> ?P & ?Q" + and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*) +fun forward_res ctxt nf st = + let fun forward_tacf [prem] = rtac (nf prem) 1 + | forward_tacf prems = + error (cat_lines + ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" :: + Display.string_of_thm ctxt st :: + "Premises:" :: map (Display.string_of_thm ctxt) prems)) + in + case Seq.pull (ALLGOALS (Misc_Legacy.METAHYPS forward_tacf) st) + of SOME(th,_) => th + | NONE => raise THM("forward_res", 0, [st]) + end; + +(*Are any of the logical connectives in "bs" present in the term?*) +fun has_conns bs = + let fun has (Const _) = false + | has (Const(@{const_name Trueprop},_) $ p) = has p + | has (Const(@{const_name Not},_) $ p) = has p + | has (Const(@{const_name HOL.disj},_) $ p $ q) = member (op =) bs @{const_name HOL.disj} orelse has p orelse has q + | has (Const(@{const_name HOL.conj},_) $ p $ q) = member (op =) bs @{const_name HOL.conj} orelse has p orelse has q + | has (Const(@{const_name All},_) $ Abs(_,_,p)) = member (op =) bs @{const_name All} orelse has p + | has (Const(@{const_name Ex},_) $ Abs(_,_,p)) = member (op =) bs @{const_name Ex} orelse has p + | has _ = false + in has end; + + +(**** Clause handling ****) + +fun literals (Const(@{const_name Trueprop},_) $ P) = literals P + | literals (Const(@{const_name HOL.disj},_) $ P $ Q) = literals P @ literals Q + | literals (Const(@{const_name Not},_) $ P) = [(false,P)] + | literals P = [(true,P)]; + +(*number of literals in a term*) +val nliterals = length o literals; + + +(*** Tautology Checking ***) + +fun signed_lits_aux (Const (@{const_name HOL.disj}, _) $ P $ Q) (poslits, neglits) = + signed_lits_aux Q (signed_lits_aux P (poslits, neglits)) + | signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits) + | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits); + +fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]); + +(*Literals like X=X are tautologous*) +fun taut_poslit (Const(@{const_name HOL.eq},_) $ t $ u) = t aconv u + | taut_poslit (Const(@{const_name True},_)) = true + | taut_poslit _ = false; + +fun is_taut th = + let val (poslits,neglits) = signed_lits th + in exists taut_poslit poslits + orelse + exists (member (op aconv) neglits) (HOLogic.false_const :: poslits) + end + handle TERM _ => false; (*probably dest_Trueprop on a weird theorem*) + + +(*** To remove trivial negated equality literals from clauses ***) + +(*They are typically functional reflexivity axioms and are the converses of + injectivity equivalences*) + +val not_refl_disj_D = @{thm meson_not_refl_disj_D}; + +(*Is either term a Var that does not properly occur in the other term?*) +fun eliminable (t as Var _, u) = t aconv u orelse not (Logic.occs(t,u)) + | eliminable (u, t as Var _) = t aconv u orelse not (Logic.occs(t,u)) + | eliminable _ = false; + +fun refl_clause_aux 0 th = th + | refl_clause_aux n th = + case HOLogic.dest_Trueprop (concl_of th) of + (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _) => + refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*) + | (Const (@{const_name HOL.disj}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ t $ u)) $ _) => + if eliminable(t,u) + then refl_clause_aux (n-1) (th RS not_refl_disj_D) (*Var inequation: delete*) + else refl_clause_aux (n-1) (th RS disj_comm) (*not between Vars: ignore*) + | (Const (@{const_name HOL.disj}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm) + | _ => (*not a disjunction*) th; + +fun notequal_lits_count (Const (@{const_name HOL.disj}, _) $ P $ Q) = + notequal_lits_count P + notequal_lits_count Q + | notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _)) = 1 + | notequal_lits_count _ = 0; + +(*Simplify a clause by applying reflexivity to its negated equality literals*) +fun refl_clause th = + let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (concl_of th)) + in zero_var_indexes (refl_clause_aux neqs th) end + handle TERM _ => th; (*probably dest_Trueprop on a weird theorem*) + + +(*** Removal of duplicate literals ***) + +(*Forward proof, passing extra assumptions as theorems to the tactic*) +fun forward_res2 nf hyps st = + case Seq.pull + (REPEAT + (Misc_Legacy.METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1) + st) + of SOME(th,_) => th + | NONE => raise THM("forward_res2", 0, [st]); + +(*Remove duplicates in P|Q by assuming ~P in Q + rls (initially []) accumulates assumptions of the form P==>False*) +fun nodups_aux ctxt rls th = nodups_aux ctxt rls (th RS disj_assoc) + handle THM _ => tryres(th,rls) + handle THM _ => tryres(forward_res2 (nodups_aux ctxt) rls (th RS disj_forward2), + [disj_FalseD1, disj_FalseD2, asm_rl]) + handle THM _ => th; + +(*Remove duplicate literals, if there are any*) +fun nodups ctxt th = + if has_duplicates (op =) (literals (prop_of th)) + then nodups_aux ctxt [] th + else th; + + +(*** The basic CNF transformation ***) + +fun estimated_num_clauses bound t = + let + fun sum x y = if x < bound andalso y < bound then x+y else bound + fun prod x y = if x < bound andalso y < bound then x*y else bound + + (*Estimate the number of clauses in order to detect infeasible theorems*) + fun signed_nclauses b (Const(@{const_name Trueprop},_) $ t) = signed_nclauses b t + | signed_nclauses b (Const(@{const_name Not},_) $ t) = signed_nclauses (not b) t + | signed_nclauses b (Const(@{const_name HOL.conj},_) $ t $ u) = + if b then sum (signed_nclauses b t) (signed_nclauses b u) + else prod (signed_nclauses b t) (signed_nclauses b u) + | signed_nclauses b (Const(@{const_name HOL.disj},_) $ t $ u) = + if b then prod (signed_nclauses b t) (signed_nclauses b u) + else sum (signed_nclauses b t) (signed_nclauses b u) + | signed_nclauses b (Const(@{const_name HOL.implies},_) $ t $ u) = + if b then prod (signed_nclauses (not b) t) (signed_nclauses b u) + else sum (signed_nclauses (not b) t) (signed_nclauses b u) + | signed_nclauses b (Const(@{const_name HOL.eq}, Type ("fun", [T, _])) $ t $ u) = + if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*) + if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u)) + (prod (signed_nclauses (not b) u) (signed_nclauses b t)) + else sum (prod (signed_nclauses b t) (signed_nclauses b u)) + (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u)) + else 1 + | signed_nclauses b (Const(@{const_name Ex}, _) $ Abs (_,_,t)) = signed_nclauses b t + | signed_nclauses b (Const(@{const_name All},_) $ Abs (_,_,t)) = signed_nclauses b t + | signed_nclauses _ _ = 1; (* literal *) + in signed_nclauses true t end + +fun has_too_many_clauses ctxt t = + let val max_cl = Config.get ctxt max_clauses in + estimated_num_clauses (max_cl + 1) t > max_cl + end + +(*Replaces universally quantified variables by FREE variables -- because + assumptions may not contain scheme variables. Later, generalize using Variable.export. *) +local + val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec)))); + val spec_varT = #T (Thm.rep_cterm spec_var); + fun name_of (Const (@{const_name All}, _) $ Abs(x,_,_)) = x | name_of _ = Name.uu; +in + fun freeze_spec th ctxt = + let + val cert = Thm.cterm_of (ProofContext.theory_of ctxt); + val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (concl_of th))] ctxt; + val spec' = Thm.instantiate ([], [(spec_var, cert (Free (x, spec_varT)))]) spec; + in (th RS spec', ctxt') end +end; + +(*Used with METAHYPS below. There is one assumption, which gets bound to prem + and then normalized via function nf. The normal form is given to resolve_tac, + instantiate a Boolean variable created by resolution with disj_forward. Since + (nf prem) returns a LIST of theorems, we can backtrack to get all combinations.*) +fun resop nf [prem] = resolve_tac (nf prem) 1; + +(* Any need to extend this list with "HOL.type_class", "HOL.eq_class", + and "Pure.term"? *) +val has_meta_conn = exists_Const (member (op =) ["==", "==>", "=simp=>", "all", "prop"] o #1); + +fun apply_skolem_theorem (th, rls) = + let + fun tryall [] = raise THM ("apply_skolem_theorem", 0, th::rls) + | tryall (rl :: rls) = + first_order_resolve th rl handle THM _ => tryall rls + in tryall rls end + +(* Conjunctive normal form, adding clauses from th in front of ths (for foldr). + Strips universal quantifiers and breaks up conjunctions. + Eliminates existential quantifiers using Skolemization theorems. *) +fun cnf old_skolem_ths ctxt (th, ths) = + let val ctxtr = Unsynchronized.ref ctxt (* FIXME ??? *) + fun cnf_aux (th,ths) = + if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*) + else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (prop_of th)) + then nodups ctxt th :: ths (*no work to do, terminate*) + else case head_of (HOLogic.dest_Trueprop (concl_of th)) of + Const (@{const_name HOL.conj}, _) => (*conjunction*) + cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths)) + | Const (@{const_name All}, _) => (*universal quantifier*) + let val (th',ctxt') = freeze_spec th (!ctxtr) + in ctxtr := ctxt'; cnf_aux (th', ths) end + | Const (@{const_name Ex}, _) => + (*existential quantifier: Insert Skolem functions*) + cnf_aux (apply_skolem_theorem (th, old_skolem_ths), ths) + | Const (@{const_name HOL.disj}, _) => + (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using + all combinations of converting P, Q to CNF.*) + let val tac = + Misc_Legacy.METAHYPS (resop cnf_nil) 1 THEN + (fn st' => st' |> Misc_Legacy.METAHYPS (resop cnf_nil) 1) + in Seq.list_of (tac (th RS disj_forward)) @ ths end + | _ => nodups ctxt th :: ths (*no work to do*) + and cnf_nil th = cnf_aux (th,[]) + val cls = + if has_too_many_clauses ctxt (concl_of th) + then (trace_msg (fn () => "cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths) + else cnf_aux (th,ths) + in (cls, !ctxtr) end; + +fun make_cnf old_skolem_ths th ctxt = cnf old_skolem_ths ctxt (th, []) + +(*Generalization, removal of redundant equalities, removal of tautologies.*) +fun finish_cnf ths = filter (not o is_taut) (map refl_clause ths); + + +(**** Generation of contrapositives ****) + +fun is_left (Const (@{const_name Trueprop}, _) $ + (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _)) = true + | is_left _ = false; + +(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) +fun assoc_right th = + if is_left (prop_of th) then assoc_right (th RS disj_assoc) + else th; + +(*Must check for negative literal first!*) +val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; + +(*For ordinary resolution. *) +val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule']; + +(*Create a goal or support clause, conclusing False*) +fun make_goal th = (*Must check for negative literal first!*) + make_goal (tryres(th, clause_rules)) + handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); + +(*Sort clauses by number of literals*) +fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2); + +fun sort_clauses ths = sort (make_ord fewerlits) ths; + +fun has_bool @{typ bool} = true + | has_bool (Type (_, Ts)) = exists has_bool Ts + | has_bool _ = false + +fun has_fun (Type (@{type_name fun}, _)) = true + | has_fun (Type (_, Ts)) = exists has_fun Ts + | has_fun _ = false + +(*Is the string the name of a connective? Really only | and Not can remain, + since this code expects to be called on a clause form.*) +val is_conn = member (op =) + [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj}, + @{const_name HOL.implies}, @{const_name Not}, + @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex}]; + +(*True if the term contains a function--not a logical connective--where the type + of any argument contains bool.*) +val has_bool_arg_const = + exists_Const + (fn (c,T) => not(is_conn c) andalso exists has_bool (binder_types T)); + +(*A higher-order instance of a first-order constant? Example is the definition of + one, 1, at a function type in theory Function_Algebras.*) +fun higher_inst_const thy (c,T) = + case binder_types T of + [] => false (*not a function type, OK*) + | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts; + +(*Returns false if any Vars in the theorem mention type bool. + Also rejects functions whose arguments are Booleans or other functions.*) +fun is_fol_term thy t = + Term.is_first_order ["all", @{const_name All}, @{const_name Ex}] t andalso + not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T + | _ => false) t orelse + has_bool_arg_const t orelse + exists_Const (higher_inst_const thy) t orelse + has_meta_conn t); + +fun rigid t = not (is_Var (head_of t)); + +fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name HOL.disj}, _) $ t $ _)) = rigid t + | ok4horn (Const (@{const_name Trueprop},_) $ t) = rigid t + | ok4horn _ = false; + +(*Create a meta-level Horn clause*) +fun make_horn crules th = + if ok4horn (concl_of th) + then make_horn crules (tryres(th,crules)) handle THM _ => th + else th; + +(*Generate Horn clauses for all contrapositives of a clause. The input, th, + is a HOL disjunction.*) +fun add_contras crules th hcs = + let fun rots (0,_) = hcs + | rots (k,th) = zero_var_indexes (make_horn crules th) :: + rots(k-1, assoc_right (th RS disj_comm)) + in case nliterals(prop_of th) of + 1 => th::hcs + | n => rots(n, assoc_right th) + end; + +(*Use "theorem naming" to label the clauses*) +fun name_thms label = + let fun name1 th (k, ths) = + (k-1, Thm.put_name_hint (label ^ string_of_int k) th :: ths) + in fn ths => #2 (fold_rev name1 ths (length ths, [])) end; + +(*Is the given disjunction an all-negative support clause?*) +fun is_negative th = forall (not o #1) (literals (prop_of th)); + +val neg_clauses = filter is_negative; + + +(***** MESON PROOF PROCEDURE *****) + +fun rhyps (Const("==>",_) $ (Const(@{const_name Trueprop},_) $ A) $ phi, + As) = rhyps(phi, A::As) + | rhyps (_, As) = As; + +(** Detecting repeated assumptions in a subgoal **) + +(*The stringtree detects repeated assumptions.*) +fun ins_term t net = Net.insert_term (op aconv) (t, t) net; + +(*detects repetitions in a list of terms*) +fun has_reps [] = false + | has_reps [_] = false + | has_reps [t,u] = (t aconv u) + | has_reps ts = (fold ins_term ts Net.empty; false) handle Net.INSERT => true; + +(*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) +fun TRYING_eq_assume_tac 0 st = Seq.single st + | TRYING_eq_assume_tac i st = + TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st) + handle THM _ => TRYING_eq_assume_tac (i-1) st; + +fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (nprems_of st) st; + +(*Loop checking: FAIL if trying to prove the same thing twice + -- if *ANY* subgoal has repeated literals*) +fun check_tac st = + if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st) + then Seq.empty else Seq.single st; + + +(* net_resolve_tac actually made it slower... *) +fun prolog_step_tac horns i = + (assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN + TRYALL_eq_assume_tac; + +(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) +fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz; + +fun size_of_subgoals st = fold_rev addconcl (prems_of st) 0; + + +(*Negation Normal Form*) +val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, + not_impD, not_iffD, not_allD, not_exD, not_notD]; + +fun ok4nnf (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ t)) = rigid t + | ok4nnf (Const (@{const_name Trueprop},_) $ t) = rigid t + | ok4nnf _ = false; + +fun make_nnf1 ctxt th = + if ok4nnf (concl_of th) + then make_nnf1 ctxt (tryres(th, nnf_rls)) + handle THM ("tryres", _, _) => + forward_res ctxt (make_nnf1 ctxt) + (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward])) + handle THM ("tryres", _, _) => th + else th + +(*The simplification removes defined quantifiers and occurrences of True and False. + nnf_ss also includes the one-point simprocs, + which are needed to avoid the various one-point theorems from generating junk clauses.*) +val nnf_simps = + @{thms simp_implies_def Ex1_def Ball_def Bex_def if_True if_False if_cancel + if_eq_cancel cases_simp} +val nnf_extra_simps = @{thms split_ifs ex_simps all_simps simp_thms} + +val nnf_ss = + HOL_basic_ss addsimps nnf_extra_simps + addsimprocs [defALL_regroup,defEX_regroup, @{simproc neq}, @{simproc let_simp}]; + +val presimplify = + rewrite_rule (map safe_mk_meta_eq nnf_simps) #> simplify nnf_ss + +fun make_nnf ctxt th = case prems_of th of + [] => th |> presimplify |> make_nnf1 ctxt + | _ => raise THM ("make_nnf: premises in argument", 0, [th]); + +(* Pull existential quantifiers to front. This accomplishes Skolemization for + clauses that arise from a subgoal. *) +fun skolemize_with_choice_thms ctxt choice_ths = + let + fun aux th = + if not (has_conns [@{const_name Ex}] (prop_of th)) then + th + else + tryres (th, choice_ths @ + [conj_exD1, conj_exD2, disj_exD, disj_exD1, disj_exD2]) + |> aux + handle THM ("tryres", _, _) => + tryres (th, [conj_forward, disj_forward, all_forward]) + |> forward_res ctxt aux + |> aux + handle THM ("tryres", _, _) => + rename_bvs_RS th ex_forward + |> forward_res ctxt aux + in aux o make_nnf ctxt end + +fun skolemize ctxt = skolemize_with_choice_thms ctxt (Meson_Choices.get ctxt) + +(* "RS" can fail if "unify_search_bound" is too small. *) +fun try_skolemize ctxt th = + try (skolemize ctxt) th + |> tap (fn NONE => trace_msg (fn () => "Failed to skolemize " ^ + Display.string_of_thm ctxt th) + | _ => ()) + +fun add_clauses th cls = + let val ctxt0 = Variable.global_thm_context th + val (cnfs, ctxt) = make_cnf [] th ctxt0 + in Variable.export ctxt ctxt0 cnfs @ cls end; + +(*Make clauses from a list of theorems, previously Skolemized and put into nnf. + The resulting clauses are HOL disjunctions.*) +fun make_clauses_unsorted ths = fold_rev add_clauses ths []; +val make_clauses = sort_clauses o make_clauses_unsorted; + +(*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*) +fun make_horns ths = + name_thms "Horn#" + (distinct Thm.eq_thm_prop (fold_rev (add_contras clause_rules) ths [])); + +(*Could simply use nprems_of, which would count remaining subgoals -- no + discrimination as to their size! With BEST_FIRST, fails for problem 41.*) + +fun best_prolog_tac sizef horns = + BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1); + +fun depth_prolog_tac horns = + DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1); + +(*Return all negative clauses, as possible goal clauses*) +fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls)); + +fun skolemize_prems_tac ctxt prems = + cut_facts_tac (map_filter (try_skolemize ctxt) prems) THEN' REPEAT o etac exE + +(*Basis of all meson-tactics. Supplies cltac with clauses: HOL disjunctions. + Function mkcl converts theorems to clauses.*) +fun MESON preskolem_tac mkcl cltac ctxt i st = + SELECT_GOAL + (EVERY [Object_Logic.atomize_prems_tac 1, + rtac ccontr 1, + preskolem_tac, + Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} => + EVERY1 [skolemize_prems_tac ctxt negs, + Subgoal.FOCUS (cltac o mkcl o #prems) ctxt']) ctxt 1]) i st + handle THM _ => no_tac st; (*probably from make_meta_clause, not first-order*) + + +(** Best-first search versions **) + +(*ths is a list of additional clauses (HOL disjunctions) to use.*) +fun best_meson_tac sizef = + MESON all_tac make_clauses + (fn cls => + THEN_BEST_FIRST (resolve_tac (gocls cls) 1) + (has_fewer_prems 1, sizef) + (prolog_step_tac (make_horns cls) 1)); + +(*First, breaks the goal into independent units*) +fun safe_best_meson_tac ctxt = + SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN + TRYALL (best_meson_tac size_of_subgoals ctxt)); + +(** Depth-first search version **) + +val depth_meson_tac = + MESON all_tac make_clauses + (fn cls => EVERY [resolve_tac (gocls cls) 1, depth_prolog_tac (make_horns cls)]); + + +(** Iterative deepening version **) + +(*This version does only one inference per call; + having only one eq_assume_tac speeds it up!*) +fun prolog_step_tac' horns = + let val (horn0s, _) = (*0 subgoals vs 1 or more*) + take_prefix Thm.no_prems horns + val nrtac = net_resolve_tac horns + in fn i => eq_assume_tac i ORELSE + match_tac horn0s i ORELSE (*no backtracking if unit MATCHES*) + ((assume_tac i APPEND nrtac i) THEN check_tac) + end; + +fun iter_deepen_prolog_tac horns = + ITER_DEEPEN iter_deepen_limit (has_fewer_prems 1) (prolog_step_tac' horns); + +fun iter_deepen_meson_tac ctxt ths = ctxt |> MESON all_tac make_clauses + (fn cls => + (case (gocls (cls @ ths)) of + [] => no_tac (*no goal clauses*) + | goes => + let + val horns = make_horns (cls @ ths) + val _ = trace_msg (fn () => + cat_lines ("meson method called:" :: + map (Display.string_of_thm ctxt) (cls @ ths) @ + ["clauses:"] @ map (Display.string_of_thm ctxt) horns)) + in + THEN_ITER_DEEPEN iter_deepen_limit + (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns) + end)); + +fun meson_tac ctxt ths = + SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN TRYALL (iter_deepen_meson_tac ctxt ths)); + + +(**** Code to support ordinary resolution, rather than Model Elimination ****) + +(*Convert a list of clauses (disjunctions) to meta-level clauses (==>), + with no contrapositives, for ordinary resolution.*) + +(*Rules to convert the head literal into a negated assumption. If the head + literal is already negated, then using notEfalse instead of notEfalse' + prevents a double negation.*) +val notEfalse = read_instantiate @{context} [(("R", 0), "False")] notE; +val notEfalse' = rotate_prems 1 notEfalse; + +fun negated_asm_of_head th = + th RS notEfalse handle THM _ => th RS notEfalse'; + +(*Converting one theorem from a disjunction to a meta-level clause*) +fun make_meta_clause th = + let val (fth,thaw) = Drule.legacy_freeze_thaw_robust th + in + (zero_var_indexes o Thm.varifyT_global o thaw 0 o + negated_asm_of_head o make_horn resolution_clause_rules) fth + end; + +fun make_meta_clauses ths = + name_thms "MClause#" + (distinct Thm.eq_thm_prop (map make_meta_clause ths)); + +end; diff -r 6e9aff5ee9b5 -r 1f01c9b2b76b src/HOL/Tools/Meson/meson_clausify.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Meson/meson_clausify.ML Mon Oct 04 21:37:42 2010 +0200 @@ -0,0 +1,376 @@ +(* Title: HOL/Tools/Sledgehammer/meson_clausify.ML + Author: Jia Meng, Cambridge University Computer Laboratory and NICTA + Author: Jasmin Blanchette, TU Muenchen + +Transformation of axiom rules (elim/intro/etc) into CNF forms. +*) + +signature MESON_CLAUSIFY = +sig + val new_skolem_var_prefix : string + val extensionalize_theorem : thm -> thm + val introduce_combinators_in_cterm : cterm -> thm + val introduce_combinators_in_theorem : thm -> thm + val to_definitional_cnf_with_quantifiers : theory -> thm -> thm + val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool + val cnf_axiom : + Proof.context -> bool -> int -> thm -> (thm * term) option * thm list + val meson_general_tac : Proof.context -> thm list -> int -> tactic + val setup: theory -> theory +end; + +structure Meson_Clausify : MESON_CLAUSIFY = +struct + +(* the extra "?" helps prevent clashes *) +val new_skolem_var_prefix = "?SK" +val new_nonskolem_var_prefix = "?V" + +(**** Transformation of Elimination Rules into First-Order Formulas****) + +val cfalse = cterm_of @{theory HOL} HOLogic.false_const; +val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop HOLogic.false_const); + +(* Converts an elim-rule into an equivalent theorem that does not have the + predicate variable. Leaves other theorems unchanged. We simply instantiate + the conclusion variable to False. (Cf. "transform_elim_term" in + "Sledgehammer_Util".) *) +fun transform_elim_theorem th = + case concl_of th of (*conclusion variable*) + @{const Trueprop} $ (v as Var (_, @{typ bool})) => + Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th + | v as Var(_, @{typ prop}) => + Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th + | _ => th + + +(**** SKOLEMIZATION BY INFERENCE (lcp) ****) + +fun mk_old_skolem_term_wrapper t = + let val T = fastype_of t in + Const (@{const_name skolem}, T --> T) $ t + end + +fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') + | beta_eta_in_abs_body t = Envir.beta_eta_contract t + +(*Traverse a theorem, accumulating Skolem function definitions.*) +fun old_skolem_defs th = + let + fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (_, T, p))) rhss = + (*Existential: declare a Skolem function, then insert into body and continue*) + let + val args = OldTerm.term_frees body + (* Forms a lambda-abstraction over the formal parameters *) + val rhs = + list_abs_free (map dest_Free args, + HOLogic.choice_const T $ beta_eta_in_abs_body body) + |> mk_old_skolem_term_wrapper + val comb = list_comb (rhs, args) + in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end + | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) rhss = + (*Universal quant: insert a free variable into body and continue*) + let val fname = Name.variant (OldTerm.add_term_names (p,[])) a + in dec_sko (subst_bound (Free(fname,T), p)) rhss end + | dec_sko (@{const conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q + | dec_sko (@{const disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q + | dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss + | dec_sko _ rhss = rhss + in dec_sko (prop_of th) [] end; + + +(**** REPLACING ABSTRACTIONS BY COMBINATORS ****) + +val fun_cong_all = @{thm fun_eq_iff [THEN iffD1]} + +(* Removes the lambdas from an equation of the form "t = (%x. u)". + (Cf. "extensionalize_term" in "Sledgehammer_Translate".) *) +fun extensionalize_theorem th = + case prop_of th of + _ $ (Const (@{const_name HOL.eq}, Type (_, [Type (@{type_name fun}, _), _])) + $ _ $ Abs _) => extensionalize_theorem (th RS fun_cong_all) + | _ => th + +fun is_quasi_lambda_free (Const (@{const_name skolem}, _) $ _) = true + | is_quasi_lambda_free (t1 $ t2) = + is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 + | is_quasi_lambda_free (Abs _) = false + | is_quasi_lambda_free _ = true + +val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B})); +val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C})); +val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S})); + +(* FIXME: Requires more use of cterm constructors. *) +fun abstract ct = + let + val thy = theory_of_cterm ct + val Abs(x,_,body) = term_of ct + val Type(@{type_name fun}, [xT,bodyT]) = typ_of (ctyp_of_term ct) + val cxT = ctyp_of thy xT + val cbodyT = ctyp_of thy bodyT + fun makeK () = + instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] + @{thm abs_K} + in + case body of + Const _ => makeK() + | Free _ => makeK() + | Var _ => makeK() (*though Var isn't expected*) + | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*) + | rator$rand => + if loose_bvar1 (rator,0) then (*C or S*) + if loose_bvar1 (rand,0) then (*S*) + let val crator = cterm_of thy (Abs(x,xT,rator)) + val crand = cterm_of thy (Abs(x,xT,rand)) + val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S} + val (_,rhs) = Thm.dest_equals (cprop_of abs_S') + in + Thm.transitive abs_S' (Conv.binop_conv abstract rhs) + end + else (*C*) + let val crator = cterm_of thy (Abs(x,xT,rator)) + val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C} + val (_,rhs) = Thm.dest_equals (cprop_of abs_C') + in + Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs) + end + else if loose_bvar1 (rand,0) then (*B or eta*) + if rand = Bound 0 then Thm.eta_conversion ct + else (*B*) + let val crand = cterm_of thy (Abs(x,xT,rand)) + val crator = cterm_of thy rator + val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B} + val (_,rhs) = Thm.dest_equals (cprop_of abs_B') + in Thm.transitive abs_B' (Conv.arg_conv abstract rhs) end + else makeK() + | _ => raise Fail "abstract: Bad term" + end; + +(* Traverse a theorem, remplacing lambda-abstractions with combinators. *) +fun introduce_combinators_in_cterm ct = + if is_quasi_lambda_free (term_of ct) then + Thm.reflexive ct + else case term_of ct of + Abs _ => + let + val (cv, cta) = Thm.dest_abs NONE ct + val (v, _) = dest_Free (term_of cv) + val u_th = introduce_combinators_in_cterm cta + val cu = Thm.rhs_of u_th + val comb_eq = abstract (Thm.cabs cv cu) + in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end + | _ $ _ => + let val (ct1, ct2) = Thm.dest_comb ct in + Thm.combination (introduce_combinators_in_cterm ct1) + (introduce_combinators_in_cterm ct2) + end + +fun introduce_combinators_in_theorem th = + if is_quasi_lambda_free (prop_of th) then + th + else + let + val th = Drule.eta_contraction_rule th + val eqth = introduce_combinators_in_cterm (cprop_of th) + in Thm.equal_elim eqth th end + handle THM (msg, _, _) => + (warning ("Error in the combinator translation of " ^ + Display.string_of_thm_without_context th ^ + "\nException message: " ^ msg ^ "."); + (* A type variable of sort "{}" will make abstraction fail. *) + TrueI) + +(*cterms are used throughout for efficiency*) +val cTrueprop = cterm_of @{theory HOL} HOLogic.Trueprop; + +(*Given an abstraction over n variables, replace the bound variables by free + ones. Return the body, along with the list of free variables.*) +fun c_variant_abs_multi (ct0, vars) = + let val (cv,ct) = Thm.dest_abs NONE ct0 + in c_variant_abs_multi (ct, cv::vars) end + handle CTERM _ => (ct0, rev vars); + +val skolem_def_raw = @{thms skolem_def_raw} + +(* Given the definition of a Skolem function, return a theorem to replace + an existential formula by a use of that function. + Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) +fun old_skolem_theorem_from_def thy rhs0 = + let + val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> cterm_of thy + val rhs' = rhs |> Thm.dest_comb |> snd + val (ch, frees) = c_variant_abs_multi (rhs', []) + val (hilbert, cabs) = ch |> Thm.dest_comb |>> term_of + val T = + case hilbert of + Const (@{const_name Eps}, Type (@{type_name fun}, [_, T])) => T + | _ => raise TERM ("old_skolem_theorem_from_def: expected \"Eps\"", + [hilbert]) + val cex = cterm_of thy (HOLogic.exists_const T) + val ex_tm = Thm.capply cTrueprop (Thm.capply cex cabs) + val conc = + Drule.list_comb (rhs, frees) + |> Drule.beta_conv cabs |> Thm.capply cTrueprop + fun tacf [prem] = + rewrite_goals_tac skolem_def_raw + THEN rtac ((prem |> rewrite_rule skolem_def_raw) RS @{thm someI_ex}) 1 + in + Goal.prove_internal [ex_tm] conc tacf + |> forall_intr_list frees + |> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) + |> Thm.varifyT_global + end + +fun to_definitional_cnf_with_quantifiers thy th = + let + val eqth = cnf.make_cnfx_thm thy (HOLogic.dest_Trueprop (prop_of th)) + val eqth = eqth RS @{thm eq_reflection} + val eqth = eqth RS @{thm TruepropI} + in Thm.equal_elim eqth th end + +fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s = + (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^ + "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ + string_of_int index_no ^ "_" ^ s + +fun cluster_of_zapped_var_name s = + let val get_int = the o Int.fromString o nth (space_explode "_" s) in + ((get_int 1, (get_int 2, get_int 3)), + String.isPrefix new_skolem_var_prefix s) + end + +fun zap (cluster as (cluster_no, cluster_skolem)) index_no pos ct = + ct + |> (case term_of ct of + Const (s, _) $ Abs (s', _, _) => + if s = @{const_name all} orelse s = @{const_name All} orelse + s = @{const_name Ex} then + let + val skolem = (pos = (s = @{const_name Ex})) + val (cluster, index_no) = + if skolem = cluster_skolem then (cluster, index_no) + else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0) + in + Thm.dest_comb #> snd + #> Thm.dest_abs (SOME (zapped_var_name cluster index_no s')) + #> snd #> zap cluster (index_no + 1) pos + end + else + Conv.all_conv + | Const (s, _) $ _ $ _ => + if s = @{const_name "==>"} orelse s = @{const_name implies} then + Conv.combination_conv (Conv.arg_conv (zap cluster index_no (not pos))) + (zap cluster index_no pos) + else if s = @{const_name conj} orelse s = @{const_name disj} then + Conv.combination_conv (Conv.arg_conv (zap cluster index_no pos)) + (zap cluster index_no pos) + else + Conv.all_conv + | Const (s, _) $ _ => + if s = @{const_name Trueprop} then + Conv.arg_conv (zap cluster index_no pos) + else if s = @{const_name Not} then + Conv.arg_conv (zap cluster index_no (not pos)) + else + Conv.all_conv + | _ => Conv.all_conv) + +fun ss_only ths = MetaSimplifier.clear_ss HOL_basic_ss addsimps ths + +val no_choice = + @{prop "ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"} + |> Logic.varify_global + |> Skip_Proof.make_thm @{theory} + +(* Converts an Isabelle theorem into NNF. *) +fun nnf_axiom choice_ths new_skolemizer ax_no th ctxt = + let + val thy = ProofContext.theory_of ctxt + val th = + th |> transform_elim_theorem + |> zero_var_indexes + |> new_skolemizer ? forall_intr_vars + val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single + val th = th |> Conv.fconv_rule Object_Logic.atomize + |> extensionalize_theorem + |> Meson.make_nnf ctxt + in + if new_skolemizer then + let + fun skolemize choice_ths = + Meson.skolemize_with_choice_thms ctxt choice_ths + #> simplify (ss_only @{thms all_simps[symmetric]}) + val pull_out = + simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]}) + val (discharger_th, fully_skolemized_th) = + if null choice_ths then + th |> `I |>> pull_out ||> skolemize [no_choice] + else + th |> skolemize choice_ths |> `I + val t = + fully_skolemized_th |> cprop_of + |> zap ((ax_no, 0), true) 0 true |> Drule.export_without_context + |> cprop_of |> Thm.dest_equals |> snd |> term_of + in + if exists_subterm (fn Var ((s, _), _) => + String.isPrefix new_skolem_var_prefix s + | _ => false) t then + let + val (ct, ctxt) = + Variable.import_terms true [t] ctxt + |>> the_single |>> cterm_of thy + in (SOME (discharger_th, ct), Thm.assume ct, ctxt) end + else + (NONE, th, ctxt) + end + else + (NONE, th, ctxt) + end + +(* Convert a theorem to CNF, with additional premises due to skolemization. *) +fun cnf_axiom ctxt0 new_skolemizer ax_no th = + let + val thy = ProofContext.theory_of ctxt0 + val choice_ths = Meson_Choices.get ctxt0 + val (opt, nnf_th, ctxt) = nnf_axiom choice_ths new_skolemizer ax_no th ctxt0 + fun clausify th = + Meson.make_cnf (if new_skolemizer then + [] + else + map (old_skolem_theorem_from_def thy) + (old_skolem_defs th)) th ctxt + val (cnf_ths, ctxt) = + clausify nnf_th + |> (fn ([], _) => + clausify (to_definitional_cnf_with_quantifiers thy nnf_th) + | p => p) + fun intr_imp ct th = + Thm.instantiate ([], map (pairself (cterm_of @{theory})) + [(Var (("i", 1), @{typ nat}), + HOLogic.mk_nat ax_no)]) + @{thm skolem_COMBK_D} + RS Thm.implies_intr ct th + in + (opt |> Option.map (I #>> singleton (Variable.export ctxt ctxt0) + ##> (term_of #> HOLogic.dest_Trueprop + #> singleton (Variable.export_terms ctxt ctxt0))), + cnf_ths |> map (introduce_combinators_in_theorem + #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) + |> Variable.export ctxt ctxt0 + |> Meson.finish_cnf + |> map Thm.close_derivation) + end + handle THM _ => (NONE, []) + +fun meson_general_tac ctxt ths = + let val ctxt = Classical.put_claset HOL_cs ctxt in + Meson.meson_tac ctxt (maps (snd o cnf_axiom ctxt false 0) ths) + end + +val setup = + Method.setup @{binding meson} (Attrib.thms >> (fn ths => fn ctxt => + SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ctxt ths))) + "MESON resolution proof procedure" + +end; diff -r 6e9aff5ee9b5 -r 1f01c9b2b76b src/HOL/Tools/Sledgehammer/meson_clausify.ML --- a/src/HOL/Tools/Sledgehammer/meson_clausify.ML Mon Oct 04 20:55:55 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,376 +0,0 @@ -(* Title: HOL/Tools/Sledgehammer/meson_clausify.ML - Author: Jia Meng, Cambridge University Computer Laboratory and NICTA - Author: Jasmin Blanchette, TU Muenchen - -Transformation of axiom rules (elim/intro/etc) into CNF forms. -*) - -signature MESON_CLAUSIFY = -sig - val new_skolem_var_prefix : string - val extensionalize_theorem : thm -> thm - val introduce_combinators_in_cterm : cterm -> thm - val introduce_combinators_in_theorem : thm -> thm - val to_definitional_cnf_with_quantifiers : theory -> thm -> thm - val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool - val cnf_axiom : - Proof.context -> bool -> int -> thm -> (thm * term) option * thm list - val meson_general_tac : Proof.context -> thm list -> int -> tactic - val setup: theory -> theory -end; - -structure Meson_Clausify : MESON_CLAUSIFY = -struct - -(* the extra "?" helps prevent clashes *) -val new_skolem_var_prefix = "?SK" -val new_nonskolem_var_prefix = "?V" - -(**** Transformation of Elimination Rules into First-Order Formulas****) - -val cfalse = cterm_of @{theory HOL} HOLogic.false_const; -val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop HOLogic.false_const); - -(* Converts an elim-rule into an equivalent theorem that does not have the - predicate variable. Leaves other theorems unchanged. We simply instantiate - the conclusion variable to False. (Cf. "transform_elim_term" in - "Sledgehammer_Util".) *) -fun transform_elim_theorem th = - case concl_of th of (*conclusion variable*) - @{const Trueprop} $ (v as Var (_, @{typ bool})) => - Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th - | v as Var(_, @{typ prop}) => - Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th - | _ => th - - -(**** SKOLEMIZATION BY INFERENCE (lcp) ****) - -fun mk_old_skolem_term_wrapper t = - let val T = fastype_of t in - Const (@{const_name skolem}, T --> T) $ t - end - -fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') - | beta_eta_in_abs_body t = Envir.beta_eta_contract t - -(*Traverse a theorem, accumulating Skolem function definitions.*) -fun old_skolem_defs th = - let - fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (_, T, p))) rhss = - (*Existential: declare a Skolem function, then insert into body and continue*) - let - val args = OldTerm.term_frees body - (* Forms a lambda-abstraction over the formal parameters *) - val rhs = - list_abs_free (map dest_Free args, - HOLogic.choice_const T $ beta_eta_in_abs_body body) - |> mk_old_skolem_term_wrapper - val comb = list_comb (rhs, args) - in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end - | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) rhss = - (*Universal quant: insert a free variable into body and continue*) - let val fname = Name.variant (OldTerm.add_term_names (p,[])) a - in dec_sko (subst_bound (Free(fname,T), p)) rhss end - | dec_sko (@{const conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q - | dec_sko (@{const disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q - | dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss - | dec_sko _ rhss = rhss - in dec_sko (prop_of th) [] end; - - -(**** REPLACING ABSTRACTIONS BY COMBINATORS ****) - -val fun_cong_all = @{thm fun_eq_iff [THEN iffD1]} - -(* Removes the lambdas from an equation of the form "t = (%x. u)". - (Cf. "extensionalize_term" in "Sledgehammer_Translate".) *) -fun extensionalize_theorem th = - case prop_of th of - _ $ (Const (@{const_name HOL.eq}, Type (_, [Type (@{type_name fun}, _), _])) - $ _ $ Abs _) => extensionalize_theorem (th RS fun_cong_all) - | _ => th - -fun is_quasi_lambda_free (Const (@{const_name skolem}, _) $ _) = true - | is_quasi_lambda_free (t1 $ t2) = - is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 - | is_quasi_lambda_free (Abs _) = false - | is_quasi_lambda_free _ = true - -val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B})); -val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C})); -val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S})); - -(* FIXME: Requires more use of cterm constructors. *) -fun abstract ct = - let - val thy = theory_of_cterm ct - val Abs(x,_,body) = term_of ct - val Type(@{type_name fun}, [xT,bodyT]) = typ_of (ctyp_of_term ct) - val cxT = ctyp_of thy xT - val cbodyT = ctyp_of thy bodyT - fun makeK () = - instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] - @{thm abs_K} - in - case body of - Const _ => makeK() - | Free _ => makeK() - | Var _ => makeK() (*though Var isn't expected*) - | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*) - | rator$rand => - if loose_bvar1 (rator,0) then (*C or S*) - if loose_bvar1 (rand,0) then (*S*) - let val crator = cterm_of thy (Abs(x,xT,rator)) - val crand = cterm_of thy (Abs(x,xT,rand)) - val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S} - val (_,rhs) = Thm.dest_equals (cprop_of abs_S') - in - Thm.transitive abs_S' (Conv.binop_conv abstract rhs) - end - else (*C*) - let val crator = cterm_of thy (Abs(x,xT,rator)) - val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C} - val (_,rhs) = Thm.dest_equals (cprop_of abs_C') - in - Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs) - end - else if loose_bvar1 (rand,0) then (*B or eta*) - if rand = Bound 0 then Thm.eta_conversion ct - else (*B*) - let val crand = cterm_of thy (Abs(x,xT,rand)) - val crator = cterm_of thy rator - val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B} - val (_,rhs) = Thm.dest_equals (cprop_of abs_B') - in Thm.transitive abs_B' (Conv.arg_conv abstract rhs) end - else makeK() - | _ => raise Fail "abstract: Bad term" - end; - -(* Traverse a theorem, remplacing lambda-abstractions with combinators. *) -fun introduce_combinators_in_cterm ct = - if is_quasi_lambda_free (term_of ct) then - Thm.reflexive ct - else case term_of ct of - Abs _ => - let - val (cv, cta) = Thm.dest_abs NONE ct - val (v, _) = dest_Free (term_of cv) - val u_th = introduce_combinators_in_cterm cta - val cu = Thm.rhs_of u_th - val comb_eq = abstract (Thm.cabs cv cu) - in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end - | _ $ _ => - let val (ct1, ct2) = Thm.dest_comb ct in - Thm.combination (introduce_combinators_in_cterm ct1) - (introduce_combinators_in_cterm ct2) - end - -fun introduce_combinators_in_theorem th = - if is_quasi_lambda_free (prop_of th) then - th - else - let - val th = Drule.eta_contraction_rule th - val eqth = introduce_combinators_in_cterm (cprop_of th) - in Thm.equal_elim eqth th end - handle THM (msg, _, _) => - (warning ("Error in the combinator translation of " ^ - Display.string_of_thm_without_context th ^ - "\nException message: " ^ msg ^ "."); - (* A type variable of sort "{}" will make abstraction fail. *) - TrueI) - -(*cterms are used throughout for efficiency*) -val cTrueprop = cterm_of @{theory HOL} HOLogic.Trueprop; - -(*Given an abstraction over n variables, replace the bound variables by free - ones. Return the body, along with the list of free variables.*) -fun c_variant_abs_multi (ct0, vars) = - let val (cv,ct) = Thm.dest_abs NONE ct0 - in c_variant_abs_multi (ct, cv::vars) end - handle CTERM _ => (ct0, rev vars); - -val skolem_def_raw = @{thms skolem_def_raw} - -(* Given the definition of a Skolem function, return a theorem to replace - an existential formula by a use of that function. - Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) -fun old_skolem_theorem_from_def thy rhs0 = - let - val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> cterm_of thy - val rhs' = rhs |> Thm.dest_comb |> snd - val (ch, frees) = c_variant_abs_multi (rhs', []) - val (hilbert, cabs) = ch |> Thm.dest_comb |>> term_of - val T = - case hilbert of - Const (@{const_name Eps}, Type (@{type_name fun}, [_, T])) => T - | _ => raise TERM ("old_skolem_theorem_from_def: expected \"Eps\"", - [hilbert]) - val cex = cterm_of thy (HOLogic.exists_const T) - val ex_tm = Thm.capply cTrueprop (Thm.capply cex cabs) - val conc = - Drule.list_comb (rhs, frees) - |> Drule.beta_conv cabs |> Thm.capply cTrueprop - fun tacf [prem] = - rewrite_goals_tac skolem_def_raw - THEN rtac ((prem |> rewrite_rule skolem_def_raw) RS @{thm someI_ex}) 1 - in - Goal.prove_internal [ex_tm] conc tacf - |> forall_intr_list frees - |> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) - |> Thm.varifyT_global - end - -fun to_definitional_cnf_with_quantifiers thy th = - let - val eqth = cnf.make_cnfx_thm thy (HOLogic.dest_Trueprop (prop_of th)) - val eqth = eqth RS @{thm eq_reflection} - val eqth = eqth RS @{thm TruepropI} - in Thm.equal_elim eqth th end - -fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s = - (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^ - "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ - string_of_int index_no ^ "_" ^ s - -fun cluster_of_zapped_var_name s = - let val get_int = the o Int.fromString o nth (space_explode "_" s) in - ((get_int 1, (get_int 2, get_int 3)), - String.isPrefix new_skolem_var_prefix s) - end - -fun zap (cluster as (cluster_no, cluster_skolem)) index_no pos ct = - ct - |> (case term_of ct of - Const (s, _) $ Abs (s', _, _) => - if s = @{const_name all} orelse s = @{const_name All} orelse - s = @{const_name Ex} then - let - val skolem = (pos = (s = @{const_name Ex})) - val (cluster, index_no) = - if skolem = cluster_skolem then (cluster, index_no) - else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0) - in - Thm.dest_comb #> snd - #> Thm.dest_abs (SOME (zapped_var_name cluster index_no s')) - #> snd #> zap cluster (index_no + 1) pos - end - else - Conv.all_conv - | Const (s, _) $ _ $ _ => - if s = @{const_name "==>"} orelse s = @{const_name implies} then - Conv.combination_conv (Conv.arg_conv (zap cluster index_no (not pos))) - (zap cluster index_no pos) - else if s = @{const_name conj} orelse s = @{const_name disj} then - Conv.combination_conv (Conv.arg_conv (zap cluster index_no pos)) - (zap cluster index_no pos) - else - Conv.all_conv - | Const (s, _) $ _ => - if s = @{const_name Trueprop} then - Conv.arg_conv (zap cluster index_no pos) - else if s = @{const_name Not} then - Conv.arg_conv (zap cluster index_no (not pos)) - else - Conv.all_conv - | _ => Conv.all_conv) - -fun ss_only ths = MetaSimplifier.clear_ss HOL_basic_ss addsimps ths - -val no_choice = - @{prop "ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"} - |> Logic.varify_global - |> Skip_Proof.make_thm @{theory} - -(* Converts an Isabelle theorem into NNF. *) -fun nnf_axiom choice_ths new_skolemizer ax_no th ctxt = - let - val thy = ProofContext.theory_of ctxt - val th = - th |> transform_elim_theorem - |> zero_var_indexes - |> new_skolemizer ? forall_intr_vars - val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single - val th = th |> Conv.fconv_rule Object_Logic.atomize - |> extensionalize_theorem - |> Meson.make_nnf ctxt - in - if new_skolemizer then - let - fun skolemize choice_ths = - Meson.skolemize_with_choice_thms ctxt choice_ths - #> simplify (ss_only @{thms all_simps[symmetric]}) - val pull_out = - simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]}) - val (discharger_th, fully_skolemized_th) = - if null choice_ths then - th |> `I |>> pull_out ||> skolemize [no_choice] - else - th |> skolemize choice_ths |> `I - val t = - fully_skolemized_th |> cprop_of - |> zap ((ax_no, 0), true) 0 true |> Drule.export_without_context - |> cprop_of |> Thm.dest_equals |> snd |> term_of - in - if exists_subterm (fn Var ((s, _), _) => - String.isPrefix new_skolem_var_prefix s - | _ => false) t then - let - val (ct, ctxt) = - Variable.import_terms true [t] ctxt - |>> the_single |>> cterm_of thy - in (SOME (discharger_th, ct), Thm.assume ct, ctxt) end - else - (NONE, th, ctxt) - end - else - (NONE, th, ctxt) - end - -(* Convert a theorem to CNF, with additional premises due to skolemization. *) -fun cnf_axiom ctxt0 new_skolemizer ax_no th = - let - val thy = ProofContext.theory_of ctxt0 - val choice_ths = Meson_Choices.get ctxt0 - val (opt, nnf_th, ctxt) = nnf_axiom choice_ths new_skolemizer ax_no th ctxt0 - fun clausify th = - Meson.make_cnf (if new_skolemizer then - [] - else - map (old_skolem_theorem_from_def thy) - (old_skolem_defs th)) th ctxt - val (cnf_ths, ctxt) = - clausify nnf_th - |> (fn ([], _) => - clausify (to_definitional_cnf_with_quantifiers thy nnf_th) - | p => p) - fun intr_imp ct th = - Thm.instantiate ([], map (pairself (cterm_of @{theory})) - [(Var (("i", 1), @{typ nat}), - HOLogic.mk_nat ax_no)]) - @{thm skolem_COMBK_D} - RS Thm.implies_intr ct th - in - (opt |> Option.map (I #>> singleton (Variable.export ctxt ctxt0) - ##> (term_of #> HOLogic.dest_Trueprop - #> singleton (Variable.export_terms ctxt ctxt0))), - cnf_ths |> map (introduce_combinators_in_theorem - #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) - |> Variable.export ctxt ctxt0 - |> Meson.finish_cnf - |> map Thm.close_derivation) - end - handle THM _ => (NONE, []) - -fun meson_general_tac ctxt ths = - let val ctxt = Classical.put_claset HOL_cs ctxt in - Meson.meson_tac ctxt (maps (snd o cnf_axiom ctxt false 0) ths) - end - -val setup = - Method.setup @{binding meson} (Attrib.thms >> (fn ths => fn ctxt => - SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ctxt ths))) - "MESON resolution proof procedure" - -end; diff -r 6e9aff5ee9b5 -r 1f01c9b2b76b src/HOL/Tools/meson.ML --- a/src/HOL/Tools/meson.ML Mon Oct 04 20:55:55 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,712 +0,0 @@ -(* Title: HOL/Tools/meson.ML - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - -The MESON resolution proof procedure for HOL. -When making clauses, avoids using the rewriter -- instead uses RS recursively. -*) - -signature MESON = -sig - val trace: bool Unsynchronized.ref - val term_pair_of: indexname * (typ * 'a) -> term * 'a - val size_of_subgoals: thm -> int - val has_too_many_clauses: Proof.context -> term -> bool - val make_cnf: thm list -> thm -> Proof.context -> thm list * Proof.context - val finish_cnf: thm list -> thm list - val presimplify: thm -> thm - val make_nnf: Proof.context -> thm -> thm - val skolemize_with_choice_thms : Proof.context -> thm list -> thm -> thm - val skolemize : Proof.context -> thm -> thm - val is_fol_term: theory -> term -> bool - val make_clauses_unsorted: thm list -> thm list - val make_clauses: thm list -> thm list - val make_horns: thm list -> thm list - val best_prolog_tac: (thm -> int) -> thm list -> tactic - val depth_prolog_tac: thm list -> tactic - val gocls: thm list -> thm list - val skolemize_prems_tac : Proof.context -> thm list -> int -> tactic - val MESON: - tactic -> (thm list -> thm list) -> (thm list -> tactic) -> Proof.context - -> int -> tactic - val best_meson_tac: (thm -> int) -> Proof.context -> int -> tactic - val safe_best_meson_tac: Proof.context -> int -> tactic - val depth_meson_tac: Proof.context -> int -> tactic - val prolog_step_tac': thm list -> int -> tactic - val iter_deepen_prolog_tac: thm list -> tactic - val iter_deepen_meson_tac: Proof.context -> thm list -> int -> tactic - val make_meta_clause: thm -> thm - val make_meta_clauses: thm list -> thm list - val meson_tac: Proof.context -> thm list -> int -> tactic - val setup: theory -> theory -end - -structure Meson : MESON = -struct - -val trace = Unsynchronized.ref false; -fun trace_msg msg = if ! trace then tracing (msg ()) else (); - -val max_clauses_default = 60; -val (max_clauses, setup) = Attrib.config_int "meson_max_clauses" (K max_clauses_default); - -(*No known example (on 1-5-2007) needs even thirty*) -val iter_deepen_limit = 50; - -val disj_forward = @{thm disj_forward}; -val disj_forward2 = @{thm disj_forward2}; -val make_pos_rule = @{thm make_pos_rule}; -val make_pos_rule' = @{thm make_pos_rule'}; -val make_pos_goal = @{thm make_pos_goal}; -val make_neg_rule = @{thm make_neg_rule}; -val make_neg_rule' = @{thm make_neg_rule'}; -val make_neg_goal = @{thm make_neg_goal}; -val conj_forward = @{thm conj_forward}; -val all_forward = @{thm all_forward}; -val ex_forward = @{thm ex_forward}; - -val not_conjD = @{thm meson_not_conjD}; -val not_disjD = @{thm meson_not_disjD}; -val not_notD = @{thm meson_not_notD}; -val not_allD = @{thm meson_not_allD}; -val not_exD = @{thm meson_not_exD}; -val imp_to_disjD = @{thm meson_imp_to_disjD}; -val not_impD = @{thm meson_not_impD}; -val iff_to_disjD = @{thm meson_iff_to_disjD}; -val not_iffD = @{thm meson_not_iffD}; -val conj_exD1 = @{thm meson_conj_exD1}; -val conj_exD2 = @{thm meson_conj_exD2}; -val disj_exD = @{thm meson_disj_exD}; -val disj_exD1 = @{thm meson_disj_exD1}; -val disj_exD2 = @{thm meson_disj_exD2}; -val disj_assoc = @{thm meson_disj_assoc}; -val disj_comm = @{thm meson_disj_comm}; -val disj_FalseD1 = @{thm meson_disj_FalseD1}; -val disj_FalseD2 = @{thm meson_disj_FalseD2}; - - -(**** Operators for forward proof ****) - - -(** First-order Resolution **) - -fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t); - -(*FIXME: currently does not "rename variables apart"*) -fun first_order_resolve thA thB = - (case - try (fn () => - let val thy = theory_of_thm thA - val tmA = concl_of thA - val Const("==>",_) $ tmB $ _ = prop_of thB - val tenv = - Pattern.first_order_match thy (tmB, tmA) - (Vartab.empty, Vartab.empty) |> snd - val ct_pairs = map (pairself (cterm_of thy) o term_pair_of) (Vartab.dest tenv) - in thA RS (cterm_instantiate ct_pairs thB) end) () of - SOME th => th - | NONE => raise THM ("first_order_resolve", 0, [thA, thB])) - -(* Applying "choice" swaps the bound variable names. We tweak - "Thm.rename_boundvars"'s input to get the desired names. *) -fun fix_bounds (_ $ (Const (@{const_name Ex}, _) - $ Abs (_, _, Const (@{const_name All}, _) $ _))) - (t0 $ (Const (@{const_name All}, T1) - $ Abs (a1, T1', Const (@{const_name Ex}, T2) - $ Abs (a2, T2', t')))) = - t0 $ (Const (@{const_name All}, T1) - $ Abs (a2, T1', Const (@{const_name Ex}, T2) $ Abs (a1, T2', t'))) - | fix_bounds _ t = t - -(* Hack to make it less likely that we lose our precious bound variable names in - "rename_bvs_RS" below, because of a clash. *) -val protect_prefix = "_" - -fun protect_bounds (t $ u) = protect_bounds t $ protect_bounds u - | protect_bounds (Abs (s, T, t')) = - Abs (protect_prefix ^ s, T, protect_bounds t') - | protect_bounds t = t - -(* Forward proof while preserving bound variables names*) -fun rename_bvs_RS th rl = - let - val t = concl_of th - val r = concl_of rl - val th' = th RS Thm.rename_boundvars r (protect_bounds r) rl - val t' = concl_of th' - in Thm.rename_boundvars t' (fix_bounds t' t) th' end - -(*raises exception if no rules apply*) -fun tryres (th, rls) = - let fun tryall [] = raise THM("tryres", 0, th::rls) - | tryall (rl::rls) = (rename_bvs_RS th rl handle THM _ => tryall rls) - in tryall rls end; - -(*Permits forward proof from rules that discharge assumptions. The supplied proof state st, - e.g. from conj_forward, should have the form - "[| P' ==> ?P; Q' ==> ?Q |] ==> ?P & ?Q" - and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*) -fun forward_res ctxt nf st = - let fun forward_tacf [prem] = rtac (nf prem) 1 - | forward_tacf prems = - error (cat_lines - ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" :: - Display.string_of_thm ctxt st :: - "Premises:" :: map (Display.string_of_thm ctxt) prems)) - in - case Seq.pull (ALLGOALS (Misc_Legacy.METAHYPS forward_tacf) st) - of SOME(th,_) => th - | NONE => raise THM("forward_res", 0, [st]) - end; - -(*Are any of the logical connectives in "bs" present in the term?*) -fun has_conns bs = - let fun has (Const _) = false - | has (Const(@{const_name Trueprop},_) $ p) = has p - | has (Const(@{const_name Not},_) $ p) = has p - | has (Const(@{const_name HOL.disj},_) $ p $ q) = member (op =) bs @{const_name HOL.disj} orelse has p orelse has q - | has (Const(@{const_name HOL.conj},_) $ p $ q) = member (op =) bs @{const_name HOL.conj} orelse has p orelse has q - | has (Const(@{const_name All},_) $ Abs(_,_,p)) = member (op =) bs @{const_name All} orelse has p - | has (Const(@{const_name Ex},_) $ Abs(_,_,p)) = member (op =) bs @{const_name Ex} orelse has p - | has _ = false - in has end; - - -(**** Clause handling ****) - -fun literals (Const(@{const_name Trueprop},_) $ P) = literals P - | literals (Const(@{const_name HOL.disj},_) $ P $ Q) = literals P @ literals Q - | literals (Const(@{const_name Not},_) $ P) = [(false,P)] - | literals P = [(true,P)]; - -(*number of literals in a term*) -val nliterals = length o literals; - - -(*** Tautology Checking ***) - -fun signed_lits_aux (Const (@{const_name HOL.disj}, _) $ P $ Q) (poslits, neglits) = - signed_lits_aux Q (signed_lits_aux P (poslits, neglits)) - | signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits) - | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits); - -fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]); - -(*Literals like X=X are tautologous*) -fun taut_poslit (Const(@{const_name HOL.eq},_) $ t $ u) = t aconv u - | taut_poslit (Const(@{const_name True},_)) = true - | taut_poslit _ = false; - -fun is_taut th = - let val (poslits,neglits) = signed_lits th - in exists taut_poslit poslits - orelse - exists (member (op aconv) neglits) (HOLogic.false_const :: poslits) - end - handle TERM _ => false; (*probably dest_Trueprop on a weird theorem*) - - -(*** To remove trivial negated equality literals from clauses ***) - -(*They are typically functional reflexivity axioms and are the converses of - injectivity equivalences*) - -val not_refl_disj_D = @{thm meson_not_refl_disj_D}; - -(*Is either term a Var that does not properly occur in the other term?*) -fun eliminable (t as Var _, u) = t aconv u orelse not (Logic.occs(t,u)) - | eliminable (u, t as Var _) = t aconv u orelse not (Logic.occs(t,u)) - | eliminable _ = false; - -fun refl_clause_aux 0 th = th - | refl_clause_aux n th = - case HOLogic.dest_Trueprop (concl_of th) of - (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _) => - refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*) - | (Const (@{const_name HOL.disj}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ t $ u)) $ _) => - if eliminable(t,u) - then refl_clause_aux (n-1) (th RS not_refl_disj_D) (*Var inequation: delete*) - else refl_clause_aux (n-1) (th RS disj_comm) (*not between Vars: ignore*) - | (Const (@{const_name HOL.disj}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm) - | _ => (*not a disjunction*) th; - -fun notequal_lits_count (Const (@{const_name HOL.disj}, _) $ P $ Q) = - notequal_lits_count P + notequal_lits_count Q - | notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _)) = 1 - | notequal_lits_count _ = 0; - -(*Simplify a clause by applying reflexivity to its negated equality literals*) -fun refl_clause th = - let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (concl_of th)) - in zero_var_indexes (refl_clause_aux neqs th) end - handle TERM _ => th; (*probably dest_Trueprop on a weird theorem*) - - -(*** Removal of duplicate literals ***) - -(*Forward proof, passing extra assumptions as theorems to the tactic*) -fun forward_res2 nf hyps st = - case Seq.pull - (REPEAT - (Misc_Legacy.METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1) - st) - of SOME(th,_) => th - | NONE => raise THM("forward_res2", 0, [st]); - -(*Remove duplicates in P|Q by assuming ~P in Q - rls (initially []) accumulates assumptions of the form P==>False*) -fun nodups_aux ctxt rls th = nodups_aux ctxt rls (th RS disj_assoc) - handle THM _ => tryres(th,rls) - handle THM _ => tryres(forward_res2 (nodups_aux ctxt) rls (th RS disj_forward2), - [disj_FalseD1, disj_FalseD2, asm_rl]) - handle THM _ => th; - -(*Remove duplicate literals, if there are any*) -fun nodups ctxt th = - if has_duplicates (op =) (literals (prop_of th)) - then nodups_aux ctxt [] th - else th; - - -(*** The basic CNF transformation ***) - -fun estimated_num_clauses bound t = - let - fun sum x y = if x < bound andalso y < bound then x+y else bound - fun prod x y = if x < bound andalso y < bound then x*y else bound - - (*Estimate the number of clauses in order to detect infeasible theorems*) - fun signed_nclauses b (Const(@{const_name Trueprop},_) $ t) = signed_nclauses b t - | signed_nclauses b (Const(@{const_name Not},_) $ t) = signed_nclauses (not b) t - | signed_nclauses b (Const(@{const_name HOL.conj},_) $ t $ u) = - if b then sum (signed_nclauses b t) (signed_nclauses b u) - else prod (signed_nclauses b t) (signed_nclauses b u) - | signed_nclauses b (Const(@{const_name HOL.disj},_) $ t $ u) = - if b then prod (signed_nclauses b t) (signed_nclauses b u) - else sum (signed_nclauses b t) (signed_nclauses b u) - | signed_nclauses b (Const(@{const_name HOL.implies},_) $ t $ u) = - if b then prod (signed_nclauses (not b) t) (signed_nclauses b u) - else sum (signed_nclauses (not b) t) (signed_nclauses b u) - | signed_nclauses b (Const(@{const_name HOL.eq}, Type ("fun", [T, _])) $ t $ u) = - if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*) - if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u)) - (prod (signed_nclauses (not b) u) (signed_nclauses b t)) - else sum (prod (signed_nclauses b t) (signed_nclauses b u)) - (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u)) - else 1 - | signed_nclauses b (Const(@{const_name Ex}, _) $ Abs (_,_,t)) = signed_nclauses b t - | signed_nclauses b (Const(@{const_name All},_) $ Abs (_,_,t)) = signed_nclauses b t - | signed_nclauses _ _ = 1; (* literal *) - in signed_nclauses true t end - -fun has_too_many_clauses ctxt t = - let val max_cl = Config.get ctxt max_clauses in - estimated_num_clauses (max_cl + 1) t > max_cl - end - -(*Replaces universally quantified variables by FREE variables -- because - assumptions may not contain scheme variables. Later, generalize using Variable.export. *) -local - val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec)))); - val spec_varT = #T (Thm.rep_cterm spec_var); - fun name_of (Const (@{const_name All}, _) $ Abs(x,_,_)) = x | name_of _ = Name.uu; -in - fun freeze_spec th ctxt = - let - val cert = Thm.cterm_of (ProofContext.theory_of ctxt); - val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (concl_of th))] ctxt; - val spec' = Thm.instantiate ([], [(spec_var, cert (Free (x, spec_varT)))]) spec; - in (th RS spec', ctxt') end -end; - -(*Used with METAHYPS below. There is one assumption, which gets bound to prem - and then normalized via function nf. The normal form is given to resolve_tac, - instantiate a Boolean variable created by resolution with disj_forward. Since - (nf prem) returns a LIST of theorems, we can backtrack to get all combinations.*) -fun resop nf [prem] = resolve_tac (nf prem) 1; - -(* Any need to extend this list with "HOL.type_class", "HOL.eq_class", - and "Pure.term"? *) -val has_meta_conn = exists_Const (member (op =) ["==", "==>", "=simp=>", "all", "prop"] o #1); - -fun apply_skolem_theorem (th, rls) = - let - fun tryall [] = raise THM ("apply_skolem_theorem", 0, th::rls) - | tryall (rl :: rls) = - first_order_resolve th rl handle THM _ => tryall rls - in tryall rls end - -(* Conjunctive normal form, adding clauses from th in front of ths (for foldr). - Strips universal quantifiers and breaks up conjunctions. - Eliminates existential quantifiers using Skolemization theorems. *) -fun cnf old_skolem_ths ctxt (th, ths) = - let val ctxtr = Unsynchronized.ref ctxt (* FIXME ??? *) - fun cnf_aux (th,ths) = - if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*) - else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (prop_of th)) - then nodups ctxt th :: ths (*no work to do, terminate*) - else case head_of (HOLogic.dest_Trueprop (concl_of th)) of - Const (@{const_name HOL.conj}, _) => (*conjunction*) - cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths)) - | Const (@{const_name All}, _) => (*universal quantifier*) - let val (th',ctxt') = freeze_spec th (!ctxtr) - in ctxtr := ctxt'; cnf_aux (th', ths) end - | Const (@{const_name Ex}, _) => - (*existential quantifier: Insert Skolem functions*) - cnf_aux (apply_skolem_theorem (th, old_skolem_ths), ths) - | Const (@{const_name HOL.disj}, _) => - (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using - all combinations of converting P, Q to CNF.*) - let val tac = - Misc_Legacy.METAHYPS (resop cnf_nil) 1 THEN - (fn st' => st' |> Misc_Legacy.METAHYPS (resop cnf_nil) 1) - in Seq.list_of (tac (th RS disj_forward)) @ ths end - | _ => nodups ctxt th :: ths (*no work to do*) - and cnf_nil th = cnf_aux (th,[]) - val cls = - if has_too_many_clauses ctxt (concl_of th) - then (trace_msg (fn () => "cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths) - else cnf_aux (th,ths) - in (cls, !ctxtr) end; - -fun make_cnf old_skolem_ths th ctxt = cnf old_skolem_ths ctxt (th, []) - -(*Generalization, removal of redundant equalities, removal of tautologies.*) -fun finish_cnf ths = filter (not o is_taut) (map refl_clause ths); - - -(**** Generation of contrapositives ****) - -fun is_left (Const (@{const_name Trueprop}, _) $ - (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _)) = true - | is_left _ = false; - -(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) -fun assoc_right th = - if is_left (prop_of th) then assoc_right (th RS disj_assoc) - else th; - -(*Must check for negative literal first!*) -val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; - -(*For ordinary resolution. *) -val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule']; - -(*Create a goal or support clause, conclusing False*) -fun make_goal th = (*Must check for negative literal first!*) - make_goal (tryres(th, clause_rules)) - handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); - -(*Sort clauses by number of literals*) -fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2); - -fun sort_clauses ths = sort (make_ord fewerlits) ths; - -fun has_bool @{typ bool} = true - | has_bool (Type (_, Ts)) = exists has_bool Ts - | has_bool _ = false - -fun has_fun (Type (@{type_name fun}, _)) = true - | has_fun (Type (_, Ts)) = exists has_fun Ts - | has_fun _ = false - -(*Is the string the name of a connective? Really only | and Not can remain, - since this code expects to be called on a clause form.*) -val is_conn = member (op =) - [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj}, - @{const_name HOL.implies}, @{const_name Not}, - @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex}]; - -(*True if the term contains a function--not a logical connective--where the type - of any argument contains bool.*) -val has_bool_arg_const = - exists_Const - (fn (c,T) => not(is_conn c) andalso exists has_bool (binder_types T)); - -(*A higher-order instance of a first-order constant? Example is the definition of - one, 1, at a function type in theory Function_Algebras.*) -fun higher_inst_const thy (c,T) = - case binder_types T of - [] => false (*not a function type, OK*) - | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts; - -(*Returns false if any Vars in the theorem mention type bool. - Also rejects functions whose arguments are Booleans or other functions.*) -fun is_fol_term thy t = - Term.is_first_order ["all", @{const_name All}, @{const_name Ex}] t andalso - not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T - | _ => false) t orelse - has_bool_arg_const t orelse - exists_Const (higher_inst_const thy) t orelse - has_meta_conn t); - -fun rigid t = not (is_Var (head_of t)); - -fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name HOL.disj}, _) $ t $ _)) = rigid t - | ok4horn (Const (@{const_name Trueprop},_) $ t) = rigid t - | ok4horn _ = false; - -(*Create a meta-level Horn clause*) -fun make_horn crules th = - if ok4horn (concl_of th) - then make_horn crules (tryres(th,crules)) handle THM _ => th - else th; - -(*Generate Horn clauses for all contrapositives of a clause. The input, th, - is a HOL disjunction.*) -fun add_contras crules th hcs = - let fun rots (0,_) = hcs - | rots (k,th) = zero_var_indexes (make_horn crules th) :: - rots(k-1, assoc_right (th RS disj_comm)) - in case nliterals(prop_of th) of - 1 => th::hcs - | n => rots(n, assoc_right th) - end; - -(*Use "theorem naming" to label the clauses*) -fun name_thms label = - let fun name1 th (k, ths) = - (k-1, Thm.put_name_hint (label ^ string_of_int k) th :: ths) - in fn ths => #2 (fold_rev name1 ths (length ths, [])) end; - -(*Is the given disjunction an all-negative support clause?*) -fun is_negative th = forall (not o #1) (literals (prop_of th)); - -val neg_clauses = filter is_negative; - - -(***** MESON PROOF PROCEDURE *****) - -fun rhyps (Const("==>",_) $ (Const(@{const_name Trueprop},_) $ A) $ phi, - As) = rhyps(phi, A::As) - | rhyps (_, As) = As; - -(** Detecting repeated assumptions in a subgoal **) - -(*The stringtree detects repeated assumptions.*) -fun ins_term t net = Net.insert_term (op aconv) (t, t) net; - -(*detects repetitions in a list of terms*) -fun has_reps [] = false - | has_reps [_] = false - | has_reps [t,u] = (t aconv u) - | has_reps ts = (fold ins_term ts Net.empty; false) handle Net.INSERT => true; - -(*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) -fun TRYING_eq_assume_tac 0 st = Seq.single st - | TRYING_eq_assume_tac i st = - TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st) - handle THM _ => TRYING_eq_assume_tac (i-1) st; - -fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (nprems_of st) st; - -(*Loop checking: FAIL if trying to prove the same thing twice - -- if *ANY* subgoal has repeated literals*) -fun check_tac st = - if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st) - then Seq.empty else Seq.single st; - - -(* net_resolve_tac actually made it slower... *) -fun prolog_step_tac horns i = - (assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN - TRYALL_eq_assume_tac; - -(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) -fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz; - -fun size_of_subgoals st = fold_rev addconcl (prems_of st) 0; - - -(*Negation Normal Form*) -val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, - not_impD, not_iffD, not_allD, not_exD, not_notD]; - -fun ok4nnf (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ t)) = rigid t - | ok4nnf (Const (@{const_name Trueprop},_) $ t) = rigid t - | ok4nnf _ = false; - -fun make_nnf1 ctxt th = - if ok4nnf (concl_of th) - then make_nnf1 ctxt (tryres(th, nnf_rls)) - handle THM ("tryres", _, _) => - forward_res ctxt (make_nnf1 ctxt) - (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward])) - handle THM ("tryres", _, _) => th - else th - -(*The simplification removes defined quantifiers and occurrences of True and False. - nnf_ss also includes the one-point simprocs, - which are needed to avoid the various one-point theorems from generating junk clauses.*) -val nnf_simps = - @{thms simp_implies_def Ex1_def Ball_def Bex_def if_True if_False if_cancel - if_eq_cancel cases_simp} -val nnf_extra_simps = @{thms split_ifs ex_simps all_simps simp_thms} - -val nnf_ss = - HOL_basic_ss addsimps nnf_extra_simps - addsimprocs [defALL_regroup,defEX_regroup, @{simproc neq}, @{simproc let_simp}]; - -val presimplify = - rewrite_rule (map safe_mk_meta_eq nnf_simps) #> simplify nnf_ss - -fun make_nnf ctxt th = case prems_of th of - [] => th |> presimplify |> make_nnf1 ctxt - | _ => raise THM ("make_nnf: premises in argument", 0, [th]); - -(* Pull existential quantifiers to front. This accomplishes Skolemization for - clauses that arise from a subgoal. *) -fun skolemize_with_choice_thms ctxt choice_ths = - let - fun aux th = - if not (has_conns [@{const_name Ex}] (prop_of th)) then - th - else - tryres (th, choice_ths @ - [conj_exD1, conj_exD2, disj_exD, disj_exD1, disj_exD2]) - |> aux - handle THM ("tryres", _, _) => - tryres (th, [conj_forward, disj_forward, all_forward]) - |> forward_res ctxt aux - |> aux - handle THM ("tryres", _, _) => - rename_bvs_RS th ex_forward - |> forward_res ctxt aux - in aux o make_nnf ctxt end - -fun skolemize ctxt = skolemize_with_choice_thms ctxt (Meson_Choices.get ctxt) - -(* "RS" can fail if "unify_search_bound" is too small. *) -fun try_skolemize ctxt th = - try (skolemize ctxt) th - |> tap (fn NONE => trace_msg (fn () => "Failed to skolemize " ^ - Display.string_of_thm ctxt th) - | _ => ()) - -fun add_clauses th cls = - let val ctxt0 = Variable.global_thm_context th - val (cnfs, ctxt) = make_cnf [] th ctxt0 - in Variable.export ctxt ctxt0 cnfs @ cls end; - -(*Make clauses from a list of theorems, previously Skolemized and put into nnf. - The resulting clauses are HOL disjunctions.*) -fun make_clauses_unsorted ths = fold_rev add_clauses ths []; -val make_clauses = sort_clauses o make_clauses_unsorted; - -(*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*) -fun make_horns ths = - name_thms "Horn#" - (distinct Thm.eq_thm_prop (fold_rev (add_contras clause_rules) ths [])); - -(*Could simply use nprems_of, which would count remaining subgoals -- no - discrimination as to their size! With BEST_FIRST, fails for problem 41.*) - -fun best_prolog_tac sizef horns = - BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1); - -fun depth_prolog_tac horns = - DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1); - -(*Return all negative clauses, as possible goal clauses*) -fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls)); - -fun skolemize_prems_tac ctxt prems = - cut_facts_tac (map_filter (try_skolemize ctxt) prems) THEN' REPEAT o etac exE - -(*Basis of all meson-tactics. Supplies cltac with clauses: HOL disjunctions. - Function mkcl converts theorems to clauses.*) -fun MESON preskolem_tac mkcl cltac ctxt i st = - SELECT_GOAL - (EVERY [Object_Logic.atomize_prems_tac 1, - rtac ccontr 1, - preskolem_tac, - Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} => - EVERY1 [skolemize_prems_tac ctxt negs, - Subgoal.FOCUS (cltac o mkcl o #prems) ctxt']) ctxt 1]) i st - handle THM _ => no_tac st; (*probably from make_meta_clause, not first-order*) - - -(** Best-first search versions **) - -(*ths is a list of additional clauses (HOL disjunctions) to use.*) -fun best_meson_tac sizef = - MESON all_tac make_clauses - (fn cls => - THEN_BEST_FIRST (resolve_tac (gocls cls) 1) - (has_fewer_prems 1, sizef) - (prolog_step_tac (make_horns cls) 1)); - -(*First, breaks the goal into independent units*) -fun safe_best_meson_tac ctxt = - SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN - TRYALL (best_meson_tac size_of_subgoals ctxt)); - -(** Depth-first search version **) - -val depth_meson_tac = - MESON all_tac make_clauses - (fn cls => EVERY [resolve_tac (gocls cls) 1, depth_prolog_tac (make_horns cls)]); - - -(** Iterative deepening version **) - -(*This version does only one inference per call; - having only one eq_assume_tac speeds it up!*) -fun prolog_step_tac' horns = - let val (horn0s, _) = (*0 subgoals vs 1 or more*) - take_prefix Thm.no_prems horns - val nrtac = net_resolve_tac horns - in fn i => eq_assume_tac i ORELSE - match_tac horn0s i ORELSE (*no backtracking if unit MATCHES*) - ((assume_tac i APPEND nrtac i) THEN check_tac) - end; - -fun iter_deepen_prolog_tac horns = - ITER_DEEPEN iter_deepen_limit (has_fewer_prems 1) (prolog_step_tac' horns); - -fun iter_deepen_meson_tac ctxt ths = ctxt |> MESON all_tac make_clauses - (fn cls => - (case (gocls (cls @ ths)) of - [] => no_tac (*no goal clauses*) - | goes => - let - val horns = make_horns (cls @ ths) - val _ = trace_msg (fn () => - cat_lines ("meson method called:" :: - map (Display.string_of_thm ctxt) (cls @ ths) @ - ["clauses:"] @ map (Display.string_of_thm ctxt) horns)) - in - THEN_ITER_DEEPEN iter_deepen_limit - (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns) - end)); - -fun meson_tac ctxt ths = - SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN TRYALL (iter_deepen_meson_tac ctxt ths)); - - -(**** Code to support ordinary resolution, rather than Model Elimination ****) - -(*Convert a list of clauses (disjunctions) to meta-level clauses (==>), - with no contrapositives, for ordinary resolution.*) - -(*Rules to convert the head literal into a negated assumption. If the head - literal is already negated, then using notEfalse instead of notEfalse' - prevents a double negation.*) -val notEfalse = read_instantiate @{context} [(("R", 0), "False")] notE; -val notEfalse' = rotate_prems 1 notEfalse; - -fun negated_asm_of_head th = - th RS notEfalse handle THM _ => th RS notEfalse'; - -(*Converting one theorem from a disjunction to a meta-level clause*) -fun make_meta_clause th = - let val (fth,thaw) = Drule.legacy_freeze_thaw_robust th - in - (zero_var_indexes o Thm.varifyT_global o thaw 0 o - negated_asm_of_head o make_horn resolution_clause_rules) fth - end; - -fun make_meta_clauses ths = - name_thms "MClause#" - (distinct Thm.eq_thm_prop (map make_meta_clause ths)); - -end;