| author | Manuel Eberl <eberlm@in.tum.de> | 
| Wed, 07 Apr 2021 11:05:00 +0200 | |
| changeset 73537 | 56db8559eadb | 
| parent 68406 | 6beb45f6cf67 | 
| child 76786 | 50672d2d78db | 
| permissions | -rw-r--r-- | 
| 16932 | 1 | (* Title: HOL/Library/BigO.thy | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 2 | Authors: Jeremy Avigad and Kevin Donnelly | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 3 | *) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 4 | |
| 60500 | 5 | section \<open>Big O notation\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 6 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 7 | theory BigO | 
| 63485 | 8 | imports | 
| 9 | Complex_Main | |
| 10 | Function_Algebras | |
| 11 | Set_Algebras | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 12 | begin | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 13 | |
| 60500 | 14 | text \<open> | 
| 63473 | 15 | This library is designed to support asymptotic ``big O'' calculations, | 
| 16 | i.e.~reasoning with expressions of the form \<open>f = O(g)\<close> and \<open>f = g + O(h)\<close>. | |
| 17 |   An earlier version of this library is described in detail in @{cite
 | |
| 18 | "Avigad-Donnelly"}. | |
| 19 | ||
| 20 | The main changes in this version are as follows: | |
| 17199 
59c1bfc81d91
moved lemmas that require the HOL-Complex logic image to Complex/ex/BigO_Complex.thy;
 wenzelm parents: 
