rename sledgehammer config attributes
authorblanchet
Wed, 01 Sep 2010 00:07:31 +0200
changeset 38991 0e2798f30087
parent 38990 7fba3ccc755a
child 38992 542474156c66
rename sledgehammer config attributes
src/HOL/Metis_Examples/Abstraction.thy
src/HOL/Metis_Examples/BT.thy
src/HOL/Metis_Examples/BigO.thy
src/HOL/Metis_Examples/Tarski.thy
src/HOL/Tools/Sledgehammer/sledgehammer.ML
--- a/src/HOL/Metis_Examples/Abstraction.thy	Wed Sep 01 00:03:15 2010 +0200
+++ b/src/HOL/Metis_Examples/Abstraction.thy	Wed Sep 01 00:07:31 2010 +0200
@@ -21,7 +21,7 @@
   pset  :: "'a set => 'a set"
   order :: "'a set => ('a * 'a) set"
 
-declare [[ atp_problem_prefix = "Abstraction__Collect_triv" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__Collect_triv" ]]
 lemma (*Collect_triv:*) "a \<in> {x. P x} ==> P a"
 proof -
   assume "a \<in> {x. P x}"
@@ -34,11 +34,11 @@
 by (metis mem_Collect_eq)
 
 
-declare [[ atp_problem_prefix = "Abstraction__Collect_mp" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__Collect_mp" ]]
 lemma "a \<in> {x. P x --> Q x} ==> a \<in> {x. P x} ==> a \<in> {x. Q x}"
   by (metis Collect_imp_eq ComplD UnE)
 
-declare [[ atp_problem_prefix = "Abstraction__Sigma_triv" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_triv" ]]
 lemma "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a"
 proof -
   assume A1: "(a, b) \<in> Sigma A B"
@@ -51,7 +51,7 @@
 lemma Sigma_triv: "(a,b) \<in> Sigma A B ==> a \<in> A & b \<in> B a"
 by (metis SigmaD1 SigmaD2)
 
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect" ]]
 lemma "(a, b) \<in> (SIGMA x:A. {y. x = f y}) \<Longrightarrow> a \<in> A \<and> a = f b"
 (* Metis says this is satisfiable!
 by (metis CollectD SigmaD1 SigmaD2)
@@ -85,7 +85,7 @@
 qed
 
 
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_in_pp" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_in_pp" ]]
 lemma "(cl,f) \<in> CLF ==> CLF = (SIGMA cl: CL.{f. f \<in> pset cl}) ==> f \<in> pset cl"
 by (metis Collect_mem_eq SigmaD2)
 
@@ -100,7 +100,7 @@
   thus "f \<in> pset cl" by metis
 qed
 
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]]
 lemma
     "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
     f \<in> pset cl \<rightarrow> pset cl"
@@ -112,7 +112,7 @@
   thus "f \<in> pset cl \<rightarrow> pset cl" by metis
 qed
 
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Int" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Int" ]]
 lemma
     "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
    f \<in> pset cl \<inter> cl"
@@ -127,27 +127,27 @@
 qed
 
 
-declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]]
 lemma
     "(cl,f) \<in> (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
    (f \<in> pset cl \<rightarrow> pset cl)  &  (monotone f (pset cl) (order cl))"
 by auto
 
-declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]]
 lemma "(cl,f) \<in> CLF ==> 
    CLF \<subseteq> (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
    f \<in> pset cl \<inter> cl"
 by auto
 
 
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]]
 lemma "(cl,f) \<in> CLF ==> 
    CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<inter> cl}) ==>
    f \<in> pset cl \<inter> cl"
 by auto
 
 
-declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]]
 lemma 
    "(cl,f) \<in> CLF ==> 
     CLF \<subseteq> (SIGMA cl': CL. {f. f \<in> pset cl' \<rightarrow> pset cl'}) ==> 
@@ -155,7 +155,7 @@
 by fast
 
 
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]]
 lemma 
   "(cl,f) \<in> CLF ==> 
    CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl}) ==> 
@@ -163,46 +163,46 @@
 by auto
 
 
-declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]]
 lemma 
   "(cl,f) \<in> CLF ==> 
    CLF = (SIGMA cl: CL. {f. f \<in> pset cl \<rightarrow> pset cl & monotone f (pset cl) (order cl)}) ==>
    (f \<in> pset cl \<rightarrow> pset cl)  &  (monotone f (pset cl) (order cl))"
 by auto
 
