# HG changeset patch # User wenzelm # Date 1164811491 -3600 # Node ID cd0dc678a2055f8f0018c20e85386be9c3fa6c36 # Parent a3561bfe0ada6aa4619d078d29a83b030ec0c9a5 simplified method setup; diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Algebra/abstract/Ring2.thy --- a/src/HOL/Algebra/abstract/Ring2.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Algebra/abstract/Ring2.thy Wed Nov 29 15:44:51 2006 +0100 @@ -231,7 +231,7 @@ *} (* Note: r_one is not necessary in ring_ss *) method_setup ring = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (full_simp_tac ring_ss)) *} + {* Method.no_args (Method.SIMPLE_METHOD' (full_simp_tac ring_ss)) *} {* computes distributive normal form in rings *} lemmas ring_simps = diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Algebra/ringsimp.ML --- a/src/HOL/Algebra/ringsimp.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Algebra/ringsimp.ML Wed Nov 29 15:44:51 2006 +0100 @@ -60,9 +60,6 @@ fun algebra_tac ctxt = EVERY' (map (fn s => TRY o struct_tac s) (AlgebraData.get (Context.Proof ctxt))); -val method = - Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD' HEADGOAL (algebra_tac ctxt)) - (** Attribute **) @@ -87,8 +84,8 @@ val setup = AlgebraData.init #> - Method.add_methods [("algebra", method, "normalisation of algebraic structure")] #> + Method.add_methods [("algebra", Method.ctxt_args (Method.SIMPLE_METHOD' o algebra_tac), + "normalisation of algebraic structure")] #> Attrib.add_attributes [("algebra", attribute, "theorems controlling algebra method")]; - end; diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/Event.thy --- a/src/HOL/Auth/Event.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/Event.thy Wed Nov 29 15:44:51 2006 +0100 @@ -290,8 +290,7 @@ intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD]) method_setup analz_mono_contra = {* - Method.no_args - (Method.METHOD (fn facts => REPEAT_FIRST analz_mono_contra_tac)) *} + Method.no_args (Method.SIMPLE_METHOD (REPEAT_FIRST analz_mono_contra_tac)) *} "for proving theorems of the form X \ analz (knows Spy evs) --> P" subsubsection{*Useful for case analysis on whether a hash is a spoof or not*} @@ -349,8 +348,7 @@ *} method_setup synth_analz_mono_contra = {* - Method.no_args - (Method.METHOD (fn facts => REPEAT_FIRST synth_analz_mono_contra_tac)) *} + Method.no_args (Method.SIMPLE_METHOD (REPEAT_FIRST synth_analz_mono_contra_tac)) *} "for proving theorems of the form X \ synth (analz (knows Spy evs)) --> P" end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/Message.thy Wed Nov 29 15:44:51 2006 +0100 @@ -948,20 +948,17 @@ method_setup spy_analz = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_spy_analz_tac (local_clasimpset_of ctxt) 1)) *} + Method.SIMPLE_METHOD (gen_spy_analz_tac (local_clasimpset_of ctxt) 1)) *} "for proving the Fake case when analz is involved" method_setup atomic_spy_analz = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - atomic_spy_analz_tac (local_clasimpset_of ctxt) 1)) *} + Method.SIMPLE_METHOD (atomic_spy_analz_tac (local_clasimpset_of ctxt) 1)) *} "for debugging spy_analz" method_setup Fake_insert_simp = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - Fake_insert_simp_tac (local_simpset_of ctxt) 1)) *} + Method.SIMPLE_METHOD (Fake_insert_simp_tac (local_simpset_of ctxt) 1)) *} "for debugging spy_analz" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/OtwayReesBella.thy --- a/src/HOL/Auth/OtwayReesBella.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/OtwayReesBella.thy Wed Nov 29 15:44:51 2006 +0100 @@ -142,9 +142,7 @@ *} method_setup parts_explicit = {* - Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - parts_explicit_tac 1)) *} + Method.no_args (Method.SIMPLE_METHOD' parts_explicit_tac) *} "to explicitly state that some message components belong to parts knows Spy" @@ -263,8 +261,8 @@ method_setup analz_freshCryptK = {* Method.ctxt_args (fn ctxt => - (Method.METHOD - (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), + (Method.SIMPLE_METHOD + (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac analz_image_freshCryptK_lemma), ALLGOALS (asm_simp_tac (Simplifier.context ctxt analz_image_freshK_ss))]))) *} "for proving useful rewrite rule" @@ -272,8 +270,8 @@ method_setup disentangle = {* Method.no_args - (Method.METHOD - (fn facts => REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE] + (Method.SIMPLE_METHOD + (REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE] ORELSE' hyp_subst_tac))) *} "for eliminating conjunctions, disjunctions and the like" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/Public.thy --- a/src/HOL/Auth/Public.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/Public.thy Wed Nov 29 15:44:51 2006 +0100 @@ -415,8 +415,8 @@ method_setup analz_freshK = {* Method.ctxt_args (fn ctxt => - (Method.METHOD - (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), + (Method.SIMPLE_METHOD + (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac analz_image_freshK_lemma), ALLGOALS (asm_simp_tac (Simplifier.context ctxt analz_image_freshK_ss))]))) *} "for proving the Session Key Compromise theorem" @@ -449,8 +449,7 @@ method_setup possibility = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_possibility_tac (local_simpset_of ctxt))) *} + Method.SIMPLE_METHOD (gen_possibility_tac (local_simpset_of ctxt))) *} "for proving possibility theorems" end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/Shared.thy --- a/src/HOL/Auth/Shared.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/Shared.thy Wed Nov 29 15:44:51 2006 +0100 @@ -265,16 +265,15 @@ method_setup analz_freshK = {* Method.ctxt_args (fn ctxt => - (Method.METHOD - (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), + (Method.SIMPLE_METHOD + (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac analz_image_freshK_lemma), ALLGOALS (asm_simp_tac (Simplifier.context ctxt analz_image_freshK_ss))]))) *} "for proving the Session Key Compromise theorem" method_setup possibility = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_possibility_tac (local_simpset_of ctxt))) *} + Method.SIMPLE_METHOD (gen_possibility_tac (local_simpset_of ctxt))) *} "for proving possibility theorems" lemma knows_subset_knows_Cons: "knows A evs <= knows A (e # evs)" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/Smartcard/ShoupRubin.thy --- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Wed Nov 29 15:44:51 2006 +0100 @@ -774,15 +774,15 @@ *} method_setup prepare = {* - Method.no_args (Method.METHOD (fn facts => prepare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD prepare_tac) *} "to launch a few simple facts that'll help the simplifier" method_setup parts_prepare = {* - Method.no_args (Method.METHOD (fn facts => parts_prepare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD parts_prepare_tac) *} "additional facts to reason about parts" method_setup analz_prepare = {* - Method.no_args (Method.METHOD (fn facts => analz_prepare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD analz_prepare_tac) *} "additional facts to reason about analz" @@ -836,8 +836,8 @@ method_setup sc_analz_freshK = {* Method.ctxt_args (fn ctxt => - (Method.METHOD - (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), + (Method.SIMPLE_METHOD + (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac analz_image_freshK_lemma), ALLGOALS (asm_simp_tac (Simplifier.context ctxt analz_image_freshK_ss addsimps [knows_Spy_Inputs_secureM_sr_Spy, diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/Smartcard/ShoupRubinBella.thy --- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Wed Nov 29 15:44:51 2006 +0100 @@ -782,15 +782,15 @@ *} method_setup prepare = {* - Method.no_args (Method.METHOD (fn facts => prepare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD prepare_tac) *} "to launch a few simple facts that'll help the simplifier" method_setup parts_prepare = {* - Method.no_args (Method.METHOD (fn facts => parts_prepare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD parts_prepare_tac) *} "additional facts to reason about parts" method_setup analz_prepare = {* - Method.no_args (Method.METHOD (fn facts => analz_prepare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD analz_prepare_tac) *} "additional facts to reason about analz" @@ -845,8 +845,8 @@ method_setup sc_analz_freshK = {* Method.ctxt_args (fn ctxt => - (Method.METHOD - (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), + (Method.SIMPLE_METHOD + (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac analz_image_freshK_lemma), ALLGOALS (asm_simp_tac (Simplifier.context ctxt analz_image_freshK_ss addsimps [knows_Spy_Inputs_secureM_srb_Spy, diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Wed Nov 29 15:44:51 2006 +0100 @@ -442,16 +442,15 @@ method_setup analz_freshK = {* Method.ctxt_args (fn ctxt => - (Method.METHOD - (fn facts => EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), + (Method.SIMPLE_METHOD + (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac analz_image_freshK_lemma), ALLGOALS (asm_simp_tac (Simplifier.context ctxt analz_image_freshK_ss))]))) *} "for proving the Session Key Compromise theorem" method_setup possibility = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_possibility_tac (Simplifier.get_local_simpset ctxt))) *} + Method.SIMPLE_METHOD (gen_possibility_tac (Simplifier.get_local_simpset ctxt))) *} "for proving possibility theorems" lemma knows_subset_knows_Cons: "knows A evs \ knows A (e # evs)" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Hoare/Hoare.thy --- a/src/HOL/Hoare/Hoare.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Hoare/Hoare.thy Wed Nov 29 15:44:51 2006 +0100 @@ -228,14 +228,12 @@ use "hoare.ML" method_setup vcg = {* - Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL (hoare_tac (K all_tac))) *} + Method.no_args (Method.SIMPLE_METHOD' (hoare_tac (K all_tac))) *} "verification condition generator" method_setup vcg_simp = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - hoare_tac (asm_full_simp_tac (local_simpset_of ctxt))1)) *} + Method.SIMPLE_METHOD' (hoare_tac (asm_full_simp_tac (local_simpset_of ctxt)))) *} "verification condition generator plus simplification" end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Hoare/HoareAbort.thy --- a/src/HOL/Hoare/HoareAbort.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Hoare/HoareAbort.thy Wed Nov 29 15:44:51 2006 +0100 @@ -238,14 +238,12 @@ use "hoareAbort.ML" method_setup vcg = {* - Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL (hoare_tac (K all_tac))) *} + Method.no_args (Method.SIMPLE_METHOD' (hoare_tac (K all_tac))) *} "verification condition generator" method_setup vcg_simp = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - hoare_tac (asm_full_simp_tac (local_simpset_of ctxt))1)) *} + Method.SIMPLE_METHOD' (hoare_tac (asm_full_simp_tac (local_simpset_of ctxt)))) *} "verification condition generator plus simplification" (* Special syntax for guarded statements and guarded array updates: *) diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/HoareParallel/OG_Tactics.thy --- a/src/HOL/HoareParallel/OG_Tactics.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/HoareParallel/OG_Tactics.thy Wed Nov 29 15:44:51 2006 +0100 @@ -464,25 +464,21 @@ Isabelle proofs. *} method_setup oghoare = {* - Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL (oghoare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD' oghoare_tac) *} "verification condition generator for the oghoare logic" method_setup annhoare = {* - Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL (annhoare_tac)) *} + Method.no_args (Method.SIMPLE_METHOD' annhoare_tac) *} "verification condition generator for the ann_hoare logic" method_setup interfree_aux = {* - Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL (interfree_aux_tac)) *} + Method.no_args (Method.SIMPLE_METHOD' interfree_aux_tac) *} "verification condition generator for interference freedom tests" text {* Tactics useful for dealing with the generated verification conditions: *} method_setup conjI_tac = {* - Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL (conjI_Tac (K all_tac))) *} + Method.no_args (Method.SIMPLE_METHOD' (conjI_Tac (K all_tac))) *} "verification condition generator for interference freedom tests" ML {* @@ -493,8 +489,7 @@ *} method_setup disjE_tac = {* - Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL (disjE_Tac (K all_tac))) *} + Method.no_args (Method.SIMPLE_METHOD' (disjE_Tac (K all_tac))) *} "verification condition generator for interference freedom tests" end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Hyperreal/HyperDef.thy --- a/src/HOL/Hyperreal/HyperDef.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Hyperreal/HyperDef.thy Wed Nov 29 15:44:51 2006 +0100 @@ -183,14 +183,12 @@ method_setup fuf = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - fuf_tac (local_clasimpset_of ctxt) 1)) *} + Method.SIMPLE_METHOD' (fuf_tac (local_clasimpset_of ctxt))) *} "free ultrafilter tactic" method_setup ultra = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - ultra_tac (local_clasimpset_of ctxt) 1)) *} + Method.SIMPLE_METHOD' (ultra_tac (local_clasimpset_of ctxt))) *} "ultrafilter tactic" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Hyperreal/transfer.ML --- a/src/HOL/Hyperreal/transfer.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Hyperreal/transfer.ML Wed Nov 29 15:44:51 2006 +0100 @@ -107,9 +107,6 @@ (fn {intros,unfolds,refolds,consts} => {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts})) -val transfer_method = Method.thms_ctxt_args (fn ths => fn ctxt => - Method.SIMPLE_METHOD' HEADGOAL (transfer_tac ctxt ths)); - val setup = TransferData.init #> Attrib.add_attributes @@ -119,6 +116,7 @@ "declaration of transfer unfolding rule"), ("transfer_refold", Attrib.add_del_args refold_add refold_del, "declaration of transfer refolding rule")] #> - Method.add_method ("transfer", transfer_method, "transfer principle"); + Method.add_method ("transfer", Method.thms_ctxt_args (fn ths => fn ctxt => + Method.SIMPLE_METHOD' (transfer_tac ctxt ths)), "transfer principle"); end; diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Import/shuffler.ML --- a/src/HOL/Import/shuffler.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Import/shuffler.ML Wed Nov 29 15:44:51 2006 +0100 @@ -38,21 +38,21 @@ val message = if_debug writeln (*Prints exceptions readably to users*) -fun print_sign_exn_unit sign e = +fun print_sign_exn_unit sign e = case e of THM (msg,i,thms) => - (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); - List.app print_thm thms) + (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); + List.app print_thm thms) | THEORY (msg,thys) => - (writeln ("Exception THEORY raised:\n" ^ msg); - List.app (writeln o Context.str_of_thy) thys) + (writeln ("Exception THEORY raised:\n" ^ msg); + List.app (writeln o Context.str_of_thy) thys) | TERM (msg,ts) => - (writeln ("Exception TERM raised:\n" ^ msg); - List.app (writeln o Sign.string_of_term sign) ts) + (writeln ("Exception TERM raised:\n" ^ msg); + List.app (writeln o Sign.string_of_term sign) ts) | TYPE (msg,Ts,ts) => - (writeln ("Exception TYPE raised:\n" ^ msg); - List.app (writeln o Sign.string_of_typ sign) Ts; - List.app (writeln o Sign.string_of_term sign) ts) + (writeln ("Exception TYPE raised:\n" ^ msg); + List.app (writeln o Sign.string_of_typ sign) Ts; + List.app (writeln o Sign.string_of_term sign) ts) | e => raise e (*Prints an exception, then fails*) @@ -63,14 +63,14 @@ fun mk_meta_eq th = (case concl_of th of - Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th RS eq_reflection + Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th RS eq_reflection | Const("==",_) $ _ $ _ => th | _ => raise THM("Not an equality",0,[th])) handle _ => raise THM("Couldn't make meta equality",0,[th]) - + fun mk_obj_eq th = (case concl_of th of - Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th + Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th | Const("==",_) $ _ $ _ => th RS meta_eq_to_obj_eq | _ => raise THM("Not an equality",0,[th])) handle _ => raise THM("Couldn't make object equality",0,[th]) @@ -85,113 +85,113 @@ fun merge _ = Library.gen_union Thm.eq_thm fun print _ thms = Pretty.writeln (Pretty.big_list "Shuffle theorems:" - (map Display.pretty_thm thms)) + (map Display.pretty_thm thms)) end structure ShuffleData = TheoryDataFun(ShuffleDataArgs) val weaken = let - val cert = cterm_of ProtoPure.thy - val P = Free("P",propT) - val Q = Free("Q",propT) - val PQ = Logic.mk_implies(P,Q) - val PPQ = Logic.mk_implies(P,PQ) - val cP = cert P - val cQ = cert Q - val cPQ = cert PQ - val cPPQ = cert PPQ - val th1 = assume cPQ |> implies_intr_list [cPQ,cP] - val th3 = assume cP - val th4 = implies_elim_list (assume cPPQ) [th3,th3] - |> implies_intr_list [cPPQ,cP] + val cert = cterm_of ProtoPure.thy + val P = Free("P",propT) + val Q = Free("Q",propT) + val PQ = Logic.mk_implies(P,Q) + val PPQ = Logic.mk_implies(P,PQ) + val cP = cert P + val cQ = cert Q + val cPQ = cert PQ + val cPPQ = cert PPQ + val th1 = assume cPQ |> implies_intr_list [cPQ,cP] + val th3 = assume cP + val th4 = implies_elim_list (assume cPPQ) [th3,th3] + |> implies_intr_list [cPPQ,cP] in - equal_intr th4 th1 |> standard + equal_intr th4 th1 |> standard end val imp_comm = let - val cert = cterm_of ProtoPure.thy - val P = Free("P",propT) - val Q = Free("Q",propT) - val R = Free("R",propT) - val PQR = Logic.mk_implies(P,Logic.mk_implies(Q,R)) - val QPR = Logic.mk_implies(Q,Logic.mk_implies(P,R)) - val cP = cert P - val cQ = cert Q - val cPQR = cert PQR - val cQPR = cert QPR - val th1 = implies_elim_list (assume cPQR) [assume cP,assume cQ] - |> implies_intr_list [cPQR,cQ,cP] - val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP] - |> implies_intr_list [cQPR,cP,cQ] + val cert = cterm_of ProtoPure.thy + val P = Free("P",propT) + val Q = Free("Q",propT) + val R = Free("R",propT) + val PQR = Logic.mk_implies(P,Logic.mk_implies(Q,R)) + val QPR = Logic.mk_implies(Q,Logic.mk_implies(P,R)) + val cP = cert P + val cQ = cert Q + val cPQR = cert PQR + val cQPR = cert QPR + val th1 = implies_elim_list (assume cPQR) [assume cP,assume cQ] + |> implies_intr_list [cPQR,cQ,cP] + val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP] + |> implies_intr_list [cQPR,cP,cQ] in - equal_intr th1 th2 |> standard + equal_intr th1 th2 |> standard end val def_norm = let - val cert = cterm_of ProtoPure.thy - val aT = TFree("'a",[]) - val bT = TFree("'b",[]) - val v = Free("v",aT) - val P = Free("P",aT-->bT) - val Q = Free("Q",aT-->bT) - val cvPQ = cert (list_all ([("v",aT)],Logic.mk_equals(P $ Bound 0,Q $ Bound 0))) - val cPQ = cert (Logic.mk_equals(P,Q)) - val cv = cert v - val rew = assume cvPQ - |> forall_elim cv - |> abstract_rule "v" cv - val (lhs,rhs) = Logic.dest_equals(concl_of rew) - val th1 = transitive (transitive - (eta_conversion (cert lhs) |> symmetric) - rew) - (eta_conversion (cert rhs)) - |> implies_intr cvPQ - val th2 = combination (assume cPQ) (reflexive cv) - |> forall_intr cv - |> implies_intr cPQ + val cert = cterm_of ProtoPure.thy + val aT = TFree("'a",[]) + val bT = TFree("'b",[]) + val v = Free("v",aT) + val P = Free("P",aT-->bT) + val Q = Free("Q",aT-->bT) + val cvPQ = cert (list_all ([("v",aT)],Logic.mk_equals(P $ Bound 0,Q $ Bound 0))) + val cPQ = cert (Logic.mk_equals(P,Q)) + val cv = cert v + val rew = assume cvPQ + |> forall_elim cv + |> abstract_rule "v" cv + val (lhs,rhs) = Logic.dest_equals(concl_of rew) + val th1 = transitive (transitive + (eta_conversion (cert lhs) |> symmetric) + rew) + (eta_conversion (cert rhs)) + |> implies_intr cvPQ + val th2 = combination (assume cPQ) (reflexive cv) + |> forall_intr cv + |> implies_intr cPQ in - equal_intr th1 th2 |> standard + equal_intr th1 th2 |> standard end val all_comm = let - val cert = cterm_of ProtoPure.thy - val xT = TFree("'a",[]) - val yT = TFree("'b",[]) - val P = Free("P",xT-->yT-->propT) - val lhs = all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0)))) - val rhs = all yT $ (Abs("y",yT,all xT $ (Abs("x",xT,P $ Bound 0 $ Bound 1)))) - val cl = cert lhs - val cr = cert rhs - val cx = cert (Free("x",xT)) - val cy = cert (Free("y",yT)) - val th1 = assume cr - |> forall_elim_list [cy,cx] - |> forall_intr_list [cx,cy] - |> implies_intr cr - val th2 = assume cl - |> forall_elim_list [cx,cy] - |> forall_intr_list [cy,cx] - |> implies_intr cl + val cert = cterm_of ProtoPure.thy + val xT = TFree("'a",[]) + val yT = TFree("'b",[]) + val P = Free("P",xT-->yT-->propT) + val lhs = all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0)))) + val rhs = all yT $ (Abs("y",yT,all xT $ (Abs("x",xT,P $ Bound 0 $ Bound 1)))) + val cl = cert lhs + val cr = cert rhs + val cx = cert (Free("x",xT)) + val cy = cert (Free("y",yT)) + val th1 = assume cr + |> forall_elim_list [cy,cx] + |> forall_intr_list [cx,cy] + |> implies_intr cr + val th2 = assume cl + |> forall_elim_list [cx,cy] + |> forall_intr_list [cy,cx] + |> implies_intr cl in - equal_intr th1 th2 |> standard + equal_intr th1 th2 |> standard end val equiv_comm = let - val cert = cterm_of ProtoPure.thy - val T = TFree("'a",[]) - val t = Free("t",T) - val u = Free("u",T) - val ctu = cert (Logic.mk_equals(t,u)) - val cut = cert (Logic.mk_equals(u,t)) - val th1 = assume ctu |> symmetric |> implies_intr ctu - val th2 = assume cut |> symmetric |> implies_intr cut + val cert = cterm_of ProtoPure.thy + val T = TFree("'a",[]) + val t = Free("t",T) + val u = Free("u",T) + val ctu = cert (Logic.mk_equals(t,u)) + val cut = cert (Logic.mk_equals(u,t)) + val th1 = assume ctu |> symmetric |> implies_intr ctu + val th2 = assume cut |> symmetric |> implies_intr cut in - equal_intr th1 th2 |> standard + equal_intr th1 th2 |> standard end (* This simplification procedure rewrites !!x y. P x y @@ -203,82 +203,82 @@ exception RESULT of int fun find_bound n (Bound i) = if i = n then raise RESULT 0 - else if i = n+1 then raise RESULT 1 - else () + else if i = n+1 then raise RESULT 1 + else () | find_bound n (t $ u) = (find_bound n t; find_bound n u) | find_bound n (Abs(_,_,t)) = find_bound (n+1) t | find_bound _ _ = () fun swap_bound n (Bound i) = if i = n then Bound (n+1) - else if i = n+1 then Bound n - else Bound i + else if i = n+1 then Bound n + else Bound i | swap_bound n (t $ u) = (swap_bound n t $ swap_bound n u) | swap_bound n (Abs(x,xT,t)) = Abs(x,xT,swap_bound (n+1) t) | swap_bound n t = t fun rew_th thy (xv as (x,xT)) (yv as (y,yT)) t = let - val lhs = list_all ([xv,yv],t) - val rhs = list_all ([yv,xv],swap_bound 0 t) - val rew = Logic.mk_equals (lhs,rhs) - val init = trivial (cterm_of thy rew) + val lhs = list_all ([xv,yv],t) + val rhs = list_all ([yv,xv],swap_bound 0 t) + val rew = Logic.mk_equals (lhs,rhs) + val init = trivial (cterm_of thy rew) in - (all_comm RS init handle e => (message "rew_th"; OldGoals.print_exn e)) + (all_comm RS init handle e => (message "rew_th"; OldGoals.print_exn e)) end fun quant_rewrite thy assumes (t as Const("all",T1) $ (Abs(x,xT,Const("all",T2) $ Abs(y,yT,body)))) = let - val res = (find_bound 0 body;2) handle RESULT i => i + val res = (find_bound 0 body;2) handle RESULT i => i in - case res of - 0 => SOME (rew_th thy (x,xT) (y,yT) body) - | 1 => if string_ord(y,x) = LESS - then - let - val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body))) - val t_th = reflexive (cterm_of thy t) - val newt_th = reflexive (cterm_of thy newt) - in - SOME (transitive t_th newt_th) - end - else NONE - | _ => error "norm_term (quant_rewrite) internal error" + case res of + 0 => SOME (rew_th thy (x,xT) (y,yT) body) + | 1 => if string_ord(y,x) = LESS + then + let + val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body))) + val t_th = reflexive (cterm_of thy t) + val newt_th = reflexive (cterm_of thy newt) + in + SOME (transitive t_th newt_th) + end + else NONE + | _ => error "norm_term (quant_rewrite) internal error" end | quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; NONE) fun freeze_thaw_term t = let - val tvars = term_tvars t - val tfree_names = add_term_tfree_names(t,[]) - val (type_inst,_) = - Library.foldl (fn ((inst,used),(w as (v,_),S)) => - let - val v' = Name.variant used v - in - ((w,TFree(v',S))::inst,v'::used) - end) - (([],tfree_names),tvars) - val t' = subst_TVars type_inst t + val tvars = term_tvars t + val tfree_names = add_term_tfree_names(t,[]) + val (type_inst,_) = + Library.foldl (fn ((inst,used),(w as (v,_),S)) => + let + val v' = Name.variant used v + in + ((w,TFree(v',S))::inst,v'::used) + end) + (([],tfree_names),tvars) + val t' = subst_TVars type_inst t in - (t',map (fn (w,TFree(v,S)) => (v,TVar(w,S)) - | _ => error "Internal error in Shuffler.freeze_thaw") type_inst) + (t',map (fn (w,TFree(v,S)) => (v,TVar(w,S)) + | _ => error "Internal error in Shuffler.freeze_thaw") type_inst) end fun inst_tfrees thy [] thm = thm - | inst_tfrees thy ((name,U)::rest) thm = + | inst_tfrees thy ((name,U)::rest) thm = let - val cU = ctyp_of thy U - val tfrees = add_term_tfrees (prop_of thm,[]) - val (rens, thm') = Thm.varifyT' + val cU = ctyp_of thy U + val tfrees = add_term_tfrees (prop_of thm,[]) + val (rens, thm') = Thm.varifyT' (remove (op = o apsnd fst) name tfrees) thm - val mid = - case rens of - [] => thm' - | [((_, S), idx)] => instantiate + val mid = + case rens of + [] => thm' + | [((_, S), idx)] => instantiate ([(ctyp_of thy (TVar (idx, S)), cU)], []) thm' - | _ => error "Shuffler.inst_tfrees internal error" + | _ => error "Shuffler.inst_tfrees internal error" in - inst_tfrees thy rest mid + inst_tfrees thy rest mid end fun is_Abs (Abs _) = true @@ -286,65 +286,65 @@ fun eta_redex (t $ Bound 0) = let - fun free n (Bound i) = i = n - | free n (t $ u) = free n t orelse free n u - | free n (Abs(_,_,t)) = free (n+1) t - | free n _ = false + fun free n (Bound i) = i = n + | free n (t $ u) = free n t orelse free n u + | free n (Abs(_,_,t)) = free (n+1) t + | free n _ = false in - not (free 0 t) + not (free 0 t) end | eta_redex _ = false fun eta_contract thy assumes origt = let - val (typet,Tinst) = freeze_thaw_term origt - val (init,thaw) = freeze_thaw (reflexive (cterm_of thy typet)) - val final = inst_tfrees thy Tinst o thaw - val t = #1 (Logic.dest_equals (prop_of init)) - val _ = - let - val lhs = #1 (Logic.dest_equals (prop_of (final init))) - in - if not (lhs aconv origt) - then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)"; - writeln (string_of_cterm (cterm_of thy origt)); - writeln (string_of_cterm (cterm_of thy lhs)); - writeln (string_of_cterm (cterm_of thy typet)); - writeln (string_of_cterm (cterm_of thy t)); - app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of thy T)))) Tinst; - writeln "done") - else () - end + val (typet,Tinst) = freeze_thaw_term origt + val (init,thaw) = freeze_thaw (reflexive (cterm_of thy typet)) + val final = inst_tfrees thy Tinst o thaw + val t = #1 (Logic.dest_equals (prop_of init)) + val _ = + let + val lhs = #1 (Logic.dest_equals (prop_of (final init))) + in + if not (lhs aconv origt) + then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)"; + writeln (string_of_cterm (cterm_of thy origt)); + writeln (string_of_cterm (cterm_of thy lhs)); + writeln (string_of_cterm (cterm_of thy typet)); + writeln (string_of_cterm (cterm_of thy t)); + app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of thy T)))) Tinst; + writeln "done") + else () + end in - case t of - Const("all",_) $ (Abs(x,xT,Const("==",eqT) $ P $ Q)) => - ((if eta_redex P andalso eta_redex Q - then - let - val cert = cterm_of thy - val v = Free(Name.variant (add_term_free_names(t,[])) "v",xT) - val cv = cert v - val ct = cert t - val th = (assume ct) - |> forall_elim cv - |> abstract_rule x cv - val ext_th = eta_conversion (cert (Abs(x,xT,P))) - val th' = transitive (symmetric ext_th) th - val cu = cert (prop_of th') - val uth = combination (assume cu) (reflexive cv) - val uth' = (beta_conversion false (cert (Abs(x,xT,Q) $ v))) - |> transitive uth - |> forall_intr cv - |> implies_intr cu - val rew_th = equal_intr (th' |> implies_intr ct) uth' - val res = final rew_th - val lhs = (#1 (Logic.dest_equals (prop_of res))) - in - SOME res - end - else NONE) - handle e => OldGoals.print_exn e) - | _ => NONE + case t of + Const("all",_) $ (Abs(x,xT,Const("==",eqT) $ P $ Q)) => + ((if eta_redex P andalso eta_redex Q + then + let + val cert = cterm_of thy + val v = Free(Name.variant (add_term_free_names(t,[])) "v",xT) + val cv = cert v + val ct = cert t + val th = (assume ct) + |> forall_elim cv + |> abstract_rule x cv + val ext_th = eta_conversion (cert (Abs(x,xT,P))) + val th' = transitive (symmetric ext_th) th + val cu = cert (prop_of th') + val uth = combination (assume cu) (reflexive cv) + val uth' = (beta_conversion false (cert (Abs(x,xT,Q) $ v))) + |> transitive uth + |> forall_intr cv + |> implies_intr cu + val rew_th = equal_intr (th' |> implies_intr ct) uth' + val res = final rew_th + val lhs = (#1 (Logic.dest_equals (prop_of res))) + in + SOME res + end + else NONE) + handle e => OldGoals.print_exn e) + | _ => NONE end fun beta_fun thy assume t = @@ -354,62 +354,62 @@ fun equals_fun thy assume t = case t of - Const("op ==",_) $ u $ v => if Term.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE + Const("op ==",_) $ u $ v => if Term.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE | _ => NONE fun eta_expand thy assumes origt = let - val (typet,Tinst) = freeze_thaw_term origt - val (init,thaw) = freeze_thaw (reflexive (cterm_of thy typet)) - val final = inst_tfrees thy Tinst o thaw - val t = #1 (Logic.dest_equals (prop_of init)) - val _ = - let - val lhs = #1 (Logic.dest_equals (prop_of (final init))) - in - if not (lhs aconv origt) - then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)"; - writeln (string_of_cterm (cterm_of thy origt)); - writeln (string_of_cterm (cterm_of thy lhs)); - writeln (string_of_cterm (cterm_of thy typet)); - writeln (string_of_cterm (cterm_of thy t)); - app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of thy T)))) Tinst; - writeln "done") - else () - end + val (typet,Tinst) = freeze_thaw_term origt + val (init,thaw) = freeze_thaw (reflexive (cterm_of thy typet)) + val final = inst_tfrees thy Tinst o thaw + val t = #1 (Logic.dest_equals (prop_of init)) + val _ = + let + val lhs = #1 (Logic.dest_equals (prop_of (final init))) + in + if not (lhs aconv origt) + then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)"; + writeln (string_of_cterm (cterm_of thy origt)); + writeln (string_of_cterm (cterm_of thy lhs)); + writeln (string_of_cterm (cterm_of thy typet)); + writeln (string_of_cterm (cterm_of thy t)); + app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of thy T)))) Tinst; + writeln "done") + else () + end in - case t of - Const("==",T) $ P $ Q => - if is_Abs P orelse is_Abs Q - then (case domain_type T of - Type("fun",[aT,bT]) => - let - val cert = cterm_of thy - val vname = Name.variant (add_term_free_names(t,[])) "v" - val v = Free(vname,aT) - val cv = cert v - val ct = cert t - val th1 = (combination (assume ct) (reflexive cv)) - |> forall_intr cv - |> implies_intr ct - val concl = cert (concl_of th1) - val th2 = (assume concl) - |> forall_elim cv - |> abstract_rule vname cv - val (lhs,rhs) = Logic.dest_equals (prop_of th2) - val elhs = eta_conversion (cert lhs) - val erhs = eta_conversion (cert rhs) - val th2' = transitive - (transitive (symmetric elhs) th2) - erhs - val res = equal_intr th1 (th2' |> implies_intr concl) - val res' = final res - in - SOME res' - end - | _ => NONE) - else NONE - | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of thy t))); NONE) + case t of + Const("==",T) $ P $ Q => + if is_Abs P orelse is_Abs Q + then (case domain_type T of + Type("fun",[aT,bT]) => + let + val cert = cterm_of thy + val vname = Name.variant (add_term_free_names(t,[])) "v" + val v = Free(vname,aT) + val cv = cert v + val ct = cert t + val th1 = (combination (assume ct) (reflexive cv)) + |> forall_intr cv + |> implies_intr ct + val concl = cert (concl_of th1) + val th2 = (assume concl) + |> forall_elim cv + |> abstract_rule vname cv + val (lhs,rhs) = Logic.dest_equals (prop_of th2) + val elhs = eta_conversion (cert lhs) + val erhs = eta_conversion (cert rhs) + val th2' = transitive + (transitive (symmetric elhs) th2) + erhs + val res = equal_intr th1 (th2' |> implies_intr concl) + val res' = final res + in + SOME res' + end + | _ => NONE) + else NONE + | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of thy t))); NONE) end handle e => (writeln "eta_expand internal error"; OldGoals.print_exn e) @@ -424,32 +424,32 @@ val S' = mk_free "S'" xT in fun beta_simproc thy = Simplifier.simproc_i - thy - "Beta-contraction" - [Abs("x",xT,Q) $ S] - beta_fun + thy + "Beta-contraction" + [Abs("x",xT,Q) $ S] + beta_fun fun equals_simproc thy = Simplifier.simproc_i - thy - "Ordered rewriting of meta equalities" - [Const("op ==",xT) $ S $ S'] - equals_fun + thy + "Ordered rewriting of meta equalities" + [Const("op ==",xT) $ S $ S'] + equals_fun fun quant_simproc thy = Simplifier.simproc_i - thy - "Ordered rewriting of nested quantifiers" - [all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0))))] - quant_rewrite + thy + "Ordered rewriting of nested quantifiers" + [all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0))))] + quant_rewrite fun eta_expand_simproc thy = Simplifier.simproc_i - thy - "Smart eta-expansion by equivalences" - [Logic.mk_equals(Q,R)] - eta_expand + thy + "Smart eta-expansion by equivalences" + [Logic.mk_equals(Q,R)] + eta_expand fun eta_contract_simproc thy = Simplifier.simproc_i - thy - "Smart handling of eta-contractions" - [all xT $ (Abs("x",xT,Logic.mk_equals(Q $ Bound 0,R $ Bound 0)))] - eta_contract + thy + "Smart handling of eta-contractions" + [all xT $ (Abs("x",xT,Logic.mk_equals(Q $ Bound 0,R $ Bound 0)))] + eta_contract end (* Disambiguates the names of bound variables in a term, returning t @@ -457,29 +457,29 @@ fun disamb_bound thy t = let - - fun F (t $ u,idx) = - let - val (t',idx') = F (t,idx) - val (u',idx'') = F (u,idx') - in - (t' $ u',idx'') - end - | F (Abs(x,xT,t),idx) = - let - val x' = "x" ^ (LargeInt.toString idx) (* amazing *) - val (t',idx') = F (t,idx+1) - in - (Abs(x',xT,t'),idx') - end - | F arg = arg - val (t',_) = F (t,0) - val ct = cterm_of thy t - val ct' = cterm_of thy t' - val res = transitive (reflexive ct) (reflexive ct') - val _ = message ("disamb_term: " ^ (string_of_thm res)) + + fun F (t $ u,idx) = + let + val (t',idx') = F (t,idx) + val (u',idx'') = F (u,idx') + in + (t' $ u',idx'') + end + | F (Abs(x,xT,t),idx) = + let + val x' = "x" ^ (LargeInt.toString idx) (* amazing *) + val (t',idx') = F (t,idx+1) + in + (Abs(x',xT,t'),idx') + end + | F arg = arg + val (t',_) = F (t,0) + val ct = cterm_of thy t + val ct' = cterm_of thy t' + val res = transitive (reflexive ct) (reflexive ct') + val _ = message ("disamb_term: " ^ (string_of_thm res)) in - res + res end (* Transforms a term t to some normal form t', returning the theorem t @@ -488,25 +488,25 @@ fun norm_term thy t = let - val norms = ShuffleData.get thy - val ss = Simplifier.theory_context thy empty_ss + val norms = ShuffleData.get thy + val ss = Simplifier.theory_context thy empty_ss setmksimps single - addsimps (map (Thm.transfer thy) norms) + addsimps (map (Thm.transfer thy) norms) addsimprocs [quant_simproc thy, eta_expand_simproc thy,eta_contract_simproc thy] - fun chain f th = - let + fun chain f th = + let val rhs = snd (dest_equals (cprop_of th)) - in - transitive th (f rhs) - end - val th = + in + transitive th (f rhs) + end + val th = t |> disamb_bound thy - |> chain (Simplifier.full_rewrite ss) + |> chain (Simplifier.full_rewrite ss) |> chain eta_conversion - |> strip_shyps - val _ = message ("norm_term: " ^ (string_of_thm th)) + |> strip_shyps + val _ = message ("norm_term: " ^ (string_of_thm th)) in - th + th end handle e => (writeln "norm_term internal error"; print_sign_exn thy e) @@ -516,11 +516,11 @@ fun close_thm th = let - val thy = sign_of_thm th - val c = prop_of th - val vars = add_term_frees (c,add_term_vars(c,[])) + val thy = sign_of_thm th + val c = prop_of th + val vars = add_term_frees (c,add_term_vars(c,[])) in - Drule.forall_intr_list (map (cterm_of thy) vars) th + Drule.forall_intr_list (map (cterm_of thy) vars) th end handle e => (writeln "close_thm internal error"; OldGoals.print_exn e) @@ -528,9 +528,9 @@ fun norm_thm thy th = let - val c = prop_of th + val c = prop_of th in - equal_elim (norm_term thy c) th + equal_elim (norm_term thy c) th end (* make_equal thy t u tries to construct the theorem t == u under the @@ -539,43 +539,43 @@ fun make_equal thy t u = let - val t_is_t' = norm_term thy t - val u_is_u' = norm_term thy u - val th = transitive t_is_t' (symmetric u_is_u') - val _ = message ("make_equal: SOME " ^ (string_of_thm th)) + val t_is_t' = norm_term thy t + val u_is_u' = norm_term thy u + val th = transitive t_is_t' (symmetric u_is_u') + val _ = message ("make_equal: SOME " ^ (string_of_thm th)) in - SOME th + SOME th end handle e as THM _ => (message "make_equal: NONE";NONE) - + fun match_consts ignore t (* th *) = let - fun add_consts (Const (c, _), cs) = - if c mem_string ignore - then cs - else insert (op =) c cs - | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs)) - | add_consts (Abs (_, _, t), cs) = add_consts (t, cs) - | add_consts (_, cs) = cs - val t_consts = add_consts(t,[]) + fun add_consts (Const (c, _), cs) = + if c mem_string ignore + then cs + else insert (op =) c cs + | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs)) + | add_consts (Abs (_, _, t), cs) = add_consts (t, cs) + | add_consts (_, cs) = cs + val t_consts = add_consts(t,[]) in fn (name,th) => - let - val th_consts = add_consts(prop_of th,[]) - in - eq_set(t_consts,th_consts) - end + let + val th_consts = add_consts(prop_of th,[]) + in + eq_set(t_consts,th_consts) + end end - + val collect_ignored = fold_rev (fn thm => fn cs => - let - val (lhs,rhs) = Logic.dest_equals (prop_of thm) - val ignore_lhs = term_consts lhs \\ term_consts rhs - val ignore_rhs = term_consts rhs \\ term_consts lhs - in - fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs) - end) + let + val (lhs,rhs) = Logic.dest_equals (prop_of thm) + val ignore_lhs = term_consts lhs \\ term_consts rhs + val ignore_rhs = term_consts rhs \\ term_consts lhs + in + fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs) + end) (* set_prop t thms tries to make a theorem with the proposition t from one of the theorems thms, by shuffling the propositions around. If it @@ -583,67 +583,67 @@ fun set_prop thy t = let - val vars = add_term_frees (t,add_term_vars (t,[])) - val closed_t = Library.foldr (fn (v, body) => + val vars = add_term_frees (t,add_term_vars (t,[])) + val closed_t = Library.foldr (fn (v, body) => let val vT = type_of v in all vT $ (Abs ("x", vT, abstract_over (v, body))) end) (vars, t) - val rew_th = norm_term thy closed_t - val rhs = snd (dest_equals (cprop_of rew_th)) + val rew_th = norm_term thy closed_t + val rhs = snd (dest_equals (cprop_of rew_th)) - val shuffles = ShuffleData.get thy - fun process [] = NONE - | process ((name,th)::thms) = - let - val norm_th = Thm.varifyT (norm_thm thy (close_thm (Thm.transfer thy th))) - val triv_th = trivial rhs - val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th)) - val mod_th = case Seq.pull (bicompose false (*true*) (false,norm_th,0) 1 triv_th) of - SOME(th,_) => SOME th - | NONE => NONE - in - case mod_th of - SOME mod_th => - let - val closed_th = equal_elim (symmetric rew_th) mod_th - in - message ("Shuffler.set_prop succeeded by " ^ name); - SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th) - end - | NONE => process thms - end - handle e as THM _ => process thms + val shuffles = ShuffleData.get thy + fun process [] = NONE + | process ((name,th)::thms) = + let + val norm_th = Thm.varifyT (norm_thm thy (close_thm (Thm.transfer thy th))) + val triv_th = trivial rhs + val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th)) + val mod_th = case Seq.pull (bicompose false (*true*) (false,norm_th,0) 1 triv_th) of + SOME(th,_) => SOME th + | NONE => NONE + in + case mod_th of + SOME mod_th => + let + val closed_th = equal_elim (symmetric rew_th) mod_th + in + message ("Shuffler.set_prop succeeded by " ^ name); + SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th) + end + | NONE => process thms + end + handle e as THM _ => process thms in - fn thms => - case process thms of - res as SOME (name,th) => if (prop_of th) aconv t - then res - else error "Internal error in set_prop" - | NONE => NONE + fn thms => + case process thms of + res as SOME (name,th) => if (prop_of th) aconv t + then res + else error "Internal error in set_prop" + | NONE => NONE end handle e => (writeln "set_prop internal error"; OldGoals.print_exn e) fun find_potential thy t = let - val shuffles = ShuffleData.get thy - val ignored = collect_ignored shuffles [] - val rel_consts = term_consts t \\ ignored - val pot_thms = PureThy.thms_containing_consts thy rel_consts + val shuffles = ShuffleData.get thy + val ignored = collect_ignored shuffles [] + val rel_consts = term_consts t \\ ignored + val pot_thms = PureThy.thms_containing_consts thy rel_consts in - List.filter (match_consts ignored t) pot_thms + List.filter (match_consts ignored t) pot_thms end fun gen_shuffle_tac thy search thms i st = let - val _ = message ("Shuffling " ^ (string_of_thm st)) - val t = List.nth(prems_of st,i-1) - val set = set_prop thy t - fun process_tac thms st = - case set thms of - SOME (_,th) => Seq.of_list (compose (th,i,st)) - | NONE => Seq.empty + val _ = message ("Shuffling " ^ (string_of_thm st)) + val t = List.nth(prems_of st,i-1) + val set = set_prop thy t + fun process_tac thms st = + case set thms of + SOME (_,th) => Seq.of_list (compose (th,i,st)) + | NONE => Seq.empty in - (process_tac thms APPEND (if search - then process_tac (find_potential thy t) - else no_tac)) st + (process_tac thms APPEND (if search + then process_tac (find_potential thy t) + else no_tac)) st end fun shuffle_tac thms i st = @@ -654,29 +654,29 @@ fun shuffle_meth (thms:thm list) ctxt = let - val thy = ProofContext.theory_of ctxt + val thy = ProofContext.theory_of ctxt in - Method.SIMPLE_METHOD' HEADGOAL (gen_shuffle_tac thy false (map (pair "") thms)) + Method.SIMPLE_METHOD' (gen_shuffle_tac thy false (map (pair "") thms)) end fun search_meth ctxt = let - val thy = ProofContext.theory_of ctxt - val prems = Assumption.prems_of ctxt + val thy = ProofContext.theory_of ctxt + val prems = Assumption.prems_of ctxt in - Method.SIMPLE_METHOD' HEADGOAL (gen_shuffle_tac thy true (map (pair "premise") prems)) + Method.SIMPLE_METHOD' (gen_shuffle_tac thy true (map (pair "premise") prems)) end val print_shuffles = ShuffleData.print fun add_shuffle_rule thm thy = let - val shuffles = ShuffleData.get thy + val shuffles = ShuffleData.get thy in - if exists (curry Thm.eq_thm thm) shuffles - then (warning ((string_of_thm thm) ^ " already known to the shuffler"); - thy) - else ShuffleData.put (thm::shuffles) thy + if exists (curry Thm.eq_thm thm) shuffles + then (warning ((string_of_thm thm) ^ " already known to the shuffler"); + thy) + else ShuffleData.put (thm::shuffles) thy end val shuffle_attr = Thm.declaration_attribute (fn th => Context.mapping (add_shuffle_rule th) I); diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Integ/presburger.ML --- a/src/HOL/Integ/presburger.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Integ/presburger.ML Wed Nov 29 15:44:51 2006 +0100 @@ -10,10 +10,9 @@ call the procedure with the parameter abs. *) -signature PRESBURGER = +signature PRESBURGER = sig val presburger_tac : bool -> bool -> int -> tactic - val presburger_method : bool -> bool -> int -> Proof.method val setup : theory -> theory val trace : bool ref end; @@ -25,7 +24,7 @@ fun trace_msg s = if !trace then tracing s else (); (*-----------------------------------------------------------------*) -(*cooper_pp: provefunction for the one-exstance quantifier elimination*) +(*cooper_pp: provefunction for the one-existance quantifier elimination*) (* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*) (*-----------------------------------------------------------------*) @@ -36,35 +35,35 @@ "pred_Pls","pred_Min","pred_1","pred_0", "succ_Pls", "succ_Min", "succ_1", "succ_0", "add_Pls", "add_Min", "add_BIT_0", "add_BIT_10", - "add_BIT_11", "minus_Pls", "minus_Min", "minus_1", - "minus_0", "mult_Pls", "mult_Min", "mult_num1", "mult_num0", + "add_BIT_11", "minus_Pls", "minus_Min", "minus_1", + "minus_0", "mult_Pls", "mult_Min", "mult_num1", "mult_num0", "add_Pls_right", "add_Min_right"]; - val intarithrel = - (map thm ["int_eq_number_of_eq","int_neg_number_of_BIT", - "int_le_number_of_eq","int_iszero_number_of_0", - "int_less_number_of_eq_neg"]) @ - (map (fn s => thm s RS thm "lift_bool") - ["int_iszero_number_of_Pls","int_iszero_number_of_1", - "int_neg_number_of_Min"])@ - (map (fn s => thm s RS thm "nlift_bool") - ["int_nonzero_number_of_Min","int_not_neg_number_of_Pls"]); - + val intarithrel = + (map thm ["int_eq_number_of_eq","int_neg_number_of_BIT", + "int_le_number_of_eq","int_iszero_number_of_0", + "int_less_number_of_eq_neg"]) @ + (map (fn s => thm s RS thm "lift_bool") + ["int_iszero_number_of_Pls","int_iszero_number_of_1", + "int_neg_number_of_Min"])@ + (map (fn s => thm s RS thm "nlift_bool") + ["int_nonzero_number_of_Min","int_not_neg_number_of_Pls"]); + val intarith = map thm ["int_number_of_add_sym", "int_number_of_minus_sym", - "int_number_of_diff_sym", "int_number_of_mult_sym"]; + "int_number_of_diff_sym", "int_number_of_mult_sym"]; val natarith = map thm ["add_nat_number_of", "diff_nat_number_of", - "mult_nat_number_of", "eq_nat_number_of", - "less_nat_number_of"] -val powerarith = - (map thm ["nat_number_of", "zpower_number_of_even", - "zpower_Pls", "zpower_Min"]) @ - [(Tactic.simplify true [thm "zero_eq_Numeral0_nring", - thm "one_eq_Numeral1_nring"] + "mult_nat_number_of", "eq_nat_number_of", + "less_nat_number_of"] +val powerarith = + (map thm ["nat_number_of", "zpower_number_of_even", + "zpower_Pls", "zpower_Min"]) @ + [(Tactic.simplify true [thm "zero_eq_Numeral0_nring", + thm "one_eq_Numeral1_nring"] (thm "zpower_number_of_odd"))] -val comp_arith = binarith @ intarith @ intarithrel @ natarith - @ powerarith @[thm"not_false_eq_true", thm "not_true_eq_false"]; +val comp_arith = binarith @ intarith @ intarithrel @ natarith + @ powerarith @[thm"not_false_eq_true", thm "not_true_eq_false"]; -fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = +fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = let val (xn1,p1) = Syntax.variant_abs (xn,xT,p) in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1) end; @@ -140,7 +139,7 @@ ("Divides.mod", iT --> iT --> iT), ("HOL.plus", iT --> iT --> iT), ("HOL.minus", iT --> iT --> iT), - ("HOL.times", iT --> iT --> iT), + ("HOL.times", iT --> iT --> iT), ("HOL.abs", iT --> iT), ("HOL.uminus", iT --> iT), ("HOL.max", iT --> iT --> iT), @@ -154,7 +153,7 @@ ("Divides.mod", nT --> nT --> nT), ("HOL.plus", nT --> nT --> nT), ("HOL.minus", nT --> nT --> nT), - ("HOL.times", nT --> nT --> nT), + ("HOL.times", nT --> nT --> nT), ("Suc", nT --> nT), ("HOL.max", nT --> nT --> nT), ("HOL.min", nT --> nT --> nT), @@ -186,12 +185,12 @@ fun getfuncs fm = case strip_comb fm of (Free (_, T), ts as _ :: _) => - if body_type T mem [iT, nT] - andalso not (ts = []) andalso forall (null o loose_bnos) ts + if body_type T mem [iT, nT] + andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] else Library.foldl op union ([], map getfuncs ts) | (Var (_, T), ts as _ :: _) => - if body_type T mem [iT, nT] + if body_type T mem [iT, nT] andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] else Library.foldl op union ([], map getfuncs ts) | (Const (s, T), ts) => @@ -202,7 +201,7 @@ | _ => []; -fun abstract_pres sg fm = +fun abstract_pres sg fm = foldr (fn (t, u) => let val T = fastype_of t in all T $ Abs ("x", T, abstract_over (t, u)) end) @@ -214,11 +213,11 @@ It returns true if there is a subterm coresponding to the application of a function on a bounded variable. - Function applications are allowed only for well predefined functions a + Function applications are allowed only for well predefined functions a consts*) fun has_free_funcs fm = case strip_comb fm of - (Free (_, T), ts as _ :: _) => + (Free (_, T), ts as _ :: _) => if (body_type T mem [iT,nT]) andalso (not (T mem [iT,nT])) then true else exists (fn x => x) (map has_free_funcs ts) @@ -233,24 +232,24 @@ (*returns true if the formula is relevant for presburger arithmetic tactic The constants occuring in term t should be a subset of the allowed_consts - There also should be no occurences of application of functions on bounded - variables. Whenever this function will be used, it will be ensured that t - will not contain subterms with function symbols that could have been + There also should be no occurences of application of functions on bounded + variables. Whenever this function will be used, it will be ensured that t + will not contain subterms with function symbols that could have been abstracted over.*) - -fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso + +fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso map (fn i => snd (List.nth (ps, i))) (loose_bnos t) @ map (snd o dest_Free) (term_frees t) @ map (snd o dest_Var) (term_vars t) subset [iT, nT] andalso not (has_free_funcs t); -fun prepare_for_presburger sg q fm = +fun prepare_for_presburger sg q fm = let val ps = Logic.strip_params fm val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) - val _ = if relevant (rev ps) c then () + val _ = if relevant (rev ps) c then () else (trace_msg ("Conclusion is not a presburger term:\n" ^ Sign.string_of_term sg c); raise CooperDec.COOPER) fun mk_all ((s, T), (P,n)) = @@ -275,7 +274,7 @@ (* the presburger tactic*) -(* Parameters : q = flag for quantify ofer free variables ; +(* Parameters : q = flag for quantify ofer free variables ; a = flag for abstracting over function occurences i = subgoal *) @@ -289,18 +288,18 @@ val (t,np,nh) = prepare_for_presburger sg q g' (* Some simpsets for dealing with mod div abs and nat*) val mod_div_simpset = Simplifier.theory_context sg HOL_basic_ss - addsimps [refl,nat_mod_add_eq, nat_mod_add_left_eq, - nat_mod_add_right_eq, int_mod_add_eq, - int_mod_add_right_eq, int_mod_add_left_eq, - nat_div_add_eq, int_div_add_eq, - mod_self, zmod_self, - DIVISION_BY_ZERO_MOD,DIVISION_BY_ZERO_DIV, - ZDIVISION_BY_ZERO_MOD,ZDIVISION_BY_ZERO_DIV, - zdiv_zero,zmod_zero,div_0,mod_0, - zdiv_1,zmod_1,div_1,mod_1, - Suc_plus1] - addsimps add_ac - addsimprocs [cancel_div_mod_proc] + addsimps [refl,nat_mod_add_eq, nat_mod_add_left_eq, + nat_mod_add_right_eq, int_mod_add_eq, + int_mod_add_right_eq, int_mod_add_left_eq, + nat_div_add_eq, int_div_add_eq, + mod_self, zmod_self, + DIVISION_BY_ZERO_MOD,DIVISION_BY_ZERO_DIV, + ZDIVISION_BY_ZERO_MOD,ZDIVISION_BY_ZERO_DIV, + zdiv_zero,zmod_zero,div_0,mod_0, + zdiv_1,zmod_1,div_1,mod_1, + Suc_plus1] + addsimps add_ac + addsimprocs [cancel_div_mod_proc] val simpset0 = HOL_basic_ss addsimps [mod_div_equality', Suc_plus1] addsimps comp_arith @@ -328,47 +327,43 @@ (* The result of the quantifier elimination *) val (th, tac) = case (prop_of pre_thm) of Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => - let val pth = + let val pth = (* If quick_and_dirty then run without proof generation as oracle*) - if !quick_and_dirty + if !quick_and_dirty then presburger_oracle sg (Pattern.eta_long [] t1) (* -assume (cterm_of sg - (HOLogic.mk_Trueprop(HOLogic.mk_eq(t1,CooperDec.integer_qelim (Pattern.eta_long [] t1))))) +assume (cterm_of sg + (HOLogic.mk_Trueprop(HOLogic.mk_eq(t1,CooperDec.integer_qelim (Pattern.eta_long [] t1))))) *) - else tmproof_of_int_qelim sg (Pattern.eta_long [] t1) - in + else tmproof_of_int_qelim sg (Pattern.eta_long [] t1) + in (trace_msg ("calling procedure with term:\n" ^ Sign.string_of_term sg t1); ((pth RS iffD2) RS pre_thm, assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) end | _ => (pre_thm, assm_tac i) - in (rtac (((mp_step nh) o (spec_step np)) th) i + in (rtac (((mp_step nh) o (spec_step np)) th) i THEN tac) st end handle Subscript => no_tac st | CooperDec.COOPER => no_tac st); -fun presburger_args meth = - let val parse_flag = +val presburger_meth = + let val parse_flag = Args.$$$ "no_quantify" >> K (apfst (K false)) || Args.$$$ "no_abs" >> K (apsnd (K false)); in - Method.simple_args - (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> - curry (Library.foldl op |>) (true, true)) - (fn (q,a) => fn _ => meth q a 1) + Method.simple_args + (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> + curry (Library.foldl op |>) (true, true)) + (fn (q,a) => K (Method.SIMPLE_METHOD' (presburger_tac q a))) end; -fun presburger_method q a i = Method.METHOD (fn facts => - Method.insert_tac facts 1 THEN presburger_tac q a i) - val presburger_arith_tac = mk_arith_tactic "presburger" (fn i => fn st => (warning "Trying full Presburger arithmetic ..."; presburger_tac true true i st)); val setup = - Method.add_method ("presburger", - presburger_args presburger_method, + Method.add_method ("presburger", presburger_meth, "decision procedure for Presburger arithmetic") #> arith_tactic_add presburger_arith_tac; diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Isar_examples/Hoare.thy --- a/src/HOL/Isar_examples/Hoare.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Isar_examples/Hoare.thy Wed Nov 29 15:44:51 2006 +0100 @@ -447,7 +447,7 @@ method_setup hoare = {* Method.no_args - (Method.SIMPLE_METHOD' HEADGOAL + (Method.SIMPLE_METHOD' (hoare_tac (simp_tac (HOL_basic_ss addsimps [thm "Record.K_record_apply"] )))) *} "verification condition generator for Hoare logic" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Library/comm_ring.ML --- a/src/HOL/Library/comm_ring.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Library/comm_ring.ML Wed Nov 29 15:44:51 2006 +0100 @@ -127,7 +127,7 @@ end); val comm_ring_meth = - Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD' HEADGOAL (comm_ring_tac ctxt)); + Method.ctxt_args (Method.SIMPLE_METHOD' o comm_ring_tac); val setup = Method.add_method ("comm_ring", comm_ring_meth, diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Nominal/nominal_permeq.ML --- a/src/HOL/Nominal/nominal_permeq.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Nominal/nominal_permeq.ML Wed Nov 29 15:44:51 2006 +0100 @@ -332,7 +332,7 @@ fun simp_meth_setup tac = Method.only_sectioned_args (Simplifier.simp_modifiers' @ Splitter.split_modifiers) - (Method.SIMPLE_METHOD' HEADGOAL o tac o local_simpset_of); + (Method.SIMPLE_METHOD' o tac o local_simpset_of); val perm_eq_meth = simp_meth_setup (perm_simp_tac NO_DEBUG_tac); val perm_eq_meth_debug = simp_meth_setup (perm_simp_tac DEBUG_tac); diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Prolog/HOHH.thy --- a/src/HOL/Prolog/HOHH.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Prolog/HOHH.thy Wed Nov 29 15:44:51 2006 +0100 @@ -11,7 +11,7 @@ begin method_setup ptac = - {* Method.thms_args (Method.SIMPLE_METHOD' HEADGOAL o Prolog.ptac) *} + {* Method.thms_args (Method.SIMPLE_METHOD' o Prolog.ptac) *} "Basic Lambda Prolog interpreter" method_setup prolog = diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Real/ferrante_rackoff.ML --- a/src/HOL/Real/ferrante_rackoff.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Real/ferrante_rackoff.ML Wed Nov 29 15:44:51 2006 +0100 @@ -69,9 +69,9 @@ fun spec_step n th = if n = 0 then th else spec_step (n - 1) th RS spec ; fun mp_step n th = if n = 0 then th else mp_step (n - 1) th RS mp; -local - val context_ss = simpset_of (the_context ()) -in fun ferrack_tac q i = ObjectLogic.atomize_tac i +val context_ss = simpset_of (the_context ()); + +fun ferrack_tac q i = ObjectLogic.atomize_tac i THEN REPEAT_DETERM (split_tac [split_min, split_max,abs_split] i) THEN (fn st => let @@ -103,21 +103,19 @@ in (rtac (((mp_step nh) o (spec_step np)) th) i THEN tac) st end handle Subscript => no_tac st | Ferrante_Rackoff_Proof.FAILURE _ => no_tac st); -end; (*local*) -fun ferrack_args meth = - let - val parse_flag = Args.$$$ "no_quantify" >> (K (K false)); - in - Method.simple_args - (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> - curry (Library.foldl op |>) true) - (fn q => fn _ => meth q 1) +val ferrack_meth = + let + val parse_flag = Args.$$$ "no_quantify" >> (K (K false)); + in + Method.simple_args + (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> + curry (Library.foldl op |>) true) + (fn q => K (Method.SIMPLE_METHOD' (ferrack_tac q))) end; val setup = - Method.add_method ("ferrack", - ferrack_args (Method.SIMPLE_METHOD oo ferrack_tac), - "LCF-proof-producing decision procedure for linear real arithmetic"); + Method.add_method ("ferrack", ferrack_meth, + "LCF-proof-producing decision procedure for linear real arithmetic"); -end +end; diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/SAT.thy --- a/src/HOL/SAT.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/SAT.thy Wed Nov 29 15:44:51 2006 +0100 @@ -17,7 +17,7 @@ begin text {* \medskip Late package setup: default values for refute, see - also theory @{text Refute}. *} + also theory @{theory Refute}. *} refute_params ["itself"=1, @@ -30,10 +30,10 @@ ML {* structure sat = SATFunc(structure cnf = cnf); *} -method_setup sat = {* Method.no_args (Method.SIMPLE_METHOD (sat.sat_tac 1)) *} +method_setup sat = {* Method.no_args (Method.SIMPLE_METHOD' sat.sat_tac) *} "SAT solver" -method_setup satx = {* Method.no_args (Method.SIMPLE_METHOD (sat.satx_tac 1)) *} +method_setup satx = {* Method.no_args (Method.SIMPLE_METHOD' sat.satx_tac) *} "SAT solver (with definitional CNF)" end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/SET-Protocol/EventSET.thy --- a/src/HOL/SET-Protocol/EventSET.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/SET-Protocol/EventSET.thy Wed Nov 29 15:44:51 2006 +0100 @@ -188,8 +188,7 @@ *} method_setup analz_mono_contra = {* - Method.no_args - (Method.METHOD (fn facts => REPEAT_FIRST analz_mono_contra_tac)) *} + Method.no_args (Method.SIMPLE_METHOD (REPEAT_FIRST analz_mono_contra_tac)) *} "for proving theorems of the form X \ analz (knows Spy evs) --> P" end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/SET-Protocol/MessageSET.thy --- a/src/HOL/SET-Protocol/MessageSET.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/SET-Protocol/MessageSET.thy Wed Nov 29 15:44:51 2006 +0100 @@ -935,21 +935,17 @@ method_setup spy_analz = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_spy_analz_tac (local_clasimpset_of ctxt) 1))*} + Method.SIMPLE_METHOD' (gen_spy_analz_tac (local_clasimpset_of ctxt))) *} "for proving the Fake case when analz is involved" method_setup atomic_spy_analz = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - atomic_spy_analz_tac (local_clasimpset_of ctxt) 1))*} + Method.SIMPLE_METHOD' (atomic_spy_analz_tac (local_clasimpset_of ctxt))) *} "for debugging spy_analz" method_setup Fake_insert_simp = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - Fake_insert_simp_tac (local_simpset_of ctxt) 1))*} + Method.SIMPLE_METHOD' (Fake_insert_simp_tac (local_simpset_of ctxt))) *} "for debugging spy_analz" - end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/SET-Protocol/PublicSET.thy --- a/src/HOL/SET-Protocol/PublicSET.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/SET-Protocol/PublicSET.thy Wed Nov 29 15:44:51 2006 +0100 @@ -373,8 +373,7 @@ method_setup possibility = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_possibility_tac (local_simpset_of ctxt))) *} + Method.SIMPLE_METHOD (gen_possibility_tac (local_simpset_of ctxt))) *} "for proving possibility theorems" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Tools/Presburger/presburger.ML --- a/src/HOL/Tools/Presburger/presburger.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Tools/Presburger/presburger.ML Wed Nov 29 15:44:51 2006 +0100 @@ -10,10 +10,9 @@ call the procedure with the parameter abs. *) -signature PRESBURGER = +signature PRESBURGER = sig val presburger_tac : bool -> bool -> int -> tactic - val presburger_method : bool -> bool -> int -> Proof.method val setup : theory -> theory val trace : bool ref end; @@ -25,7 +24,7 @@ fun trace_msg s = if !trace then tracing s else (); (*-----------------------------------------------------------------*) -(*cooper_pp: provefunction for the one-exstance quantifier elimination*) +(*cooper_pp: provefunction for the one-existance quantifier elimination*) (* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*) (*-----------------------------------------------------------------*) @@ -36,35 +35,35 @@ "pred_Pls","pred_Min","pred_1","pred_0", "succ_Pls", "succ_Min", "succ_1", "succ_0", "add_Pls", "add_Min", "add_BIT_0", "add_BIT_10", - "add_BIT_11", "minus_Pls", "minus_Min", "minus_1", - "minus_0", "mult_Pls", "mult_Min", "mult_num1", "mult_num0", + "add_BIT_11", "minus_Pls", "minus_Min", "minus_1", + "minus_0", "mult_Pls", "mult_Min", "mult_num1", "mult_num0", "add_Pls_right", "add_Min_right"]; - val intarithrel = - (map thm ["int_eq_number_of_eq","int_neg_number_of_BIT", - "int_le_number_of_eq","int_iszero_number_of_0", - "int_less_number_of_eq_neg"]) @ - (map (fn s => thm s RS thm "lift_bool") - ["int_iszero_number_of_Pls","int_iszero_number_of_1", - "int_neg_number_of_Min"])@ - (map (fn s => thm s RS thm "nlift_bool") - ["int_nonzero_number_of_Min","int_not_neg_number_of_Pls"]); - + val intarithrel = + (map thm ["int_eq_number_of_eq","int_neg_number_of_BIT", + "int_le_number_of_eq","int_iszero_number_of_0", + "int_less_number_of_eq_neg"]) @ + (map (fn s => thm s RS thm "lift_bool") + ["int_iszero_number_of_Pls","int_iszero_number_of_1", + "int_neg_number_of_Min"])@ + (map (fn s => thm s RS thm "nlift_bool") + ["int_nonzero_number_of_Min","int_not_neg_number_of_Pls"]); + val intarith = map thm ["int_number_of_add_sym", "int_number_of_minus_sym", - "int_number_of_diff_sym", "int_number_of_mult_sym"]; + "int_number_of_diff_sym", "int_number_of_mult_sym"]; val natarith = map thm ["add_nat_number_of", "diff_nat_number_of", - "mult_nat_number_of", "eq_nat_number_of", - "less_nat_number_of"] -val powerarith = - (map thm ["nat_number_of", "zpower_number_of_even", - "zpower_Pls", "zpower_Min"]) @ - [(Tactic.simplify true [thm "zero_eq_Numeral0_nring", - thm "one_eq_Numeral1_nring"] + "mult_nat_number_of", "eq_nat_number_of", + "less_nat_number_of"] +val powerarith = + (map thm ["nat_number_of", "zpower_number_of_even", + "zpower_Pls", "zpower_Min"]) @ + [(Tactic.simplify true [thm "zero_eq_Numeral0_nring", + thm "one_eq_Numeral1_nring"] (thm "zpower_number_of_odd"))] -val comp_arith = binarith @ intarith @ intarithrel @ natarith - @ powerarith @[thm"not_false_eq_true", thm "not_true_eq_false"]; +val comp_arith = binarith @ intarith @ intarithrel @ natarith + @ powerarith @[thm"not_false_eq_true", thm "not_true_eq_false"]; -fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = +fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = let val (xn1,p1) = Syntax.variant_abs (xn,xT,p) in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1) end; @@ -140,7 +139,7 @@ ("Divides.mod", iT --> iT --> iT), ("HOL.plus", iT --> iT --> iT), ("HOL.minus", iT --> iT --> iT), - ("HOL.times", iT --> iT --> iT), + ("HOL.times", iT --> iT --> iT), ("HOL.abs", iT --> iT), ("HOL.uminus", iT --> iT), ("HOL.max", iT --> iT --> iT), @@ -154,7 +153,7 @@ ("Divides.mod", nT --> nT --> nT), ("HOL.plus", nT --> nT --> nT), ("HOL.minus", nT --> nT --> nT), - ("HOL.times", nT --> nT --> nT), + ("HOL.times", nT --> nT --> nT), ("Suc", nT --> nT), ("HOL.max", nT --> nT --> nT), ("HOL.min", nT --> nT --> nT), @@ -186,12 +185,12 @@ fun getfuncs fm = case strip_comb fm of (Free (_, T), ts as _ :: _) => - if body_type T mem [iT, nT] - andalso not (ts = []) andalso forall (null o loose_bnos) ts + if body_type T mem [iT, nT] + andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] else Library.foldl op union ([], map getfuncs ts) | (Var (_, T), ts as _ :: _) => - if body_type T mem [iT, nT] + if body_type T mem [iT, nT] andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] else Library.foldl op union ([], map getfuncs ts) | (Const (s, T), ts) => @@ -202,7 +201,7 @@ | _ => []; -fun abstract_pres sg fm = +fun abstract_pres sg fm = foldr (fn (t, u) => let val T = fastype_of t in all T $ Abs ("x", T, abstract_over (t, u)) end) @@ -214,11 +213,11 @@ It returns true if there is a subterm coresponding to the application of a function on a bounded variable. - Function applications are allowed only for well predefined functions a + Function applications are allowed only for well predefined functions a consts*) fun has_free_funcs fm = case strip_comb fm of - (Free (_, T), ts as _ :: _) => + (Free (_, T), ts as _ :: _) => if (body_type T mem [iT,nT]) andalso (not (T mem [iT,nT])) then true else exists (fn x => x) (map has_free_funcs ts) @@ -233,24 +232,24 @@ (*returns true if the formula is relevant for presburger arithmetic tactic The constants occuring in term t should be a subset of the allowed_consts - There also should be no occurences of application of functions on bounded - variables. Whenever this function will be used, it will be ensured that t - will not contain subterms with function symbols that could have been + There also should be no occurences of application of functions on bounded + variables. Whenever this function will be used, it will be ensured that t + will not contain subterms with function symbols that could have been abstracted over.*) - -fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso + +fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso map (fn i => snd (List.nth (ps, i))) (loose_bnos t) @ map (snd o dest_Free) (term_frees t) @ map (snd o dest_Var) (term_vars t) subset [iT, nT] andalso not (has_free_funcs t); -fun prepare_for_presburger sg q fm = +fun prepare_for_presburger sg q fm = let val ps = Logic.strip_params fm val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) - val _ = if relevant (rev ps) c then () + val _ = if relevant (rev ps) c then () else (trace_msg ("Conclusion is not a presburger term:\n" ^ Sign.string_of_term sg c); raise CooperDec.COOPER) fun mk_all ((s, T), (P,n)) = @@ -275,7 +274,7 @@ (* the presburger tactic*) -(* Parameters : q = flag for quantify ofer free variables ; +(* Parameters : q = flag for quantify ofer free variables ; a = flag for abstracting over function occurences i = subgoal *) @@ -289,18 +288,18 @@ val (t,np,nh) = prepare_for_presburger sg q g' (* Some simpsets for dealing with mod div abs and nat*) val mod_div_simpset = Simplifier.theory_context sg HOL_basic_ss - addsimps [refl,nat_mod_add_eq, nat_mod_add_left_eq, - nat_mod_add_right_eq, int_mod_add_eq, - int_mod_add_right_eq, int_mod_add_left_eq, - nat_div_add_eq, int_div_add_eq, - mod_self, zmod_self, - DIVISION_BY_ZERO_MOD,DIVISION_BY_ZERO_DIV, - ZDIVISION_BY_ZERO_MOD,ZDIVISION_BY_ZERO_DIV, - zdiv_zero,zmod_zero,div_0,mod_0, - zdiv_1,zmod_1,div_1,mod_1, - Suc_plus1] - addsimps add_ac - addsimprocs [cancel_div_mod_proc] + addsimps [refl,nat_mod_add_eq, nat_mod_add_left_eq, + nat_mod_add_right_eq, int_mod_add_eq, + int_mod_add_right_eq, int_mod_add_left_eq, + nat_div_add_eq, int_div_add_eq, + mod_self, zmod_self, + DIVISION_BY_ZERO_MOD,DIVISION_BY_ZERO_DIV, + ZDIVISION_BY_ZERO_MOD,ZDIVISION_BY_ZERO_DIV, + zdiv_zero,zmod_zero,div_0,mod_0, + zdiv_1,zmod_1,div_1,mod_1, + Suc_plus1] + addsimps add_ac + addsimprocs [cancel_div_mod_proc] val simpset0 = HOL_basic_ss addsimps [mod_div_equality', Suc_plus1] addsimps comp_arith @@ -328,47 +327,43 @@ (* The result of the quantifier elimination *) val (th, tac) = case (prop_of pre_thm) of Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => - let val pth = + let val pth = (* If quick_and_dirty then run without proof generation as oracle*) - if !quick_and_dirty + if !quick_and_dirty then presburger_oracle sg (Pattern.eta_long [] t1) (* -assume (cterm_of sg - (HOLogic.mk_Trueprop(HOLogic.mk_eq(t1,CooperDec.integer_qelim (Pattern.eta_long [] t1))))) +assume (cterm_of sg + (HOLogic.mk_Trueprop(HOLogic.mk_eq(t1,CooperDec.integer_qelim (Pattern.eta_long [] t1))))) *) - else tmproof_of_int_qelim sg (Pattern.eta_long [] t1) - in + else tmproof_of_int_qelim sg (Pattern.eta_long [] t1) + in (trace_msg ("calling procedure with term:\n" ^ Sign.string_of_term sg t1); ((pth RS iffD2) RS pre_thm, assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) end | _ => (pre_thm, assm_tac i) - in (rtac (((mp_step nh) o (spec_step np)) th) i + in (rtac (((mp_step nh) o (spec_step np)) th) i THEN tac) st end handle Subscript => no_tac st | CooperDec.COOPER => no_tac st); -fun presburger_args meth = - let val parse_flag = +val presburger_meth = + let val parse_flag = Args.$$$ "no_quantify" >> K (apfst (K false)) || Args.$$$ "no_abs" >> K (apsnd (K false)); in - Method.simple_args - (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> - curry (Library.foldl op |>) (true, true)) - (fn (q,a) => fn _ => meth q a 1) + Method.simple_args + (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> + curry (Library.foldl op |>) (true, true)) + (fn (q,a) => K (Method.SIMPLE_METHOD' (presburger_tac q a))) end; -fun presburger_method q a i = Method.METHOD (fn facts => - Method.insert_tac facts 1 THEN presburger_tac q a i) - val presburger_arith_tac = mk_arith_tactic "presburger" (fn i => fn st => (warning "Trying full Presburger arithmetic ..."; presburger_tac true true i st)); val setup = - Method.add_method ("presburger", - presburger_args presburger_method, + Method.add_method ("presburger", presburger_meth, "decision procedure for Presburger arithmetic") #> arith_tactic_add presburger_arith_tac; diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Tools/function_package/auto_term.ML --- a/src/HOL/Tools/function_package/auto_term.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Tools/function_package/auto_term.ML Wed Nov 29 15:44:51 2006 +0100 @@ -6,34 +6,33 @@ Method "relation" to commence a termination proof using a user-specified relation. *) - signature FUNDEF_RELATION = sig - val relation_meth : Proof.context -> term -> Method.method - + val relation_tac: Proof.context -> term -> int -> tactic val setup: theory -> theory end structure FundefRelation : FUNDEF_RELATION = struct -fun relation_meth ctxt rel = +fun relation_tac ctxt rel = let val cert = Thm.cterm_of (ProofContext.theory_of ctxt) - + val rel' = cert (singleton (Variable.polymorphic ctxt) rel) val rule = FundefCommon.get_termination_rule ctxt |> the |> Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) - + val prem = #1 (Logic.dest_implies (Thm.prop_of rule)) val Rvar = cert (Var (the_single (Term.add_vars prem []))) - in - Method.SIMPLE_METHOD (rtac (Drule.cterm_instantiate [(Rvar, rel')] rule) 1) - end - + in rtac (Drule.cterm_instantiate [(Rvar, rel')] rule) end + +fun relation_meth src = + Method.syntax Args.term src + #> (fn (ctxt, rel) => Method.SIMPLE_METHOD' (relation_tac ctxt rel)) val setup = Method.add_methods - [("relation", uncurry relation_meth oo Method.syntax Args.term, + [("relation", relation_meth, "proves termination using a user-specified wellfounded relation")] end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Tools/function_package/fundef_datatype.ML --- a/src/HOL/Tools/function_package/fundef_datatype.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Tools/function_package/fundef_datatype.ML Wed Nov 29 15:44:51 2006 +0100 @@ -161,20 +161,20 @@ handle COMPLETENESS => Seq.empty -val pat_completeness = Method.SIMPLE_METHOD (pat_complete_tac 1) +val pat_completeness = Method.SIMPLE_METHOD' pat_complete_tac val by_pat_completeness_simp = Proof.global_terminal_proof (Method.Basic (K pat_completeness), - SOME (Method.Source (Args.src (("simp_all", []), Position.none)))) - (* FIXME avoid dynamic scoping of method name! *) + SOME (Method.Source_i (Args.src (("HOL.simp_all", []), Position.none)))) fun termination_by_lexicographic_order name = FundefPackage.setup_termination_proof (SOME name) #> Proof.global_terminal_proof (Method.Basic LexicographicOrder.lexicographic_order, NONE) val setup = - Method.add_methods [("pat_completeness", Method.no_args pat_completeness, "Completeness prover for datatype patterns")] + Method.add_methods [("pat_completeness", Method.no_args pat_completeness, + "Completeness prover for datatype patterns")] diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Tools/meson.ML --- a/src/HOL/Tools/meson.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Tools/meson.ML Wed Nov 29 15:44:51 2006 +0100 @@ -666,15 +666,11 @@ (*No CHANGED_PROP here, since these always appear in the preamble*) -val skolemize_meth = Method.SIMPLE_METHOD' HEADGOAL skolemize_tac; - -val make_clauses_meth = Method.SIMPLE_METHOD' HEADGOAL make_clauses_tac; - val skolemize_setup = Method.add_methods - [("skolemize", Method.no_args skolemize_meth, + [("skolemize", Method.no_args (Method.SIMPLE_METHOD' skolemize_tac), "Skolemization into existential quantifiers"), - ("make_clauses", Method.no_args make_clauses_meth, + ("make_clauses", Method.no_args (Method.SIMPLE_METHOD' make_clauses_tac), "Conversion to !!-quantified meta-level clauses")]; end; diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Tools/res_atp.ML --- a/src/HOL/Tools/res_atp.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Tools/res_atp.ML Wed Nov 29 15:44:51 2006 +0100 @@ -15,7 +15,7 @@ val problem_name: string ref val time_limit: int ref val set_prover: string -> unit - + datatype mode = Auto | Fol | Hol val linkup_logic_mode : mode ref val write_subgoal_file: bool -> mode -> Proof.context -> thm list -> thm list -> int -> string @@ -65,22 +65,22 @@ (********************************************************************) (*** background linkup ***) -val call_atp = ref false; +val call_atp = ref false; val hook_count = ref 0; val time_limit = ref 60; -val prover = ref ""; +val prover = ref ""; -fun set_prover atp = +fun set_prover atp = case String.map Char.toLower atp of - "e" => + "e" => (ReduceAxiomsN.max_new := 100; ReduceAxiomsN.theory_const := false; prover := "E") - | "spass" => + | "spass" => (ReduceAxiomsN.max_new := 40; ReduceAxiomsN.theory_const := true; prover := "spass") - | "vampire" => + | "vampire" => (ReduceAxiomsN.max_new := 60; ReduceAxiomsN.theory_const := false; prover := "vampire") @@ -94,17 +94,17 @@ val problem_name = ref "prob"; (*Return the path to a "helper" like SPASS or tptp2X, first checking that - it exists. FIXME: modify to use Path primitives and move to some central place.*) + it exists. FIXME: modify to use Path primitives and move to some central place.*) fun helper_path evar base = case getenv evar of "" => error ("Isabelle environment variable " ^ evar ^ " not defined") - | home => + | home => let val path = home ^ "/" ^ base - in if File.exists (File.unpack_platform_path path) then path - else error ("Could not find the file " ^ path) - end; + in if File.exists (File.unpack_platform_path path) then path + else error ("Could not find the file " ^ path) + end; -fun probfile_nosuffix _ = +fun probfile_nosuffix _ = if !destdir = "" then File.platform_path (File.tmp_path (Path.basic (!problem_name))) else if File.exists (File.unpack_platform_path (!destdir)) then !destdir ^ "/" ^ !problem_name @@ -118,15 +118,15 @@ val eprover_time = ref 60; val spass_time = ref 60; -fun run_vampire time = +fun run_vampire time = if (time >0) then vampire_time:= time else vampire_time:=60; -fun run_eprover time = +fun run_eprover time = if (time > 0) then eprover_time:= time else eprover_time:=60; -fun run_spass time = +fun run_spass time = if (time > 0) then spass_time:=time else spass_time:=60; @@ -141,24 +141,24 @@ val hol_const_types_only = ResHolClause.const_types_only; val hol_no_types = ResHolClause.no_types; fun hol_typ_level () = ResHolClause.find_typ_level (); -fun is_typed_hol () = +fun is_typed_hol () = let val tp_level = hol_typ_level() in - not (tp_level = ResHolClause.T_NONE) + not (tp_level = ResHolClause.T_NONE) end; fun atp_input_file () = - let val file = !problem_name + let val file = !problem_name in - if !destdir = "" then File.platform_path (File.tmp_path (Path.basic file)) - else if File.exists (File.unpack_platform_path (!destdir)) - then !destdir ^ "/" ^ file - else error ("No such directory: " ^ !destdir) + if !destdir = "" then File.platform_path (File.tmp_path (Path.basic file)) + else if File.exists (File.unpack_platform_path (!destdir)) + then !destdir ^ "/" ^ file + else error ("No such directory: " ^ !destdir) end; val include_all = ref true; val include_simpset = ref false; -val include_claset = ref false; +val include_claset = ref false; val include_atpset = ref true; (*Tests show that follow_defs gives VERY poor results with "include_all"*) @@ -200,7 +200,7 @@ fun upgrade_lg HOLC _ = HOLC | upgrade_lg HOL HOLC = HOLC | upgrade_lg HOL _ = HOL - | upgrade_lg FOL lg = lg; + | upgrade_lg FOL lg = lg; (* check types *) fun has_bool_hfn (Type("bool",_)) = true @@ -211,34 +211,34 @@ fun is_hol_fn tp = let val (targs,tr) = strip_type tp in - exists (has_bool_hfn) (tr::targs) + exists (has_bool_hfn) (tr::targs) end; fun is_hol_pred tp = let val (targs,tr) = strip_type tp in - exists (has_bool_hfn) targs + exists (has_bool_hfn) targs end; exception FN_LG of term; -fun fn_lg (t as Const(f,tp)) (lg,seen) = - if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) - | fn_lg (t as Free(f,tp)) (lg,seen) = - if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) +fun fn_lg (t as Const(f,tp)) (lg,seen) = + if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) + | fn_lg (t as Free(f,tp)) (lg,seen) = + if is_hol_fn tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg, insert (op =) t seen) | fn_lg (t as Var(f,tp)) (lg,seen) = if is_hol_fn tp then (upgrade_lg HOL lg,insert (op =) t seen) else (lg,insert (op =) t seen) | fn_lg (t as Abs(_,_,_)) (lg,seen) = (upgrade_lg HOLC lg,insert (op =) t seen) - | fn_lg f _ = raise FN_LG(f); + | fn_lg f _ = raise FN_LG(f); fun term_lg [] (lg,seen) = (lg,seen) | term_lg (tm::tms) (FOL,seen) = let val (f,args) = strip_comb tm - val (lg',seen') = if f mem seen then (FOL,seen) - else fn_lg f (FOL,seen) + val (lg',seen') = if f mem seen then (FOL,seen) + else fn_lg f (FOL,seen) in - if is_fol_logic lg' then () + if is_fol_logic lg' then () else Output.debug ("Found a HOL term: " ^ Display.raw_string_of_term f); term_lg (args@tms) (lg',seen') end @@ -246,11 +246,11 @@ exception PRED_LG of term; -fun pred_lg (t as Const(P,tp)) (lg,seen)= - if is_hol_pred tp then (upgrade_lg HOL lg, insert (op =) t seen) - else (lg,insert (op =) t seen) +fun pred_lg (t as Const(P,tp)) (lg,seen)= + if is_hol_pred tp then (upgrade_lg HOL lg, insert (op =) t seen) + else (lg,insert (op =) t seen) | pred_lg (t as Free(P,tp)) (lg,seen) = - if is_hol_pred tp then (upgrade_lg HOL lg, insert (op =) t seen) + if is_hol_pred tp then (upgrade_lg HOL lg, insert (op =) t seen) else (lg,insert (op =) t seen) | pred_lg (t as Var(_,_)) (lg,seen) = (upgrade_lg HOL lg, insert (op =) t seen) | pred_lg P _ = raise PRED_LG(P); @@ -259,21 +259,21 @@ fun lit_lg (Const("Not",_) $ P) (lg,seen) = lit_lg P (lg,seen) | lit_lg P (lg,seen) = let val (pred,args) = strip_comb P - val (lg',seen') = if pred mem seen then (lg,seen) - else pred_lg pred (lg,seen) + val (lg',seen') = if pred mem seen then (lg,seen) + else pred_lg pred (lg,seen) in - if is_fol_logic lg' then () - else Output.debug ("Found a HOL predicate: " ^ Display.raw_string_of_term pred); - term_lg args (lg',seen') + if is_fol_logic lg' then () + else Output.debug ("Found a HOL predicate: " ^ Display.raw_string_of_term pred); + term_lg args (lg',seen') end; fun lits_lg [] (lg,seen) = (lg,seen) | lits_lg (lit::lits) (FOL,seen) = let val (lg,seen') = lit_lg lit (FOL,seen) in - if is_fol_logic lg then () - else Output.debug ("Found a HOL literal: " ^ Display.raw_string_of_term lit); - lits_lg lits (lg,seen') + if is_fol_logic lg then () + else Output.debug ("Found a HOL literal: " ^ Display.raw_string_of_term lit); + lits_lg lits (lg,seen') end | lits_lg lits (lg,seen) = (lg,seen); @@ -288,18 +288,18 @@ fun logic_of_clauses [] (lg,seen) = (lg,seen) | logic_of_clauses (cls::clss) (FOL,seen) = let val (lg,seen') = logic_of_clause cls (FOL,seen) - val _ = + val _ = if is_fol_logic lg then () else Output.debug ("Found a HOL clause: " ^ Display.raw_string_of_term cls) in - logic_of_clauses clss (lg,seen') + logic_of_clauses clss (lg,seen') end | logic_of_clauses (cls::clss) (lg,seen) = (lg,seen); fun problem_logic_goals_aux [] (lg,seen) = lg - | problem_logic_goals_aux (subgoal::subgoals) (lg,seen) = + | problem_logic_goals_aux (subgoal::subgoals) (lg,seen) = problem_logic_goals_aux subgoals (logic_of_clauses subgoal (lg,seen)); - + fun problem_logic_goals subgoals = problem_logic_goals_aux subgoals (FOL,[]); fun is_fol_thms ths = ((fst(logic_of_clauses (map prop_of ths) (FOL,[]))) = FOL); @@ -311,9 +311,9 @@ (*** white list and black list of lemmas ***) (*The rule subsetI is frequently omitted by the relevance filter.*) -val whitelist = ref [subsetI]; - -(*Names of theorems (not theorem lists! See multi_blacklist below) to be banned. +val whitelist = ref [subsetI]; + +(*Names of theorems (not theorem lists! See multi_blacklist below) to be banned. These theorems typically produce clauses that are prolific (match too many equality or membership literals) and relate to seldom-used facts. Some duplicate other rules. @@ -447,7 +447,7 @@ (*These might be prolific but are probably OK, and min and max are basic. - "Orderings.max_less_iff_conj", + "Orderings.max_less_iff_conj", "Orderings.min_less_iff_conj", "Orderings.min_max.below_inf.below_inf_conv", "Orderings.min_max.below_sup.above_sup_conv", @@ -463,25 +463,25 @@ exception HASH_CLAUSE and HASH_STRING; (*Catches (for deletion) theorems automatically generated from other theorems*) -fun insert_suffixed_names ht x = - (Polyhash.insert ht (x^"_iff1", ()); - Polyhash.insert ht (x^"_iff2", ()); - Polyhash.insert ht (x^"_dest", ())); +fun insert_suffixed_names ht x = + (Polyhash.insert ht (x^"_iff1", ()); + Polyhash.insert ht (x^"_iff2", ()); + Polyhash.insert ht (x^"_dest", ())); (*Reject theorems with names like "List.filter.filter_list_def" or "Accessible_Part.acc.defs", as these are definitions arising from packages. FIXME: this will also block definitions within locales*) fun is_package_def a = - length (NameSpace.unpack a) > 2 andalso + length (NameSpace.unpack a) > 2 andalso String.isSuffix "_def" a orelse String.isSuffix "_defs" a; -fun make_banned_test xs = +fun make_banned_test xs = let val ht = Polyhash.mkTable (Polyhash.hash_string, op =) (6000, HASH_STRING) - fun banned s = + fun banned s = isSome (Polyhash.peek ht s) orelse is_package_def s in app (fn x => Polyhash.insert ht (x,())) (!blacklist); - app (insert_suffixed_names ht) (!blacklist @ xs); + app (insert_suffixed_names ht) (!blacklist @ xs); banned end; @@ -509,40 +509,40 @@ (*Use a hash table to eliminate duplicates from xs. Argument is a list of (thm * (string * int)) tuples. The theorems are hashed into the table. *) -fun make_unique xs = +fun make_unique xs = let val ht = mk_clause_table 7000 in Output.debug("make_unique gets " ^ Int.toString (length xs) ^ " clauses"); - app (ignore o Polyhash.peekInsert ht) xs; + app (ignore o Polyhash.peekInsert ht) xs; Polyhash.listItems ht end; (*Remove existing axiom clauses from the conjecture clauses, as this can dramatically boost an ATP's performance (for some reason)*) -fun subtract_cls c_clauses ax_clauses = +fun subtract_cls c_clauses ax_clauses = let val ht = mk_clause_table 2200 fun known x = isSome (Polyhash.peek ht x) in - app (ignore o Polyhash.peekInsert ht) ax_clauses; - filter (not o known) c_clauses + app (ignore o Polyhash.peekInsert ht) ax_clauses; + filter (not o known) c_clauses end; -(*Filter axiom clauses, but keep supplied clauses and clauses in whitelist. +(*Filter axiom clauses, but keep supplied clauses and clauses in whitelist. Duplicates are removed later.*) fun get_relevant_clauses thy cls_thms white_cls goals = white_cls @ (ReduceAxiomsN.relevance_filter thy cls_thms goals); fun display_thms [] = () - | display_thms ((name,thm)::nthms) = + | display_thms ((name,thm)::nthms) = let val nthm = name ^ ": " ^ (string_of_thm thm) in Output.debug nthm; display_thms nthms end; - + fun all_valid_thms ctxt = PureThy.thms_containing (ProofContext.theory_of ctxt) ([], []) @ filter (ProofContext.valid_thms ctxt) (FactIndex.find (ProofContext.fact_index_of ctxt) ([], [])); -fun multi_name a (th, (n,pairs)) = +fun multi_name a (th, (n,pairs)) = (n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs) fun add_multi_names_aux ((a, []), pairs) = pairs @@ -557,7 +557,7 @@ ["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm"]; (*Ignore blacklisted theorem lists*) -fun add_multi_names ((a, ths), pairs) = +fun add_multi_names ((a, ths), pairs) = if a mem_string multi_blacklist orelse (Sign.base_name a) mem_string multi_base_blacklist then pairs else add_multi_names_aux ((a, ths), pairs); @@ -565,7 +565,7 @@ fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a; (*The single theorems go BEFORE the multiple ones*) -fun name_thm_pairs ctxt = +fun name_thm_pairs ctxt = let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt) in foldl add_multi_names (foldl add_multi_names [] mults) singles end; @@ -573,40 +573,40 @@ | check_named (_,th) = true; (* get lemmas from claset, simpset, atpset and extra supplied rules *) -fun get_clasimp_atp_lemmas ctxt user_thms = +fun get_clasimp_atp_lemmas ctxt user_thms = let val included_thms = - if !include_all - then (tap (fn ths => Output.debug - ("Including all " ^ Int.toString (length ths) ^ " theorems")) - (name_thm_pairs ctxt)) - else - let val claset_thms = - if !include_claset then ResAxioms.claset_rules_of ctxt - else [] - val simpset_thms = - if !include_simpset then ResAxioms.simpset_rules_of ctxt - else [] - val atpset_thms = - if !include_atpset then ResAxioms.atpset_rules_of ctxt - else [] - val _ = if !Output.show_debug_msgs - then (Output.debug "ATP theorems: "; display_thms atpset_thms) - else () - in claset_thms @ simpset_thms @ atpset_thms end - val user_rules = filter check_named + if !include_all + then (tap (fn ths => Output.debug + ("Including all " ^ Int.toString (length ths) ^ " theorems")) + (name_thm_pairs ctxt)) + else + let val claset_thms = + if !include_claset then ResAxioms.claset_rules_of ctxt + else [] + val simpset_thms = + if !include_simpset then ResAxioms.simpset_rules_of ctxt + else [] + val atpset_thms = + if !include_atpset then ResAxioms.atpset_rules_of ctxt + else [] + val _ = if !Output.show_debug_msgs + then (Output.debug "ATP theorems: "; display_thms atpset_thms) + else () + in claset_thms @ simpset_thms @ atpset_thms end + val user_rules = filter check_named (map (ResAxioms.pairname) - (if null user_thms then !whitelist else user_thms)) + (if null user_thms then !whitelist else user_thms)) in (filter check_named included_thms, user_rules) end; (*Remove lemmas that are banned from the backlist. Also remove duplicates. *) -fun blacklist_filter ths = - if !run_blacklist_filter then +fun blacklist_filter ths = + if !run_blacklist_filter then let val _ = Output.debug("blacklist filter gets " ^ Int.toString (length ths) ^ " theorems") val banned = make_banned_test (map #1 ths) - fun ok (a,_) = not (banned a) - val okthms = filter ok ths + fun ok (a,_) = not (banned a) + val okthms = filter ok ths val _ = Output.debug("...and returns " ^ Int.toString (length okthms)) in okthms end else ths; @@ -653,7 +653,7 @@ (* ATP invocation methods setup *) (***************************************************************) -fun cnf_hyps_thms ctxt = +fun cnf_hyps_thms ctxt = let val ths = Assumption.prems_of ctxt in fold (fold (insert Thm.eq_thm) o ResAxioms.skolem_thm) ths [] end; @@ -664,15 +664,15 @@ (*Ensures that no higher-order theorems "leak out"*) fun restrict_to_logic logic cls = - if is_fol_logic logic then filter (Meson.is_fol_term o prop_of o fst) cls - else cls; + if is_fol_logic logic then filter (Meson.is_fol_term o prop_of o fst) cls + else cls; (**** Predicates to detect unwanted clauses (prolific or likely to cause unsoundness) ****) (** Too general means, positive equality literal with a variable X as one operand, when X does not occur properly in the other operand. This rules out clearly inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **) - + fun occurs ix = let fun occ(Var (jx,_)) = (ix=jx) | occ(t1$t2) = occ t1 orelse occ t2 @@ -685,7 +685,7 @@ (*Unwanted equalities include (1) those between a variable that does not properly occur in the second operand, (2) those between a variable and a record, since these seem to be prolific "cases" thms -*) +*) fun too_general_eqterms (Var (ix,T), t) = not (occurs ix t) orelse is_recordtype T | too_general_eqterms _ = false; @@ -698,7 +698,7 @@ | is_taut _ = false; (*True if the term contains a variable whose (atomic) type is in the given list.*) -fun has_typed_var tycons = +fun has_typed_var tycons = let fun var_tycon (Var (_, Type(a,_))) = a mem_string tycons | var_tycon _ = false in exists var_tycon o term_vars end; @@ -713,72 +713,67 @@ filter (not o unwanted o prop_of o fst) cls; fun tptp_writer logic goals filename (axioms,classrels,arities) user_lemmas = - if is_fol_logic logic + if is_fol_logic logic then ResClause.tptp_write_file goals filename (axioms, classrels, arities) else ResHolClause.tptp_write_file goals filename (axioms, classrels, arities) user_lemmas; fun dfg_writer logic goals filename (axioms,classrels,arities) user_lemmas = - if is_fol_logic logic + if is_fol_logic logic then ResClause.dfg_write_file goals filename (axioms, classrels, arities) else ResHolClause.dfg_write_file goals filename (axioms, classrels, arities) user_lemmas; (*Called by the oracle-based methods declared in res_atp_methods.ML*) fun write_subgoal_file dfg mode ctxt conjectures user_thms n = - let val conj_cls = make_clauses conjectures + let val conj_cls = make_clauses conjectures |> ResAxioms.assume_abstract_list |> Meson.finish_cnf - val hyp_cls = cnf_hyps_thms ctxt - val goal_cls = conj_cls@hyp_cls - val goal_tms = map prop_of goal_cls - val logic = case mode of + val hyp_cls = cnf_hyps_thms ctxt + val goal_cls = conj_cls@hyp_cls + val goal_tms = map prop_of goal_cls + val logic = case mode of Auto => problem_logic_goals [goal_tms] - | Fol => FOL - | Hol => HOL - val (included_thms,user_rules) = get_clasimp_atp_lemmas ctxt user_thms - val cla_simp_atp_clauses = included_thms |> blacklist_filter - |> ResAxioms.cnf_rules_pairs |> make_unique - |> restrict_to_logic logic + | Fol => FOL + | Hol => HOL + val (included_thms,user_rules) = get_clasimp_atp_lemmas ctxt user_thms + val cla_simp_atp_clauses = included_thms |> blacklist_filter + |> ResAxioms.cnf_rules_pairs |> make_unique + |> restrict_to_logic logic |> remove_unwanted_clauses - val user_cls = ResAxioms.cnf_rules_pairs user_rules - val thy = ProofContext.theory_of ctxt - val axclauses = make_unique (get_relevant_clauses thy cla_simp_atp_clauses user_cls goal_tms) - val keep_types = if is_fol_logic logic then !fol_keep_types else is_typed_hol () + val user_cls = ResAxioms.cnf_rules_pairs user_rules + val thy = ProofContext.theory_of ctxt + val axclauses = make_unique (get_relevant_clauses thy cla_simp_atp_clauses user_cls goal_tms) + val keep_types = if is_fol_logic logic then !fol_keep_types else is_typed_hol () val subs = tfree_classes_of_terms goal_tms and axtms = map (prop_of o #1) axclauses val supers = tvar_classes_of_terms axtms and tycons = type_consts_of_terms thy (goal_tms@axtms) (*TFrees in conjecture clauses; TVars in axiom clauses*) - val classrel_clauses = + val classrel_clauses = if keep_types then ResClause.make_classrel_clauses thy subs supers else [] - val arity_clauses = if keep_types then ResClause.arity_clause_thy thy tycons supers else [] - val writer = if dfg then dfg_writer else tptp_writer - and file = atp_input_file() - and user_lemmas_names = map #1 user_rules + val arity_clauses = if keep_types then ResClause.arity_clause_thy thy tycons supers else [] + val writer = if dfg then dfg_writer else tptp_writer + and file = atp_input_file() + and user_lemmas_names = map #1 user_rules in - writer logic goal_cls file (axclauses,classrel_clauses,arity_clauses) user_lemmas_names; - Output.debug ("Writing to " ^ file); - file + writer logic goal_cls file (axclauses,classrel_clauses,arity_clauses) user_lemmas_names; + Output.debug ("Writing to " ^ file); + file end; (**** remove tmp files ****) -fun cond_rm_tmp file = - if !Output.show_debug_msgs orelse !destdir <> "" then Output.debug "ATP input kept..." +fun cond_rm_tmp file = + if !Output.show_debug_msgs orelse !destdir <> "" then Output.debug "ATP input kept..." else OS.FileSys.remove file; (****** setup ATPs as Isabelle methods ******) -fun atp_meth' tac ths ctxt = - Method.SIMPLE_METHOD' HEADGOAL - (tac ctxt ths); -fun atp_meth tac ths ctxt = +fun atp_meth tac ths ctxt = let val thy = ProofContext.theory_of ctxt - val _ = ResClause.init thy - val _ = ResHolClause.init thy - in - atp_meth' tac ths ctxt - end; + val _ = ResClause.init thy + val _ = ResHolClause.init thy + in Method.SIMPLE_METHOD' (tac ctxt ths) end; fun atp_method tac = Method.thms_ctxt_args (atp_meth tac); @@ -802,25 +797,25 @@ let val spass = helper_path "SPASS_HOME" "SPASS" val sopts = "-Auto%-SOS=1%-PGiven=0%-PProblem=0%-Splits=0%-FullRed=0%-DocProof%-TimeLimit=" ^ time - in + in ("spass", spass, sopts, probfile) :: make_atp_list xs (n+1) end else if !prover = "vampire" - then + then let val vampire = helper_path "VAMPIRE_HOME" "vampire" val vopts = "--mode casc%-t " ^ time (*what about -m 100000?*) in ("vampire", vampire, vopts, probfile) :: make_atp_list xs (n+1) end - else if !prover = "E" - then - let val Eprover = helper_path "E_HOME" "eproof" - in - ("E", Eprover, - "--tstp-in%-l5%-xAutoDev%-tAutoDev%--silent%--cpu-limit=" ^ time, probfile) :: - make_atp_list xs (n+1) - end - else error ("Invalid prover name: " ^ !prover) + else if !prover = "E" + then + let val Eprover = helper_path "E_HOME" "eproof" + in + ("E", Eprover, + "--tstp-in%-l5%-xAutoDev%-tAutoDev%--silent%--cpu-limit=" ^ time, probfile) :: + make_atp_list xs (n+1) + end + else error ("Invalid prover name: " ^ !prover) end val atp_list = make_atp_list sg_terms 1 @@ -828,7 +823,7 @@ Watcher.callResProvers(childout,atp_list); Output.debug "Sent commands to watcher!" end - + fun trace_array fname = let val path = File.unpack_platform_path fname in Array.app (File.append path o (fn s => s ^ "\n")) end; @@ -839,7 +834,7 @@ val st = Seq.hd (EVERY' tacs n th) val negs = Option.valOf (metahyps_thms n st) in make_clauses negs |> ResAxioms.assume_abstract_list |> Meson.finish_cnf end; - + (*We write out problem files for each subgoal. Argument probfile generates filenames, and allows the suppression of the suffix "_1" in problem-generation mode. FIXME: does not cope with &&, and it isn't easy because one could have multiple @@ -852,12 +847,12 @@ | get_neg_subgoals (gl::gls) n = neg_clauses th n :: get_neg_subgoals gls (n+1) val goal_cls = get_neg_subgoals goals 1 val logic = case !linkup_logic_mode of - Auto => problem_logic_goals (map ((map prop_of)) goal_cls) - | Fol => FOL - | Hol => HOL + Auto => problem_logic_goals (map ((map prop_of)) goal_cls) + | Fol => FOL + | Hol => HOL val (included_thms,white_thms) = get_clasimp_atp_lemmas ctxt [] val included_cls = included_thms |> blacklist_filter - |> ResAxioms.cnf_rules_pairs |> make_unique + |> ResAxioms.cnf_rules_pairs |> make_unique |> restrict_to_logic logic |> remove_unwanted_clauses val _ = Output.debug ("clauses = " ^ Int.toString(length included_cls)) @@ -866,11 +861,11 @@ val axcls_list = map (fn ngcls => get_relevant_clauses thy included_cls white_cls (map prop_of ngcls)) goal_cls val _ = app (fn axcls => Output.debug ("filtered clauses = " ^ Int.toString(length axcls))) axcls_list - val keep_types = if is_fol_logic logic then !ResClause.keep_types + val keep_types = if is_fol_logic logic then !ResClause.keep_types else is_typed_hol () - val writer = if !prover = "spass" then dfg_writer else tptp_writer + val writer = if !prover = "spass" then dfg_writer else tptp_writer fun write_all [] [] _ = [] - | write_all (ccls::ccls_list) (axcls::axcls_list) k = + | write_all (ccls::ccls_list) (axcls::axcls_list) k = let val fname = probfile k val axcls = make_unique axcls val ccls = subtract_cls ccls axcls @@ -880,7 +875,7 @@ and supers = tvar_classes_of_terms axtms and tycons = type_consts_of_terms thy (ccltms@axtms) (*TFrees in conjecture clauses; TVars in axiom clauses*) - val classrel_clauses = + val classrel_clauses = if keep_types then ResClause.make_classrel_clauses thy subs supers else [] val _ = Output.debug ("classrel clauses = " ^ Int.toString (length classrel_clauses)) @@ -889,7 +884,7 @@ val _ = Output.debug ("arity clauses = " ^ Int.toString (length arity_clauses)) val clnames = writer logic ccls fname (axcls,classrel_clauses,arity_clauses) [] val thm_names = Array.fromList clnames - val _ = if !Output.show_debug_msgs + val _ = if !Output.show_debug_msgs then trace_array (fname ^ "_thm_names") thm_names else () in (thm_names,fname) :: write_all ccls_list axcls_list (k+1) end val (thm_names_list, filenames) = ListPair.unzip (write_all goal_cls axcls_list 1) @@ -897,16 +892,16 @@ (filenames, thm_names_list) end; -val last_watcher_pid = ref (NONE : (TextIO.instream * TextIO.outstream * +val last_watcher_pid = ref (NONE : (TextIO.instream * TextIO.outstream * Posix.Process.pid * string list) option); fun kill_last_watcher () = - (case !last_watcher_pid of + (case !last_watcher_pid of NONE => () - | SOME (_, _, pid, files) => - (Output.debug ("Killing old watcher, pid = " ^ string_of_pid pid); - Watcher.killWatcher pid; - ignore (map (try cond_rm_tmp) files))) + | SOME (_, _, pid, files) => + (Output.debug ("Killing old watcher, pid = " ^ string_of_pid pid); + Watcher.killWatcher pid; + ignore (map (try cond_rm_tmp) files))) handle OS.SysErr _ => Output.debug "Attempt to kill watcher failed"; (*writes out the current clasimpset to a tptp file; @@ -920,7 +915,7 @@ val (childin, childout, pid) = Watcher.createWatcher (th, thm_names_list) in last_watcher_pid := SOME (childin, childout, pid, files); - Output.debug ("problem files: " ^ space_implode ", " files); + Output.debug ("problem files: " ^ space_implode ", " files); Output.debug ("pid: " ^ string_of_pid pid); watcher_call_provers (sign_of_thm th) (Thm.prems_of th) (childin, childout, pid) end; @@ -928,17 +923,17 @@ val isar_atp = setmp print_mode [] isar_atp_body; (*For ML scripts, and primarily, for debugging*) -fun callatp () = +fun callatp () = let val th = topthm() val ctxt = ProofContext.init (theory_of_thm th) in isar_atp_body (ctxt, th) end; -val isar_atp_writeonly = setmp print_mode [] +val isar_atp_writeonly = setmp print_mode [] (fn (ctxt,th) => if Thm.no_prems th then () - else - let val probfile = if Thm.nprems_of th = 1 then probfile_nosuffix - else prob_pathname + else + let val probfile = if Thm.nprems_of th = 1 then probfile_nosuffix + else prob_pathname in ignore (write_problem_files probfile (ctxt,th)) end); @@ -947,9 +942,9 @@ fun invoke_atp_ml (ctxt, goal) = let val thy = ProofContext.theory_of ctxt; in - Output.debug ("subgoals in isar_atp:\n" ^ - Pretty.string_of (ProofContext.pretty_term ctxt - (Logic.mk_conjunction_list (Thm.prems_of goal)))); + Output.debug ("subgoals in isar_atp:\n" ^ + Pretty.string_of (ProofContext.pretty_term ctxt + (Logic.mk_conjunction_list (Thm.prems_of goal)))); Output.debug ("current theory: " ^ Context.theory_name thy); inc hook_count; Output.debug ("in hook for time: " ^ Int.toString (!hook_count)); @@ -965,9 +960,9 @@ in invoke_atp_ml (ctxt, goal) end); val call_atpP = - OuterSyntax.command - "ProofGeneral.call_atp" - "call automatic theorem provers" + OuterSyntax.command + "ProofGeneral.call_atp" + "call automatic theorem provers" OuterKeyword.diag (Scan.succeed invoke_atp); diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/Tools/res_axioms.ML --- a/src/HOL/Tools/res_axioms.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/Tools/res_axioms.ML Wed Nov 29 15:44:51 2006 +0100 @@ -608,14 +608,10 @@ fun cnf_rules_of_ths ths = List.concat (map skolem_use_cache_thm ths); -fun meson_meth ths ctxt = - Method.SIMPLE_METHOD' HEADGOAL - (CHANGED_PROP o Meson.meson_claset_tac (cnf_rules_of_ths ths) HOL_cs); - -val meson_method_setup = - Method.add_methods - [("meson", Method.thms_ctxt_args meson_meth, - "MESON resolution proof procedure")]; +val meson_method_setup = Method.add_methods + [("meson", Method.thms_args (fn ths => + Method.SIMPLE_METHOD' (CHANGED_PROP o Meson.meson_claset_tac (cnf_rules_of_ths ths) HOL_cs)), + "MESON resolution proof procedure")]; (** Attribute for converting a theorem into clauses **) diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/UNITY/Simple/NSP_Bad.thy --- a/src/HOL/UNITY/Simple/NSP_Bad.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/UNITY/Simple/NSP_Bad.thy Wed Nov 29 15:44:51 2006 +0100 @@ -129,8 +129,7 @@ method_setup ns_induct = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - ns_induct_tac (local_clasimpset_of ctxt) 1)) *} + Method.SIMPLE_METHOD' (ns_induct_tac (local_clasimpset_of ctxt))) *} "for inductive reasoning about the Needham-Schroeder protocol" text{*Converts invariants into statements about reachable states*} diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/UNITY/UNITY_Main.thy --- a/src/HOL/UNITY/UNITY_Main.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/UNITY/UNITY_Main.thy Wed Nov 29 15:44:51 2006 +0100 @@ -11,13 +11,12 @@ method_setup safety = {* Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_constrains_tac (local_clasimpset_of ctxt) 1)) *} + Method.SIMPLE_METHOD' (gen_constrains_tac (local_clasimpset_of ctxt))) *} "for proving safety properties" method_setup ensures_tac = {* fn args => fn ctxt => - Method.goal_args' (Scan.lift Args.name) + Method.goal_args' (Scan.lift Args.name) (gen_ensures_tac (local_clasimpset_of ctxt)) args ctxt *} "for proving progress properties" diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/ex/Reflection.thy --- a/src/HOL/ex/Reflection.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/ex/Reflection.thy Wed Nov 29 15:44:51 2006 +0100 @@ -18,14 +18,14 @@ fn src => Method.syntax (Attrib.thms -- Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")"))) src #> - (fn (ctxt, (eqs,to)) => Method.SIMPLE_METHOD' HEADGOAL (Reflection.genreify_tac ctxt eqs to)) + (fn (ctxt, (eqs,to)) => Method.SIMPLE_METHOD' (Reflection.genreify_tac ctxt eqs to)) *} "partial automatic reification" method_setup reflection = {* fn src => Method.syntax (Attrib.thms -- Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")"))) src #> - (fn (ctxt, (ths,to)) => Method.SIMPLE_METHOD' HEADGOAL (Reflection.reflection_tac ctxt ths to)) + (fn (ctxt, (ths,to)) => Method.SIMPLE_METHOD' (Reflection.reflection_tac ctxt ths to)) *} "reflection method" end diff -r a3561bfe0ada -r cd0dc678a205 src/HOL/ex/SAT_Examples.thy --- a/src/HOL/ex/SAT_Examples.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/HOL/ex/SAT_Examples.thy Wed Nov 29 15:44:51 2006 +0100 @@ -82,7 +82,7 @@ ML {* reset sat.trace_sat; *} ML {* reset quick_and_dirty; *} -method_setup rawsat = {* Method.no_args (Method.SIMPLE_METHOD (sat.rawsat_tac 1)) *} +method_setup rawsat = {* Method.no_args (Method.SIMPLE_METHOD' sat.rawsat_tac) *} "SAT solver (no preprocessing)" (* ML {* Toplevel.profiling := 1; *} *) diff -r a3561bfe0ada -r cd0dc678a205 src/Provers/eqsubst.ML --- a/src/Provers/eqsubst.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/Provers/eqsubst.ML Wed Nov 29 15:44:51 2006 +0100 @@ -407,9 +407,7 @@ (* inthms are the given arguments in Isar, and treated as eqstep with the first one, then the second etc *) fun eqsubst_meth ctxt occL inthms = - Method.METHOD - (fn facts => - HEADGOAL (Method.insert_tac facts THEN' eqsubst_tac ctxt occL inthms)); + Method.SIMPLE_METHOD' (eqsubst_tac ctxt occL inthms); (* apply a substitution inside assumption j, keeps asm in the same place *) fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) = @@ -523,9 +521,7 @@ (* inthms are the given arguments in Isar, and treated as eqstep with the first one, then the second etc *) fun eqsubst_asm_meth ctxt occL inthms = - Method.METHOD - (fn facts => - HEADGOAL (Method.insert_tac facts THEN' eqsubst_asm_tac ctxt occL inthms )); + Method.SIMPLE_METHOD' (eqsubst_asm_tac ctxt occL inthms); (* syntax for options, given "(asm)" will give back true, without gives back false *) diff -r a3561bfe0ada -r cd0dc678a205 src/Provers/hypsubst.ML --- a/src/Provers/hypsubst.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/Provers/hypsubst.ML Wed Nov 29 15:44:51 2006 +0100 @@ -224,18 +224,12 @@ in CHANGED_GOAL (rtac (th' RS ssubst)) end; -(* proof methods *) - -val subst_meth = Method.thm_args (Method.SIMPLE_METHOD' HEADGOAL o stac); -val hyp_subst_meth = - Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o hyp_subst_tac)); - - (* theory setup *) val hypsubst_setup = Method.add_methods - [("hypsubst", hyp_subst_meth, "substitution using an assumption (improper)"), - ("simplesubst", subst_meth, "simple substitution")]; + [("hypsubst", Method.no_args (Method.SIMPLE_METHOD' (CHANGED_PROP o hyp_subst_tac)), + "substitution using an assumption (improper)"), + ("simplesubst", Method.thm_args (Method.SIMPLE_METHOD' o stac), "simple substitution")]; end; diff -r a3561bfe0ada -r cd0dc678a205 src/Provers/splitter.ML --- a/src/Provers/splitter.ML Wed Nov 29 15:44:46 2006 +0100 +++ b/src/Provers/splitter.ML Wed Nov 29 15:44:51 2006 +0100 @@ -477,7 +477,7 @@ fun split_meth src = Method.syntax Attrib.thms src - #> (fn (_, ths) => Method.SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o gen_split_tac ths)); + #> (fn (_, ths) => Method.SIMPLE_METHOD' (CHANGED_PROP o gen_split_tac ths)); (* theory setup *) diff -r a3561bfe0ada -r cd0dc678a205 src/Sequents/ILL.thy --- a/src/Sequents/ILL.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/Sequents/ILL.thy Wed Nov 29 15:44:51 2006 +0100 @@ -146,8 +146,8 @@ *} method_setup best_lazy = -{* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (best_tac lazy_cs)) *} -"lazy classical reasoning" + {* Method.no_args (Method.SIMPLE_METHOD' (best_tac lazy_cs)) *} + "lazy classical reasoning" lemma aux_impl: "$F, $G |- A ==> $F, !(A -o B), $G |- B" apply (rule derelict) @@ -282,10 +282,10 @@ method_setup best_safe = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (best_tac safe_cs)) *} "" + {* Method.no_args (Method.SIMPLE_METHOD' (best_tac safe_cs)) *} "" method_setup best_power = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (best_tac power_cs)) *} "" + {* Method.no_args (Method.SIMPLE_METHOD' (best_tac power_cs)) *} "" (* Some examples from Troelstra and van Dalen *) diff -r a3561bfe0ada -r cd0dc678a205 src/Sequents/LK0.thy --- a/src/Sequents/LK0.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/Sequents/LK0.thy Wed Nov 29 15:44:51 2006 +0100 @@ -248,23 +248,23 @@ *} method_setup fast_prop = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (fast_tac prop_pack)) *} + {* Method.no_args (Method.SIMPLE_METHOD' (fast_tac prop_pack)) *} "propositional reasoning" method_setup fast = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (fast_tac LK_pack)) *} + {* Method.no_args (Method.SIMPLE_METHOD' (fast_tac LK_pack)) *} "classical reasoning" method_setup fast_dup = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (fast_tac LK_dup_pack)) *} + {* Method.no_args (Method.SIMPLE_METHOD' (fast_tac LK_dup_pack)) *} "classical reasoning" method_setup best = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (best_tac LK_pack)) *} + {* Method.no_args (Method.SIMPLE_METHOD' (best_tac LK_pack)) *} "classical reasoning" method_setup best_dup = - {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (best_tac LK_dup_pack)) *} + {* Method.no_args (Method.SIMPLE_METHOD' (best_tac LK_dup_pack)) *} "classical reasoning" diff -r a3561bfe0ada -r cd0dc678a205 src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy Wed Nov 29 15:44:46 2006 +0100 +++ b/src/ZF/UNITY/Constrains.thy Wed Nov 29 15:44:51 2006 +0100 @@ -565,11 +565,9 @@ *} method_setup safety = {* - Method.ctxt_args (fn ctxt => - Method.METHOD (fn facts => - gen_constrains_tac (local_clasimpset_of ctxt) 1)) *} - "for proving safety properties" + Method.ctxt_args (fn ctxt => + Method.SIMPLE_METHOD' (gen_constrains_tac (local_clasimpset_of ctxt))) *} + "for proving safety properties" end -