# HG changeset patch # User wenzelm # Date 1468412895 -7200 # Node ID 151bb79536a75eafe5c00dbb6c5246b03a888cd6 # Parent f3781c5fb03f445926406000d61ccca7cdb51d49 misc tuning and modernization; diff -r f3781c5fb03f -r 151bb79536a7 src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy Tue Jul 12 22:54:37 2016 +0200 +++ b/src/HOL/Library/BigO.thy Wed Jul 13 14:28:15 2016 +0200 @@ -9,29 +9,28 @@ begin text \ -This library is designed to support asymptotic ``big O'' calculations, -i.e.~reasoning with expressions of the form $f = O(g)$ and $f = g + -O(h)$. An earlier version of this library is described in detail in -@{cite "Avigad-Donnelly"}. + This library is designed to support asymptotic ``big O'' calculations, + i.e.~reasoning with expressions of the form \f = O(g)\ and \f = g + O(h)\. + An earlier version of this library is described in detail in @{cite + "Avigad-Donnelly"}. + + The main changes in this version are as follows: -The main changes in this version are as follows: -\begin{itemize} -\item We have eliminated the \O\ operator on sets. (Most uses of this seem - to be inessential.) -\item We no longer use \+\ as output syntax for \+o\ -\item Lemmas involving \sumr\ have been replaced by more general lemmas - involving `\setsum\. -\item The library has been expanded, with e.g.~support for expressions of - the form \f < g + O(h)\. -\end{itemize} + \<^item> We have eliminated the \O\ operator on sets. (Most uses of this seem + to be inessential.) + \<^item> We no longer use \+\ as output syntax for \+o\ + \<^item> Lemmas involving \sumr\ have been replaced by more general lemmas + involving `\setsum\. + \<^item> The library has been expanded, with e.g.~support for expressions of + the form \f < g + O(h)\. -Note also since the Big O library includes rules that demonstrate set -inclusion, to use the automated reasoners effectively with the library -one should redeclare the theorem \subsetI\ as an intro rule, -rather than as an \intro!\ rule, for example, using -\<^theory_text>\declare subsetI [del, intro]\. + Note also since the Big O library includes rules that demonstrate set + inclusion, to use the automated reasoners effectively with the library one + should redeclare the theorem \subsetI\ as an intro rule, rather than as an + \intro!\ rule, for example, using \<^theory_text>\declare subsetI [del, intro]\. \ + subsection \Definitions\ definition bigo :: "('a \ 'b::linordered_idom) \ ('a \ 'b) set" ("(1O'(_'))") @@ -42,16 +41,16 @@ (\c. 0 < c \ (\x. \h x\ \ c * \f x\))" apply auto apply (case_tac "c = 0") - apply simp - apply (rule_tac x = "1" in exI) - apply simp + apply simp + apply (rule_tac x = "1" in exI) + apply simp apply (rule_tac x = "\c\" in exI) apply auto apply (subgoal_tac "c * \f x\ \ \c\ * \f x\") - apply (erule_tac x = x in allE) - apply force + apply (erule_tac x = x in allE) + apply force apply (rule mult_right_mono) - apply (rule abs_ge_self) + apply (rule abs_ge_self) apply (rule abs_ge_zero) done @@ -62,19 +61,19 @@ apply (auto simp add: bigo_alt_def) apply (rule_tac x = "ca * c" in exI) apply (rule conjI) - apply simp + apply simp apply (rule allI) apply (drule_tac x = "xa" in spec)+ apply (subgoal_tac "ca * \f xa\ \ ca * (c * \g xa\)") - apply (erule order_trans) - apply (simp add: ac_simps) + apply (erule order_trans) + apply (simp add: ac_simps) apply (rule mult_left_mono, assumption) apply (rule order_less_imp_le, assumption) done lemma bigo_refl [intro]: "f \ O(f)" - apply(auto simp add: bigo_def) - apply(rule_tac x = 1 in exI) + apply (auto simp add: bigo_def) + apply (rule_tac x = 1 in exI) apply simp done @@ -93,15 +92,15 @@ apply auto apply (simp add: ring_distribs func_plus) apply (rule order_trans) - apply (rule abs_triangle_ineq) + apply (rule abs_triangle_ineq) apply (rule add_mono) - apply force + apply force apply force done lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" apply (rule equalityI) - apply (rule bigo_plus_self_subset) + apply (rule bigo_plus_self_subset) apply (rule set_zero_plus2) apply (rule bigo_zero) done @@ -112,73 +111,73 @@ apply (subst bigo_pos_const [symmetric])+ apply (rule_tac x = "\n. if \g n\ \ \f n\ then x n else 0" in exI) 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 (erule_tac x = xa in allE) - apply (erule order_trans) - apply (simp) - apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") - apply (erule order_trans) - apply (simp add: ring_distribs) - apply (rule mult_left_mono) - apply (simp add: abs_triangle_ineq) - apply (simp add: order_less_le) + apply (rule_tac x = "c + c" in exI) + apply (clarsimp) + apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \f xa\") + apply (erule_tac x = xa in allE) + apply (erule order_trans) + apply (simp) + apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") + apply (erule order_trans) + apply (simp add: ring_distribs) + apply (rule mult_left_mono) + apply (simp add: abs_triangle_ineq) + 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) - apply (rule_tac x = "c + c" in exI) - apply auto + apply (rule_tac x = "c + c" in exI) + apply auto apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \g xa\") - apply (erule_tac x = xa in allE) - apply (erule order_trans) - apply simp + apply (erule_tac x = xa in allE) + apply (erule order_trans) + apply simp apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") - apply (erule order_trans) - apply (simp add: ring_distribs) + apply (erule order_trans) + apply (simp add: ring_distribs) apply (rule mult_left_mono) - apply (rule abs_triangle_ineq) + apply (rule abs_triangle_ineq) apply (simp add: order_less_le) done lemma bigo_plus_subset2 [intro]: "A \ O(f) \ B \ O(f) \ A + B \ O(f)" apply (subgoal_tac "A + B \ O(f) + O(f)") - apply (erule order_trans) - apply simp + apply (erule order_trans) + apply simp apply (auto del: subsetI simp del: bigo_plus_idemp) done lemma bigo_plus_eq: "\x. 0 \ f x \ \x. 0 \ g x \ O(f + g) = O(f) + O(g)" apply (rule equalityI) - apply (rule bigo_plus_subset) + apply (rule bigo_plus_subset) 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 (erule order_less_le_trans) - apply assumption - apply (rule max.cobounded1) + apply (subgoal_tac "c \ max c ca") + apply (erule order_less_le_trans) + apply assumption + 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 (rule add_mono) - apply (subgoal_tac "c * f xa \ max c ca * f xa") - apply force - apply (rule mult_right_mono) - apply (rule max.cobounded1) - apply assumption - apply (subgoal_tac "ca * g xa \ max c ca * g xa") - apply force - apply (rule mult_right_mono) - apply (rule max.cobounded2) - apply assumption - apply (rule abs_triangle_ineq) + 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 (rule add_mono) + apply (subgoal_tac "c * f xa \ max c ca * f xa") + apply force + apply (rule mult_right_mono) + apply (rule max.cobounded1) + apply assumption + apply (subgoal_tac "ca * g xa \ max c ca * g xa") + apply force + apply (rule mult_right_mono) + apply (rule max.cobounded2) + apply assumption + apply (rule abs_triangle_ineq) apply (rule add_nonneg_nonneg) - apply assumption+ + apply assumption+ done lemma bigo_bounded_alt: "\x. 0 \ f x \ \x. f x \ c * g x \ f \ O(g)" @@ -197,7 +196,7 @@ 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 (auto simp add: fun_Compl_def func_plus) apply (drule_tac x = x in spec)+ apply force done @@ -218,8 +217,8 @@ 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_abs2) apply (rule bigo_elt_subset) apply (rule bigo_abs) done @@ -229,13 +228,13 @@ apply (rule set_minus_imp_plus) apply (subst fun_diff_def) proof - - assume a: "f - g \ 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\)" apply (rule bigo_elt_subset) apply (rule bigo_bounded) - apply force + apply force apply (rule allI) apply (rule abs_triangle_ineq3) done @@ -244,23 +243,23 @@ apply (subst fun_diff_def) apply (rule bigo_abs) done - also from a have "\ \ O(h)" + also from * have "\ \ O(h)" by (rule bigo_elt_subset) finally show "(\x. \f x\ - \g x\) \ O(h)". qed lemma bigo_abs5: "f =o O(g) \ (\x. \f x\) =o O(g)" - by (unfold bigo_def, auto) + by (auto simp: bigo_def) -lemma bigo_elt_subset2 [intro]: "f \ g +o O(h) \ O(f) \ O(g) + O(h)" +lemma bigo_elt_subset2 [intro]: + assumes *: "f \ g +o O(h)" + shows "O(f) \ O(g) + O(h)" proof - - assume "f \ g +o O(h)" - also have "\ \ O(g) + O(h)" + note * + also have "g +o O(h) \ O(g) + O(h)" by (auto del: subsetI) also have "\ = O(\x. \g x\) + O(\x. \h x\)" - apply (subst bigo_abs3 [symmetric])+ - apply (rule refl) - done + by (subst bigo_abs3 [symmetric])+ (rule refl) also have "\ = O((\x. \g x\) + (\x. \h x\))" by (rule bigo_plus_eq [symmetric]) auto finally have "f \ \" . @@ -280,11 +279,11 @@ 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 (erule ssubst) + apply (subst abs_mult) + apply (rule mult_mono) + apply assumption+ + apply auto apply (simp add: ac_simps abs_mult) done @@ -294,14 +293,14 @@ 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 (force simp add: ac_simps) apply (rule mult_left_mono, assumption) apply (rule abs_ge_zero) done lemma bigo_mult3: "f \ O(h) \ g \ O(j) \ f * g \ O(h * j)" apply (rule subsetD) - apply (rule bigo_mult) + apply (rule bigo_mult) apply (erule set_times_intro, assumption) done @@ -309,7 +308,7 @@ 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) + apply (auto simp add: algebra_simps) done lemma bigo_mult5: @@ -339,28 +338,25 @@ finally show "h \ f *o O(g)" . qed -lemma bigo_mult6: - fixes f :: "'a \ 'b::linordered_field" - shows "\x. f x \ 0 \ O(f * g) = f *o O(g)" +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 (erule bigo_mult5) apply (rule bigo_mult2) done -lemma bigo_mult7: - fixes f :: "'a \ 'b::linordered_field" - shows "\x. f x \ 0 \ O(f * g) \ O(f) * O(g)" +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 assumption apply (rule set_times_mono3) apply (rule bigo_refl) done -lemma bigo_mult8: - fixes f :: "'a \ 'b::linordered_field" - shows "\x. f x \ 0 \ O(f * g) = O(f) * O(g)" +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 (erule bigo_mult7) apply (rule bigo_mult) done @@ -377,65 +373,63 @@ lemma bigo_minus3: "O(- f) = O(f)" by (auto simp add: bigo_def fun_Compl_def) -lemma bigo_plus_absorb_lemma1: "f \ O(g) \ f +o O(g) \ O(g)" +lemma bigo_plus_absorb_lemma1: + assumes *: "f \ O(g)" + shows "f +o O(g) \ O(g)" proof - - assume a: "f \ O(g)" - show "f +o O(g) \ O(g)" + 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 - - have "f \ O(f)" by auto - then have "f +o O(g) \ O(f) + O(g)" + from * have "O(f) \ O(g)" by (auto del: subsetI) - also have "\ \ O(g) + O(g)" - proof - - from a 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 . + then show ?thesis + by (auto del: subsetI) qed + also have "\ \ O(g)" by simp + finally show ?thesis . qed -lemma bigo_plus_absorb_lemma2: "f \ O(g) \ O(g) \ f +o O(g)" +lemma bigo_plus_absorb_lemma2: + assumes *: "f \ O(g)" + shows "O(g) \ f +o O(g)" proof - - assume a: "f \ O(g)" - show "O(g) \ f +o O(g)" - proof - - from a have "- f \ O(g)" - by auto - then have "- f +o O(g) \ O(g)" - by (elim bigo_plus_absorb_lemma1) - then have "f +o (- f +o O(g)) \ f +o O(g)" - by auto - also have "f +o (- f +o O(g)) = O(g)" - by (simp add: set_plus_rearranges) - finally show ?thesis . - qed + from * have "- f \ O(g)" + by auto + then have "- f +o O(g) \ O(g)" + by (elim bigo_plus_absorb_lemma1) + then have "f +o (- f +o O(g)) \ f +o O(g)" + by auto + also have "f +o (- f +o O(g)) = O(g)" + by (simp add: set_plus_rearranges) + finally show ?thesis . 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_lemma1) apply (erule bigo_plus_absorb_lemma2) done 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+ + apply force+ done 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 (erule ssubst) + apply (rule bigo_minus) + apply (subst set_minus_plus) + apply assumption apply (simp add: ac_simps) done lemma bigo_add_commute: "f \ g +o O(h) \ g \ f +o O(h)" apply (rule iffI) - apply (erule bigo_add_commute_imp)+ + apply (erule bigo_add_commute_imp)+ done lemma bigo_const1: "(\x. c) \ O(\x. 1)" @@ -446,27 +440,24 @@ apply (rule bigo_const1) done -lemma bigo_const3: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ (\x. 1) \ O(\x. c)" +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 add: abs_mult [symmetric]) done -lemma bigo_const4: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(\x. 1) \ O(\x. c)" +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 -lemma bigo_const [simp]: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(\x. c) = O(\x. 1)" +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_const2) apply (rule bigo_const4) apply assumption done @@ -482,37 +473,33 @@ apply (rule bigo_const_mult1) done -lemma bigo_const_mult3: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ f \ O(\x. c * f x)" +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 -lemma bigo_const_mult4: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(f) \ O(\x. c * f x)" +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 -lemma bigo_const_mult [simp]: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(\x. c * f x) = O(f)" +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 (rule bigo_const_mult2) apply (erule bigo_const_mult4) done -lemma bigo_const_mult5 [simp]: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ (\x. c) *o O(f) = O(f)" +lemma bigo_const_mult5 [simp]: "c \ 0 \ (\x. c) *o O(f) = O(f)" + for c :: "'a::linordered_field" apply (auto del: subsetI) - apply (rule order_trans) - apply (rule bigo_mult2) - apply (simp add: func_times) + 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) @@ -525,18 +512,19 @@ 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) + apply (erule ssubst) + apply (subst abs_mult) + apply (rule mult_left_mono) + apply (erule spec) + apply simp + apply (simp add: ac_simps) done -lemma bigo_const_mult7 [intro]: "f =o O(g) \ (\x. c * f x) =o O(g)" +lemma bigo_const_mult7 [intro]: + assumes *: "f =o O(g)" + shows "(\x. c * f x) =o O(g)" proof - - assume "f =o O(g)" - then have "(\x. c) * f =o (\x. c) *o O(g)" + from * have "(\x. c) * f =o (\x. c) *o O(g)" by auto also have "(\x. c) * f = (\x. c * f x)" by (simp add: func_times) @@ -546,10 +534,9 @@ qed lemma bigo_compose1: "f =o O(g) \ (\x. f (k x)) =o O(\x. g (k x))" - unfolding bigo_def by auto + 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))" +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) @@ -564,21 +551,21 @@ apply (auto simp add: bigo_def) apply (rule_tac x = "\c\" in exI) apply (subst abs_of_nonneg) back back - apply (rule setsum_nonneg) - apply force + apply (rule setsum_nonneg) + apply force apply (subst setsum_right_distrib) apply (rule allI) apply (rule order_trans) - apply (rule setsum_abs) + apply (rule setsum_abs) apply (rule setsum_mono) apply (rule order_trans) - apply (drule spec)+ - apply (drule bspec)+ - apply assumption+ - apply (drule bspec) - apply assumption+ + apply (drule spec)+ + apply (drule bspec)+ + apply assumption+ + apply (drule bspec) + apply assumption+ apply (rule mult_right_mono) - apply (rule abs_ge_self) + apply (rule abs_ge_self) apply force done @@ -586,7 +573,7 @@ \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_setsum_main) - apply force + apply force apply clarsimp apply (rule_tac x = c in exI) apply force @@ -600,8 +587,8 @@ lemma bigo_setsum3: "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_setsum1) - apply (rule allI)+ - apply (rule abs_ge_zero) + apply (rule allI)+ + apply (rule abs_ge_zero) apply (unfold bigo_def) apply auto apply (rule_tac x = c in exI) @@ -609,7 +596,7 @@ apply (subst abs_mult)+ apply (subst mult.left_commute) apply (rule mult_left_mono) - apply (erule spec) + apply (erule spec) apply (rule abs_ge_zero) done @@ -632,13 +619,13 @@ 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_setsum3) + apply (erule ssubst) + apply (erule bigo_setsum3) apply (rule ext) apply (rule setsum.cong) - apply (rule refl) + apply (rule refl) apply (subst abs_of_nonneg) - apply auto + apply auto done lemma bigo_setsum6: "f =o g +o O(h) \ \x y. 0 \ l x y \ @@ -651,9 +638,9 @@ apply (subst setsum_subtractf [symmetric]) apply (subst right_diff_distrib [symmetric]) apply (rule bigo_setsum5) - apply (subst fun_diff_def [symmetric]) - apply (drule set_plus_imp_minus) - apply auto + apply (subst fun_diff_def [symmetric]) + apply (drule set_plus_imp_minus) + apply auto done @@ -662,25 +649,24 @@ 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+ + 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+ + apply assumption+ done -lemma bigo_useful_const_mult: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ (\x. c) * f =o O(h) \ f =o O(h)" +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 "(\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 (erule ssubst) + apply (erule set_times_intro2) apply (simp add: func_times) done @@ -690,10 +676,10 @@ apply (rule_tac x = c in exI) apply auto apply (case_tac "x = 0") - apply simp + apply simp apply (subgoal_tac "x = Suc (x - 1)") - apply (erule ssubst) back - apply (erule spec) + apply (erule ssubst) back + apply (erule spec) apply simp done @@ -702,10 +688,10 @@ 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 (subst fun_diff_def) + apply (subst fun_diff_def [symmetric]) + apply (rule set_plus_imp_minus) + apply simp apply (simp add: fun_diff_def) done @@ -721,7 +707,7 @@ apply (rule_tac x = c in exI) apply (rule allI) apply (rule order_trans) - apply (erule spec)+ + apply (erule spec)+ done lemma bigo_lesseq2: "f =o O(h) \ \x. \g x\ \ f x \ g =o O(h)" @@ -729,7 +715,7 @@ apply (rule allI) apply (drule_tac x = x in spec) apply (rule order_trans) - apply assumption + apply assumption apply (rule abs_ge_self) done @@ -737,7 +723,7 @@ apply (erule bigo_lesseq2) apply (rule allI) apply (subst abs_of_nonneg) - apply (erule spec)+ + apply (erule spec)+ done lemma bigo_lesseq4: "f =o O(h) \ @@ -745,75 +731,72 @@ apply (erule bigo_lesseq1) apply (rule allI) apply (subst abs_of_nonneg) - apply (erule spec)+ + apply (erule spec)+ done 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 (erule ssubst) + apply (rule bigo_zero) apply (unfold func_zero) apply (rule ext) apply (simp split: split_max) done -lemma bigo_lesso2: "f =o g +o O(h) \ - \x. 0 \ k x \ \x. k x \ f x \ k \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 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) + prefer 2 + apply (rule abs_ge_zero) apply (simp add: algebra_simps) done -lemma bigo_lesso3: "f =o g +o O(h) \ - \x. 0 \ k x \ \x. g x \ k x \ f \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 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) + prefer 2 + apply (rule abs_ge_zero) apply (simp add: algebra_simps) done -lemma bigo_lesso4: - fixes k :: "'a \ 'b::linordered_field" - shows "f g =o h +o O(k) \ 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 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) @@ -826,7 +809,7 @@ 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 (clarsimp simp add: algebra_simps) apply (rule abs_of_nonneg) apply (rule max.cobounded2) done @@ -834,39 +817,41 @@ lemma lesso_add: "f k (f + k) g \ 0 \ f \ (0::real)" +lemma bigo_LIMSEQ1: "f =o O(g) \ g \ 0 \ f \ 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 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 assumption apply (rule order_le_less_trans) - apply assumption + 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 (subgoal_tac "c * \g n\ < c * (r / c)") + apply assumption + apply (erule mult_strict_left_mono) + apply assumption apply simp done -lemma bigo_LIMSEQ2: "f =o g +o O(h) \ h \ 0 \ f \ a \ g \ (a::real)" +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 assumption apply (simp only: fun_diff_def) apply (erule Lim_transform2) apply assumption diff -r f3781c5fb03f -r 151bb79536a7 src/HOL/Library/Set_Algebras.thy --- a/src/HOL/Library/Set_Algebras.thy Tue Jul 12 22:54:37 2016 +0200 +++ b/src/HOL/Library/Set_Algebras.thy Wed Jul 13 14:28:15 2016 +0200 @@ -1,5 +1,7 @@ (* Title: HOL/Library/Set_Algebras.thy - Author: Jeremy Avigad and Kevin Donnelly; Florian Haftmann, TUM + Author: Jeremy Avigad + Author: Kevin Donnelly + Author: Florian Haftmann, TUM *) section \Algebraic operations on sets\ @@ -11,14 +13,14 @@ text \ This library lifts operations like addition and multiplication to sets. It was designed to support asymptotic calculations. See the - comments at the top of theory \BigO\. + comments at the top of @{file "BigO.thy"}. \ instantiation set :: (plus) plus begin -definition plus_set :: "'a::plus set \ 'a set \ 'a set" where - set_plus_def: "A + B = {c. \a\A. \b\B. c = a + b}" +definition plus_set :: "'a::plus set \ 'a set \ 'a set" + where set_plus_def: "A + B = {c. \a\A. \b\B. c = a + b}" instance .. @@ -27,8 +29,8 @@ instantiation set :: (times) times begin -definition times_set :: "'a::times set \ 'a set \ 'a set" where - set_times_def: "A * B = {c. \a\A. \b\B. c = a * b}" +definition times_set :: "'a::times set \ 'a set \ 'a set" + where set_times_def: "A * B = {c. \a\A. \b\B. c = a * b}" instance .. @@ -37,8 +39,7 @@ instantiation set :: (zero) zero begin -definition - set_zero[simp]: "(0::'a::zero set) = {0}" +definition set_zero[simp]: "(0::'a::zero set) = {0}" instance .. @@ -47,21 +48,20 @@ instantiation set :: (one) one begin -definition - set_one[simp]: "(1::'a::one set) = {1}" +definition set_one[simp]: "(1::'a::one set) = {1}" instance .. end -definition elt_set_plus :: "'a::plus \ 'a set \ 'a set" (infixl "+o" 70) where - "a +o B = {c. \b\B. c = a + b}" +definition elt_set_plus :: "'a::plus \ 'a set \ 'a set" (infixl "+o" 70) + where "a +o B = {c. \b\B. c = a + b}" -definition elt_set_times :: "'a::times \ 'a set \ 'a set" (infixl "*o" 80) where - "a *o B = {c. \b\B. c = a * b}" +definition elt_set_times :: "'a::times \ 'a set \ 'a set" (infixl "*o" 80) + where "a *o B = {c. \b\B. c = a * b}" -abbreviation (input) elt_set_eq :: "'a \ 'a set \ bool" (infix "=o" 50) where - "x =o A \ x \ A" +abbreviation (input) elt_set_eq :: "'a \ 'a set \ bool" (infix "=o" 50) + where "x =o A \ x \ A" instance set :: (semigroup_add) semigroup_add by standard (force simp add: set_plus_def add.assoc) @@ -98,19 +98,21 @@ lemma set_plus_intro2 [intro]: "b \ C \ a + b \ a +o C" by (auto simp add: elt_set_plus_def) -lemma set_plus_rearrange: - "((a::'a::comm_monoid_add) +o C) + (b +o D) = (a + b) +o (C + D)" +lemma set_plus_rearrange: "(a +o C) + (b +o D) = (a + b) +o (C + D)" + for a b :: "'a::comm_monoid_add" apply (auto simp add: elt_set_plus_def set_plus_def ac_simps) apply (rule_tac x = "ba + bb" in exI) - apply (auto simp add: ac_simps) + apply (auto simp add: ac_simps) apply (rule_tac x = "aa + a" in exI) apply (auto simp add: ac_simps) done -lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = (a + b) +o C" +lemma set_plus_rearrange2: "a +o (b +o C) = (a + b) +o C" + for a b :: "'a::semigroup_add" by (auto simp add: elt_set_plus_def add.assoc) -lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C = a +o (B + C)" +lemma set_plus_rearrange3: "(a +o B) + C = a +o (B + C)" + for a :: "'a::semigroup_add" apply (auto simp add: elt_set_plus_def set_plus_def) apply (blast intro: ac_simps) apply (rule_tac x = "a + aa" in exI) @@ -121,7 +123,8 @@ apply (auto simp add: ac_simps) done -theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) = a +o (C + D)" +theorem set_plus_rearrange4: "C + (a +o D) = a +o (C + D)" + for a :: "'a::comm_monoid_add" apply (auto simp add: elt_set_plus_def set_plus_def ac_simps) apply (rule_tac x = "aa + ba" in exI) apply (auto simp add: ac_simps) @@ -133,13 +136,15 @@ lemma set_plus_mono [intro!]: "C \ D \ a +o C \ a +o D" by (auto simp add: elt_set_plus_def) -lemma set_plus_mono2 [intro]: "(C::'a::plus set) \ D \ E \ F \ C + E \ D + F" +lemma set_plus_mono2 [intro]: "C \ D \ E \ F \ C + E \ D + F" + for C D E F :: "'a::plus set" by (auto simp add: set_plus_def) lemma set_plus_mono3 [intro]: "a \ C \ a +o D \ C + D" by (auto simp add: elt_set_plus_def set_plus_def) -lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) \ C \ a +o D \ D + C" +lemma set_plus_mono4 [intro]: "a \ C \ a +o D \ D + C" + for a :: "'a::comm_monoid_add" by (auto simp add: elt_set_plus_def set_plus_def ac_simps) lemma set_plus_mono5: "a \ C \ B \ D \ a +o B \ C + D" @@ -166,33 +171,45 @@ apply auto done -lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C \ x \ a +o D \ x \ D + C" +lemma set_plus_mono4_b: "a \ C \ x \ a +o D \ x \ D + C" + for a x :: "'a::comm_monoid_add" apply (frule set_plus_mono4) apply auto done -lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" +lemma set_zero_plus [simp]: "0 +o C = C" + for C :: "'a::comm_monoid_add set" by (auto simp add: elt_set_plus_def) -lemma set_zero_plus2: "(0::'a::comm_monoid_add) \ A \ B \ A + B" +lemma set_zero_plus2: "0 \ A \ B \ A + B" + for A B :: "'a::comm_monoid_add set" apply (auto simp add: set_plus_def) apply (rule_tac x = 0 in bexI) apply (rule_tac x = x in bexI) apply (auto simp add: ac_simps) done -lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C \ (a - b) \ C" +lemma set_plus_imp_minus: "a \ b +o C \ a - b \ C" + for a b :: "'a::ab_group_add" by (auto simp add: elt_set_plus_def ac_simps) -lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C \ a \ b +o C" +lemma set_minus_imp_plus: "a - b \ C \ a \ b +o C" + for a b :: "'a::ab_group_add" apply (auto simp add: elt_set_plus_def ac_simps) apply (subgoal_tac "a = (a + - b) + b") - apply (rule bexI, assumption) - apply (auto simp add: ac_simps) + apply (rule bexI) + apply assumption + apply (auto simp add: ac_simps) done -lemma set_minus_plus: "(a::'a::ab_group_add) - b \ C \ a \ b +o C" - by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus) +lemma set_minus_plus: "a - b \ C \ a \ b +o C" + for a b :: "'a::ab_group_add" + apply (rule iffI) + apply (rule set_minus_imp_plus) + apply assumption + apply (rule set_plus_imp_minus) + apply assumption + done lemma set_times_intro [intro]: "a \ C \ b \ D \ a * b \ C * D" by (auto simp add: set_times_def) @@ -205,8 +222,8 @@ lemma set_times_intro2 [intro!]: "b \ C \ a * b \ a *o C" by (auto simp add: elt_set_times_def) -lemma set_times_rearrange: - "((a::'a::comm_monoid_mult) *o C) * (b *o D) = (a * b) *o (C * D)" +lemma set_times_rearrange: "(a *o C) * (b *o D) = (a * b) *o (C * D)" + for a b :: "'a::comm_monoid_mult" apply (auto simp add: elt_set_times_def set_times_def) apply (rule_tac x = "ba * bb" in exI) apply (auto simp add: ac_simps) @@ -214,12 +231,12 @@ apply (auto simp add: ac_simps) done -lemma set_times_rearrange2: - "(a::'a::semigroup_mult) *o (b *o C) = (a * b) *o C" +lemma set_times_rearrange2: "a *o (b *o C) = (a * b) *o C" + for a b :: "'a::semigroup_mult" by (auto simp add: elt_set_times_def mult.assoc) -lemma set_times_rearrange3: - "((a::'a::semigroup_mult) *o B) * C = a *o (B * C)" +lemma set_times_rearrange3: "(a *o B) * C = a *o (B * C)" + for a :: "'a::semigroup_mult" apply (auto simp add: elt_set_times_def set_times_def) apply (blast intro: ac_simps) apply (rule_tac x = "a * aa" in exI) @@ -230,8 +247,8 @@ apply (auto simp add: ac_simps) done -theorem set_times_rearrange4: - "C * ((a::'a::comm_monoid_mult) *o D) = a *o (C * D)" +theorem set_times_rearrange4: "C * (a *o D) = a *o (C * D)" + for a :: "'a::comm_monoid_mult" apply (auto simp add: elt_set_times_def set_times_def ac_simps) apply (rule_tac x = "aa * ba" in exI) apply (auto simp add: ac_simps) @@ -243,13 +260,15 @@ lemma set_times_mono [intro]: "C \ D \ a *o C \ a *o D" by (auto simp add: elt_set_times_def) -lemma set_times_mono2 [intro]: "(C::'a::times set) \ D \ E \ F \ C * E \ D * F" +lemma set_times_mono2 [intro]: "C \ D \ E \ F \ C * E \ D * F" + for C D E F :: "'a::times set" by (auto simp add: set_times_def) lemma set_times_mono3 [intro]: "a \ C \ a *o D \ C * D" by (auto simp add: elt_set_times_def set_times_def) -lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C \ a *o D \ D * C" +lemma set_times_mono4 [intro]: "a \ C \ a *o D \ D * C" + for a :: "'a::comm_monoid_mult" by (auto simp add: elt_set_times_def set_times_def ac_simps) lemma set_times_mono5: "a \ C \ B \ D \ a *o B \ C * D" @@ -276,30 +295,31 @@ apply auto done -lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) \ C \ x \ a *o D \ x \ D * C" +lemma set_times_mono4_b: "a \ C \ x \ a *o D \ x \ D * C" + for a x :: "'a::comm_monoid_mult" apply (frule set_times_mono4) apply auto done -lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" +lemma set_one_times [simp]: "1 *o C = C" + for C :: "'a::comm_monoid_mult set" by (auto simp add: elt_set_times_def) -lemma set_times_plus_distrib: - "(a::'a::semiring) *o (b +o C) = (a * b) +o (a *o C)" +lemma set_times_plus_distrib: "a *o (b +o C) = (a * b) +o (a *o C)" + for a b :: "'a::semiring" by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs) -lemma set_times_plus_distrib2: - "(a::'a::semiring) *o (B + C) = (a *o B) + (a *o C)" +lemma set_times_plus_distrib2: "a *o (B + C) = (a *o B) + (a *o C)" + for a :: "'a::semiring" apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) apply blast apply (rule_tac x = "b + bb" in exI) apply (auto simp add: ring_distribs) done -lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) * D \ a *o D + C * D" - apply (auto simp add: - elt_set_plus_def elt_set_times_def set_times_def - set_plus_def ring_distribs) +lemma set_times_plus_distrib3: "(a +o C) * D \ a *o D + C * D" + for a :: "'a::semiring" + apply (auto simp: elt_set_plus_def elt_set_times_def set_times_def set_plus_def ring_distribs) apply auto done @@ -307,23 +327,25 @@ set_times_plus_distrib set_times_plus_distrib2 -lemma set_neg_intro: "(a::'a::ring_1) \ (- 1) *o C \ - a \ C" +lemma set_neg_intro: "a \ (- 1) *o C \ - a \ C" + for a :: "'a::ring_1" by (auto simp add: elt_set_times_def) -lemma set_neg_intro2: "(a::'a::ring_1) \ C \ - a \ (- 1) *o C" +lemma set_neg_intro2: "a \ C \ - a \ (- 1) *o C" + for a :: "'a::ring_1" by (auto simp add: elt_set_times_def) lemma set_plus_image: "S + T = (\(x, y). x + y) ` (S \ T)" - unfolding set_plus_def by (fastforce simp: image_iff) + by (fastforce simp: set_plus_def image_iff) lemma set_times_image: "S * T = (\(x, y). x * y) ` (S \ T)" - unfolding set_times_def by (fastforce simp: image_iff) + by (fastforce simp: set_times_def image_iff) lemma finite_set_plus: "finite s \ finite t \ finite (s + t)" - unfolding set_plus_image by simp + by (simp add: set_plus_image) lemma finite_set_times: "finite s \ finite t \ finite (s * t)" - unfolding set_times_image by simp + by (simp add: set_times_image) lemma set_setsum_alt: assumes fin: "finite I"