-declare [[ atp_problem_prefix = "Abstraction__map_eq_zipA" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__map_eq_zipA" ]]
 lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)"
 apply (induct xs)
  apply (metis map_is_Nil_conv zip.simps(1))
 by auto
 
-declare [[ atp_problem_prefix = "Abstraction__map_eq_zipB" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__map_eq_zipB" ]]
 lemma "map (%w. (w -> w, w \<times> w)) xs = 
        zip (map (%w. w -> w) xs) (map (%w. w \<times> w) xs)"
 apply (induct xs)
  apply (metis Nil_is_map_conv zip_Nil)
 by auto
 
-declare [[ atp_problem_prefix = "Abstraction__image_evenA" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__image_evenA" ]]
 lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\<forall>x. even x --> Suc(f x) \<in> A)"
 by (metis Collect_def image_subset_iff mem_def)
 
-declare [[ atp_problem_prefix = "Abstraction__image_evenB" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__image_evenB" ]]
 lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A 
        ==> (\<forall>x. even x --> f (f (Suc(f x))) \<in> A)";
 by (metis Collect_def imageI image_image image_subset_iff mem_def)
 
-declare [[ atp_problem_prefix = "Abstraction__image_curry" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__image_curry" ]]
 lemma "f \<in> (%u v. b \<times> u \<times> v) ` A ==> \<forall>u v. P (b \<times> u \<times> v) ==> P(f y)" 
 (*sledgehammer*)
 by auto
 
-declare [[ atp_problem_prefix = "Abstraction__image_TimesA" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesA" ]]
 lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A \<times> B) = (f`A) \<times> (g`B)"
 (*sledgehammer*)
 apply (rule equalityI)
 (***Even the two inclusions are far too difficult
-using [[ atp_problem_prefix = "Abstraction__image_TimesA_simpler"]]
+using [[ sledgehammer_problem_prefix = "Abstraction__image_TimesA_simpler"]]
 ***)
 apply (rule subsetI)
 apply (erule imageE)
@@ -225,13 +225,13 @@
 (*Given the difficulty of the previous problem, these two are probably
 impossible*)
 
-declare [[ atp_problem_prefix = "Abstraction__image_TimesB" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesB" ]]
 lemma image_TimesB:
     "(%(x,y,z). (f x, g y, h z)) ` (A \<times> B \<times> C) = (f`A) \<times> (g`B) \<times> (h`C)"
 (*sledgehammer*)
 by force
 
-declare [[ atp_problem_prefix = "Abstraction__image_TimesC" ]]
+declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesC" ]]
 lemma image_TimesC:
     "(%(x,y). (x \<rightarrow> x, y \<times> y)) ` (A \<times> B) = 
      ((%x. x \<rightarrow> x) ` A) \<times> ((%y. y \<times> y) ` B)" 
--- a/src/HOL/Metis_Examples/BT.thy	Wed Sep 01 00:03:15 2010 +0200
+++ b/src/HOL/Metis_Examples/BT.thy	Wed Sep 01 00:07:31 2010 +0200
@@ -65,7 +65,7 @@
 
 text {* \medskip BT simplification *}
 
-declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]]
+declare [[ sledgehammer_problem_prefix = "BT__n_leaves_reflect" ]]
 
 lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t"
 proof (induct t)
@@ -81,7 +81,7 @@
     by (metis n_leaves.simps(2) nat_add_commute reflect.simps(2))
 qed
 
-declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]]
+declare [[ sledgehammer_problem_prefix = "BT__n_nodes_reflect" ]]
 
 lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t"
 proof (induct t)
@@ -91,7 +91,7 @@
     by (metis add_commute n_nodes.simps(2) reflect.simps(2))
 qed
 
-declare [[ atp_problem_prefix = "BT__depth_reflect" ]]
+declare [[ sledgehammer_problem_prefix = "BT__depth_reflect" ]]
 
 lemma depth_reflect: "depth (reflect t) = depth t"
 apply (induct t)
@@ -102,14 +102,14 @@
 The famous relationship between the numbers of leaves and nodes.
 *}
 
-declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]]
+declare [[ sledgehammer_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))
 by auto
 
-declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]]
+declare [[ sledgehammer_problem_prefix = "BT__reflect_reflect_ident" ]]
 
 lemma reflect_reflect_ident: "reflect (reflect t) = t"
 apply (induct t)
@@ -127,7 +127,7 @@
   thus "reflect (reflect (Br a t1 t2)) = Br a t1 t2" by blast
 qed
 
-declare [[ atp_problem_prefix = "BT__bt_map_ident" ]]
+declare [[ sledgehammer_problem_prefix = "BT__bt_map_ident" ]]
 
 lemma bt_map_ident: "bt_map (%x. x) = (%y. y)"
 apply (rule ext) 
@@ -135,35 +135,35 @@
  apply (metis bt_map.simps(1))
 by (metis bt_map.simps(2))
 
-declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]]
+declare [[ sledgehammer_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))
 by (metis appnd.simps(2) bt_map.simps(2))
 