16961diff
changeset | 21 | |
| 63473 | 22 | \<^item> We have eliminated the \<open>O\<close> operator on sets. (Most uses of this seem | 
| 23 | to be inessential.) | |
| 24 | \<^item> We no longer use \<open>+\<close> as output syntax for \<open>+o\<close> | |
| 25 | \<^item> Lemmas involving \<open>sumr\<close> have been replaced by more general lemmas | |
| 64267 | 26 | involving `\<open>sum\<close>. | 
| 63473 | 27 | \<^item> The library has been expanded, with e.g.~support for expressions of | 
| 28 | the form \<open>f < g + O(h)\<close>. | |
| 17199 
59c1bfc81d91
moved lemmas that require the HOL-Complex logic image to Complex/ex/BigO_Complex.thy;
 wenzelm parents: 
16961diff
changeset | 29 | |
| 63473 | 30 | Note also since the Big O library includes rules that demonstrate set | 
| 31 | inclusion, to use the automated reasoners effectively with the library one | |
| 32 | should redeclare the theorem \<open>subsetI\<close> as an intro rule, rather than as an | |
| 33 | \<open>intro!\<close> rule, for example, using \<^theory_text>\<open>declare subsetI [del, intro]\<close>. | |
| 60500 | 34 | \<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 35 | |
| 63473 | 36 | |
| 60500 | 37 | subsection \<open>Definitions\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 38 | |
| 55821 | 39 | definition bigo :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) set"  ("(1O'(_'))")
 | 
| 61945 | 40 |   where "O(f:: 'a \<Rightarrow> 'b) = {h. \<exists>c. \<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>}"
 | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 41 | |
| 55821 | 42 | lemma bigo_pos_const: | 
| 61945 | 43 | "(\<exists>c::'a::linordered_idom. \<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) \<longleftrightarrow> | 
| 44 | (\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 45 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 46 | apply (case_tac "c = 0") | 
| 63473 | 47 | apply simp | 
| 48 | apply (rule_tac x = "1" in exI) | |
| 49 | apply simp | |
| 61945 | 50 | apply (rule_tac x = "\<bar>c\<bar>" in exI) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 51 | apply auto | 
| 61945 | 52 | apply (subgoal_tac "c * \<bar>f x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>") | 
| 63473 | 53 | apply (erule_tac x = x in allE) | 
| 54 | apply force | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 55 | apply (rule mult_right_mono) | 
| 63473 | 56 | apply (rule abs_ge_self) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 57 | apply (rule abs_ge_zero) | 
| 22665 | 58 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 59 | |
| 61945 | 60 | lemma bigo_alt_def: "O(f) = {h. \<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>)}"
 | 
| 22665 | 61 | by (auto simp add: bigo_def bigo_pos_const) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 62 | |
| 55821 | 63 | lemma bigo_elt_subset [intro]: "f \<in> O(g) \<Longrightarrow> O(f) \<le> O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 64 | apply (auto simp add: bigo_alt_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 65 | apply (rule_tac x = "ca * c" in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 66 | apply (rule conjI) | 
| 63473 | 67 | apply simp | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 68 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 69 | apply (drule_tac x = "xa" in spec)+ | 
| 61945 | 70 | apply (subgoal_tac "ca * \<bar>f xa\<bar> \<le> ca * (c * \<bar>g xa\<bar>)") | 
| 63473 | 71 | apply (erule order_trans) | 
| 72 | apply (simp add: ac_simps) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 73 | apply (rule mult_left_mono, assumption) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 74 | apply (rule order_less_imp_le, assumption) | 
| 22665 | 75 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 76 | |
| 55821 | 77 | lemma bigo_refl [intro]: "f \<in> O(f)" | 
| 63473 | 78 | apply (auto simp add: bigo_def) | 
| 79 | apply (rule_tac x = 1 in exI) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 80 | apply simp | 
| 22665 | 81 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 82 | |
| 55821 | 83 | lemma bigo_zero: "0 \<in> O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 84 | apply (auto simp add: bigo_def func_zero) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 85 | apply (rule_tac x = 0 in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 86 | apply auto | 
| 22665 | 87 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 88 | |
| 55821 | 89 | lemma bigo_zero2: "O(\<lambda>x. 0) = {\<lambda>x. 0}"
 | 
| 90 | by (auto simp add: bigo_def) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 91 | |
| 55821 | 92 | lemma bigo_plus_self_subset [intro]: "O(f) + O(f) \<subseteq> O(f)" | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 93 | apply (auto simp add: bigo_alt_def set_plus_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 94 | apply (rule_tac x = "c + ca" in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 95 | apply auto | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23413diff
changeset | 96 | apply (simp add: ring_distribs func_plus) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 97 | apply (rule order_trans) | 
| 63473 | 98 | apply (rule abs_triangle_ineq) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 99 | apply (rule add_mono) | 
| 63473 | 100 | apply force | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 101 | apply force | 
| 55821 | 102 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 103 | |
| 47445 
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
 krauss parents: 
47108diff
changeset | 104 | lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 105 | apply (rule equalityI) | 
| 63473 | 106 | apply (rule bigo_plus_self_subset) | 
| 55821 | 107 | apply (rule set_zero_plus2) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 108 | apply (rule bigo_zero) | 
| 22665 | 109 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 110 | |
| 55821 | 111 | lemma bigo_plus_subset [intro]: "O(f + g) \<subseteq> O(f) + O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 112 | apply (rule subsetI) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 113 | apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 114 | apply (subst bigo_pos_const [symmetric])+ | 
| 61945 | 115 | apply (rule_tac x = "\<lambda>n. if \<bar>g n\<bar> \<le> \<bar>f n\<bar> then x n else 0" in exI) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 116 | apply (rule conjI) | 
| 63473 | 117 | apply (rule_tac x = "c + c" in exI) | 
| 118 | apply (clarsimp) | |
| 119 | apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> (c + c) * \<bar>f xa\<bar>") | |
| 120 | apply (erule_tac x = xa in allE) | |
| 121 | apply (erule order_trans) | |
| 122 | apply (simp) | |
| 123 | apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> c * (\<bar>f xa\<bar> + \<bar>g xa\<bar>)") | |
| 124 | apply (erule order_trans) | |
| 125 | apply (simp add: ring_distribs) | |
| 126 | apply (rule mult_left_mono) | |
| 127 | apply (simp add: abs_triangle_ineq) | |
| 128 | apply (simp add: order_less_le) | |
| 61945 | 129 | apply (rule_tac x = "\<lambda>n. if \<bar>f n\<bar> < \<bar>g n\<bar> then x n else 0" in exI) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 130 | apply (rule conjI) | 
| 63473 | 131 | apply (rule_tac x = "c + c" in exI) | 
| 132 | apply auto | |
| 61945 | 133 | apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> (c + c) * \<bar>g xa\<bar>") | 
| 63473 | 134 | apply (erule_tac x = xa in allE) | 
| 135 | apply (erule order_trans) | |
| 136 | apply simp | |
| 61945 | 137 | apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> c * (\<bar>f xa\<bar> + \<bar>g xa\<bar>)") | 
| 63473 | 138 | apply (erule order_trans) | 
| 139 | apply (simp add: ring_distribs) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 140 | apply (rule mult_left_mono) | 
| 63473 | 141 | apply (rule abs_triangle_ineq) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 142 | apply (simp add: order_less_le) | 
| 22665 | 143 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 144 | |
| 55821 | 145 | lemma bigo_plus_subset2 [intro]: "A \<subseteq> O(f) \<Longrightarrow> B \<subseteq> O(f) \<Longrightarrow> A + B \<subseteq> O(f)" | 
| 146 | apply (subgoal_tac "A + B \<subseteq> O(f) + O(f)") | |
| 63473 | 147 | apply (erule order_trans) | 
| 148 | apply simp | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 149 | apply (auto del: subsetI simp del: bigo_plus_idemp) | 
| 22665 | 150 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 151 | |
| 55821 | 152 | lemma bigo_plus_eq: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. 0 \<le> g x \<Longrightarrow> O(f + g) = O(f) + O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 153 | apply (rule equalityI) | 
| 63473 | 154 | apply (rule bigo_plus_subset) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 155 | apply (simp add: bigo_alt_def set_plus_def func_plus) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 156 | apply clarify | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 157 | apply (rule_tac x = "max c ca" in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 158 | apply (rule conjI) | 
| 63473 | 159 | apply (subgoal_tac "c \<le> max c ca") | 
| 160 | apply (erule order_less_le_trans) | |
| 161 | apply assumption | |
| 162 | apply (rule max.cobounded1) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 163 | apply clarify | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 164 | apply (drule_tac x = "xa" in spec)+ | 
| 55821 | 165 | apply (subgoal_tac "0 \<le> f xa + g xa") | 
| 63473 | 166 | apply (simp add: ring_distribs) | 
| 167 | apply (subgoal_tac "\<bar>a xa + b xa\<bar> \<le> \<bar>a xa\<bar> + \<bar>b xa\<bar>") | |
| 168 | apply (subgoal_tac "\<bar>a xa\<bar> + \<bar>b xa\<bar> \<le> max c ca * f xa + max c ca * g xa") | |
| 169 | apply force | |
| 170 | apply (rule add_mono) | |
| 171 | apply (subgoal_tac "c * f xa \<le> max c ca * f xa") | |
| 172 | apply force | |
| 173 | apply (rule mult_right_mono) | |
| 174 | apply (rule max.cobounded1) | |
| 175 | apply assumption | |
| 176 | apply (subgoal_tac "ca * g xa \<le> max c ca * g xa") | |
| 177 | apply force | |
| 178 | apply (rule mult_right_mono) | |
| 179 | apply (rule max.cobounded2) | |
| 180 | apply assumption | |
| 181 | apply (rule abs_triangle_ineq) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 182 | apply (rule add_nonneg_nonneg) | 
| 63473 | 183 | apply assumption+ | 
| 22665 | 184 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 185 | |
| 55821 | 186 | lemma bigo_bounded_alt: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> c * g x \<Longrightarrow> f \<in> O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 187 | apply (auto simp add: bigo_def) | 
| 61945 | 188 | apply (rule_tac x = "\<bar>c\<bar>" in exI) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 189 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 190 | apply (drule_tac x = x in spec)+ | 
| 68406 | 191 | apply (simp flip: abs_mult) | 
| 22665 | 192 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 193 | |
| 55821 | 194 | lemma bigo_bounded: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> g x \<Longrightarrow> f \<in> O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 195 | apply (erule bigo_bounded_alt [of f 1 g]) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 196 | apply simp | 
| 22665 | 197 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 198 | |
| 55821 | 199 | lemma bigo_bounded2: "\<forall>x. lb x \<le> f x \<Longrightarrow> \<forall>x. f x \<le> lb x + g x \<Longrightarrow> f \<in> lb +o O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 200 | apply (rule set_minus_imp_plus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 201 | apply (rule bigo_bounded) | 
| 63473 | 202 | apply (auto simp add: fun_Compl_def func_plus) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 203 | apply (drule_tac x = x in spec)+ | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 204 | apply force | 
| 22665 | 205 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 206 | |
| 61945 | 207 | lemma bigo_abs: "(\<lambda>x. \<bar>f x\<bar>) =o O(f)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 208 | apply (unfold bigo_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 209 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 210 | apply (rule_tac x = 1 in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 211 | apply auto | 
| 22665 | 212 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 213 | |
| 61945 | 214 | lemma bigo_abs2: "f =o O(\<lambda>x. \<bar>f x\<bar>)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 215 | apply (unfold bigo_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 216 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 217 | apply (rule_tac x = 1 in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 218 | apply auto | 
| 22665 | 219 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 220 | |
| 61945 | 221 | lemma bigo_abs3: "O(f) = O(\<lambda>x. \<bar>f x\<bar>)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 222 | apply (rule equalityI) | 
| 63473 | 223 | apply (rule bigo_elt_subset) | 
| 224 | apply (rule bigo_abs2) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 225 | apply (rule bigo_elt_subset) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 226 | apply (rule bigo_abs) | 
| 22665 | 227 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 228 | |
| 61945 | 229 | lemma bigo_abs4: "f =o g +o O(h) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) =o (\<lambda>x. \<bar>g x\<bar>) +o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 230 | apply (drule set_plus_imp_minus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 231 | apply (rule set_minus_imp_plus) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 232 | apply (subst fun_diff_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 233 | proof - | 
| 63473 | 234 | assume *: "f - g \<in> O(h)" | 
| 61945 | 235 | have "(\<lambda>x. \<bar>f x\<bar> - \<bar>g x\<bar>) =o O(\<lambda>x. \<bar>\<bar>f x\<bar> - \<bar>g x\<bar>\<bar>)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 236 | by (rule bigo_abs2) | 
| 61945 | 237 | also have "\<dots> \<subseteq> O(\<lambda>x. \<bar>f x - g x\<bar>)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 238 | apply (rule bigo_elt_subset) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 239 | apply (rule bigo_bounded) | 
| 63473 | 240 | apply force | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 241 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 242 | apply (rule abs_triangle_ineq3) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 243 | done | 
| 55821 | 244 | also have "\<dots> \<subseteq> O(f - g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 245 | apply (rule bigo_elt_subset) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 246 | apply (subst fun_diff_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 247 | apply (rule bigo_abs) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 248 | done | 
| 63473 | 249 | also from * have "\<dots> \<subseteq> O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 250 | by (rule bigo_elt_subset) | 
| 61945 | 251 | finally show "(\<lambda>x. \<bar>f x\<bar> - \<bar>g x\<bar>) \<in> O(h)". | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 252 | qed | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 253 | |
| 61945 | 254 | lemma bigo_abs5: "f =o O(g) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) =o O(g)" | 
| 63473 | 255 | by (auto simp: bigo_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 256 | |
| 63473 | 257 | lemma bigo_elt_subset2 [intro]: | 
| 258 | assumes *: "f \<in> g +o O(h)" | |
| 259 | shows "O(f) \<subseteq> O(g) + O(h)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 260 | proof - | 
| 63473 | 261 | note * | 
| 262 | also have "g +o O(h) \<subseteq> O(g) + O(h)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 263 | by (auto del: subsetI) | 
| 61945 | 264 | also have "\<dots> = O(\<lambda>x. \<bar>g x\<bar>) + O(\<lambda>x. \<bar>h x\<bar>)" | 
| 63473 | 265 | by (subst bigo_abs3 [symmetric])+ (rule refl) | 
| 61945 | 266 | also have "\<dots> = O((\<lambda>x. \<bar>g x\<bar>) + (\<lambda>x. \<bar>h x\<bar>))" | 
| 55821 | 267 | by (rule bigo_plus_eq [symmetric]) auto | 
| 268 | finally have "f \<in> \<dots>" . | |
| 269 | then have "O(f) \<subseteq> \<dots>" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 270 | by (elim bigo_elt_subset) | 
| 61945 | 271 | also have "\<dots> = O(\<lambda>x. \<bar>g x\<bar>) + O(\<lambda>x. \<bar>h x\<bar>)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 272 | by (rule bigo_plus_eq, auto) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 273 | finally show ?thesis | 
| 68406 | 274 | by (simp flip: bigo_abs3) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 275 | qed | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 276 | |
| 55821 | 277 | lemma bigo_mult [intro]: "O(f)*O(g) \<subseteq> O(f * g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 278 | apply (rule subsetI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 279 | apply (subst bigo_def) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 280 | apply (auto simp add: bigo_alt_def set_times_def func_times) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 281 | apply (rule_tac x = "c * ca" in exI) | 
| 55821 | 282 | apply (rule allI) | 
| 283 | apply (erule_tac x = x in allE)+ | |
| 61945 | 284 | apply (subgoal_tac "c * ca * \<bar>f x * g x\<bar> = (c * \<bar>f x\<bar>) * (ca * \<bar>g x\<bar>)") | 
| 63473 | 285 | apply (erule ssubst) | 
| 286 | apply (subst abs_mult) | |
| 287 | apply (rule mult_mono) | |
| 288 | apply assumption+ | |
| 289 | apply auto | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 290 | apply (simp add: ac_simps abs_mult) | 
| 22665 | 291 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 292 | |
| 55821 | 293 | lemma bigo_mult2 [intro]: "f *o O(g) \<subseteq> O(f * g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 294 | apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 295 | apply (rule_tac x = c in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 296 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 297 | apply (drule_tac x = x in spec) | 
| 61945 | 298 | apply (subgoal_tac "\<bar>f x\<bar> * \<bar>b x\<bar> \<le> \<bar>f x\<bar> * (c * \<bar>g x\<bar>)") | 
| 63473 | 299 | apply (force simp add: ac_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 300 | apply (rule mult_left_mono, assumption) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 301 | apply (rule abs_ge_zero) | 
| 22665 | 302 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 303 | |
| 55821 | 304 | lemma bigo_mult3: "f \<in> O(h) \<Longrightarrow> g \<in> O(j) \<Longrightarrow> f * g \<in> O(h * j)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 305 | apply (rule subsetD) | 
| 63473 | 306 | apply (rule bigo_mult) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 307 | apply (erule set_times_intro, assumption) | 
| 22665 | 308 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 309 | |
| 55821 | 310 | lemma bigo_mult4 [intro]: "f \<in> k +o O(h) \<Longrightarrow> g * f \<in> (g * k) +o O(g * h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 311 | apply (drule set_plus_imp_minus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 312 | apply (rule set_minus_imp_plus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 313 | apply (drule bigo_mult3 [where g = g and j = g]) | 
| 63473 | 314 | apply (auto simp add: algebra_simps) | 
| 22665 | 315 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 316 | |
| 41528 | 317 | lemma bigo_mult5: | 
| 55821 | 318 | fixes f :: "'a \<Rightarrow> 'b::linordered_field" | 
| 319 | assumes "\<forall>x. f x \<noteq> 0" | |
| 320 | shows "O(f * g) \<subseteq> f *o O(g)" | |
| 41528 | 321 | proof | 
| 322 | fix h | |
| 55821 | 323 | assume "h \<in> O(f * g)" | 
| 324 | then have "(\<lambda>x. 1 / (f x)) * h \<in> (\<lambda>x. 1 / f x) *o O(f * g)" | |
| 41528 | 325 | by auto | 
| 55821 | 326 | also have "\<dots> \<subseteq> O((\<lambda>x. 1 / f x) * (f * g))" | 
| 41528 | 327 | by (rule bigo_mult2) | 
| 55821 | 328 | also have "(\<lambda>x. 1 / f x) * (f * g) = g" | 
| 329 | apply (simp add: func_times) | |
| 41528 | 330 | apply (rule ext) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 331 | apply (simp add: assms nonzero_divide_eq_eq ac_simps) | 
| 41528 | 332 | done | 
| 55821 | 333 | finally have "(\<lambda>x. (1::'b) / f x) * h \<in> O(g)" . | 
| 334 | then have "f * ((\<lambda>x. (1::'b) / f x) * h) \<in> f *o O(g)" | |
| 41528 | 335 | by auto | 
| 55821 | 336 | also have "f * ((\<lambda>x. (1::'b) / f x) * h) = h" | 
| 337 | apply (simp add: func_times) | |
| 41528 | 338 | apply (rule ext) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 339 | apply (simp add: assms nonzero_divide_eq_eq ac_simps) | 
| 41528 | 340 | done | 
| 55821 | 341 | finally show "h \<in> f *o O(g)" . | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 342 | qed | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 343 | |
| 63473 | 344 | lemma bigo_mult6: "\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = f *o O(g)" | 
| 345 | for f :: "'a \<Rightarrow> 'b::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 346 | apply (rule equalityI) | 
| 63473 | 347 | apply (erule bigo_mult5) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 348 | apply (rule bigo_mult2) | 
| 22665 | 349 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 350 | |
| 63473 | 351 | lemma bigo_mult7: "\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<subseteq> O(f) * O(g)" | 
| 352 | for f :: "'a \<Rightarrow> 'b::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 353 | apply (subst bigo_mult6) | 
| 63473 | 354 | apply assumption | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 355 | apply (rule set_times_mono3) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 356 | apply (rule bigo_refl) | 
| 22665 | 357 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 358 | |
| 63473 | 359 | lemma bigo_mult8: "\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f) * O(g)" | 
| 360 | for f :: "'a \<Rightarrow> 'b::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 361 | apply (rule equalityI) | 
| 63473 | 362 | apply (erule bigo_mult7) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 363 | apply (rule bigo_mult) | 
| 22665 | 364 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 365 | |
| 55821 | 366 | lemma bigo_minus [intro]: "f \<in> O(g) \<Longrightarrow> - f \<in> O(g)" | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 367 | by (auto simp add: bigo_def fun_Compl_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 368 | |
| 55821 | 369 | lemma bigo_minus2: "f \<in> g +o O(h) \<Longrightarrow> - f \<in> -g +o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 370 | apply (rule set_minus_imp_plus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 371 | apply (drule set_plus_imp_minus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 372 | apply (drule bigo_minus) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
47445diff
changeset | 373 | apply simp | 
| 22665 | 374 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 375 | |
| 55821 | 376 | lemma bigo_minus3: "O(- f) = O(f)" | 
| 41528 | 377 | by (auto simp add: bigo_def fun_Compl_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 378 | |
| 63473 | 379 | lemma bigo_plus_absorb_lemma1: | 
| 380 | assumes *: "f \<in> O(g)" | |
| 381 | shows "f +o O(g) \<subseteq> O(g)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 382 | proof - | 
| 63473 | 383 | have "f \<in> O(f)" by auto | 
| 384 | then have "f +o O(g) \<subseteq> O(f) + O(g)" | |
| 385 | by (auto del: subsetI) | |
| 386 | also have "\<dots> \<subseteq> O(g) + O(g)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 387 | proof - | 
| 63473 | 388 | from * have "O(f) \<subseteq> O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 389 | by (auto del: subsetI) | 
| 63473 | 390 | then show ?thesis | 
| 391 | by (auto del: subsetI) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 392 | qed | 
| 63473 | 393 | also have "\<dots> \<subseteq> O(g)" by simp | 
| 394 | finally show ?thesis . | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 395 | qed | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 396 | |
| 63473 | 397 | lemma bigo_plus_absorb_lemma2: | 
| 398 | assumes *: "f \<in> O(g)" | |
| 399 | shows "O(g) \<subseteq> f +o O(g)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 400 | proof - | 
| 63473 | 401 | from * have "- f \<in> O(g)" | 
| 402 | by auto | |
| 403 | then have "- f +o O(g) \<subseteq> O(g)" | |
| 404 | by (elim bigo_plus_absorb_lemma1) | |
| 405 | then have "f +o (- f +o O(g)) \<subseteq> f +o O(g)" | |
| 406 | by auto | |
| 407 | also have "f +o (- f +o O(g)) = O(g)" | |
| 408 | by (simp add: set_plus_rearranges) | |
| 409 | finally show ?thesis . | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 410 | qed | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 411 | |
| 55821 | 412 | lemma bigo_plus_absorb [simp]: "f \<in> O(g) \<Longrightarrow> f +o O(g) = O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 413 | apply (rule equalityI) | 
| 63473 | 414 | apply (erule bigo_plus_absorb_lemma1) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 415 | apply (erule bigo_plus_absorb_lemma2) | 
| 22665 | 416 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 417 | |
| 55821 | 418 | lemma bigo_plus_absorb2 [intro]: "f \<in> O(g) \<Longrightarrow> A \<subseteq> O(g) \<Longrightarrow> f +o A \<subseteq> O(g)" | 
| 419 | apply (subgoal_tac "f +o A \<subseteq> f +o O(g)") | |
| 63473 | 420 | apply force+ | 
| 22665 | 421 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 422 | |
| 55821 | 423 | lemma bigo_add_commute_imp: "f \<in> g +o O(h) \<Longrightarrow> g \<in> f +o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 424 | apply (subst set_minus_plus [symmetric]) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 425 | apply (subgoal_tac "g - f = - (f - g)") | 
| 63473 | 426 | apply (erule ssubst) | 
| 427 | apply (rule bigo_minus) | |
| 428 | apply (subst set_minus_plus) | |
| 429 | apply assumption | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 430 | apply (simp add: ac_simps) | 
| 22665 | 431 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 432 | |
| 55821 | 433 | lemma bigo_add_commute: "f \<in> g +o O(h) \<longleftrightarrow> g \<in> f +o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 434 | apply (rule iffI) | 
| 63473 | 435 | apply (erule bigo_add_commute_imp)+ | 
| 22665 | 436 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 437 | |
| 55821 | 438 | lemma bigo_const1: "(\<lambda>x. c) \<in> O(\<lambda>x. 1)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 439 | by (auto simp add: bigo_def ac_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 440 | |
| 55821 | 441 | lemma bigo_const2 [intro]: "O(\<lambda>x. c) \<subseteq> O(\<lambda>x. 1)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 442 | apply (rule bigo_elt_subset) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 443 | apply (rule bigo_const1) | 
| 22665 | 444 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 445 | |
| 63473 | 446 | lemma bigo_const3: "c \<noteq> 0 \<Longrightarrow> (\<lambda>x. 1) \<in> O(\<lambda>x. c)" | 
| 447 | for c :: "'a::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 448 | apply (simp add: bigo_def) | 
| 61945 | 449 | apply (rule_tac x = "\<bar>inverse c\<bar>" in exI) | 
| 68406 | 450 | apply (simp flip: abs_mult) | 
| 22665 | 451 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 452 | |
| 63473 | 453 | lemma bigo_const4: "c \<noteq> 0 \<Longrightarrow> O(\<lambda>x. 1) \<subseteq> O(\<lambda>x. c)" | 
| 454 | for c :: "'a::linordered_field" | |
| 55821 | 455 | apply (rule bigo_elt_subset) | 
| 456 | apply (rule bigo_const3) | |
| 457 | apply assumption | |
| 458 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 459 | |
| 63473 | 460 | lemma bigo_const [simp]: "c \<noteq> 0 \<Longrightarrow> O(\<lambda>x. c) = O(\<lambda>x. 1)" | 
| 461 | for c :: "'a::linordered_field" | |
| 55821 | 462 | apply (rule equalityI) | 
| 63473 | 463 | apply (rule bigo_const2) | 
| 55821 | 464 | apply (rule bigo_const4) | 
| 465 | apply assumption | |
| 466 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 467 | |
| 55821 | 468 | lemma bigo_const_mult1: "(\<lambda>x. c * f x) \<in> O(f)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 469 | apply (simp add: bigo_def) | 
| 61945 | 470 | apply (rule_tac x = "\<bar>c\<bar>" in exI) | 
| 68406 | 471 | apply (auto simp flip: abs_mult) | 
| 22665 | 472 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 473 | |
| 55821 | 474 | lemma bigo_const_mult2: "O(\<lambda>x. c * f x) \<subseteq> O(f)" | 
| 475 | apply (rule bigo_elt_subset) | |
| 476 | apply (rule bigo_const_mult1) | |
| 477 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 478 | |
| 63473 | 479 | lemma bigo_const_mult3: "c \<noteq> 0 \<Longrightarrow> f \<in> O(\<lambda>x. c * f x)" | 
| 480 | for c :: "'a::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 481 | apply (simp add: bigo_def) | 
| 61945 | 482 | apply (rule_tac x = "\<bar>inverse c\<bar>" in exI) | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
58881diff
changeset | 483 | apply (simp add: abs_mult mult.assoc [symmetric]) | 
| 22665 | 484 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 485 | |
| 63473 | 486 | lemma bigo_const_mult4: "c \<noteq> 0 \<Longrightarrow> O(f) \<subseteq> O(\<lambda>x. c * f x)" | 
| 487 | for c :: "'a::linordered_field" | |
| 55821 | 488 | apply (rule bigo_elt_subset) | 
| 489 | apply (rule bigo_const_mult3) | |
| 490 | apply assumption | |
| 491 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 492 | |
| 63473 | 493 | lemma bigo_const_mult [simp]: "c \<noteq> 0 \<Longrightarrow> O(\<lambda>x. c * f x) = O(f)" | 
| 494 | for c :: "'a::linordered_field" | |
| 55821 | 495 | apply (rule equalityI) | 
| 63473 | 496 | apply (rule bigo_const_mult2) | 
| 55821 | 497 | apply (erule bigo_const_mult4) | 
| 498 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 499 | |
| 63473 | 500 | lemma bigo_const_mult5 [simp]: "c \<noteq> 0 \<Longrightarrow> (\<lambda>x. c) *o O(f) = O(f)" | 
| 501 | for c :: "'a::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 502 | apply (auto del: subsetI) | 
| 63473 | 503 | apply (rule order_trans) | 
| 504 | apply (rule bigo_mult2) | |
| 505 | apply (simp add: func_times) | |
| 41528 | 506 | apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) | 
| 55821 | 507 | apply (rule_tac x = "\<lambda>y. inverse c * x y" in exI) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 508 | apply (simp add: mult.assoc [symmetric] abs_mult) | 
| 61945 | 509 | apply (rule_tac x = "\<bar>inverse c\<bar> * ca" in exI) | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
58881diff
changeset | 510 | apply auto | 
| 22665 | 511 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 512 | |
| 55821 | 513 | lemma bigo_const_mult6 [intro]: "(\<lambda>x. c) *o O(f) \<subseteq> O(f)" | 
| 41528 | 514 | apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) | 
| 61945 | 515 | apply (rule_tac x = "ca * \<bar>c\<bar>" in exI) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 516 | apply (rule allI) | 
| 61945 | 517 | apply (subgoal_tac "ca * \<bar>c\<bar> * \<bar>f x\<bar> = \<bar>c\<bar> * (ca * \<bar>f x\<bar>)") | 
| 63473 | 518 | apply (erule ssubst) | 
| 519 | apply (subst abs_mult) | |
| 520 | apply (rule mult_left_mono) | |
| 521 | apply (erule spec) | |
| 522 | apply simp | |
| 523 | apply (simp add: ac_simps) | |
| 22665 | 524 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 525 | |
| 63473 | 526 | lemma bigo_const_mult7 [intro]: | 
| 527 | assumes *: "f =o O(g)" | |
| 528 | shows "(\<lambda>x. c * f x) =o O(g)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 529 | proof - | 
| 63473 | 530 | from * have "(\<lambda>x. c) * f =o (\<lambda>x. c) *o O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 531 | by auto | 
| 55821 | 532 | also have "(\<lambda>x. c) * f = (\<lambda>x. c * f x)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 533 | by (simp add: func_times) | 
| 55821 | 534 | also have "(\<lambda>x. c) *o O(g) \<subseteq> O(g)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 535 | by (auto del: subsetI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 536 | finally show ?thesis . | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 537 | qed | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 538 | |
| 55821 | 539 | lemma bigo_compose1: "f =o O(g) \<Longrightarrow> (\<lambda>x. f (k x)) =o O(\<lambda>x. g (k x))" | 
| 63473 | 540 | by (auto simp: bigo_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 541 | |
| 63473 | 542 | lemma bigo_compose2: "f =o g +o O(h) \<Longrightarrow> (\<lambda>x. f (k x)) =o (\<lambda>x. g (k x)) +o O(\<lambda>x. h(k x))" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
47445diff
changeset | 543 | apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus) | 
| 55821 | 544 | apply (drule bigo_compose1) | 
| 545 | apply (simp add: fun_diff_def) | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
47445diff
changeset | 546 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 547 | |
| 22665 | 548 | |
| 64267 | 549 | subsection \<open>Sum\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 550 | |
| 64267 | 551 | lemma bigo_sum_main: "\<forall>x. \<forall>y \<in> A x. 0 \<le> h x y \<Longrightarrow> | 
| 61945 | 552 | \<exists>c. \<forall>x. \<forall>y \<in> A x. \<bar>f x y\<bar> \<le> c * h x y \<Longrightarrow> | 
| 55821 | 553 | (\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 554 | apply (auto simp add: bigo_def) | 
| 61945 | 555 | apply (rule_tac x = "\<bar>c\<bar>" in exI) | 
| 17199 
59c1bfc81d91
moved lemmas that require the HOL-Complex logic image to Complex/ex/BigO_Complex.thy;
 wenzelm parents: 
16961diff
changeset | 556 | apply (subst abs_of_nonneg) back back | 
| 64267 | 557 | apply (rule sum_nonneg) | 
| 63473 | 558 | apply force | 
| 64267 | 559 | apply (subst sum_distrib_left) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 560 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 561 | apply (rule order_trans) | 
| 64267 | 562 | apply (rule sum_abs) | 
| 563 | apply (rule sum_mono) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 564 | apply (rule order_trans) | 
| 63473 | 565 | apply (drule spec)+ | 
| 566 | apply (drule bspec)+ | |
| 567 | apply assumption+ | |
| 568 | apply (drule bspec) | |
| 569 | apply assumption+ | |
| 55821 | 570 | apply (rule mult_right_mono) | 
| 63473 | 571 | apply (rule abs_ge_self) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 572 | apply force | 
| 22665 | 573 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 574 | |
| 64267 | 575 | lemma bigo_sum1: "\<forall>x y. 0 \<le> h x y \<Longrightarrow> | 
| 61945 | 576 | \<exists>c. \<forall>x y. \<bar>f x y\<bar> \<le> c * h x y \<Longrightarrow> | 
| 55821 | 577 | (\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)" | 
| 64267 | 578 | apply (rule bigo_sum_main) | 
| 63473 | 579 | apply force | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 580 | apply clarsimp | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 581 | apply (rule_tac x = c in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 582 | apply force | 
| 22665 | 583 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 584 | |
| 64267 | 585 | lemma bigo_sum2: "\<forall>y. 0 \<le> h y \<Longrightarrow> | 
| 61945 | 586 | \<exists>c. \<forall>y. \<bar>f y\<bar> \<le> c * (h y) \<Longrightarrow> | 
| 55821 | 587 | (\<lambda>x. \<Sum>y \<in> A x. f y) =o O(\<lambda>x. \<Sum>y \<in> A x. h y)" | 
| 64267 | 588 | by (rule bigo_sum1) auto | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 589 | |
| 64267 | 590 | lemma bigo_sum3: "f =o O(h) \<Longrightarrow> | 
| 61945 | 591 | (\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)" | 
| 64267 | 592 | apply (rule bigo_sum1) | 
| 63473 | 593 | apply (rule allI)+ | 
| 594 | apply (rule abs_ge_zero) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 595 | apply (unfold bigo_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 596 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 597 | apply (rule_tac x = c in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 598 | apply (rule allI)+ | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 599 | apply (subst abs_mult)+ | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 600 | apply (subst mult.left_commute) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 601 | apply (rule mult_left_mono) | 
| 63473 | 602 | apply (erule spec) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 603 | apply (rule abs_ge_zero) | 
| 22665 | 604 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 605 | |
| 64267 | 606 | lemma bigo_sum4: "f =o g +o O(h) \<Longrightarrow> | 
| 55821 | 607 | (\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o | 
| 608 | (\<lambda>x. \<Sum>y \<in> A x. l x y * g (k x y)) +o | |
| 61945 | 609 | O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 610 | apply (rule set_minus_imp_plus) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 611 | apply (subst fun_diff_def) | 
| 64267 | 612 | apply (subst sum_subtractf [symmetric]) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 613 | apply (subst right_diff_distrib [symmetric]) | 
| 64267 | 614 | apply (rule bigo_sum3) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 615 | apply (subst fun_diff_def [symmetric]) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 616 | apply (erule set_plus_imp_minus) | 
| 22665 | 617 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 618 | |
| 64267 | 619 | lemma bigo_sum5: "f =o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow> | 
| 55821 | 620 | \<forall>x. 0 \<le> h x \<Longrightarrow> | 
| 621 | (\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o | |
| 622 | O(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y))" | |
| 623 | apply (subgoal_tac "(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y)) = | |
| 61945 | 624 | (\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)") | 
| 63473 | 625 | apply (erule ssubst) | 
| 64267 | 626 | apply (erule bigo_sum3) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 627 | apply (rule ext) | 
| 64267 | 628 | apply (rule sum.cong) | 
| 63473 | 629 | apply (rule refl) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 630 | apply (subst abs_of_nonneg) | 
| 63473 | 631 | apply auto | 
| 22665 | 632 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 633 | |
| 64267 | 634 | lemma bigo_sum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow> | 
| 55821 | 635 | \<forall>x. 0 \<le> h x \<Longrightarrow> | 
| 636 | (\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o | |
| 637 | (\<lambda>x. \<Sum>y \<in> A x. l x y * g (k x y)) +o | |
| 638 | O(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y))" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 639 | apply (rule set_minus_imp_plus) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 640 | apply (subst fun_diff_def) | 
| 64267 | 641 | apply (subst sum_subtractf [symmetric]) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 642 | apply (subst right_diff_distrib [symmetric]) | 
| 64267 | 643 | apply (rule bigo_sum5) | 
| 63473 | 644 | apply (subst fun_diff_def [symmetric]) | 
| 645 | apply (drule set_plus_imp_minus) | |
| 646 | apply auto | |
| 22665 | 647 | done | 
| 648 | ||
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 649 | |
| 60500 | 650 | subsection \<open>Misc useful stuff\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 651 | |
| 55821 | 652 | lemma bigo_useful_intro: "A \<subseteq> O(f) \<Longrightarrow> B \<subseteq> O(f) \<Longrightarrow> A + B \<subseteq> O(f)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 653 | apply (subst bigo_plus_idemp [symmetric]) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 654 | apply (rule set_plus_mono2) | 
| 63473 | 655 | apply assumption+ | 
| 22665 | 656 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 657 | |
| 55821 | 658 | lemma bigo_useful_add: "f =o O(h) \<Longrightarrow> g =o O(h) \<Longrightarrow> f + g =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 659 | apply (subst bigo_plus_idemp [symmetric]) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 660 | apply (rule set_plus_intro) | 
| 63473 | 661 | apply assumption+ | 
| 22665 | 662 | done | 
| 55821 | 663 | |
| 63473 | 664 | lemma bigo_useful_const_mult: "c \<noteq> 0 \<Longrightarrow> (\<lambda>x. c) * f =o O(h) \<Longrightarrow> f =o O(h)" | 
| 665 | for c :: "'a::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 666 | apply (rule subsetD) | 
| 63473 | 667 | apply (subgoal_tac "(\<lambda>x. 1 / c) *o O(h) \<subseteq> O(h)") | 
| 668 | apply assumption | |
| 669 | apply (rule bigo_const_mult6) | |
| 55821 | 670 | apply (subgoal_tac "f = (\<lambda>x. 1 / c) * ((\<lambda>x. c) * f)") | 
| 63473 | 671 | apply (erule ssubst) | 
| 672 | apply (erule set_times_intro2) | |
| 23413 
5caa2710dd5b
tuned laws for cancellation in divisions for fields.
 nipkow parents: 
23373diff
changeset | 673 | apply (simp add: func_times) | 
| 22665 | 674 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 675 | |
| 55821 | 676 | lemma bigo_fix: "(\<lambda>x::nat. f (x + 1)) =o O(\<lambda>x. h (x + 1)) \<Longrightarrow> f 0 = 0 \<Longrightarrow> f =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 677 | apply (simp add: bigo_alt_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 678 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 679 | apply (rule_tac x = c in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 680 | apply auto | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 681 | apply (case_tac "x = 0") | 
| 63473 | 682 | apply simp | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 683 | apply (subgoal_tac "x = Suc (x - 1)") | 
| 63473 | 684 | apply (erule ssubst) back | 
| 685 | apply (erule spec) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 686 | apply simp | 
| 22665 | 687 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 688 | |
| 55821 | 689 | lemma bigo_fix2: | 
| 690 | "(\<lambda>x. f ((x::nat) + 1)) =o (\<lambda>x. g(x + 1)) +o O(\<lambda>x. h(x + 1)) \<Longrightarrow> | |
| 691 | f 0 = g 0 \<Longrightarrow> f =o g +o O(h)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 692 | apply (rule set_minus_imp_plus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 693 | apply (rule bigo_fix) | 
| 63473 | 694 | apply (subst fun_diff_def) | 
| 695 | apply (subst fun_diff_def [symmetric]) | |
| 696 | apply (rule set_plus_imp_minus) | |
| 697 | apply simp | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 698 | apply (simp add: fun_diff_def) | 
| 22665 | 699 | done | 
| 700 | ||
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 701 | |
| 60500 | 702 | subsection \<open>Less than or equal to\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 703 | |
| 55821 | 704 | definition lesso :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"  (infixl "<o" 70)
 | 
| 705 | where "f <o g = (\<lambda>x. max (f x - g x) 0)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 706 | |
| 61945 | 707 | lemma bigo_lesseq1: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> \<le> \<bar>f x\<bar> \<Longrightarrow> g =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 708 | apply (unfold bigo_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 709 | apply clarsimp | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 710 | apply (rule_tac x = c in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 711 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 712 | apply (rule order_trans) | 
| 63473 | 713 | apply (erule spec)+ | 
| 22665 | 714 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 715 | |
| 61945 | 716 | lemma bigo_lesseq2: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> \<le> f x \<Longrightarrow> g =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 717 | apply (erule bigo_lesseq1) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 718 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 719 | apply (drule_tac x = x in spec) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 720 | apply (rule order_trans) | 
| 63473 | 721 | apply assumption | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 722 | apply (rule abs_ge_self) | 
| 22665 | 723 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 724 | |
| 55821 | 725 | lemma bigo_lesseq3: "f =o O(h) \<Longrightarrow> \<forall>x. 0 \<le> g x \<Longrightarrow> \<forall>x. g x \<le> f x \<Longrightarrow> g =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 726 | apply (erule bigo_lesseq2) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 727 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 728 | apply (subst abs_of_nonneg) | 
| 63473 | 729 | apply (erule spec)+ | 
| 22665 | 730 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 731 | |
| 55821 | 732 | lemma bigo_lesseq4: "f =o O(h) \<Longrightarrow> | 
| 61945 | 733 | \<forall>x. 0 \<le> g x \<Longrightarrow> \<forall>x. g x \<le> \<bar>f x\<bar> \<Longrightarrow> g =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 734 | apply (erule bigo_lesseq1) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 735 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 736 | apply (subst abs_of_nonneg) | 
| 63473 | 737 | apply (erule spec)+ | 
| 22665 | 738 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 739 | |
| 55821 | 740 | lemma bigo_lesso1: "\<forall>x. f x \<le> g x \<Longrightarrow> f <o g =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 741 | apply (unfold lesso_def) | 
| 55821 | 742 | apply (subgoal_tac "(\<lambda>x. max (f x - g x) 0) = 0") | 
| 63473 | 743 | apply (erule ssubst) | 
| 744 | apply (rule bigo_zero) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 745 | apply (unfold func_zero) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 746 | apply (rule ext) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 747 | apply (simp split: split_max) | 
| 22665 | 748 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 749 | |
| 63473 | 750 | lemma bigo_lesso2: "f =o g +o O(h) \<Longrightarrow> \<forall>x. 0 \<le> k x \<Longrightarrow> \<forall>x. k x \<le> f x \<Longrightarrow> k <o g =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 751 | apply (unfold lesso_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 752 | apply (rule bigo_lesseq4) | 
| 63473 | 753 | apply (erule set_plus_imp_minus) | 
| 754 | apply (rule allI) | |
| 755 | apply (rule max.cobounded2) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 756 | apply (rule allI) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 757 | apply (subst fun_diff_def) | 
| 55821 | 758 | apply (case_tac "0 \<le> k x - g x") | 
| 63473 | 759 | apply simp | 
| 760 | apply (subst abs_of_nonneg) | |
| 761 | apply (drule_tac x = x in spec) back | |
| 762 | apply (simp add: algebra_simps) | |
| 763 | apply (subst diff_conv_add_uminus)+ | |
| 764 | apply (rule add_right_mono) | |
| 765 | apply (erule spec) | |
| 55821 | 766 | apply (rule order_trans) | 
| 63473 | 767 | prefer 2 | 
| 768 | apply (rule abs_ge_zero) | |
| 29667 | 769 | apply (simp add: algebra_simps) | 
| 22665 | 770 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 771 | |
| 63473 | 772 | lemma bigo_lesso3: "f =o g +o O(h) \<Longrightarrow> \<forall>x. 0 \<le> k x \<Longrightarrow> \<forall>x. g x \<le> k x \<Longrightarrow> f <o k =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 773 | apply (unfold lesso_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 774 | apply (rule bigo_lesseq4) | 
| 63473 | 775 | apply (erule set_plus_imp_minus) | 
| 776 | apply (rule allI) | |
| 777 | apply (rule max.cobounded2) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 778 | apply (rule allI) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 779 | apply (subst fun_diff_def) | 
| 55821 | 780 | apply (case_tac "0 \<le> f x - k x") | 
| 63473 | 781 | apply simp | 
| 782 | apply (subst abs_of_nonneg) | |
| 783 | apply (drule_tac x = x in spec) back | |
| 784 | apply (simp add: algebra_simps) | |
| 785 | apply (subst diff_conv_add_uminus)+ | |
| 786 | apply (rule add_left_mono) | |
| 787 | apply (rule le_imp_neg_le) | |
| 788 | apply (erule spec) | |
| 55821 | 789 | apply (rule order_trans) | 
| 63473 | 790 | prefer 2 | 
| 791 | apply (rule abs_ge_zero) | |
| 29667 | 792 | apply (simp add: algebra_simps) | 
| 22665 | 793 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 794 | |
| 63473 | 795 | lemma bigo_lesso4: "f <o g =o O(k) \<Longrightarrow> g =o h +o O(k) \<Longrightarrow> f <o h =o O(k)" | 
| 796 | for k :: "'a \<Rightarrow> 'b::linordered_field" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 797 | apply (unfold lesso_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 798 | apply (drule set_plus_imp_minus) | 
| 17199 
59c1bfc81d91
moved lemmas that require the HOL-Complex logic image to Complex/ex/BigO_Complex.thy;
 wenzelm parents: 
16961diff
changeset | 799 | apply (drule bigo_abs5) back | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25592diff
changeset | 800 | apply (simp add: fun_diff_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 801 | apply (drule bigo_useful_add) | 
| 63473 | 802 | apply assumption | 
| 17199 
59c1bfc81d91
moved lemmas that require the HOL-Complex logic image to Complex/ex/BigO_Complex.thy;
 wenzelm parents: 
16961diff
changeset | 803 | apply (erule bigo_lesseq2) back | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 804 | apply (rule allI) | 
| 55821 | 805 | apply (auto simp add: func_plus fun_diff_def algebra_simps split: split_max abs_split) | 
| 22665 | 806 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 807 | |
| 61945 | 808 | lemma bigo_lesso5: "f <o g =o O(h) \<Longrightarrow> \<exists>C. \<forall>x. f x \<le> g x + C * \<bar>h x\<bar>" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 809 | apply (simp only: lesso_def bigo_alt_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 810 | apply clarsimp | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 811 | apply (rule_tac x = c in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 812 | apply (rule allI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 813 | apply (drule_tac x = x in spec) | 
| 61945 | 814 | apply (subgoal_tac "\<bar>max (f x - g x) 0\<bar> = max (f x - g x) 0") | 
| 63473 | 815 | apply (clarsimp simp add: algebra_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 816 | apply (rule abs_of_nonneg) | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
54230diff
changeset | 817 | apply (rule max.cobounded2) | 
| 22665 | 818 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 819 | |
| 55821 | 820 | lemma lesso_add: "f <o g =o O(h) \<Longrightarrow> k <o l =o O(h) \<Longrightarrow> (f + k) <o (g + l) =o O(h)" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 821 | apply (unfold lesso_def) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 822 | apply (rule bigo_lesseq3) | 
| 63473 | 823 | apply (erule bigo_useful_add) | 
| 824 | apply assumption | |
| 825 | apply (force split: split_max) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 826 | apply (auto split: split_max simp add: func_plus) | 
| 22665 | 827 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 828 | |
| 63473 | 829 | lemma bigo_LIMSEQ1: "f =o O(g) \<Longrightarrow> g \<longlonglongrightarrow> 0 \<Longrightarrow> f \<longlonglongrightarrow> 0" | 
| 830 | for f g :: "nat \<Rightarrow> real" | |
| 31337 | 831 | apply (simp add: LIMSEQ_iff bigo_alt_def) | 
| 29786 | 832 | apply clarify | 
| 833 | apply (drule_tac x = "r / c" in spec) | |
| 834 | apply (drule mp) | |
| 63473 | 835 | apply simp | 
| 29786 | 836 | apply clarify | 
| 837 | apply (rule_tac x = no in exI) | |
| 838 | apply (rule allI) | |
| 839 | apply (drule_tac x = n in spec)+ | |
| 840 | apply (rule impI) | |
| 841 | apply (drule mp) | |
| 63473 | 842 | apply assumption | 
| 29786 | 843 | apply (rule order_le_less_trans) | 
| 63473 | 844 | apply assumption | 
| 29786 | 845 | apply (rule order_less_le_trans) | 
| 63473 | 846 | apply (subgoal_tac "c * \<bar>g n\<bar> < c * (r / c)") | 
| 847 | apply assumption | |
| 848 | apply (erule mult_strict_left_mono) | |
| 849 | apply assumption | |
| 29786 | 850 | apply simp | 
| 55821 | 851 | done | 
| 29786 | 852 | |
| 63473 | 853 | lemma bigo_LIMSEQ2: "f =o g +o O(h) \<Longrightarrow> h \<longlonglongrightarrow> 0 \<Longrightarrow> f \<longlonglongrightarrow> a \<Longrightarrow> g \<longlonglongrightarrow> a" | 
| 854 | for f g h :: "nat \<Rightarrow> real" | |
| 29786 | 855 | apply (drule set_plus_imp_minus) | 
| 856 | apply (drule bigo_LIMSEQ1) | |
| 63473 | 857 | apply assumption | 
| 29786 | 858 | apply (simp only: fun_diff_def) | 
| 60142 | 859 | apply (erule Lim_transform2) | 
| 29786 | 860 | apply assumption | 
| 55821 | 861 | done | 
| 29786 | 862 | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 863 | end |