src/HOL/MetisExamples/BigO.thy
changeset 32864 a226f29d4bdc
parent 29823 0ab754d13ccd
--- 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) \<oplus> 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)\<otimes>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)) \<otimes> 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)) \<otimes> 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 <o g =o O(h)"
   apply (unfold lesso_def)
   apply (subgoal_tac "(%x. max (f x - g x) 0) = 0")
@@ -1146,7 +1146,7 @@
 done
 
 
-ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso2"*}
+declare [[ atp_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)"
@@ -1183,7 +1183,7 @@
   by (metis min_max.sup_commute min_max.inf_commute min_max.sup_inf_absorb min_max.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8)
 qed
 
-ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso3"*}
+declare [[ atp_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)"
@@ -1200,7 +1200,7 @@
   apply (simp)
   apply (subst abs_of_nonneg)
   apply (drule_tac x = x in spec) back
-ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso3_simpler"*}
+using [[ atp_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)
@@ -1220,7 +1220,7 @@
     split: split_max abs_split)
 done
 
-ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso5"*}
+declare [[ atp_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)