-declare [[ atp_problem_prefix = "BT__bt_map_compose" ]]
+declare [[ sledgehammer_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))
 by (metis bt_map.simps(2) o_eq_dest_lhs)
 
-declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]]
+declare [[ sledgehammer_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 bt_map.simps(1) reflect.simps(1))
 by (metis bt_map.simps(2) reflect.simps(2))
 
-declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]]
+declare [[ sledgehammer_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))
 by simp
 
-declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]]
+declare [[ sledgehammer_problem_prefix = "BT__inorder_bt_map" ]]
 
 lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)"
 proof (induct t)
@@ -178,21 +178,21 @@
   case (Br a t1 t2) thus ?case by simp
 qed
 
-declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]]
+declare [[ sledgehammer_problem_prefix = "BT__postorder_bt_map" ]]
 
 lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)"
 apply (induct t)
  apply (metis Nil_is_map_conv bt_map.simps(1) postorder.simps(1))
 by simp
 
-declare [[ atp_problem_prefix = "BT__depth_bt_map" ]]
+declare [[ sledgehammer_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))
 by simp
 
-declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]]
+declare [[ sledgehammer_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)
@@ -213,7 +213,7 @@
     using F1 by metis
 qed
 
-declare [[ atp_problem_prefix = "BT__preorder_reflect" ]]
+declare [[ sledgehammer_problem_prefix = "BT__preorder_reflect" ]]
 
 lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)"
 apply (induct t)
@@ -222,7 +222,7 @@
 by (metis append.simps(1) append.simps(2) postorder.simps(2) preorder.simps(2)
           reflect.simps(2) rev.simps(2) rev_append rev_swap)
 
-declare [[ atp_problem_prefix = "BT__inorder_reflect" ]]
+declare [[ sledgehammer_problem_prefix = "BT__inorder_reflect" ]]
 
 lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)"
 apply (induct t)
@@ -233,7 +233,7 @@
           reflect.simps(2) rev.simps(2) rev_append)
 *)
 
-declare [[ atp_problem_prefix = "BT__postorder_reflect" ]]
+declare [[ sledgehammer_problem_prefix = "BT__postorder_reflect" ]]
 
 lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)"
 apply (induct t)
@@ -245,14 +245,14 @@
 Analogues of the standard properties of the append function for lists.
 *}
 
-declare [[ atp_problem_prefix = "BT__appnd_assoc" ]]
+declare [[ sledgehammer_problem_prefix = "BT__appnd_assoc" ]]
 
 lemma appnd_assoc [simp]: "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)"
 apply (induct t1)
  apply (metis appnd.simps(1))
 by (metis appnd.simps(2))
 
-declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]]
+declare [[ sledgehammer_problem_prefix = "BT__appnd_Lf2" ]]
 
 lemma appnd_Lf2 [simp]: "appnd t Lf = t"
 apply (induct t)
@@ -261,14 +261,14 @@
 
 declare max_add_distrib_left [simp]
 
-declare [[ atp_problem_prefix = "BT__depth_appnd" ]]
+declare [[ sledgehammer_problem_prefix = "BT__depth_appnd" ]]
 
 lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2"
 apply (induct t1)
  apply (metis appnd.simps(1) depth.simps(1) plus_nat.simps(1))
 by simp
 
-declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]]
+declare [[ sledgehammer_problem_prefix = "BT__n_leaves_appnd" ]]
 
 lemma n_leaves_appnd [simp]:
      "n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2"
@@ -277,7 +277,7 @@
               semiring_norm(111))
 by (simp add: left_distrib)
 
-declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]]
+declare [[ sledgehammer_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)"
--- a/src/HOL/Metis_Examples/BigO.thy	Wed Sep 01 00:03:15 2010 +0200
+++ b/src/HOL/Metis_Examples/BigO.thy	Wed Sep 01 00:07:31 2010 +0200
@@ -15,7 +15,7 @@
 definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set"    ("(1O'(_'))") where
   "O(f::('a => 'b)) ==   {h. EX c. ALL x. abs (h x) <= c * abs (f x)}"
 
-declare [[ atp_problem_prefix = "BigO__bigo_pos_const" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_pos_const" ]]
 lemma bigo_pos_const: "(EX (c::'a::linordered_idom). 
     ALL x. (abs (h x)) <= (c * (abs (f x))))
       = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))"
@@ -124,7 +124,7 @@
     {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}"
 by (auto simp add: bigo_def bigo_pos_const)
 
