# HG changeset patch # User paulson # Date 1674126825 0 # Node ID 68f1fc53c8fda42e7e998078e30d46000b0a7a87 # Parent d481dc154310782d2d86ab17c07329c9b723e97d tidy up of this messy and obsolete theory diff -r d481dc154310 -r 68f1fc53c8fd src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy Mon Jan 16 22:41:00 2023 +0100 +++ b/src/HOL/Library/BigO.thy Thu Jan 19 11:13:45 2023 +0000 @@ -77,11 +77,7 @@ apply (rule conjI) apply (rule_tac x = "c + c" in exI) apply (clarsimp) - apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \f xa\") - apply (metis mult_2 order_trans) - apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") - apply auto[1] - using abs_triangle_ineq mult_le_cancel_iff2 apply blast + apply (smt (verit, ccfv_threshold) mult.commute abs_triangle_ineq add_le_cancel_left dual_order.trans mult.left_commute mult_2 mult_le_cancel_iff2) apply (simp add: order_less_le) apply (rule_tac x = "\n. if \f n\ < \g n\ then x n else 0" in exI) apply (rule conjI) @@ -101,87 +97,40 @@ apply (simp add: bigo_alt_def set_plus_def func_plus) apply clarify apply (rule_tac x = "max c ca" in exI) - apply (rule conjI) - apply (subgoal_tac "c \ max c ca") - apply linarith - apply (rule max.cobounded1) - apply clarify - apply (drule_tac x = "xa" in spec)+ - apply (subgoal_tac "0 \ f xa + g xa") - apply (simp add: ring_distribs) - apply (subgoal_tac "\a xa + b xa\ \ \a xa\ + \b xa\") - apply (subgoal_tac "\a xa\ + \b xa\ \ max c ca * f xa + max c ca * g xa") - apply force - apply (metis add_mono le_max_iff_disj max_mult_distrib_right) - using abs_triangle_ineq apply blast - using add_nonneg_nonneg by blast + by (smt (verit, del_insts) add.commute abs_triangle_ineq add_mono_thms_linordered_field(3) distrib_left less_max_iff_disj linorder_not_less max.orderE max_mult_distrib_right order_le_less) lemma bigo_bounded_alt: "\x. 0 \ f x \ \x. f x \ c * g x \ f \ O(g)" - apply (auto simp add: bigo_def) - apply (rule_tac x = "\c\" in exI) - apply auto - apply (drule_tac x = x in spec)+ - apply (simp flip: abs_mult) - done + by (simp add: bigo_def) (metis abs_mult abs_of_nonneg order_trans) lemma bigo_bounded: "\x. 0 \ f x \ \x. f x \ g x \ f \ O(g)" - apply (erule bigo_bounded_alt [of f 1 g]) - apply simp - done + by (metis bigo_bounded_alt mult_1) lemma bigo_bounded2: "\x. lb x \ f x \ \x. f x \ lb x + g x \ f \ lb +o O(g)" - apply (rule set_minus_imp_plus) - apply (rule bigo_bounded) - apply (auto simp add: fun_Compl_def func_plus) - apply (drule_tac x = x in spec)+ - apply force - done + by (simp add: add.commute bigo_bounded diff_le_eq set_minus_imp_plus) lemma bigo_abs: "(\x. \f x\) =o O(f)" - apply (unfold bigo_def) - apply auto - apply (rule_tac x = 1 in exI) - apply auto - done + by (smt (verit, del_insts) abs_abs bigo_def bigo_refl mem_Collect_eq) lemma bigo_abs2: "f =o O(\x. \f x\)" - apply (unfold bigo_def) - apply auto - apply (rule_tac x = 1 in exI) - apply auto - done + by (smt (verit, del_insts) abs_abs bigo_def bigo_refl mem_Collect_eq) lemma bigo_abs3: "O(f) = O(\x. \f x\)" - apply (rule equalityI) - apply (rule bigo_elt_subset) - apply (rule bigo_abs2) - apply (rule bigo_elt_subset) - apply (rule bigo_abs) - done + using bigo_abs bigo_abs2 bigo_elt_subset by blast -lemma bigo_abs4: "f =o g +o O(h) \ (\x. \f x\) =o (\x. \g x\) +o O(h)" - apply (drule set_plus_imp_minus) - apply (rule set_minus_imp_plus) - apply (subst fun_diff_def) +lemma bigo_abs4: assumes "f =o g +o O(h)" shows "(\x. \f x\) =o (\x. \g x\) +o O(h)" proof - - assume *: "f - g \ O(h)" - have "(\x. \f x\ - \g x\) =o O(\x. \\f x\ - \g x\\)" - by (rule bigo_abs2) - also have "\ \ O(\x. \f x - g x\)" - apply (rule bigo_elt_subset) - apply (rule bigo_bounded) - apply force - apply (rule allI) - apply (rule abs_triangle_ineq3) - done - also have "\ \ O(f - g)" - apply (rule bigo_elt_subset) - apply (subst fun_diff_def) - apply (rule bigo_abs) - done - also from * have "\ \ O(h)" - by (rule bigo_elt_subset) - finally show "(\x. \f x\ - \g x\) \ O(h)". + { assume *: "f - g \ O(h)" + have "(\x. \f x\ - \g x\) =o O(\x. \\f x\ - \g x\\)" + by (rule bigo_abs2) + also have "\ \ O(\x. \f x - g x\)" + by (simp add: abs_triangle_ineq3 bigo_bounded bigo_elt_subset) + also have "\ \ O(f - g)" + using bigo_abs3 by fastforce + also from * have "\ \ O(h)" + by (rule bigo_elt_subset) + finally have "(\x. \f x\ - \g x\) \ O(h)" . } + then show ?thesis + by (smt (verit) assms bigo_alt_def fun_diff_def mem_Collect_eq set_minus_imp_plus set_plus_imp_minus) qed lemma bigo_abs5: "f =o O(g) \ (\x. \f x\) =o O(g)" @@ -210,42 +159,18 @@ lemma bigo_mult [intro]: "O(f)*O(g) \ O(f * g)" apply (rule subsetI) apply (subst bigo_def) - apply (auto simp add: bigo_alt_def set_times_def func_times) + apply (clarsimp simp add: bigo_alt_def set_times_def func_times) apply (rule_tac x = "c * ca" in exI) - apply (rule allI) - apply (erule_tac x = x in allE)+ - apply (subgoal_tac "c * ca * \f x * g x\ = (c * \f x\) * (ca * \g x\)") - apply (erule ssubst) - apply (subst abs_mult) - apply (rule mult_mono) - apply assumption+ - apply auto - apply (simp add: ac_simps abs_mult) - done + by (smt (verit, ccfv_threshold) mult.commute mult.assoc abs_ge_zero abs_mult dual_order.trans mult_mono) 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) - apply (rule_tac x = c in exI) - apply auto - apply (drule_tac x = x in spec) - apply (subgoal_tac "\f x\ * \b x\ \ \f x\ * (c * \g x\)") - apply (force simp add: ac_simps) - apply (rule mult_left_mono, assumption) - apply (rule abs_ge_zero) - done + by (metis bigo_mult bigo_refl dual_order.trans mult.commute set_times_mono4) lemma bigo_mult3: "f \ O(h) \ g \ O(j) \ f * g \ O(h * j)" - apply (rule subsetD) - apply (rule bigo_mult) - apply (erule set_times_intro, assumption) - done + using bigo_mult mult.commute mult.commute set_times_intro subsetD by blast lemma bigo_mult4 [intro]: "f \ k +o O(h) \ g * f \ (g * k) +o O(g * h)" - apply (drule set_plus_imp_minus) - apply (rule set_minus_imp_plus) - apply (drule bigo_mult3 [where g = g and j = g]) - apply (auto simp add: algebra_simps) - done + by (metis bigo_mult3 bigo_refl left_diff_distrib' mult.commute set_minus_imp_plus set_plus_imp_minus) lemma bigo_mult5: fixes f :: "'a \ 'b::linordered_field" @@ -259,52 +184,38 @@ also have "\ \ O((\x. 1 / f x) * (f * g))" by (rule bigo_mult2) also have "(\x. 1 / f x) * (f * g) = g" - apply (simp add: func_times) - apply (rule ext) - apply (simp add: assms nonzero_divide_eq_eq ac_simps) - done + using assms by auto finally have "(\x. (1::'b) / f x) * h \ O(g)" . then have "f * ((\x. (1::'b) / f x) * h) \ f *o O(g)" by auto also have "f * ((\x. (1::'b) / f x) * h) = h" - apply (simp add: func_times) - apply (rule ext) - apply (simp add: assms nonzero_divide_eq_eq ac_simps) - done + by (simp add: assms times_fun_def) finally show "h \ f *o O(g)" . qed lemma bigo_mult6: "\x. f x \ 0 \ O(f * g) = f *o O(g)" for f :: "'a \ 'b::linordered_field" - apply (rule equalityI) - apply (erule bigo_mult5) - apply (rule bigo_mult2) - done + by (simp add: bigo_mult2 bigo_mult5 subset_antisym) lemma bigo_mult7: "\x. f x \ 0 \ O(f * g) \ O(f) * O(g)" for f :: "'a \ 'b::linordered_field" - apply (subst bigo_mult6) - apply assumption - apply (rule set_times_mono3) - apply (rule bigo_refl) - done + by (metis bigo_mult6 bigo_refl mult.commute set_times_mono4) lemma bigo_mult8: "\x. f x \ 0 \ O(f * g) = O(f) * O(g)" for f :: "'a \ 'b::linordered_field" - apply (rule equalityI) - apply (erule bigo_mult7) - apply (rule bigo_mult) - done + by (simp add: bigo_mult bigo_mult7 subset_antisym) lemma bigo_minus [intro]: "f \ O(g) \ - f \ O(g)" by (auto simp add: bigo_def fun_Compl_def) -lemma bigo_minus2: "f \ g +o O(h) \ - f \ -g +o O(h)" - apply (rule set_minus_imp_plus) - apply (drule set_plus_imp_minus) - apply (drule bigo_minus) - apply simp - done +lemma bigo_minus2: + assumes "f \ g +o O(h)" shows "- f \ -g +o O(h)" +proof - + have "- f + g \ O(h)" + by (metis assms bigo_minus minus_diff_eq set_plus_imp_minus uminus_add_conv_diff) + then show ?thesis + by (simp add: set_minus_imp_plus) +qed lemma bigo_minus3: "O(- f) = O(f)" by (auto simp add: bigo_def fun_Compl_def) @@ -312,20 +223,7 @@ lemma bigo_plus_absorb_lemma1: assumes *: "f \ O(g)" shows "f +o O(g) \ O(g)" -proof - - have "f \ O(f)" by auto - then have "f +o O(g) \ O(f) + O(g)" - by (auto del: subsetI) - also have "\ \ O(g) + O(g)" - proof - - from * have "O(f) \ O(g)" - by (auto del: subsetI) - then show ?thesis - by (auto del: subsetI) - qed - also have "\ \ O(g)" by simp - finally show ?thesis . -qed + using assms bigo_plus_idemp set_plus_mono4 by blast lemma bigo_plus_absorb_lemma2: assumes *: "f \ O(g)" @@ -343,118 +241,75 @@ qed lemma bigo_plus_absorb [simp]: "f \ O(g) \ f +o O(g) = O(g)" - apply (rule equalityI) - apply (erule bigo_plus_absorb_lemma1) - apply (erule bigo_plus_absorb_lemma2) - done + by (simp add: bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 subset_antisym) lemma bigo_plus_absorb2 [intro]: "f \ O(g) \ A \ O(g) \ f +o A \ O(g)" - apply (subgoal_tac "f +o A \ f +o O(g)") - apply force+ - done + using bigo_plus_absorb set_plus_mono by blast lemma bigo_add_commute_imp: "f \ g +o O(h) \ g \ f +o O(h)" - apply (subst set_minus_plus [symmetric]) - apply (subgoal_tac "g - f = - (f - g)") - apply (erule ssubst) - apply (rule bigo_minus) - apply (subst set_minus_plus) - apply assumption - apply (simp add: ac_simps) - done + by (metis bigo_minus minus_diff_eq set_minus_imp_plus set_plus_imp_minus) lemma bigo_add_commute: "f \ g +o O(h) \ g \ f +o O(h)" - apply (rule iffI) - apply (erule bigo_add_commute_imp)+ - done + using bigo_add_commute_imp by blast lemma bigo_const1: "(\x. c) \ O(\x. 1)" by (auto simp add: bigo_def ac_simps) lemma bigo_const2 [intro]: "O(\x. c) \ O(\x. 1)" - apply (rule bigo_elt_subset) - apply (rule bigo_const1) - done + by (metis bigo_elt_subset bigo_const1) lemma bigo_const3: "c \ 0 \ (\x. 1) \ O(\x. c)" for c :: "'a::linordered_field" - apply (simp add: bigo_def) - apply (rule_tac x = "\inverse c\" in exI) - apply (simp flip: abs_mult) - done + by (metis bigo_bounded_alt le_numeral_extra(4) nonzero_divide_eq_eq zero_less_one_class.zero_le_one) lemma bigo_const4: "c \ 0 \ O(\x. 1) \ O(\x. c)" for c :: "'a::linordered_field" - apply (rule bigo_elt_subset) - apply (rule bigo_const3) - apply assumption - done + by (metis bigo_elt_subset bigo_const3) lemma bigo_const [simp]: "c \ 0 \ O(\x. c) = O(\x. 1)" for c :: "'a::linordered_field" - apply (rule equalityI) - apply (rule bigo_const2) - apply (rule bigo_const4) - apply assumption - done + by (metis equalityI bigo_const2 bigo_const4) lemma bigo_const_mult1: "(\x. c * f x) \ O(f)" - apply (simp add: bigo_def) - apply (rule_tac x = "\c\" in exI) - apply (auto simp flip: abs_mult) - done + by (smt (z3) abs_mult bigo_def bigo_refl mem_Collect_eq mult.left_commute mult_commute_abs) lemma bigo_const_mult2: "O(\x. c * f x) \ O(f)" - apply (rule bigo_elt_subset) - apply (rule bigo_const_mult1) - done + by (metis bigo_elt_subset bigo_const_mult1) lemma bigo_const_mult3: "c \ 0 \ f \ O(\x. c * f x)" for c :: "'a::linordered_field" - apply (simp add: bigo_def) - apply (rule_tac x = "\inverse c\" in exI) - apply (simp add: abs_mult mult.assoc [symmetric]) - done + by (simp add: bigo_def) (metis abs_mult field_class.field_divide_inverse mult.commute nonzero_divide_eq_eq order_refl) lemma bigo_const_mult4: "c \ 0 \ O(f) \ O(\x. c * f x)" for c :: "'a::linordered_field" - apply (rule bigo_elt_subset) - apply (rule bigo_const_mult3) - apply assumption - done + by (simp add: bigo_const_mult3 bigo_elt_subset) lemma bigo_const_mult [simp]: "c \ 0 \ O(\x. c * f x) = O(f)" for c :: "'a::linordered_field" - apply (rule equalityI) - apply (rule bigo_const_mult2) - apply (erule bigo_const_mult4) - done + by (simp add: bigo_const_mult2 bigo_const_mult4 subset_antisym) -lemma bigo_const_mult5 [simp]: "c \ 0 \ (\x. c) *o O(f) = O(f)" +lemma bigo_const_mult5 [simp]: "(\x. c) *o O(f) = O(f)" if "c \ 0" for c :: "'a::linordered_field" - apply (auto del: subsetI) - apply (rule order_trans) - apply (rule bigo_mult2) - apply (simp add: func_times) - apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) - apply (rule_tac x = "\y. inverse c * x y" in exI) - apply (simp add: mult.assoc [symmetric] abs_mult) - apply (rule_tac x = "\inverse c\ * ca" in exI) - apply auto - done +proof + show "O(f) \ (\x. c) *o O(f)" + using that + apply (clarsimp simp add: bigo_def elt_set_times_def func_times) + apply (rule_tac x = "\y. inverse c * x y" in exI) + apply (simp add: mult.assoc [symmetric] abs_mult) + apply (rule_tac x = "\inverse c\ * ca" in exI) + apply auto + done + have "O(\x. c * f x) \ O(f)" + by (simp add: bigo_const_mult2) + then show "(\x. c) *o O(f) \ O(f)" + using order_trans[OF bigo_mult2] by (force simp add: times_fun_def) +qed + lemma bigo_const_mult6 [intro]: "(\x. c) *o O(f) \ O(f)" apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) apply (rule_tac x = "ca * \c\" in exI) - apply (rule allI) - apply (subgoal_tac "ca * \c\ * \f x\ = \c\ * (ca * \f x\)") - apply (erule ssubst) - apply (subst abs_mult) - apply (rule mult_left_mono) - apply (erule spec) - apply simp - apply (simp add: ac_simps) - done + by (smt (verit, ccfv_SIG) ab_semigroup_mult_class.mult_ac(1) abs_abs abs_le_self_iff abs_mult le_cases3 mult.commute mult_left_mono) lemma bigo_const_mult7 [intro]: assumes *: "f =o O(g)" @@ -473,47 +328,25 @@ by (auto simp: bigo_def) lemma bigo_compose2: "f =o g +o O(h) \ (\x. f (k x)) =o (\x. g (k x)) +o O(\x. h(k x))" - apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus) - apply (drule bigo_compose1) - apply (simp add: fun_diff_def) - done + by (smt (verit, best) set_minus_plus bigo_def fun_diff_def mem_Collect_eq) subsection \Sum\ -lemma bigo_sum_main: "\x. \y \ A x. 0 \ h x y \ - \c. \x. \y \ A x. \f x y\ \ c * h x y \ - (\x. \y \ A x. f x y) =o O(\x. \y \ A x. h x y)" - apply (auto simp add: bigo_def) - apply (rule_tac x = "\c\" in exI) - apply (subst abs_of_nonneg) back back - apply (rule sum_nonneg) - apply force - apply (subst sum_distrib_left) - apply (rule allI) - apply (rule order_trans) - apply (rule sum_abs) - apply (rule sum_mono) - apply (rule order_trans) - apply (drule spec)+ - apply (drule bspec)+ - apply assumption+ - apply (drule bspec) - apply assumption+ - apply (rule mult_right_mono) - apply (rule abs_ge_self) - apply force - done +lemma bigo_sum_main: + assumes "\x. \y \ A x. 0 \ h x y" and "\x. \y \ A x. \f x y\ \ c * h x y" + shows "(\x. \y \ A x. f x y) =o O(\x. \y \ A x. h x y)" +proof - + have "(\i\A x. \f x i\) \ \c\ * sum (h x) (A x)" for x + by (smt (verit, ccfv_threshold) assms abs_mult_pos abs_of_nonneg abs_of_nonpos dual_order.trans le_cases3 neg_0_le_iff_le sum_distrib_left sum_mono) + then show ?thesis + using assms by (fastforce simp add: bigo_def sum_nonneg intro: order_trans [OF sum_abs]) +qed lemma bigo_sum1: "\x y. 0 \ h x y \ \c. \x y. \f x y\ \ c * h x y \ (\x. \y \ A x. f x y) =o O(\x. \y \ A x. h x y)" - apply (rule bigo_sum_main) - apply force - apply clarsimp - apply (rule_tac x = c in exI) - apply force - done + by (metis (no_types) bigo_sum_main) lemma bigo_sum2: "\y. 0 \ h y \ \c. \y. \f y\ \ c * (h y) \ @@ -523,112 +356,52 @@ lemma bigo_sum3: "f =o O(h) \ (\x. \y \ A x. l x y * f (k x y)) =o O(\x. \y \ A x. \l x y * h (k x y)\)" apply (rule bigo_sum1) - apply (rule allI)+ - apply (rule abs_ge_zero) - apply (unfold bigo_def) - apply auto - apply (rule_tac x = c in exI) - apply (rule allI)+ - apply (subst abs_mult)+ - apply (subst mult.left_commute) - apply (rule mult_left_mono) - apply (erule spec) - apply (rule abs_ge_zero) - done + using abs_ge_zero apply blast + apply (clarsimp simp: bigo_def) + by (smt (verit, ccfv_threshold) abs_mult abs_not_less_zero mult.left_commute mult_le_cancel_left) lemma bigo_sum4: "f =o g +o O(h) \ (\x. \y \ A x. l x y * f (k x y)) =o (\x. \y \ A x. l x y * g (k x y)) +o O(\x. \y \ A x. \l x y * h (k x y)\)" - apply (rule set_minus_imp_plus) - apply (subst fun_diff_def) - apply (subst sum_subtractf [symmetric]) - apply (subst right_diff_distrib [symmetric]) - apply (rule bigo_sum3) - apply (subst fun_diff_def [symmetric]) - apply (erule set_plus_imp_minus) - done + using bigo_sum3 [of "f-g" h l k A] + apply (simp add: algebra_simps sum_subtractf) + by (smt (verit) bigo_alt_def minus_apply set_minus_imp_plus set_plus_imp_minus mem_Collect_eq) lemma bigo_sum5: "f =o O(h) \ \x y. 0 \ l x y \ \x. 0 \ h x \ (\x. \y \ A x. l x y * f (k x y)) =o O(\x. \y \ A x. l x y * h (k x y))" - apply (subgoal_tac "(\x. \y \ A x. l x y * h (k x y)) = - (\x. \y \ A x. \l x y * h (k x y)\)") - apply (erule ssubst) - apply (erule bigo_sum3) - apply (rule ext) - apply (rule sum.cong) - apply (rule refl) - apply (subst abs_of_nonneg) - apply auto - done + using bigo_sum3 [of f h l k A] by simp lemma bigo_sum6: "f =o g +o O(h) \ \x y. 0 \ l x y \ \x. 0 \ h x \ (\x. \y \ A x. l x y * f (k x y)) =o (\x. \y \ A x. l x y * g (k x y)) +o O(\x. \y \ A x. l x y * h (k x y))" - apply (rule set_minus_imp_plus) - apply (subst fun_diff_def) - apply (subst sum_subtractf [symmetric]) - apply (subst right_diff_distrib [symmetric]) - apply (rule bigo_sum5) - apply (subst fun_diff_def [symmetric]) - apply (drule set_plus_imp_minus) - apply auto - done + using bigo_sum5 [of "f-g" h l k A] + apply (simp add: algebra_simps sum_subtractf) + by (smt (verit, del_insts) bigo_alt_def set_minus_imp_plus minus_apply set_plus_imp_minus mem_Collect_eq) subsection \Misc useful stuff\ -lemma bigo_useful_intro: "A \ O(f) \ B \ O(f) \ A + B \ O(f)" - apply (subst bigo_plus_idemp [symmetric]) - apply (rule set_plus_mono2) - apply assumption+ - done - lemma bigo_useful_add: "f =o O(h) \ g =o O(h) \ f + g =o O(h)" - apply (subst bigo_plus_idemp [symmetric]) - apply (rule set_plus_intro) - apply assumption+ - done + using bigo_plus_idemp set_plus_intro by blast lemma bigo_useful_const_mult: "c \ 0 \ (\x. c) * f =o O(h) \ f =o O(h)" for c :: "'a::linordered_field" - apply (rule subsetD) - apply (subgoal_tac "(\x. 1 / c) *o O(h) \ O(h)") - apply assumption - apply (rule bigo_const_mult6) - apply (subgoal_tac "f = (\x. 1 / c) * ((\x. c) * f)") - apply (erule ssubst) - apply (erule set_times_intro2) - apply (simp add: func_times) - done + using bigo_elt_subset bigo_mult6 by fastforce lemma bigo_fix: "(\x::nat. f (x + 1)) =o O(\x. h (x + 1)) \ f 0 = 0 \ f =o O(h)" - apply (simp add: bigo_alt_def) - apply auto - apply (rule_tac x = c in exI) - apply auto - apply (case_tac "x = 0") - apply simp - apply (subgoal_tac "x = Suc (x - 1)") - apply (erule ssubst) back - apply (erule spec) - apply simp - done + by (simp add: bigo_alt_def) (metis abs_eq_0_iff abs_ge_zero abs_mult abs_of_pos not0_implies_Suc) lemma bigo_fix2: - "(\x. f ((x::nat) + 1)) =o (\x. g(x + 1)) +o O(\x. h(x + 1)) \ + "(\x. f ((x::nat) + 1)) =o (\x. g(x + 1)) +o O(\x. h(x + 1)) \ f 0 = g 0 \ f =o g +o O(h)" - apply (rule set_minus_imp_plus) - apply (rule bigo_fix) - apply (subst fun_diff_def) - apply (subst fun_diff_def [symmetric]) - apply (rule set_plus_imp_minus) - apply simp - apply (simp add: fun_diff_def) + apply (rule set_minus_imp_plus [OF bigo_fix]) + apply (smt (verit, del_insts) bigo_alt_def fun_diff_def set_plus_imp_minus mem_Collect_eq) + apply simp done @@ -638,159 +411,93 @@ where "f x. max (f x - g x) 0)" lemma bigo_lesseq1: "f =o O(h) \ \x. \g x\ \ \f x\ \ g =o O(h)" - apply (unfold bigo_def) - apply clarsimp - apply (rule_tac x = c in exI) - apply (rule allI) - apply (rule order_trans) - apply (erule spec)+ - done + by (smt (verit, del_insts) bigo_def mem_Collect_eq order_trans) lemma bigo_lesseq2: "f =o O(h) \ \x. \g x\ \ f x \ g =o O(h)" - apply (erule bigo_lesseq1) - apply (rule allI) - apply (drule_tac x = x in spec) - apply (rule order_trans) - apply assumption - apply (rule abs_ge_self) - done + by (metis (mono_tags, lifting) abs_ge_zero abs_of_nonneg bigo_lesseq1 dual_order.trans) lemma bigo_lesseq3: "f =o O(h) \ \x. 0 \ g x \ \x. g x \ f x \ g =o O(h)" - apply (erule bigo_lesseq2) - apply (rule allI) - apply (subst abs_of_nonneg) - apply (erule spec)+ - done + by (meson bigo_bounded bigo_elt_subset subsetD) -lemma bigo_lesseq4: "f =o O(h) \ - \x. 0 \ g x \ \x. g x \ \f x\ \ g =o O(h)" - apply (erule bigo_lesseq1) - apply (rule allI) - apply (subst abs_of_nonneg) - apply (erule spec)+ - done +lemma bigo_lesseq4: "f =o O(h) \ \x. 0 \ g x \ \x. g x \ \f x\ \ g =o O(h)" + by (metis abs_of_nonneg bigo_lesseq1) lemma bigo_lesso1: "\x. f x \ g x \ f x. max (f x - g x) 0) = 0") - apply (erule ssubst) - apply (rule bigo_zero) - apply (unfold func_zero) - apply (rule ext) - apply (simp split: split_max) - done + by (smt (verit, del_insts) abs_ge_zero add_0 bigo_abs3 bigo_bounded diff_le_eq lesso_def max_def order_refl) lemma bigo_lesso2: "f =o g +o O(h) \ \x. 0 \ k x \ \x. k x \ f x \ k k x - g x") - apply simp - apply (subst abs_of_nonneg) - apply (drule_tac x = x in spec) back - apply (simp add: algebra_simps) - apply (subst diff_conv_add_uminus)+ - apply (rule add_right_mono) - apply (erule spec) - apply (rule order_trans) - prefer 2 - apply (rule abs_ge_zero) - apply (simp add: algebra_simps) - done + using max.cobounded2 apply blast + by (smt (verit) abs_ge_zero abs_of_nonneg diff_ge_0_iff_ge diff_mono diff_self fun_diff_def order_refl max.coboundedI2 max_def) lemma bigo_lesso3: "f =o g +o O(h) \ \x. 0 \ k x \ \x. g x \ k x \ f f x - k x") - apply simp - apply (subst abs_of_nonneg) - apply (drule_tac x = x in spec) back - apply (simp add: algebra_simps) - apply (subst diff_conv_add_uminus)+ - apply (rule add_left_mono) - apply (rule le_imp_neg_le) - apply (erule spec) - apply (rule order_trans) - prefer 2 - apply (rule abs_ge_zero) - apply (simp add: algebra_simps) - done + using max.cobounded2 apply blast + by (smt (verit) abs_eq_iff abs_ge_zero abs_if abs_minus_le_zero diff_left_mono fun_diff_def le_max_iff_disj order.trans order_eq_refl) -lemma bigo_lesso4: "f g =o h +o O(k) \ f 'b::linordered_field" - apply (unfold lesso_def) - apply (drule set_plus_imp_minus) - apply (drule bigo_abs5) back - apply (simp add: fun_diff_def) - apply (drule bigo_useful_add) - apply assumption - apply (erule bigo_lesseq2) back - apply (rule allI) - apply (auto simp add: func_plus fun_diff_def algebra_simps split: split_max abs_split) - done +lemma bigo_lesso4: + fixes k :: "'a \ 'b::linordered_field" + assumes f: "f O(k)" + by (simp add: g set_plus_imp_minus) + then have "(\x. \g x - h x\) \ O(k)" + using bigo_abs5 by force + then have \
: "(\x. max (f x - g x) 0) + (\x. \g x - h x\) \ O(k)" + by (metis (mono_tags, lifting) bigo_lesseq1 bigo_useful_add dual_order.eq_iff f lesso_def) + have "\max (f x - h x) 0\ \ ((\x. max (f x - g x) 0) + (\x. \g x - h x\)) x" for x + by (auto simp add: func_plus fun_diff_def algebra_simps split: split_max abs_split) + then show ?thesis + by (smt (verit, ccfv_SIG) \
bigo_lesseq2 lesso_def) +qed -lemma bigo_lesso5: "f \C. \x. f x \ g x + C * \h x\" - apply (simp only: lesso_def bigo_alt_def) - apply clarsimp - apply (rule_tac x = c in exI) - apply (rule allI) - apply (drule_tac x = x in spec) - apply (subgoal_tac "\max (f x - g x) 0\ = max (f x - g x) 0") - apply (clarsimp simp add: algebra_simps) - apply (rule abs_of_nonneg) - apply (rule max.cobounded2) - done + +lemma bigo_lesso5: + assumes "f C. \x. f x \ g x + C * \h x\" +proof - + obtain c where "0 < c" and c: "\x. f x - g x \ c * \h x\" + using assms by (auto simp: lesso_def bigo_alt_def) + have "\max (f x - g x) 0\ = max (f x - g x) 0" for x + by (auto simp add: algebra_simps) + then show ?thesis + by (metis c add.commute diff_le_eq) +qed lemma lesso_add: "f k (f + k) g \ 0 \ f \ 0" +lemma bigo_LIMSEQ1: "f \ 0" if f: "f =o O(g)" and g: "g \ 0" for f g :: "nat \ real" - apply (simp add: LIMSEQ_iff bigo_alt_def) - apply clarify - apply (drule_tac x = "r / c" in spec) - apply (drule mp) - apply simp - apply clarify - apply (rule_tac x = no in exI) - apply (rule allI) - apply (drule_tac x = n in spec)+ - apply (rule impI) - apply (drule mp) - apply assumption - apply (rule order_le_less_trans) - apply assumption - apply (rule order_less_le_trans) - apply (subgoal_tac "c * \g n\ < c * (r / c)") - apply assumption - apply (erule mult_strict_left_mono) - apply assumption - apply simp - done +proof - + { fix r::real + assume "0 < r" + obtain c::real where "0 < c" and rc: "\x. \f x\ \ c * \g x\" + using f by (auto simp: LIMSEQ_iff bigo_alt_def) + with g \0 < r\ obtain no where "\n\no. \g n\ < r/c" + by (fastforce simp: LIMSEQ_iff) + then have "\no. \n\no. \f n\ < r" + by (metis \0 < c\ mult.commute order_le_less_trans pos_less_divide_eq rc) } + then show ?thesis + by (auto simp: LIMSEQ_iff) +qed -lemma bigo_LIMSEQ2: "f =o g +o O(h) \ h \ 0 \ f \ a \ g \ a" - for f g h :: "nat \ real" - apply (drule set_plus_imp_minus) - apply (drule bigo_LIMSEQ1) - apply assumption - apply (simp only: fun_diff_def) - apply (erule Lim_transform2) - apply assumption - done +lemma bigo_LIMSEQ2: + fixes f g :: "nat \ real" + assumes "f =o g +o O(h)" "h \ 0" and f: "f \ a" + shows "g \ a" +proof - + have "f - g \ 0" + using assms bigo_LIMSEQ1 set_plus_imp_minus by blast + then have "(\n. f n - g n) \ 0" + by (simp add: fun_diff_def) + then show ?thesis + using Lim_transform_eq f by blast +qed end