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