-declare [[ atp_problem_prefix = "BigO__bigo_elt_subset" ]]
+declare [[ sledgehammer_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)
@@ -142,12 +142,12 @@
 done
 
 
-declare [[ atp_problem_prefix = "BigO__bigo_refl" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_refl" ]]
 lemma bigo_refl [intro]: "f : O(f)"
 apply (auto simp add: bigo_def)
 by (metis mult_1 order_refl)
 
-declare [[ atp_problem_prefix = "BigO__bigo_zero" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_zero" ]]
 lemma bigo_zero: "0 : O(g)"
 apply (auto simp add: bigo_def func_zero)
 by (metis mult_zero_left order_refl)
@@ -230,7 +230,7 @@
   apply (auto del: subsetI simp del: bigo_plus_idemp)
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_plus_eq" ]]
+declare [[ sledgehammer_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) \<oplus> O(g)"
   apply (rule equalityI)
@@ -253,13 +253,13 @@
   apply (rule abs_triangle_ineq)
   apply (metis add_nonneg_nonneg)
   apply (rule add_mono)
-using [[ atp_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] 
+using [[ sledgehammer_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
 
-declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt" ]]
+declare [[ sledgehammer_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)
@@ -277,14 +277,14 @@
 qed
 
 text{*So here is the easier (and more natural) problem using transitivity*}
-declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]]
+declare [[ sledgehammer_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-line proof *)
 by (metis abs_ge_self abs_mult order_trans)
 
 text{*So here is the easier (and more natural) problem using transitivity*}
-declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]]
+declare [[ sledgehammer_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: structured proof *)
@@ -299,7 +299,7 @@
   apply simp
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_bounded2" ]]
+declare [[ sledgehammer_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)
@@ -314,13 +314,13 @@
   thus "(0\<Colon>'b) \<le> f x + - lb x" by (metis not_leE diff_minus less_iff_diff_less_0 less_le_not_le)
 qed
 
-declare [[ atp_problem_prefix = "BigO__bigo_abs" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_abs" ]]
 lemma bigo_abs: "(%x. abs(f x)) =o O(f)" 
 apply (unfold bigo_def)
 apply auto
 by (metis mult_1 order_refl)
 
-declare [[ atp_problem_prefix = "BigO__bigo_abs2" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_abs2" ]]
 lemma bigo_abs2: "f =o O(%x. abs(f x))"
 apply (unfold bigo_def)
 apply auto
@@ -383,7 +383,7 @@
     by (simp add: bigo_abs3 [symmetric])
 qed
 
-declare [[ atp_problem_prefix = "BigO__bigo_mult" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult" ]]
 lemma bigo_mult [intro]: "O(f)\<otimes>O(g) <= O(f * g)"
   apply (rule subsetI)
   apply (subst bigo_def)
@@ -395,7 +395,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))")
-using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]]
+using [[ sledgehammer_problem_prefix = "BigO__bigo_mult_simpler" ]]
 prefer 2 
 apply (metis mult_assoc mult_left_commute
   abs_of_pos mult_left_commute
@@ -406,14 +406,14 @@
    abs_mult has just been done *)
 by (metis abs_ge_zero mult_mono')
 
-declare [[ atp_problem_prefix = "BigO__bigo_mult2" ]]
+declare [[ sledgehammer_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)
-using [[ atp_problem_prefix = "BigO__bigo_mult2_simpler" ]]
+using [[ sledgehammer_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)
@@ -421,11 +421,11 @@
   apply (rule abs_ge_zero)
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_mult3" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult3" ]]
 lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)"
 by (metis bigo_mult set_rev_mp set_times_intro)
 
-declare [[ atp_problem_prefix = "BigO__bigo_mult4" ]]
+declare [[ sledgehammer_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)
 
@@ -459,13 +459,13 @@
   qed
 qed
 
-declare [[ atp_problem_prefix = "BigO__bigo_mult6" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult6" ]]
 lemma bigo_mult6: "ALL x. f x ~= 0 ==>
     O(f * g) = (f::'a => ('b::linordered_field)) *o O(g)"
 by (metis bigo_mult2 bigo_mult5 order_antisym)
 
 (*proof requires relaxing relevance: 2007-01-25*)
-declare [[ atp_problem_prefix = "BigO__bigo_mult7" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult7" ]]
   declare bigo_mult6 [simp]
 lemma bigo_mult7: "ALL x. f x ~= 0 ==>
     O(f * g) <= O(f::'a => ('b::linordered_field)) \<otimes> O(g)"
@@ -477,7 +477,7 @@
 done
   declare bigo_mult6 [simp del]
 
-declare [[ atp_problem_prefix = "BigO__bigo_mult8" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult8" ]]
   declare bigo_mult7[intro!]
 lemma bigo_mult8: "ALL x. f x ~= 0 ==>
     O(f * g) = O(f::'a => ('b::linordered_field)) \<otimes> O(g)"
@@ -528,7 +528,7 @@
   qed
 qed
 
-declare [[ atp_problem_prefix = "BigO__bigo_plus_absorb" ]]
+declare [[ sledgehammer_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);
 
@@ -555,7 +555,7 @@
 lemma bigo_const1: "(%x. c) : O(%x. 1)"
 by (auto simp add: bigo_def mult_ac)
 
-declare [[ atp_problem_prefix = "BigO__bigo_const2" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_const2" ]]
 lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)"
 by (metis bigo_const1 bigo_elt_subset);
 
@@ -566,7 +566,7 @@
   show "O(\<lambda>x. c) \<subseteq> O(\<lambda>x. 1)" by (metis F1 bigo_elt_subset)
 qed
 
-declare [[ atp_problem_prefix = "BigO__bigo_const3" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_const3" ]]
 lemma bigo_const3: "(c::'a::linordered_field) ~= 0 ==> (%x. 1) : O(%x. c)"
 apply (simp add: bigo_def)
 by (metis abs_eq_0 left_inverse order_refl)
@@ -578,7 +578,7 @@
     O(%x. c) = O(%x. 1)"
 by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption)
 
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult1" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult1" ]]
 lemma bigo_const_mult1: "(%x. c * f x) : O(f)"
   apply (simp add: bigo_def abs_mult)
 by (metis le_less)
@@ -586,7 +586,7 @@
 lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)"
 by (rule bigo_elt_subset, rule bigo_const_mult1)
 
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult3" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult3" ]]
 lemma bigo_const_mult3: "(c::'a::linordered_field) ~= 0 ==> f : O(%x. c * f x)"
   apply (simp add: bigo_def)
 (*sledgehammer [no luck]*)
@@ -604,7 +604,7 @@
     O(%x. c * f x) = O(f)"
 by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4)
 
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult5" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult5" ]]
 lemma bigo_const_mult5 [simp]: "(c::'a::linordered_field) ~= 0 ==> 
     (%x. c) *o O(f) = O(f)"
   apply (auto del: subsetI)
@@ -624,7 +624,7 @@
 done
 
 
-declare [[ atp_problem_prefix = "BigO__bigo_const_mult6" ]]
+declare [[ sledgehammer_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
@@ -681,7 +681,7 @@
 apply (blast intro: order_trans mult_right_mono abs_ge_self) 
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_setsum1" ]]
+declare [[ sledgehammer_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)"
@@ -698,7 +698,7 @@
       (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)"
 by (rule bigo_setsum1, auto)  
 
-declare [[ atp_problem_prefix = "BigO__bigo_setsum3" ]]
+declare [[ sledgehammer_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)))"
@@ -729,7 +729,7 @@
   apply (erule set_plus_imp_minus)
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_setsum5" ]]
+declare [[ sledgehammer_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
@@ -786,7 +786,7 @@
   apply (simp add: func_times) 
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_fix" ]]
+declare [[ sledgehammer_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)
@@ -849,7 +849,7 @@
   apply (erule spec)+
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_lesso1" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso1" ]]
 lemma bigo_lesso1: "ALL x. f x <= g x ==> f <o g =o O(h)"
 apply (unfold lesso_def)
 apply (subgoal_tac "(%x. max (f x - g x) 0) = 0")
@@ -863,7 +863,7 @@
   by (simp split: split_max)
 qed
 
-declare [[ atp_problem_prefix = "BigO__bigo_lesso2" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso2" ]]
 lemma bigo_lesso2: "f =o g +o O(h) ==>
     ALL x. 0 <= k x ==> ALL x. k x <= f x ==>
       k <o g =o O(h)"
@@ -899,7 +899,7 @@
     by (metis abs_ge_zero le_cases min_max.sup_absorb2)
 qed
 
-declare [[ atp_problem_prefix = "BigO__bigo_lesso3" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso3" ]]
 lemma bigo_lesso3: "f =o g +o O(h) ==>
     ALL x. 0 <= k x ==> ALL x. g x <= k x ==>
       f <o k =o O(h)"
@@ -916,7 +916,7 @@
   apply (simp)
   apply (subst abs_of_nonneg)
   apply (drule_tac x = x in spec) back
-using [[ atp_problem_prefix = "BigO__bigo_lesso3_simpler" ]]
+using [[ sledgehammer_problem_prefix = "BigO__bigo_lesso3_simpler" ]]
 apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6))
 apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff)
 apply (metis abs_ge_zero linorder_linear min_max.sup_absorb1 min_max.sup_commute)
@@ -936,7 +936,7 @@
     split: split_max abs_split)
 done
 
-declare [[ atp_problem_prefix = "BigO__bigo_lesso5" ]]
+declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso5" ]]
 lemma bigo_lesso5: "f <o g =o O(h) ==>
     EX C. ALL x. f x <= g x + C * abs(h x)"
   apply (simp only: lesso_def bigo_alt_def)
--- a/src/HOL/Metis_Examples/Tarski.thy	Wed Sep 01 00:03:15 2010 +0200
+++ b/src/HOL/Metis_Examples/Tarski.thy	Wed Sep 01 00:07:31 2010 +0200
@@ -409,7 +409,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.*)
-declare [[ atp_problem_prefix = "Tarski__CLF_unnamed_lemma" ]]
+declare [[ sledgehammer_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)
@@ -426,7 +426,7 @@
 by (simp add: A_def r_def)
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_CLF_dual" ]]
+declare [[ sledgehammer_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) \<in> CLF_set" 
@@ -454,7 +454,7 @@
 subsection {* lemmas for Tarski, lub *}
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]]
+declare [[ sledgehammer_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) \<in> r & x \<in> A} ==> (lub H cl, f (lub H cl)) \<in> r"
@@ -464,7 +464,7 @@
 -- {* @{text "\<forall>x:H. (x, f (lub H r)) \<in> r"} *}
 apply (rule ballI)
 (*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]]
+using [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]]
 apply (rule transE)
 -- {* instantiates @{text "(x, ?z) \<in> order cl to (x, f x)"}, *}
 -- {* because of the def of @{text H} *}
@@ -482,7 +482,7 @@
           CLF.monotone_f[rule del] CL.lub_upper[rule del] 
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]]
+declare [[ sledgehammer_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]
@@ -492,7 +492,7 @@
 apply (rule_tac t = "H" in ssubst, assumption)
 apply (rule CollectI)
 apply (rule conjI)
-using [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]]
+using [[ sledgehammer_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) 
 *)
@@ -506,7 +506,7 @@
 
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]]
+declare [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]]
 (* Single-step version fails. The conjecture clauses refer to local abstraction
 functions (Frees). *)
 lemma (in CLF) lubH_is_fixp:
@@ -557,7 +557,7 @@
      "H = {x. (x, f x) \<in> r & x \<in> A} ==> lub H cl \<in> fix f A"
 apply (simp add: fix_def)
 apply (rule conjI)
-using [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]]
+using [[ sledgehammer_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
@@ -576,7 +576,7 @@
 apply (simp_all add: P_def)
 done
 
-declare [[ atp_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]]
+declare [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]]
 lemma (in CLF) lubH_least_fixf:
      "H = {x. (x, f x) \<in> r & x \<in> A}
       ==> \<forall>L. (\<forall>y \<in> fix f A. (y,L) \<in> r) --> (lub H cl, L) \<in> r"
@@ -584,7 +584,7 @@
 done
 
 subsection {* Tarski fixpoint theorem 1, first part *}
-declare [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]]
+declare [[ sledgehammer_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) \<in> r & x \<in> A} cl"
@@ -592,7 +592,7 @@
 apply (rule sym)
 apply (simp add: P_def)
 apply (rule lubI)
-using [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]]
+using [[ sledgehammer_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
@@ -607,7 +607,7 @@
 
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]]
+declare [[ sledgehammer_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) \<in> r & x \<in> A} ==> glb H cl \<in> P"
@@ -631,13 +631,13 @@
 
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__T_thm_1_glb" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__T_thm_1_glb" ]]  (*ALL THEOREMS*)
 lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \<in> r & x \<in> 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*)
