# HG changeset patch # User boehmes # Date 1254564340 -7200 # Node ID a226f29d4bdcc0947a332906febaccde61238001 # Parent 7c8bc41ce810df5a3bb3f49fb1b52b4504649c2c re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values, eliminated unused provers, turned references into configuration values diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/ATP_Linkup.thy --- a/src/HOL/ATP_Linkup.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/ATP_Linkup.thy Sat Oct 03 12:05:40 2009 +0200 @@ -96,29 +96,26 @@ use "Tools/res_reconstruct.ML" setup ResReconstruct.setup use "Tools/res_atp.ML" +use "Tools/ATP_Manager/atp_wrapper.ML" setup AtpWrapper.setup use "Tools/ATP_Manager/atp_manager.ML" -use "Tools/ATP_Manager/atp_wrapper.ML" use "Tools/ATP_Manager/atp_minimal.ML" text {* basic provers *} -setup {* AtpManager.add_prover "spass" AtpWrapper.spass *} -setup {* AtpManager.add_prover "vampire" AtpWrapper.vampire *} -setup {* AtpManager.add_prover "e" AtpWrapper.eprover *} +setup {* AtpManager.add_prover AtpWrapper.spass *} +setup {* AtpManager.add_prover AtpWrapper.vampire *} +setup {* AtpManager.add_prover AtpWrapper.eprover *} text {* provers with stuctured output *} -setup {* AtpManager.add_prover "vampire_full" AtpWrapper.vampire_full *} -setup {* AtpManager.add_prover "e_full" AtpWrapper.eprover_full *} +setup {* AtpManager.add_prover AtpWrapper.vampire_full *} +setup {* AtpManager.add_prover AtpWrapper.eprover_full *} text {* on some problems better results *} -setup {* AtpManager.add_prover "spass_no_tc" (AtpWrapper.spass_opts 40 false) *} +setup {* AtpManager.add_prover AtpWrapper.spass_no_tc *} text {* remote provers via SystemOnTPTP *} -setup {* AtpManager.add_prover "remote_vampire" - (AtpWrapper.remote_prover_opts 60 false "" "Vampire---9") *} -setup {* AtpManager.add_prover "remote_spass" - (AtpWrapper.remote_prover_opts 40 true "-x" "SPASS---") *} -setup {* AtpManager.add_prover "remote_e" - (AtpWrapper.remote_prover_opts 100 false "" "EP---") *} +setup {* AtpManager.add_prover AtpWrapper.remote_vampire *} +setup {* AtpManager.add_prover AtpWrapper.remote_spass *} +setup {* AtpManager.add_prover AtpWrapper.remote_eprover *} diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/MetisExamples/Abstraction.thy --- a/src/HOL/MetisExamples/Abstraction.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/MetisExamples/Abstraction.thy Sat Oct 03 12:05:40 2009 +0200 @@ -22,7 +22,7 @@ pset :: "'a set => 'a set" order :: "'a set => ('a * 'a) set" -ML{*AtpWrapper.problem_name := "Abstraction__Collect_triv"*} +declare [[ atp_problem_prefix = "Abstraction__Collect_triv" ]] lemma (*Collect_triv:*) "a \ {x. P x} ==> P a" proof (neg_clausify) assume 0: "(a\'a\type) \ Collect (P\'a\type \ bool)" @@ -37,12 +37,12 @@ by (metis mem_Collect_eq) -ML{*AtpWrapper.problem_name := "Abstraction__Collect_mp"*} +declare [[ atp_problem_prefix = "Abstraction__Collect_mp" ]] lemma "a \ {x. P x --> Q x} ==> a \ {x. P x} ==> a \ {x. Q x}" by (metis CollectI Collect_imp_eq ComplD UnE mem_Collect_eq); --{*34 secs*} -ML{*AtpWrapper.problem_name := "Abstraction__Sigma_triv"*} +declare [[ atp_problem_prefix = "Abstraction__Sigma_triv" ]] lemma "(a,b) \ Sigma A B ==> a \ A & b \ B a" proof (neg_clausify) assume 0: "(a\'a\type, b\'b\type) \ Sigma (A\'a\type set) (B\'a\type \ 'b\type set)" @@ -60,7 +60,7 @@ lemma Sigma_triv: "(a,b) \ Sigma A B ==> a \ A & b \ B a" by (metis SigmaD1 SigmaD2) -ML{*AtpWrapper.problem_name := "Abstraction__Sigma_Collect"*} +declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect" ]] lemma "(a,b) \ (SIGMA x: A. {y. x = f y}) ==> a \ A & a = f b" (*???metis says this is satisfiable! by (metis CollectD SigmaD1 SigmaD2) @@ -112,7 +112,7 @@ qed -ML{*AtpWrapper.problem_name := "Abstraction__CLF_eq_in_pp"*} +declare [[ atp_problem_prefix = "Abstraction__CLF_eq_in_pp" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL.{f. f \ pset cl}) ==> f \ pset cl" by (metis Collect_mem_eq SigmaD2) @@ -131,7 +131,7 @@ by (metis 5 0) qed -ML{*AtpWrapper.problem_name := "Abstraction__Sigma_Collect_Pi"*} +declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]] lemma "(cl,f) \ (SIGMA cl: CL. {f. f \ pset cl \ pset cl}) ==> f \ pset cl \ pset cl" @@ -147,7 +147,7 @@ qed -ML{*AtpWrapper.problem_name := "Abstraction__Sigma_Collect_Int"*} +declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Int" ]] lemma "(cl,f) \ (SIGMA cl: CL. {f. f \ pset cl \ cl}) ==> f \ pset cl \ cl" @@ -170,13 +170,13 @@ qed -ML{*AtpWrapper.problem_name := "Abstraction__Sigma_Collect_Pi_mono"*} +declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]] lemma "(cl,f) \ (SIGMA cl: CL. {f. f \ pset cl \ pset cl & monotone f (pset cl) (order cl)}) ==> (f \ pset cl \ pset cl) & (monotone f (pset cl) (order cl))" by auto -ML{*AtpWrapper.problem_name := "Abstraction__CLF_subset_Collect_Int"*} +declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]] lemma "(cl,f) \ CLF ==> CLF \ (SIGMA cl: CL. {f. f \ pset cl \ cl}) ==> f \ pset cl \ cl" @@ -187,7 +187,7 @@ --{*@{text Int_def} is redundant*} *) -ML{*AtpWrapper.problem_name := "Abstraction__CLF_eq_Collect_Int"*} +declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL. {f. f \ pset cl \ cl}) ==> f \ pset cl \ cl" @@ -196,7 +196,7 @@ by (metis Collect_mem_eq Int_commute SigmaD2) *) -ML{*AtpWrapper.problem_name := "Abstraction__CLF_subset_Collect_Pi"*} +declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]] lemma "(cl,f) \ CLF ==> CLF \ (SIGMA cl': CL. {f. f \ pset cl' \ pset cl'}) ==> @@ -206,7 +206,7 @@ by (metis Collect_mem_eq SigmaD2 subsetD) *) -ML{*AtpWrapper.problem_name := "Abstraction__CLF_eq_Collect_Pi"*} +declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL. {f. f \ pset cl \ pset cl}) ==> @@ -216,21 +216,21 @@ by (metis Collect_mem_eq SigmaD2 contra_subsetD equalityE) *) -ML{*AtpWrapper.problem_name := "Abstraction__CLF_eq_Collect_Pi_mono"*} +declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL. {f. f \ pset cl \ pset cl & monotone f (pset cl) (order cl)}) ==> (f \ pset cl \ pset cl) & (monotone f (pset cl) (order cl))" by auto -ML{*AtpWrapper.problem_name := "Abstraction__map_eq_zipA"*} +declare [[ atp_problem_prefix = "Abstraction__map_eq_zipA" ]] lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)" apply (induct xs) (*sledgehammer*) apply auto done -ML{*AtpWrapper.problem_name := "Abstraction__map_eq_zipB"*} +declare [[ atp_problem_prefix = "Abstraction__map_eq_zipB" ]] lemma "map (%w. (w -> w, w \ w)) xs = zip (map (%w. w -> w) xs) (map (%w. w \ w) xs)" apply (induct xs) @@ -238,28 +238,28 @@ apply auto done -ML{*AtpWrapper.problem_name := "Abstraction__image_evenA"*} +declare [[ atp_problem_prefix = "Abstraction__image_evenA" ]] lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\x. even x --> Suc(f x) \ A)"; (*sledgehammer*) by auto -ML{*AtpWrapper.problem_name := "Abstraction__image_evenB"*} +declare [[ atp_problem_prefix = "Abstraction__image_evenB" ]] lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A ==> (\x. even x --> f (f (Suc(f x))) \ A)"; (*sledgehammer*) by auto -ML{*AtpWrapper.problem_name := "Abstraction__image_curry"*} +declare [[ atp_problem_prefix = "Abstraction__image_curry" ]] lemma "f \ (%u v. b \ u \ v) ` A ==> \u v. P (b \ u \ v) ==> P(f y)" (*sledgehammer*) by auto -ML{*AtpWrapper.problem_name := "Abstraction__image_TimesA"*} +declare [[ atp_problem_prefix = "Abstraction__image_TimesA" ]] lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A \ B) = (f`A) \ (g`B)" (*sledgehammer*) apply (rule equalityI) (***Even the two inclusions are far too difficult -ML{*AtpWrapper.problem_name := "Abstraction__image_TimesA_simpler"*} +using [[ atp_problem_prefix = "Abstraction__image_TimesA_simpler"]] ***) apply (rule subsetI) apply (erule imageE) @@ -282,13 +282,13 @@ (*Given the difficulty of the previous problem, these two are probably impossible*) -ML{*AtpWrapper.problem_name := "Abstraction__image_TimesB"*} +declare [[ atp_problem_prefix = "Abstraction__image_TimesB" ]] lemma image_TimesB: "(%(x,y,z). (f x, g y, h z)) ` (A \ B \ C) = (f`A) \ (g`B) \ (h`C)" (*sledgehammer*) by force -ML{*AtpWrapper.problem_name := "Abstraction__image_TimesC"*} +declare [[ atp_problem_prefix = "Abstraction__image_TimesC" ]] lemma image_TimesC: "(%(x,y). (x \ x, y \ y)) ` (A \ B) = ((%x. x \ x) ` A) \ ((%y. y \ y) ` B)" diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/MetisExamples/BT.thy --- a/src/HOL/MetisExamples/BT.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/MetisExamples/BT.thy Sat Oct 03 12:05:40 2009 +0200 @@ -65,21 +65,21 @@ text {* \medskip BT simplification *} -ML {*AtpWrapper.problem_name := "BT__n_leaves_reflect"*} +declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]] lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t" apply (induct t) apply (metis add_right_cancel n_leaves.simps(1) reflect.simps(1)) apply (metis add_commute n_leaves.simps(2) reflect.simps(2)) done -ML {*AtpWrapper.problem_name := "BT__n_nodes_reflect"*} +declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]] lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t" apply (induct t) apply (metis reflect.simps(1)) apply (metis n_nodes.simps(2) nat_add_commute reflect.simps(2)) done -ML {*AtpWrapper.problem_name := "BT__depth_reflect"*} +declare [[ atp_problem_prefix = "BT__depth_reflect" ]] lemma depth_reflect: "depth (reflect t) = depth t" apply (induct t) apply (metis depth.simps(1) reflect.simps(1)) @@ -90,21 +90,21 @@ The famous relationship between the numbers of leaves and nodes. *} -ML {*AtpWrapper.problem_name := "BT__n_leaves_nodes"*} +declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]] lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)" apply (induct t) apply (metis n_leaves.simps(1) n_nodes.simps(1)) apply auto done -ML {*AtpWrapper.problem_name := "BT__reflect_reflect_ident"*} +declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]] lemma reflect_reflect_ident: "reflect (reflect t) = t" apply (induct t) apply (metis add_right_cancel reflect.simps(1)); apply (metis reflect.simps(2)) done -ML {*AtpWrapper.problem_name := "BT__bt_map_ident"*} +declare [[ atp_problem_prefix = "BT__bt_map_ident" ]] lemma bt_map_ident: "bt_map (%x. x) = (%y. y)" apply (rule ext) apply (induct_tac y) @@ -115,7 +115,7 @@ done -ML {*AtpWrapper.problem_name := "BT__bt_map_appnd"*} +declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] lemma bt_map_appnd: "bt_map f (appnd t u) = appnd (bt_map f t) (bt_map f u)" apply (induct t) apply (metis appnd.simps(1) bt_map.simps(1)) @@ -123,7 +123,7 @@ done -ML {*AtpWrapper.problem_name := "BT__bt_map_compose"*} +declare [[ atp_problem_prefix = "BT__bt_map_compose" ]] lemma bt_map_compose: "bt_map (f o g) t = bt_map f (bt_map g t)" apply (induct t) apply (metis bt_map.simps(1)) @@ -133,42 +133,42 @@ done -ML {*AtpWrapper.problem_name := "BT__bt_map_reflect"*} +declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]] lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)" apply (induct t) apply (metis add_right_cancel bt_map.simps(1) reflect.simps(1)) apply (metis add_right_cancel bt_map.simps(2) reflect.simps(2)) done -ML {*AtpWrapper.problem_name := "BT__preorder_bt_map"*} +declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]] lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)" apply (induct t) apply (metis bt_map.simps(1) map.simps(1) preorder.simps(1)) apply simp done -ML {*AtpWrapper.problem_name := "BT__inorder_bt_map"*} +declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]] lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)" apply (induct t) apply (metis bt_map.simps(1) inorder.simps(1) map.simps(1)) apply simp done -ML {*AtpWrapper.problem_name := "BT__postorder_bt_map"*} +declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]] lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)" apply (induct t) apply (metis bt_map.simps(1) map.simps(1) postorder.simps(1)) apply simp done -ML {*AtpWrapper.problem_name := "BT__depth_bt_map"*} +declare [[ atp_problem_prefix = "BT__depth_bt_map" ]] lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t" apply (induct t) apply (metis bt_map.simps(1) depth.simps(1)) apply simp done -ML {*AtpWrapper.problem_name := "BT__n_leaves_bt_map"*} +declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]] lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t" apply (induct t) apply (metis One_nat_def Suc_eq_plus1 bt_map.simps(1) less_add_one less_antisym linorder_neq_iff n_leaves.simps(1)) @@ -176,21 +176,21 @@ done -ML {*AtpWrapper.problem_name := "BT__preorder_reflect"*} +declare [[ atp_problem_prefix = "BT__preorder_reflect" ]] lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)" apply (induct t) apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev_is_Nil_conv) apply (metis append_Nil Cons_eq_append_conv postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append rev_rev_ident) done -ML {*AtpWrapper.problem_name := "BT__inorder_reflect"*} +declare [[ atp_problem_prefix = "BT__inorder_reflect" ]] lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)" apply (induct t) apply (metis inorder.simps(1) reflect.simps(1) rev.simps(1)) apply simp done -ML {*AtpWrapper.problem_name := "BT__postorder_reflect"*} +declare [[ atp_problem_prefix = "BT__postorder_reflect" ]] lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)" apply (induct t) apply (metis postorder.simps(1) preorder.simps(1) reflect.simps(1) rev.simps(1)) @@ -201,7 +201,7 @@ Analogues of the standard properties of the append function for lists. *} -ML {*AtpWrapper.problem_name := "BT__appnd_assoc"*} +declare [[ atp_problem_prefix = "BT__appnd_assoc" ]] lemma appnd_assoc [simp]: "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)" apply (induct t1) @@ -209,14 +209,14 @@ apply (metis appnd.simps(2)) done -ML {*AtpWrapper.problem_name := "BT__appnd_Lf2"*} +declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]] lemma appnd_Lf2 [simp]: "appnd t Lf = t" apply (induct t) apply (metis appnd.simps(1)) apply (metis appnd.simps(2)) done -ML {*AtpWrapper.problem_name := "BT__depth_appnd"*} +declare [[ atp_problem_prefix = "BT__depth_appnd" ]] declare max_add_distrib_left [simp] lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2" apply (induct t1) @@ -224,7 +224,7 @@ apply (simp add: ); done -ML {*AtpWrapper.problem_name := "BT__n_leaves_appnd"*} +declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]] lemma n_leaves_appnd [simp]: "n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2" apply (induct t1) @@ -232,7 +232,7 @@ apply (simp add: left_distrib) done -ML {*AtpWrapper.problem_name := "BT__bt_map_appnd"*} +declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] lemma (*bt_map_appnd:*) "bt_map f (appnd t1 t2) = appnd (bt_map f t1) (bt_map f t2)" apply (induct t1) diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/MetisExamples/BigO.thy --- a/src/HOL/MetisExamples/BigO.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/MetisExamples/BigO.thy Sat Oct 03 12:05:40 2009 +0200 @@ -15,7 +15,7 @@ definition bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set" ("(1O'(_'))") where "O(f::('a => 'b)) == {h. EX c. ALL x. abs (h x) <= c * abs (f x)}" -ML_command{*AtpWrapper.problem_name := "BigO__bigo_pos_const"*} +declare [[ atp_problem_prefix = "BigO__bigo_pos_const" ]] lemma bigo_pos_const: "(EX (c::'a::ordered_idom). ALL x. (abs (h x)) <= (c * (abs (f x)))) = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" @@ -212,7 +212,7 @@ {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}" by (auto simp add: bigo_def bigo_pos_const) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_elt_subset"*} +declare [[ atp_problem_prefix = "BigO__bigo_elt_subset" ]] lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)" apply (auto simp add: bigo_alt_def) apply (rule_tac x = "ca * c" in exI) @@ -230,7 +230,7 @@ done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_refl"*} +declare [[ atp_problem_prefix = "BigO__bigo_refl" ]] lemma bigo_refl [intro]: "f : O(f)" apply(auto simp add: bigo_def) proof (neg_clausify) @@ -244,7 +244,7 @@ by (metis 0 2) qed -ML_command{*AtpWrapper.problem_name := "BigO__bigo_zero"*} +declare [[ atp_problem_prefix = "BigO__bigo_zero" ]] lemma bigo_zero: "0 : O(g)" apply (auto simp add: bigo_def func_zero) proof (neg_clausify) @@ -334,7 +334,7 @@ apply (auto del: subsetI simp del: bigo_plus_idemp) done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_plus_eq"*} +declare [[ atp_problem_prefix = "BigO__bigo_plus_eq" ]] lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> O(f + g) = O(f) \ O(g)" apply (rule equalityI) @@ -357,13 +357,13 @@ apply (rule abs_triangle_ineq) apply (metis add_nonneg_nonneg) apply (rule add_mono) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_plus_eq_simpler"*} +using [[ atp_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] (*Found by SPASS; SLOW*) apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans) apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans) done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded_alt"*} +declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt" ]] lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) @@ -423,7 +423,7 @@ text{*So here is the easier (and more natural) problem using transitivity*} -ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded_alt_trans"*} +declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (*Version 1: one-shot proof*) @@ -431,7 +431,7 @@ done text{*So here is the easier (and more natural) problem using transitivity*} -ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded_alt_trans"*} +declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (*Version 2: single-step proof*) @@ -470,7 +470,7 @@ apply simp done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded2"*} +declare [[ atp_problem_prefix = "BigO__bigo_bounded2" ]] lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==> f : lb +o O(g)" apply (rule set_minus_imp_plus) @@ -493,7 +493,7 @@ by (metis 4 2 0) qed -ML_command{*AtpWrapper.problem_name := "BigO__bigo_abs"*} +declare [[ atp_problem_prefix = "BigO__bigo_abs" ]] lemma bigo_abs: "(%x. abs(f x)) =o O(f)" apply (unfold bigo_def) apply auto @@ -508,7 +508,7 @@ by (metis 0 2) qed -ML_command{*AtpWrapper.problem_name := "BigO__bigo_abs2"*} +declare [[ atp_problem_prefix = "BigO__bigo_abs2" ]] lemma bigo_abs2: "f =o O(%x. abs(f x))" apply (unfold bigo_def) apply auto @@ -580,7 +580,7 @@ by (simp add: bigo_abs3 [symmetric]) qed -ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult"*} +declare [[ atp_problem_prefix = "BigO__bigo_mult" ]] lemma bigo_mult [intro]: "O(f)\O(g) <= O(f * g)" apply (rule subsetI) apply (subst bigo_def) @@ -592,7 +592,7 @@ apply(erule_tac x = x in allE)+ apply(subgoal_tac "c * ca * abs(f x * g x) = (c * abs(f x)) * (ca * abs(g x))") -ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult_simpler"*} +using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]] prefer 2 apply (metis mult_assoc mult_left_commute OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute @@ -657,14 +657,14 @@ qed -ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult2"*} +declare [[ atp_problem_prefix = "BigO__bigo_mult2" ]] lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) (*sledgehammer*); apply (rule_tac x = c in exI) apply clarify apply (drule_tac x = x in spec) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult2_simpler"*} +using [[ atp_problem_prefix = "BigO__bigo_mult2_simpler" ]] (*sledgehammer [no luck]*); apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))") apply (simp add: mult_ac) @@ -672,11 +672,11 @@ apply (rule abs_ge_zero) done -ML_command{*AtpWrapper.problem_name:="BigO__bigo_mult3"*} +declare [[ atp_problem_prefix = "BigO__bigo_mult3" ]] lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)" by (metis bigo_mult set_times_intro subset_iff) -ML_command{*AtpWrapper.problem_name:="BigO__bigo_mult4"*} +declare [[ atp_problem_prefix = "BigO__bigo_mult4" ]] lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)" by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) @@ -710,13 +710,13 @@ qed qed -ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult6"*} +declare [[ atp_problem_prefix = "BigO__bigo_mult6" ]] lemma bigo_mult6: "ALL x. f x ~= 0 ==> O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)" by (metis bigo_mult2 bigo_mult5 order_antisym) (*proof requires relaxing relevance: 2007-01-25*) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult7"*} +declare [[ atp_problem_prefix = "BigO__bigo_mult7" ]] declare bigo_mult6 [simp] lemma bigo_mult7: "ALL x. f x ~= 0 ==> O(f * g) <= O(f::'a => ('b::ordered_field)) \ O(g)" @@ -728,7 +728,7 @@ done declare bigo_mult6 [simp del] -ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult8"*} +declare [[ atp_problem_prefix = "BigO__bigo_mult8" ]] declare bigo_mult7[intro!] lemma bigo_mult8: "ALL x. f x ~= 0 ==> O(f * g) = O(f::'a => ('b::ordered_field)) \ O(g)" @@ -779,7 +779,7 @@ qed qed -ML_command{*AtpWrapper.problem_name:="BigO__bigo_plus_absorb"*} +declare [[ atp_problem_prefix = "BigO__bigo_plus_absorb" ]] lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)" by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff); @@ -806,7 +806,7 @@ lemma bigo_const1: "(%x. c) : O(%x. 1)" by (auto simp add: bigo_def mult_ac) -ML_command{*AtpWrapper.problem_name:="BigO__bigo_const2"*} +declare [[ atp_problem_prefix = "BigO__bigo_const2" ]] lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)" by (metis bigo_const1 bigo_elt_subset); @@ -829,7 +829,7 @@ apply (rule bigo_const1) done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_const3"*} +declare [[ atp_problem_prefix = "BigO__bigo_const3" ]] lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)" apply (simp add: bigo_def) proof (neg_clausify) @@ -853,7 +853,7 @@ O(%x. c) = O(%x. 1)" by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult1"*} +declare [[ atp_problem_prefix = "BigO__bigo_const_mult1" ]] lemma bigo_const_mult1: "(%x. c * f x) : O(f)" apply (simp add: bigo_def abs_mult) proof (neg_clausify) @@ -869,7 +869,7 @@ lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)" by (rule bigo_elt_subset, rule bigo_const_mult1) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult3"*} +declare [[ atp_problem_prefix = "BigO__bigo_const_mult3" ]] lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)" apply (simp add: bigo_def) (*sledgehammer [no luck]*); @@ -887,7 +887,7 @@ O(%x. c * f x) = O(f)" by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult5"*} +declare [[ atp_problem_prefix = "BigO__bigo_const_mult5" ]] lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==> (%x. c) *o O(f) = O(f)" apply (auto del: subsetI) @@ -907,7 +907,7 @@ done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult6"*} +declare [[ atp_problem_prefix = "BigO__bigo_const_mult6" ]] lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)" apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times @@ -964,7 +964,7 @@ apply (blast intro: order_trans mult_right_mono abs_ge_self) done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_setsum1"*} +declare [[ atp_problem_prefix = "BigO__bigo_setsum1" ]] lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> EX c. ALL x y. abs(f x y) <= c * (h x y) ==> (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" @@ -981,7 +981,7 @@ (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)" by (rule bigo_setsum1, auto) -ML_command{*AtpWrapper.problem_name := "BigO__bigo_setsum3"*} +declare [[ atp_problem_prefix = "BigO__bigo_setsum3" ]] lemma bigo_setsum3: "f =o O(h) ==> (%x. SUM y : A x. (l x y) * f(k x y)) =o O(%x. SUM y : A x. abs(l x y * h(k x y)))" @@ -1012,7 +1012,7 @@ apply (erule set_plus_imp_minus) done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_setsum5"*} +declare [[ atp_problem_prefix = "BigO__bigo_setsum5" ]] lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> ALL x. 0 <= h x ==> (%x. SUM y : A x. (l x y) * f(k x y)) =o @@ -1069,7 +1069,7 @@ apply (simp add: func_times) done -ML_command{*AtpWrapper.problem_name := "BigO__bigo_fix"*} +declare [[ atp_problem_prefix = "BigO__bigo_fix" ]] lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==> f =o O(h)" apply (simp add: bigo_alt_def) @@ -1134,7 +1134,7 @@ apply (erule spec)+ done -ML_command{*AtpWrapper.problem_name:="BigO__bigo_lesso1"*} +declare [[ atp_problem_prefix = "BigO__bigo_lesso1" ]] lemma bigo_lesso1: "ALL x. f x <= g x ==> f ALL x. 0 <= k x ==> ALL x. k x <= f x ==> k ALL x. 0 <= k x ==> ALL x. g x <= k x ==> f EX C. ALL x. f x <= g x + C * abs(h x)" apply (simp only: lesso_def bigo_alt_def) diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/MetisExamples/Message.thy --- a/src/HOL/MetisExamples/Message.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/MetisExamples/Message.thy Sat Oct 03 12:05:40 2009 +0200 @@ -77,7 +77,7 @@ | Body: "Crypt K X \ parts H ==> X \ parts H" -ML{*AtpWrapper.problem_name := "Message__parts_mono"*} +declare [[ atp_problem_prefix = "Message__parts_mono" ]] lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" apply auto apply (erule parts.induct) @@ -101,7 +101,7 @@ subsubsection{*Inverse of keys *} -ML{*AtpWrapper.problem_name := "Message__invKey_eq"*} +declare [[ atp_problem_prefix = "Message__invKey_eq" ]] lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')" by (metis invKey) @@ -202,7 +202,7 @@ apply (simp only: parts_Un) done -ML{*AtpWrapper.problem_name := "Message__parts_insert_two"*} +declare [[ atp_problem_prefix = "Message__parts_insert_two" ]] lemma parts_insert2: "parts (insert X (insert Y H)) = parts {X} \ parts {Y} \ parts H" by (metis Un_commute Un_empty_left Un_empty_right Un_insert_left Un_insert_right parts_Un) @@ -239,7 +239,7 @@ lemma parts_idem [simp]: "parts (parts H) = parts H" by blast -ML{*AtpWrapper.problem_name := "Message__parts_subset_iff"*} +declare [[ atp_problem_prefix = "Message__parts_subset_iff" ]] lemma parts_subset_iff [simp]: "(parts G \ parts H) = (G \ parts H)" apply (rule iffI) apply (metis Un_absorb1 Un_subset_iff parts_Un parts_increasing) @@ -250,7 +250,7 @@ by (blast dest: parts_mono); -ML{*AtpWrapper.problem_name := "Message__parts_cut"*} +declare [[ atp_problem_prefix = "Message__parts_cut" ]] lemma parts_cut: "[|Y\ parts(insert X G); X\ parts H|] ==> Y\ parts(G \ H)" by (metis Un_subset_iff insert_subset parts_increasing parts_trans) @@ -315,7 +315,7 @@ done -ML{*AtpWrapper.problem_name := "Message__msg_Nonce_supply"*} +declare [[ atp_problem_prefix = "Message__msg_Nonce_supply" ]] lemma msg_Nonce_supply: "\N. \n. N\n --> Nonce n \ parts {msg}" apply (induct_tac "msg") apply (simp_all add: parts_insert2) @@ -367,7 +367,7 @@ lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard] -ML{*AtpWrapper.problem_name := "Message__parts_analz"*} +declare [[ atp_problem_prefix = "Message__parts_analz" ]] lemma parts_analz [simp]: "parts (analz H) = parts H" apply (rule equalityI) apply (metis analz_subset_parts parts_subset_iff) @@ -519,7 +519,7 @@ by (drule analz_mono, blast) -ML{*AtpWrapper.problem_name := "Message__analz_cut"*} +declare [[ atp_problem_prefix = "Message__analz_cut" ]] declare analz_trans[intro] lemma analz_cut: "[| Y\ analz (insert X H); X\ analz H |] ==> Y\ analz H" (*TOO SLOW @@ -537,7 +537,7 @@ text{*A congruence rule for "analz" *} -ML{*AtpWrapper.problem_name := "Message__analz_subset_cong"*} +declare [[ atp_problem_prefix = "Message__analz_subset_cong" ]] lemma analz_subset_cong: "[| analz G \ analz G'; analz H \ analz H' |] ==> analz (G \ H) \ analz (G' \ H')" @@ -615,7 +615,7 @@ by (intro Un_least synth_mono Un_upper1 Un_upper2) -ML{*AtpWrapper.problem_name := "Message__synth_insert"*} +declare [[ atp_problem_prefix = "Message__synth_insert" ]] lemma synth_insert: "insert X (synth H) \ synth(insert X H)" by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono) @@ -637,7 +637,7 @@ lemma synth_trans: "[| X\ synth G; G \ synth H |] ==> X\ synth H" by (drule synth_mono, blast) -ML{*AtpWrapper.problem_name := "Message__synth_cut"*} +declare [[ atp_problem_prefix = "Message__synth_cut" ]] lemma synth_cut: "[| Y\ synth (insert X H); X\ synth H |] ==> Y\ synth H" (*TOO SLOW by (metis insert_absorb insert_mono insert_subset synth_idem synth_increasing synth_mono) @@ -669,7 +669,7 @@ subsubsection{*Combinations of parts, analz and synth *} -ML{*AtpWrapper.problem_name := "Message__parts_synth"*} +declare [[ atp_problem_prefix = "Message__parts_synth" ]] lemma parts_synth [simp]: "parts (synth H) = parts H \ synth H" apply (rule equalityI) apply (rule subsetI) @@ -684,14 +684,14 @@ -ML{*AtpWrapper.problem_name := "Message__analz_analz_Un"*} +declare [[ atp_problem_prefix = "Message__analz_analz_Un" ]] lemma analz_analz_Un [simp]: "analz (analz G \ H) = analz (G \ H)" apply (rule equalityI); apply (metis analz_idem analz_subset_cong order_eq_refl) apply (metis analz_increasing analz_subset_cong order_eq_refl) done -ML{*AtpWrapper.problem_name := "Message__analz_synth_Un"*} +declare [[ atp_problem_prefix = "Message__analz_synth_Un" ]] declare analz_mono [intro] analz.Fst [intro] analz.Snd [intro] Un_least [intro] lemma analz_synth_Un [simp]: "analz (synth G \ H) = analz (G \ H) \ synth G" apply (rule equalityI) @@ -705,7 +705,7 @@ done -ML{*AtpWrapper.problem_name := "Message__analz_synth"*} +declare [[ atp_problem_prefix = "Message__analz_synth" ]] lemma analz_synth [simp]: "analz (synth H) = analz H \ synth H" proof (neg_clausify) assume 0: "analz (synth H) \ analz H \ synth H" @@ -728,7 +728,7 @@ subsubsection{*For reasoning about the Fake rule in traces *} -ML{*AtpWrapper.problem_name := "Message__parts_insert_subset_Un"*} +declare [[ atp_problem_prefix = "Message__parts_insert_subset_Un" ]] lemma parts_insert_subset_Un: "X\ G ==> parts(insert X H) \ parts G \ parts H" proof (neg_clausify) assume 0: "X \ G" @@ -747,7 +747,7 @@ by (metis 6 0) qed -ML{*AtpWrapper.problem_name := "Message__Fake_parts_insert"*} +declare [[ atp_problem_prefix = "Message__Fake_parts_insert" ]] lemma Fake_parts_insert: "X \ synth (analz H) ==> parts (insert X H) \ synth (analz H) \ parts H" @@ -790,14 +790,14 @@ ==> Z \ synth (analz H) \ parts H"; by (blast dest: Fake_parts_insert [THEN subsetD, dest]) -ML{*AtpWrapper.problem_name := "Message__Fake_analz_insert"*} +declare [[ atp_problem_prefix = "Message__Fake_analz_insert" ]] declare analz_mono [intro] synth_mono [intro] lemma Fake_analz_insert: "X\ synth (analz G) ==> analz (insert X H) \ synth (analz G) \ analz (G \ H)" by (metis Un_commute Un_insert_left Un_insert_right Un_upper1 analz_analz_Un analz_mono analz_synth_Un equalityE insert_absorb order_le_less xt1(12)) -ML{*AtpWrapper.problem_name := "Message__Fake_analz_insert_simpler"*} +declare [[ atp_problem_prefix = "Message__Fake_analz_insert_simpler" ]] (*simpler problems? BUT METIS CAN'T PROVE lemma Fake_analz_insert_simpler: "X\ synth (analz G) ==> diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/MetisExamples/Tarski.thy --- a/src/HOL/MetisExamples/Tarski.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/MetisExamples/Tarski.thy Sat Oct 03 12:05:40 2009 +0200 @@ -416,7 +416,7 @@ (*never proved, 2007-01-22: Tarski__CLF_unnamed_lemma NOT PROVABLE because of the conjunction used in the definition: we don't allow reasoning with rules like conjE, which is essential here.*) -ML_command{*AtpWrapper.problem_name:="Tarski__CLF_unnamed_lemma"*} +declare [[ atp_problem_prefix = "Tarski__CLF_unnamed_lemma" ]] lemma (in CLF) [simp]: "f: pset cl -> pset cl & monotone f (pset cl) (order cl)" apply (insert f_cl) @@ -433,7 +433,7 @@ by (simp add: A_def r_def) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__CLF_CLF_dual"*} +declare [[ atp_problem_prefix = "Tarski__CLF_CLF_dual" ]] declare (in CLF) CLF_set_def [simp] CL_dualCL [simp] monotone_dual [simp] dualA_iff [simp] lemma (in CLF) CLF_dual: "(dual cl, f) \ CLF_set" @@ -461,7 +461,7 @@ subsection {* lemmas for Tarski, lub *} (*never proved, 2007-01-22*) -ML{*AtpWrapper.problem_name:="Tarski__CLF_lubH_le_flubH"*} +declare [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]] declare CL.lub_least[intro] CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.transE[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] lemma (in CLF) lubH_le_flubH: "H = {x. (x, f x) \ r & x \ A} ==> (lub H cl, f (lub H cl)) \ r" @@ -471,7 +471,7 @@ -- {* @{text "\x:H. (x, f (lub H r)) \ r"} *} apply (rule ballI) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__CLF_lubH_le_flubH_simpler"*} +using [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]] apply (rule transE) -- {* instantiates @{text "(x, ?z) \ order cl to (x, f x)"}, *} -- {* because of the def of @{text H} *} @@ -489,7 +489,7 @@ CLF.monotone_f[rule del] CL.lub_upper[rule del] (*never proved, 2007-01-22*) -ML{*AtpWrapper.problem_name:="Tarski__CLF_flubH_le_lubH"*} +declare [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]] declare CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] CLF.lubH_le_flubH[simp] @@ -499,7 +499,7 @@ apply (rule_tac t = "H" in ssubst, assumption) apply (rule CollectI) apply (rule conjI) -ML_command{*AtpWrapper.problem_name:="Tarski__CLF_flubH_le_lubH_simpler"*} +using [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]] (*??no longer terminates, with combinators apply (metis CO_refl_on lubH_le_flubH monotone_def monotone_f reflD1 reflD2) *) @@ -513,7 +513,7 @@ (*never proved, 2007-01-22*) -ML{*AtpWrapper.problem_name:="Tarski__CLF_lubH_is_fixp"*} +declare [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]] (*Single-step version fails. The conjecture clauses refer to local abstraction functions (Frees), which prevents expand_defs_tac from removing those "definitions" at the end of the proof. *) @@ -588,7 +588,7 @@ "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" apply (simp add: fix_def) apply (rule conjI) -ML_command{*AtpWrapper.problem_name:="Tarski__CLF_lubH_is_fixp_simpler"*} +using [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]] apply (metis CO_refl_on lubH_le_flubH refl_onD1) apply (metis antisymE flubH_le_lubH lubH_le_flubH) done @@ -607,7 +607,7 @@ apply (simp_all add: P_def) done -ML{*AtpWrapper.problem_name:="Tarski__CLF_lubH_least_fixf"*} +declare [[ atp_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]] lemma (in CLF) lubH_least_fixf: "H = {x. (x, f x) \ r & x \ A} ==> \L. (\y \ fix f A. (y,L) \ r) --> (lub H cl, L) \ r" @@ -615,7 +615,7 @@ done subsection {* Tarski fixpoint theorem 1, first part *} -ML{*AtpWrapper.problem_name:="Tarski__CLF_T_thm_1_lub"*} +declare [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]] declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro] CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp] lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \ r & x \ A} cl" @@ -623,7 +623,7 @@ apply (rule sym) apply (simp add: P_def) apply (rule lubI) -ML_command{*AtpWrapper.problem_name:="Tarski__CLF_T_thm_1_lub_simpler"*} +using [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]] apply (metis P_def fix_subset) apply (metis Collect_conj_eq Collect_mem_eq Int_commute Int_lower1 lub_in_lattice vimage_def) (*??no longer terminates, with combinators @@ -638,7 +638,7 @@ (*never proved, 2007-01-22*) -ML{*AtpWrapper.problem_name:="Tarski__CLF_glbH_is_fixp"*} +declare [[ atp_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]] declare glb_dual_lub[simp] PO.dualA_iff[intro] CLF.lubH_is_fixp[intro] PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp] lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \ r & x \ A} ==> glb H cl \ P" @@ -662,13 +662,13 @@ (*never proved, 2007-01-22*) -ML{*AtpWrapper.problem_name:="Tarski__T_thm_1_glb"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__T_thm_1_glb" ]] (*ALL THEOREMS*) lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \ r & x \ A} cl" (*sledgehammer;*) apply (simp add: glb_dual_lub P_def A_def r_def) apply (rule dualA_iff [THEN subst]) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__T_thm_1_glb_simpler"*} (*ALL THEOREMS*) +using [[ atp_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]] (*ALL THEOREMS*) (*sledgehammer;*) apply (simp add: CLF.T_thm_1_lub [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff) @@ -677,13 +677,13 @@ subsection {* interval *} -ML{*AtpWrapper.problem_name:="Tarski__rel_imp_elem"*} +declare [[ atp_problem_prefix = "Tarski__rel_imp_elem" ]] declare (in CLF) CO_refl_on[simp] refl_on_def [simp] lemma (in CLF) rel_imp_elem: "(x, y) \ r ==> x \ A" by (metis CO_refl_on refl_onD1) declare (in CLF) CO_refl_on[simp del] refl_on_def [simp del] -ML{*AtpWrapper.problem_name:="Tarski__interval_subset"*} +declare [[ atp_problem_prefix = "Tarski__interval_subset" ]] declare (in CLF) rel_imp_elem[intro] declare interval_def [simp] lemma (in CLF) interval_subset: "[| a \ A; b \ A |] ==> interval r a b \ A" @@ -718,7 +718,7 @@ "[| a \ A; b \ A; S \ interval r a b |]==> S \ A" by (simp add: subset_trans [OF _ interval_subset]) -ML{*AtpWrapper.problem_name:="Tarski__L_in_interval"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__L_in_interval" ]] (*ALL THEOREMS*) lemma (in CLF) L_in_interval: "[| a \ A; b \ A; S \ interval r a b; S \ {}; isLub S cl L; interval r a b \ {} |] ==> L \ interval r a b" @@ -737,7 +737,7 @@ done (*never proved, 2007-01-22*) -ML{*AtpWrapper.problem_name:="Tarski__G_in_interval"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__G_in_interval" ]] (*ALL THEOREMS*) lemma (in CLF) G_in_interval: "[| a \ A; b \ A; interval r a b \ {}; S \ interval r a b; isGlb S cl G; S \ {} |] ==> G \ interval r a b" @@ -746,7 +746,7 @@ dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub) done -ML{*AtpWrapper.problem_name:="Tarski__intervalPO"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__intervalPO" ]] (*ALL THEOREMS*) lemma (in CLF) intervalPO: "[| a \ A; b \ A; interval r a b \ {} |] ==> (| pset = interval r a b, order = induced (interval r a b) r |) @@ -819,7 +819,7 @@ lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual] (*never proved, 2007-01-22*) -ML{*AtpWrapper.problem_name:="Tarski__interval_is_sublattice"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__interval_is_sublattice" ]] (*ALL THEOREMS*) lemma (in CLF) interval_is_sublattice: "[| a \ A; b \ A; interval r a b \ {} |] ==> interval r a b <<= cl" @@ -827,7 +827,7 @@ apply (rule sublatticeI) apply (simp add: interval_subset) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__interval_is_sublattice_simpler"*} +using [[ atp_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]] (*sledgehammer *) apply (rule CompleteLatticeI) apply (simp add: intervalPO) @@ -846,7 +846,7 @@ lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)" by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) -ML_command{*AtpWrapper.problem_name:="Tarski__Bot_in_lattice"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) lemma (in CLF) Bot_in_lattice: "Bot cl \ A" (*sledgehammer; *) apply (simp add: Bot_def least_def) @@ -856,12 +856,12 @@ done (*first proved 2007-01-25 after relaxing relevance*) -ML_command{*AtpWrapper.problem_name:="Tarski__Top_in_lattice"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__Top_in_lattice" ]] (*ALL THEOREMS*) lemma (in CLF) Top_in_lattice: "Top cl \ A" (*sledgehammer;*) apply (simp add: Top_dual_Bot A_def) (*first proved 2007-01-25 after relaxing relevance*) -ML_command{*AtpWrapper.problem_name:="Tarski__Top_in_lattice_simpler"*} (*ALL THEOREMS*) +using [[ atp_problem_prefix = "Tarski__Top_in_lattice_simpler" ]] (*ALL THEOREMS*) (*sledgehammer*) apply (rule dualA_iff [THEN subst]) apply (blast intro!: CLF.Bot_in_lattice [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualPO CL_dualCL CLF_dual) @@ -876,7 +876,7 @@ done (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__Bot_prop"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__Bot_prop" ]] (*ALL THEOREMS*) lemma (in CLF) Bot_prop: "x \ A ==> (Bot cl, x) \ r" (*sledgehammer*) apply (simp add: Bot_dual_Top r_def) @@ -885,12 +885,12 @@ dualA_iff A_def dualPO CL_dualCL CLF_dual) done -ML_command{*AtpWrapper.problem_name:="Tarski__Bot_in_lattice"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) lemma (in CLF) Top_intv_not_empty: "x \ A ==> interval r x (Top cl) \ {}" apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE) done -ML_command{*AtpWrapper.problem_name:="Tarski__Bot_intv_not_empty"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__Bot_intv_not_empty" ]] (*ALL THEOREMS*) lemma (in CLF) Bot_intv_not_empty: "x \ A ==> interval r (Bot cl) x \ {}" apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem) done @@ -902,7 +902,7 @@ by (simp add: P_def fix_subset po_subset_po) (*first proved 2007-01-25 after relaxing relevance*) -ML_command{*AtpWrapper.problem_name:="Tarski__Y_subset_A"*} +declare [[ atp_problem_prefix = "Tarski__Y_subset_A" ]] declare (in Tarski) P_def[simp] Y_ss [simp] declare fix_subset [intro] subset_trans [intro] lemma (in Tarski) Y_subset_A: "Y \ A" @@ -918,7 +918,7 @@ by (rule Y_subset_A [THEN lub_in_lattice]) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__lubY_le_flubY"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__lubY_le_flubY" ]] (*ALL THEOREMS*) lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \ r" (*sledgehammer*) apply (rule lub_least) @@ -927,12 +927,12 @@ apply (rule lubY_in_A) -- {* @{text "Y \ P ==> f x = x"} *} apply (rule ballI) -ML_command{*AtpWrapper.problem_name:="Tarski__lubY_le_flubY_simpler"*} (*ALL THEOREMS*) +using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]] (*ALL THEOREMS*) (*sledgehammer *) apply (rule_tac t = "x" in fix_imp_eq [THEN subst]) apply (erule Y_ss [simplified P_def, THEN subsetD]) -- {* @{text "reduce (f x, f (lub Y cl)) \ r to (x, lub Y cl) \ r"} by monotonicity *} -ML_command{*AtpWrapper.problem_name:="Tarski__lubY_le_flubY_simplest"*} (*ALL THEOREMS*) +using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]] (*ALL THEOREMS*) (*sledgehammer*) apply (rule_tac f = "f" in monotoneE) apply (rule monotone_f) @@ -942,7 +942,7 @@ done (*first proved 2007-01-25 after relaxing relevance*) -ML_command{*AtpWrapper.problem_name:="Tarski__intY1_subset"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__intY1_subset" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_subset: "intY1 \ A" (*sledgehammer*) apply (unfold intY1_def) @@ -954,7 +954,7 @@ lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD] (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__intY1_f_closed"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__intY1_f_closed" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_f_closed: "x \ intY1 \ f x \ intY1" (*sledgehammer*) apply (simp add: intY1_def interval_def) @@ -962,7 +962,7 @@ apply (rule transE) apply (rule lubY_le_flubY) -- {* @{text "(f (lub Y cl), f x) \ r"} *} -ML_command{*AtpWrapper.problem_name:="Tarski__intY1_f_closed_simpler"*} (*ALL THEOREMS*) +using [[ atp_problem_prefix = "Tarski__intY1_f_closed_simpler" ]] (*ALL THEOREMS*) (*sledgehammer [has been proved before now...]*) apply (rule_tac f=f in monotoneE) apply (rule monotone_f) @@ -975,13 +975,13 @@ apply (simp add: intY1_def interval_def intY1_elem) done -ML_command{*AtpWrapper.problem_name:="Tarski__intY1_func"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__intY1_func" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_func: "(%x: intY1. f x) \ intY1 -> intY1" apply (rule restrict_in_funcset) apply (metis intY1_f_closed restrict_in_funcset) done -ML_command{*AtpWrapper.problem_name:="Tarski__intY1_mono"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__intY1_mono" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_mono: "monotone (%x: intY1. f x) intY1 (induced intY1 r)" (*sledgehammer *) @@ -990,7 +990,7 @@ done (*proof requires relaxing relevance: 2007-01-25*) -ML_command{*AtpWrapper.problem_name:="Tarski__intY1_is_cl"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__intY1_is_cl" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_is_cl: "(| pset = intY1, order = induced intY1 r |) \ CompleteLattice" (*sledgehammer*) @@ -1003,7 +1003,7 @@ done (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__v_in_P"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__v_in_P" ]] (*ALL THEOREMS*) lemma (in Tarski) v_in_P: "v \ P" (*sledgehammer*) apply (unfold P_def) @@ -1013,7 +1013,7 @@ v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono) done -ML_command{*AtpWrapper.problem_name:="Tarski__z_in_interval"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__z_in_interval" ]] (*ALL THEOREMS*) lemma (in Tarski) z_in_interval: "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> z \ intY1" (*sledgehammer *) @@ -1027,14 +1027,14 @@ apply (simp add: induced_def) done -ML_command{*AtpWrapper.problem_name:="Tarski__fz_in_int_rel"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__fz_in_int_rel" ]] (*ALL THEOREMS*) lemma (in Tarski) f'z_in_int_rel: "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> ((%x: intY1. f x) z, z) \ induced intY1 r" apply (metis P_def acc_def fix_imp_eq fix_subset indI reflE restrict_apply subset_eq z_in_interval) done (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__tarski_full_lemma"*} (*ALL THEOREMS*) +declare [[ atp_problem_prefix = "Tarski__tarski_full_lemma" ]] (*ALL THEOREMS*) lemma (in Tarski) tarski_full_lemma: "\L. isLub Y (| pset = P, order = induced P r |) L" apply (rule_tac x = "v" in exI) @@ -1064,12 +1064,12 @@ prefer 2 apply (simp add: v_in_P) apply (unfold v_def) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__tarski_full_lemma_simpler"*} +using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]] (*sledgehammer*) apply (rule indE) apply (rule_tac [2] intY1_subset) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__tarski_full_lemma_simplest"*} +using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]] (*sledgehammer*) apply (rule CL.glb_lower [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) apply (simp add: CL_imp_PO intY1_is_cl) @@ -1087,7 +1087,7 @@ (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__Tarski_full"*} +declare [[ atp_problem_prefix = "Tarski__Tarski_full" ]] declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp] Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro] CompleteLatticeI_simp [intro] @@ -1097,7 +1097,7 @@ apply (rule CompleteLatticeI_simp) apply (rule fixf_po, clarify) (*never proved, 2007-01-22*) -ML_command{*AtpWrapper.problem_name:="Tarski__Tarski_full_simpler"*} +using [[ atp_problem_prefix = "Tarski__Tarski_full_simpler" ]] (*sledgehammer*) apply (simp add: P_def A_def r_def) apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro, diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/MetisExamples/TransClosure.thy --- a/src/HOL/MetisExamples/TransClosure.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/MetisExamples/TransClosure.thy Sat Oct 03 12:05:40 2009 +0200 @@ -22,7 +22,7 @@ consts f:: "addr \ val" -ML {*AtpWrapper.problem_name := "TransClosure__test"*} +declare [[ atp_problem_prefix = "TransClosure__test" ]] lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" by (metis Transitive_Closure.rtrancl_into_rtrancl converse_rtranclE trancl_reflcl) @@ -51,7 +51,7 @@ by (metis 10 3) qed -ML {*AtpWrapper.problem_name := "TransClosure__test_simpler"*} +declare [[ atp_problem_prefix = "TransClosure__test_simpler" ]] lemma "\ f c = Intg x; \ y. f b = Intg y \ y \ x; (a,b) \ R\<^sup>*; (b,c) \ R\<^sup>* \ \ \ c. (b,c) \ R \ (a,c) \ R\<^sup>*" apply (erule_tac x="b" in converse_rtranclE) diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/MetisExamples/set.thy --- a/src/HOL/MetisExamples/set.thy Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/MetisExamples/set.thy Sat Oct 03 12:05:40 2009 +0200 @@ -197,7 +197,7 @@ by (metis 11 6 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0) qed -ML {*AtpWrapper.problem_name := "set__equal_union"*} +declare [[ atp_problem_prefix = "set__equal_union" ]] lemma (*equal_union: *) "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" @@ -205,12 +205,12 @@ by (metis Un_least Un_upper1 Un_upper2 set_eq_subset) -ML {*AtpWrapper.problem_name := "set__equal_inter"*} +declare [[ atp_problem_prefix = "set__equal_inter" ]] lemma "(X = Y \ Z) = (X \ Y \ X \ Z \ (\V. V \ Y \ V \ Z \ V \ X))" by (metis Int_greatest Int_lower1 Int_lower2 set_eq_subset) -ML {*AtpWrapper.problem_name := "set__fixedpoint"*} +declare [[ atp_problem_prefix = "set__fixedpoint" ]] lemma fixedpoint: "\!x. f (g x) = x \ \!y. g (f y) = y" by metis @@ -229,7 +229,7 @@ by (metis 4 0) qed -ML {*AtpWrapper.problem_name := "set__singleton_example"*} +declare [[ atp_problem_prefix = "set__singleton_example" ]] lemma (*singleton_example_2:*) "\x \ S. \S \ x \ \z. S \ {z}" by (metis Set.subsetI Union_upper insertCI set_eq_subset) @@ -259,7 +259,7 @@ 293-314. *} -ML {*AtpWrapper.problem_name := "set__Bledsoe_Fung"*} +declare [[ atp_problem_prefix = "set__Bledsoe_Fung" ]] (*Notes: 1, the numbering doesn't completely agree with the paper. 2, we must rename set variables to avoid type clashes.*) lemma "\B. (\x \ B. x \ (0::int))" diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Sat Oct 03 12:05:40 2009 +0200 @@ -291,39 +291,27 @@ local -fun safe init done f x = - let - val y = init x - val z = Exn.capture f y - val _ = done y - in Exn.release z end - -fun init_sh NONE = !AtpWrapper.destdir - | init_sh (SOME path) = - let - (* Warning: we implicitly assume single-threaded execution here! *) - val old = !AtpWrapper.destdir - val _ = AtpWrapper.destdir := path - in old end - -fun done_sh path = AtpWrapper.destdir := path - datatype sh_result = SH_OK of int * int * string list | SH_FAIL of int * int | SH_ERROR -fun run_sh (prover_name, prover) hard_timeout timeout st _ = +fun run_sh prover hard_timeout timeout dir st = let - val atp = prover timeout NONE NONE prover_name 1 + val (ctxt, goal) = Proof.get_goal st + val ctxt' = ctxt |> is_some dir ? Config.put AtpWrapper.destdir (the dir) + val atp = prover (AtpWrapper.atp_problem_of_goal + (AtpManager.get_full_types ()) 1 (ctxt', goal)) + val time_limit = (case hard_timeout of NONE => I | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs)) - val ((success, (message, thm_names), time_atp, _, _, _), time_isa) = - time_limit (Mirabelle.cpu_time atp) (Proof.get_goal st) + val (AtpWrapper.Prover_Result {success, message, theorem_names, + runtime=time_atp, ...}, time_isa) = + time_limit (Mirabelle.cpu_time atp) timeout in - if success then (message, SH_OK (time_isa, time_atp, thm_names)) + if success then (message, SH_OK (time_isa, time_atp, theorem_names)) else (message, SH_FAIL(time_isa, time_atp)) end handle ResHolClause.TOO_TRIVIAL => ("trivial", SH_OK (0, 0, [])) @@ -348,13 +336,12 @@ fun run_sledgehammer args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) = let val _ = change_data id inc_sh_calls - val atp as (prover_name, _) = get_atp (Proof.theory_of st) args + val (prover_name, prover) = get_atp (Proof.theory_of st) args val dir = AList.lookup (op =) args keepK val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30) val hard_timeout = AList.lookup (op =) args prover_hard_timeoutK |> Option.map (fst o read_int o explode) - val (msg, result) = safe init_sh done_sh - (run_sh atp hard_timeout timeout st) dir + val (msg, result) = run_sh prover hard_timeout timeout dir st in case result of SH_OK (time_isa, time_atp, names) => diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/Tools/ATP_Manager/atp_manager.ML --- a/src/HOL/Tools/ATP_Manager/atp_manager.ML Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML Sat Oct 03 12:05:40 2009 +0200 @@ -21,13 +21,9 @@ val kill: unit -> unit val info: unit -> unit val messages: int option -> unit - type prover = int -> (thm * (string * int)) list option -> - (thm * (string * int)) list option -> string -> int -> - Proof.context * (thm list * thm) -> - bool * (string * string list) * int * string * string vector * (thm * (string * int)) list - val add_prover: string -> prover -> theory -> theory + val add_prover: string * AtpWrapper.prover -> theory -> theory val print_provers: theory -> unit - val get_prover: string -> theory -> prover option + val get_prover: string -> theory -> AtpWrapper.prover option val sledgehammer: string list -> Proof.state -> unit end; @@ -302,16 +298,11 @@ (* named provers *) -type prover = int -> (thm * (string * int)) list option -> - (thm * (string * int)) list option -> string -> int -> - Proof.context * (thm list * thm) -> - bool * (string * string list) * int * string * string vector * (thm * (string * int)) list - fun err_dup_prover name = error ("Duplicate prover: " ^ quote name); structure Provers = TheoryDataFun ( - type T = (prover * stamp) Symtab.table + type T = (AtpWrapper.prover * stamp) Symtab.table val empty = Symtab.empty val copy = I val extend = I @@ -319,7 +310,7 @@ handle Symtab.DUP dup => err_dup_prover dup ); -fun add_prover name prover thy = +fun add_prover (name, prover) thy = Provers.map (Symtab.update_new (name, (prover, stamp ()))) thy handle Symtab.DUP dup => err_dup_prover dup; @@ -344,9 +335,11 @@ val _ = SimpleThread.fork true (fn () => let val _ = register birthtime deadtime (Thread.self (), desc) + val problem = AtpWrapper.atp_problem_of_goal (get_full_types ()) i + (Proof.get_goal proof_state) val result = - let val (success, (message, _), _, _, _, _) = - prover (get_timeout ()) NONE NONE name i (Proof.get_goal proof_state) + let val AtpWrapper.Prover_Result {success, message, ...} = + prover problem (get_timeout ()) in (success, message) end handle ResHolClause.TOO_TRIVIAL => (true, "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis") diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/Tools/ATP_Manager/atp_minimal.ML --- a/src/HOL/Tools/ATP_Manager/atp_minimal.ML Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/Tools/ATP_Manager/atp_minimal.ML Sat Oct 03 12:05:40 2009 +0200 @@ -6,7 +6,7 @@ signature ATP_MINIMAL = sig - val minimalize: AtpManager.prover -> string -> int -> Proof.state -> + val minimalize: AtpWrapper.prover -> string -> int -> Proof.state -> (string * thm list) list -> ((string * thm list) list * int) option * string end @@ -97,7 +97,11 @@ ("# Cannot determine problem status within resource limit", Timeout), ("Error", Error)] -fun produce_answer (success, message, _, result_string, thm_name_vec, filtered) = +fun produce_answer result = + let + val AtpWrapper.Prover_Result {success, message, proof=result_string, + internal_thm_names=thm_name_vec, filtered_clauses=filtered, ...} = result + in if success then (Success (Vector.foldr op:: [] thm_name_vec, filtered), result_string) else @@ -110,19 +114,23 @@ else (Failure, result_string) end - + end (* wrapper for calling external prover *) -fun sh_test_thms prover prover_name time_limit subgoalno state filtered name_thms_pairs = +fun sh_test_thms prover time_limit subgoalno state filtered name_thms_pairs = let val _ = println ("Testing " ^ (length_string name_thms_pairs) ^ " theorems... ") val name_thm_pairs = maps (fn (n, ths) => map (pair n) ths) name_thms_pairs val _ = debug_fn (fn () => print_names name_thm_pairs) val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs - val (result, proof) = - produce_answer - (prover time_limit (SOME axclauses) filtered prover_name subgoalno (Proof.get_goal state)) + val problem = AtpWrapper.ATP_Problem { + with_full_types = AtpManager.get_full_types (), + subgoal = subgoalno, + goal = Proof.get_goal state, + axiom_clauses = SOME axclauses, + filtered_clauses = filtered } + val (result, proof) = produce_answer (prover problem time_limit) val _ = println (string_of_result result) val _ = debug proof in @@ -140,7 +148,7 @@ val _ = debug_fn (fn () => app (fn (n, tl) => (debug n; app (fn t => debug (" " ^ Display.string_of_thm (Proof.context_of state) t)) tl)) name_thms_pairs) - val test_thms_fun = sh_test_thms prover prover_name time_limit 1 state + val test_thms_fun = sh_test_thms prover time_limit 1 state fun test_thms filtered thms = case test_thms_fun filtered thms of (Success _, _) => true | _ => false val answer' = pair and answer'' = pair NONE diff -r 7c8bc41ce810 -r a226f29d4bdc src/HOL/Tools/ATP_Manager/atp_wrapper.ML --- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Fri Oct 02 04:44:56 2009 +0200 +++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Sat Oct 03 12:05:40 2009 +0200 @@ -6,25 +6,46 @@ signature ATP_WRAPPER = sig - val destdir: string Unsynchronized.ref - val problem_name: string Unsynchronized.ref - val tptp_prover_opts_full: int -> bool -> bool -> Path.T * string -> AtpManager.prover - val tptp_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover - val tptp_prover: Path.T * string -> AtpManager.prover - val full_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover - val full_prover: Path.T * string -> AtpManager.prover - val vampire_opts: int -> bool -> AtpManager.prover - val vampire: AtpManager.prover - val vampire_opts_full: int -> bool -> AtpManager.prover - val vampire_full: AtpManager.prover - val eprover_opts: int -> bool -> AtpManager.prover - val eprover: AtpManager.prover - val eprover_opts_full: int -> bool -> AtpManager.prover - val eprover_full: AtpManager.prover - val spass_opts: int -> bool -> AtpManager.prover - val spass: AtpManager.prover - val remote_prover_opts: int -> bool -> string -> string -> AtpManager.prover - val remote_prover: string -> string -> AtpManager.prover + (*hooks for problem files*) + val destdir: string Config.T + val problem_prefix: string Config.T + val setup: theory -> theory + + (*prover configuration, problem format, and prover result*) + datatype prover_config = Prover_Config of { + command: Path.T, + arguments: int -> string, + max_new_clauses: int, + insert_theory_const: bool, + emit_structured_proof: bool } + datatype atp_problem = ATP_Problem of { + with_full_types: bool, + subgoal: int, + goal: Proof.context * (thm list * thm), + axiom_clauses: (thm * (string * int)) list option, + filtered_clauses: (thm * (string * int)) list option } + val atp_problem_of_goal: bool -> int -> Proof.context * (thm list * thm) -> + atp_problem + datatype prover_result = Prover_Result of { + success: bool, + message: string, + theorem_names: string list, + runtime: int, + proof: string, + internal_thm_names: string Vector.vector, + filtered_clauses: (thm * (string * int)) list } + type prover = atp_problem -> int -> prover_result + + (*common provers*) + val vampire: string * prover + val vampire_full: string * prover + val eprover: string * prover + val eprover_full: string * prover + val spass: string * prover + val spass_no_tc: string * prover + val remote_vampire: string * prover + val remote_eprover: string * prover + val remote_spass: string * prover val refresh_systems: unit -> unit end; @@ -33,10 +54,50 @@ (** generic ATP wrapper **) -(* global hooks for writing problemfiles *) +(* hooks for writing problem files *) + +val (destdir, destdir_setup) = Attrib.config_string "atp_destdir" ""; + (*Empty string means create files in Isabelle's temporary files directory.*) + +val (problem_prefix, problem_prefix_setup) = + Attrib.config_string "atp_problem_prefix" "prob"; + +val setup = destdir_setup #> problem_prefix_setup; + + +(* prover configuration, problem format, and prover result *) + +datatype prover_config = Prover_Config of { + command: Path.T, + arguments: int -> string, + max_new_clauses: int, + insert_theory_const: bool, + emit_structured_proof: bool } -val destdir = Unsynchronized.ref ""; (*Empty means write files to /tmp*) -val problem_name = Unsynchronized.ref "prob"; +datatype atp_problem = ATP_Problem of { + with_full_types: bool, + subgoal: int, + goal: Proof.context * (thm list * thm), + axiom_clauses: (thm * (string * int)) list option, + filtered_clauses: (thm * (string * int)) list option } + +fun atp_problem_of_goal with_full_types subgoal goal = ATP_Problem { + with_full_types = with_full_types, + subgoal = subgoal, + goal = goal, + axiom_clauses = NONE, + filtered_clauses = NONE } + +datatype prover_result = Prover_Result of { + success: bool, + message: string, + theorem_names: string list, (* relevant theorems *) + runtime: int, (* user time of the ATP, in milliseconds *) + proof: string, + internal_thm_names: string Vector.vector, + filtered_clauses: (thm * (string * int)) list } + +type prover = atp_problem -> int -> prover_result (* basic template *) @@ -47,20 +108,9 @@ |> Exn.release |> tap (after path) -fun external_prover relevance_filter preparer writer (cmd, args) find_failure produce_answer - timeout axiom_clauses filtered_clauses name subgoalno goal = +fun external_prover relevance_filter preparer writer cmd args find_failure produce_answer + axiom_clauses filtered_clauses name subgoalno goal = let - (* path to unique problem file *) - val destdir' = ! destdir - val problem_name' = ! problem_name - fun prob_pathname nr = - let val probfile = Path.basic (problem_name' ^ serial_string () ^ "_" ^ string_of_int nr) - in if destdir' = "" then File.tmp_path probfile - else if File.exists (Path.explode (destdir')) - then Path.append (Path.explode (destdir')) probfile - else error ("No such directory: " ^ destdir') - end - (* get clauses and prepare them for writing *) val (ctxt, (chain_ths, th)) = goal val thy = ProofContext.theory_of ctxt @@ -78,6 +128,17 @@ val (thm_names, clauses) = preparer goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy + (* path to unique problem file *) + val destdir' = Config.get ctxt destdir + val problem_prefix' = Config.get ctxt problem_prefix + fun prob_pathname nr = + let val probfile = Path.basic (problem_prefix' ^ serial_string () ^ "_" ^ string_of_int nr) + in if destdir' = "" then File.tmp_path probfile + else if File.exists (Path.explode (destdir')) + then Path.append (Path.explode (destdir')) probfile + else error ("No such directory: " ^ destdir') + end + (* write out problem file and call prover *) fun cmd_line probfile = "TIMEFORMAT='%3U'; { time " ^ space_implode " " [File.shell_path cmd, args, File.platform_path probfile] ^ " ; } 2>&1" @@ -110,97 +171,100 @@ (* check for success and print out some information on failure *) val failure = find_failure proof val success = rc = 0 andalso is_none failure - val message = + val (message, real_thm_names) = if is_some failure then ("External prover failed.", []) else if rc <> 0 then ("External prover failed: " ^ proof, []) else apfst (fn s => "Try this command: " ^ s) (produce_answer name (proof, thm_names, conj_pos, ctxt, th, subgoalno)) val _ = Output.debug (fn () => "Sledgehammer response (rc = " ^ string_of_int rc ^ "):\n" ^ proof) - in (success, message, time, proof, thm_names, the_filtered_clauses) end; - + in + Prover_Result {success=success, message=message, + theorem_names=real_thm_names, runtime=time, proof=proof, + internal_thm_names = thm_names, filtered_clauses=the_filtered_clauses} + end -(** common provers **) - (* generic TPTP-based provers *) -fun tptp_prover_opts_full max_new theory_const full command timeout ax_clauses fcls name n goal = - external_prover - (ResAtp.get_relevant max_new theory_const) - (ResAtp.prepare_clauses false) - (ResHolClause.tptp_write_file (AtpManager.get_full_types())) - command - ResReconstruct.find_failure - (if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list false) - timeout ax_clauses fcls name n goal; +fun gen_tptp_prover (name, prover_config) problem timeout = + let + val Prover_Config {max_new_clauses, insert_theory_const, + emit_structured_proof, command, arguments} = prover_config + val ATP_Problem {with_full_types, subgoal, goal, axiom_clauses, + filtered_clauses} = problem + in + external_prover + (ResAtp.get_relevant max_new_clauses insert_theory_const) + (ResAtp.prepare_clauses false) + (ResHolClause.tptp_write_file with_full_types) + command + (arguments timeout) + ResReconstruct.find_failure + (if emit_structured_proof then ResReconstruct.structured_proof + else ResReconstruct.lemma_list false) + axiom_clauses + filtered_clauses + name + subgoal + goal + end -(*arbitrary ATP with TPTP input/output and problemfile as last argument*) -fun tptp_prover_opts max_new theory_const = - tptp_prover_opts_full max_new theory_const false; - -fun tptp_prover x = tptp_prover_opts 60 true x; +fun tptp_prover (name, config) = (name, gen_tptp_prover (name, config)) -(*for structured proofs: prover must support TSTP*) -fun full_prover_opts max_new theory_const = - tptp_prover_opts_full max_new theory_const true; -fun full_prover x = full_prover_opts 60 true x; - +(** common provers **) (* Vampire *) (*NB: Vampire does not work without explicit timelimit*) -fun vampire_opts max_new theory_const timeout = tptp_prover_opts - max_new theory_const - (Path.explode "$VAMPIRE_HOME/vampire", - ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout)) - timeout; - -val vampire = vampire_opts 60 false; +val vampire_max_new_clauses = 60 +val vampire_insert_theory_const = false -fun vampire_opts_full max_new theory_const timeout = full_prover_opts - max_new theory_const - (Path.explode "$VAMPIRE_HOME/vampire", - ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout)) - timeout; +fun vampire_prover_config full = Prover_Config { + command = Path.explode "$VAMPIRE_HOME/vampire", + arguments = (fn timeout => "--output_syntax tptp --mode casc" ^ + " -t " ^ string_of_int timeout), + max_new_clauses = vampire_max_new_clauses, + insert_theory_const = vampire_insert_theory_const, + emit_structured_proof = full } -val vampire_full = vampire_opts_full 60 false; +val vampire = tptp_prover ("vampire", vampire_prover_config false) +val vampire_full = tptp_prover ("vampire_full", vampire_prover_config true) (* E prover *) -fun eprover_opts max_new theory_const timeout = tptp_prover_opts - max_new theory_const - (Path.explode "$E_HOME/eproof", - "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout) - timeout; - -val eprover = eprover_opts 100 false; +val eprover_max_new_clauses = 100 +val eprover_insert_theory_const = false -fun eprover_opts_full max_new theory_const timeout = full_prover_opts - max_new theory_const - (Path.explode "$E_HOME/eproof", - "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout) - timeout; +fun eprover_config full = Prover_Config { + command = Path.explode "$E_HOME/eproof", + arguments = (fn timeout => "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev" ^ + " --silent --cpu-limit=" ^ string_of_int timeout), + max_new_clauses = eprover_max_new_clauses, + insert_theory_const = eprover_insert_theory_const, + emit_structured_proof = full } -val eprover_full = eprover_opts_full 100 false; +val eprover = tptp_prover ("e", eprover_config false) +val eprover_full = tptp_prover ("e_full", eprover_config true) (* SPASS *) -fun spass_opts max_new theory_const timeout ax_clauses fcls name n goal = external_prover - (ResAtp.get_relevant max_new theory_const) - (ResAtp.prepare_clauses true) - (ResHolClause.dfg_write_file (AtpManager.get_full_types())) - (Path.explode "$SPASS_HOME/SPASS", - "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^ - string_of_int timeout) - ResReconstruct.find_failure - (ResReconstruct.lemma_list true) - timeout ax_clauses fcls name n goal; +val spass_max_new_clauses = 40 +val spass_insert_theory_const = true -val spass = spass_opts 40 true; +fun spass_config insert_theory_const = Prover_Config { + command = Path.explode "$SPASS_HOME/SPASS", + arguments = (fn timeout => "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0" ^ + " -FullRed=0 -DocProof -TimeLimit=" ^ string_of_int timeout), + max_new_clauses = spass_max_new_clauses, + insert_theory_const = insert_theory_const, + emit_structured_proof = false } + +val spass = tptp_prover ("spass", spass_config spass_insert_theory_const) +val spass_no_tc = tptp_prover ("spass_no_tc", spass_config false) (* remote prover invocation via SystemOnTPTP *) @@ -220,20 +284,29 @@ get_systems ()); fun get_system prefix = Synchronized.change_result systems (fn systems => - let val systems = if null systems then get_systems() else systems - in (find_first (String.isPrefix prefix) systems, systems) end); + (if null systems then get_systems () else systems) + |> ` (find_first (String.isPrefix prefix))); + +fun get_the_system prefix = + (case get_system prefix of + NONE => error ("No system like " ^ quote prefix ^ " at SystemOnTPTP") + | SOME sys => sys) -fun remote_prover_opts max_new theory_const args prover_prefix timeout = - let val sys = - case get_system prover_prefix of - NONE => error ("No system like " ^ quote prover_prefix ^ " at SystemOnTPTP") - | SOME sys => sys - in tptp_prover_opts max_new theory_const - (Path.explode "$ISABELLE_ATP_MANAGER/SystemOnTPTP", - args ^ " -t " ^ string_of_int timeout ^ " -s " ^ sys) timeout - end; +fun remote_prover_config prover_prefix args max_new insert_tc = Prover_Config { + command = Path.explode "$ISABELLE_ATP_MANAGER/SystemOnTPTP", + arguments = (fn timeout => args ^ " -t " ^ string_of_int timeout ^ " -s " ^ + get_the_system prover_prefix), + max_new_clauses = max_new, + insert_theory_const = insert_tc, + emit_structured_proof = false } -val remote_prover = remote_prover_opts 60 false; +val remote_vampire = tptp_prover ("remote_vampire", remote_prover_config + "Vampire---9" "" vampire_max_new_clauses vampire_insert_theory_const) + +val remote_eprover = tptp_prover ("remote_e", remote_prover_config + "EP---" "" eprover_max_new_clauses eprover_insert_theory_const) + +val remote_spass = tptp_prover ("remote_spass", remote_prover_config + "SPASS---" "-x" spass_max_new_clauses spass_insert_theory_const) end; -