-using [[ atp_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]]  (*ALL THEOREMS*)
+using [[ sledgehammer_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)
@@ -646,13 +646,13 @@
 subsection {* interval *}
 
 
-declare [[ atp_problem_prefix = "Tarski__rel_imp_elem" ]]
+declare [[ sledgehammer_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) \<in> r ==> x \<in> A"
 by (metis CO_refl_on refl_onD1)
   declare (in CLF) CO_refl_on[simp del]  refl_on_def [simp del]
 
-declare [[ atp_problem_prefix = "Tarski__interval_subset" ]]
+declare [[ sledgehammer_problem_prefix = "Tarski__interval_subset" ]]
   declare (in CLF) rel_imp_elem[intro] 
   declare interval_def [simp]
 lemma (in CLF) interval_subset: "[| a \<in> A; b \<in> A |] ==> interval r a b \<subseteq> A"
@@ -687,7 +687,7 @@
      "[| a \<in> A; b \<in> A; S \<subseteq> interval r a b |]==> S \<subseteq> A"
 by (simp add: subset_trans [OF _ interval_subset])
 
-declare [[ atp_problem_prefix = "Tarski__L_in_interval" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__L_in_interval" ]]  (*ALL THEOREMS*)
 lemma (in CLF) L_in_interval:
      "[| a \<in> A; b \<in> A; S \<subseteq> interval r a b;
          S \<noteq> {}; isLub S cl L; interval r a b \<noteq> {} |] ==> L \<in> interval r a b" 
@@ -706,7 +706,7 @@
 done
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__G_in_interval" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__G_in_interval" ]]  (*ALL THEOREMS*)
 lemma (in CLF) G_in_interval:
      "[| a \<in> A; b \<in> A; interval r a b \<noteq> {}; S \<subseteq> interval r a b; isGlb S cl G;
          S \<noteq> {} |] ==> G \<in> interval r a b"
@@ -715,7 +715,7 @@
                  dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub)
 done
 
-declare [[ atp_problem_prefix = "Tarski__intervalPO" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__intervalPO" ]]  (*ALL THEOREMS*)
 lemma (in CLF) intervalPO:
      "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
       ==> (| pset = interval r a b, order = induced (interval r a b) r |)
@@ -783,7 +783,7 @@
 lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual]
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__interval_is_sublattice" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__interval_is_sublattice" ]]  (*ALL THEOREMS*)
 lemma (in CLF) interval_is_sublattice:
      "[| a \<in> A; b \<in> A; interval r a b \<noteq> {} |]
         ==> interval r a b <<= cl"
@@ -791,7 +791,7 @@
 apply (rule sublatticeI)
 apply (simp add: interval_subset)
 (*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]]
+using [[ sledgehammer_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]]
 (*sledgehammer *)
 apply (rule CompleteLatticeI)
 apply (simp add: intervalPO)
@@ -810,7 +810,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)
 
-declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
 lemma (in CLF) Bot_in_lattice: "Bot cl \<in> A"
 (*sledgehammer; *)
 apply (simp add: Bot_def least_def)
@@ -820,12 +820,12 @@
 done
 
 (*first proved 2007-01-25 after relaxing relevance*)
-declare [[ atp_problem_prefix = "Tarski__Top_in_lattice" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__Top_in_lattice" ]]  (*ALL THEOREMS*)
 lemma (in CLF) Top_in_lattice: "Top cl \<in> A"
 (*sledgehammer;*)
 apply (simp add: Top_dual_Bot A_def)
 (*first proved 2007-01-25 after relaxing relevance*)
-using [[ atp_problem_prefix = "Tarski__Top_in_lattice_simpler" ]]  (*ALL THEOREMS*)
+using [[ sledgehammer_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)
@@ -840,7 +840,7 @@
 done
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__Bot_prop" ]]  (*ALL THEOREMS*) 
+declare [[ sledgehammer_problem_prefix = "Tarski__Bot_prop" ]]  (*ALL THEOREMS*) 
 lemma (in CLF) Bot_prop: "x \<in> A ==> (Bot cl, x) \<in> r"
 (*sledgehammer*) 
 apply (simp add: Bot_dual_Top r_def)
@@ -849,12 +849,12 @@
                  dualA_iff A_def dualPO CL_dualCL CLF_dual)
 done
 
-declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__Bot_in_lattice" ]]  (*ALL THEOREMS*)
 lemma (in CLF) Top_intv_not_empty: "x \<in> A  ==> interval r x (Top cl) \<noteq> {}" 
 apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE)
 done
 
-declare [[ atp_problem_prefix = "Tarski__Bot_intv_not_empty" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__Bot_intv_not_empty" ]]  (*ALL THEOREMS*)
 lemma (in CLF) Bot_intv_not_empty: "x \<in> A ==> interval r (Bot cl) x \<noteq> {}" 
 apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem)
 done
@@ -866,7 +866,7 @@
 by (simp add: P_def fix_subset po_subset_po)
 
 (*first proved 2007-01-25 after relaxing relevance*)
-declare [[ atp_problem_prefix = "Tarski__Y_subset_A" ]]
+declare [[ sledgehammer_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 \<subseteq> A"
@@ -882,7 +882,7 @@
   by (rule Y_subset_A [THEN lub_in_lattice])
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__lubY_le_flubY" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__lubY_le_flubY" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \<in> r"
 (*sledgehammer*) 
 apply (rule lub_least)
@@ -891,12 +891,12 @@
 apply (rule lubY_in_A)
 -- {* @{text "Y \<subseteq> P ==> f x = x"} *}
 apply (rule ballI)
-using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]]  (*ALL THEOREMS*)
+using [[ sledgehammer_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)) \<in> r to (x, lub Y cl) \<in> r"} by monotonicity *}
-using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]]  (*ALL THEOREMS*)
+using [[ sledgehammer_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]]  (*ALL THEOREMS*)
 (*sledgehammer*)
 apply (rule_tac f = "f" in monotoneE)
 apply (rule monotone_f)
@@ -906,7 +906,7 @@
 done
 
 (*first proved 2007-01-25 after relaxing relevance*)
-declare [[ atp_problem_prefix = "Tarski__intY1_subset" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__intY1_subset" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) intY1_subset: "intY1 \<subseteq> A"
 (*sledgehammer*) 
 apply (unfold intY1_def)
@@ -918,7 +918,7 @@
 lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD]
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__intY1_f_closed" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__intY1_f_closed" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) intY1_f_closed: "x \<in> intY1 \<Longrightarrow> f x \<in> intY1"
 (*sledgehammer*) 
 apply (simp add: intY1_def  interval_def)
@@ -926,7 +926,7 @@
 apply (rule transE)
 apply (rule lubY_le_flubY)
 -- {* @{text "(f (lub Y cl), f x) \<in> r"} *}
-using [[ atp_problem_prefix = "Tarski__intY1_f_closed_simpler" ]]  (*ALL THEOREMS*)
+using [[ sledgehammer_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)
@@ -939,13 +939,13 @@
 apply (simp add: intY1_def interval_def  intY1_elem)
 done
 
-declare [[ atp_problem_prefix = "Tarski__intY1_func" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__intY1_func" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) intY1_func: "(%x: intY1. f x) \<in> intY1 -> intY1"
 apply (rule restrict_in_funcset)
 apply (metis intY1_f_closed restrict_in_funcset)
 done
 
-declare [[ atp_problem_prefix = "Tarski__intY1_mono" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__intY1_mono" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) intY1_mono:
      "monotone (%x: intY1. f x) intY1 (induced intY1 r)"
 (*sledgehammer *)
@@ -954,7 +954,7 @@
 done
 
 (*proof requires relaxing relevance: 2007-01-25*)
-declare [[ atp_problem_prefix = "Tarski__intY1_is_cl" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__intY1_is_cl" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) intY1_is_cl:
     "(| pset = intY1, order = induced intY1 r |) \<in> CompleteLattice"
 (*sledgehammer*) 
@@ -967,7 +967,7 @@
 done
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__v_in_P" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__v_in_P" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) v_in_P: "v \<in> P"
 (*sledgehammer*) 
 apply (unfold P_def)
@@ -977,7 +977,7 @@
                  v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono)
 done
 
-declare [[ atp_problem_prefix = "Tarski__z_in_interval" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__z_in_interval" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) z_in_interval:
      "[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |] ==> z \<in> intY1"
 (*sledgehammer *)
@@ -991,14 +991,14 @@
 apply (simp add: induced_def)
 done
 
-declare [[ atp_problem_prefix = "Tarski__fz_in_int_rel" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__fz_in_int_rel" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) f'z_in_int_rel: "[| z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r |]
       ==> ((%x: intY1. f x) z, z) \<in> 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*)
-declare [[ atp_problem_prefix = "Tarski__tarski_full_lemma" ]]  (*ALL THEOREMS*)
+declare [[ sledgehammer_problem_prefix = "Tarski__tarski_full_lemma" ]]  (*ALL THEOREMS*)
 lemma (in Tarski) tarski_full_lemma:
      "\<exists>L. isLub Y (| pset = P, order = induced P r |) L"
 apply (rule_tac x = "v" in exI)
@@ -1028,12 +1028,12 @@
  prefer 2 apply (simp add: v_in_P)
 apply (unfold v_def)
 (*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]]
+using [[ sledgehammer_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]]
 (*sledgehammer*) 
 apply (rule indE)
 apply (rule_tac [2] intY1_subset)
 (*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]]
+using [[ sledgehammer_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)
@@ -1051,7 +1051,7 @@
 
 
 (*never proved, 2007-01-22*)
-declare [[ atp_problem_prefix = "Tarski__Tarski_full" ]]
+declare [[ sledgehammer_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]
@@ -1061,7 +1061,7 @@
 apply (rule CompleteLatticeI_simp)
 apply (rule fixf_po, clarify)
 (*never proved, 2007-01-22*)
-using [[ atp_problem_prefix = "Tarski__Tarski_full_simpler" ]]
+using [[ sledgehammer_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,
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Wed Sep 01 00:03:15 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Wed Sep 01 00:07:31 2010 +0200
@@ -118,14 +118,15 @@
 
 (* configuration attributes *)
 
-val (dest_dir, dest_dir_setup) = Attrib.config_string "atp_dest_dir" (K "");
-  (*Empty string means create files in Isabelle's temporary files directory.*)
+val (dest_dir, dest_dir_setup) =
+  Attrib.config_string "sledgehammer_dest_dir" (K "");
+  (* Empty string means create files in Isabelle's temporary files directory. *)
 
 val (problem_prefix, problem_prefix_setup) =
-  Attrib.config_string "atp_problem_prefix" (K "prob");
+  Attrib.config_string "sledgehammer_problem_prefix" (K "prob");
 
 val (measure_runtime, measure_runtime_setup) =
-  Attrib.config_bool "atp_measure_runtime" (K false);
+  Attrib.config_bool "sledgehammer_measure_runtime" (K false);
 
 fun with_path cleanup after f path =
   Exn.capture f path