author | wenzelm |
Tue, 09 May 2023 19:47:11 +0200 | |
changeset 78006 | 2587b492664a |
parent 77003 | ab905b5bb206 |
child 79559 | cd9ede8488af |
permissions | -rw-r--r-- |
77003 | 1 |
(* Title: HOL/ex/BigO.thy |
2 |
Authors: Jeremy Avigad and Kevin Donnelly; proofs tidied by LCP |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
3 |
*) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
4 |
|
60500 | 5 |
section \<open>Big O notation\<close> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
6 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
7 |
theory BigO |
63485 | 8 |
imports |
9 |
Complex_Main |
|
77003 | 10 |
"HOL-Library.Function_Algebras" |
11 |
"HOL-Library.Set_Algebras" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
12 |
begin |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
13 |
|
60500 | 14 |
text \<open> |
63473 | 15 |
This library is designed to support asymptotic ``big O'' calculations, |
16 |
i.e.~reasoning with expressions of the form \<open>f = O(g)\<close> and \<open>f = g + O(h)\<close>. |
|
76987 | 17 |
An earlier version of this library is described in detail in \<^cite>\<open>"Avigad-Donnelly"\<close>. |
63473 | 18 |
|
19 |
The main changes in this version are as follows: |
|
17199
59c1bfc81d91
moved lemmas that require the HOL-Complex logic image to Complex/ex/BigO_Complex.thy;
wenzelm
parents:
16961
diff
changeset
|
20 |
|
63473 | 21 |
\<^item> We have eliminated the \<open>O\<close> operator on sets. (Most uses of this seem |
22 |
to be inessential.) |
|
23 |
\<^item> We no longer use \<open>+\<close> as output syntax for \<open>+o\<close> |
|
24 |
\<^item> Lemmas involving \<open>sumr\<close> have been replaced by more general lemmas |
|
64267 | 25 |
involving `\<open>sum\<close>. |
63473 | 26 |
\<^item> The library has been expanded, with e.g.~support for expressions of |
27 |
the form \<open>f < g + O(h)\<close>. |
|
17199
59c1bfc81d91
moved lemmas that require the HOL-Complex logic image to Complex/ex/BigO_Complex.thy;
wenzelm
parents:
16961
diff
changeset
|
28 |
|
63473 | 29 |
Note also since the Big O library includes rules that demonstrate set |
30 |
inclusion, to use the automated reasoners effectively with the library one |
|
31 |
should redeclare the theorem \<open>subsetI\<close> as an intro rule, rather than as an |
|
32 |
\<open>intro!\<close> rule, for example, using \<^theory_text>\<open>declare subsetI [del, intro]\<close>. |
|
60500 | 33 |
\<close> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
34 |
|
63473 | 35 |
|
60500 | 36 |
subsection \<open>Definitions\<close> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
37 |
|
55821 | 38 |
definition bigo :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(1O'(_'))") |
61945 | 39 |
where "O(f:: 'a \<Rightarrow> 'b) = {h. \<exists>c. \<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>}" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
40 |
|
55821 | 41 |
lemma bigo_pos_const: |
61945 | 42 |
"(\<exists>c::'a::linordered_idom. \<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) \<longleftrightarrow> |
43 |
(\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))" |
|
77003 | 44 |
by (metis (no_types, opaque_lifting) abs_ge_zero abs_not_less_zero abs_of_nonneg dual_order.trans |
76786 | 45 |
mult_1 zero_less_abs_iff zero_less_mult_iff zero_less_one) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
46 |
|
61945 | 47 |
lemma bigo_alt_def: "O(f) = {h. \<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>)}" |
22665 | 48 |
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
|
49 |
|
55821 | 50 |
lemma bigo_elt_subset [intro]: "f \<in> O(g) \<Longrightarrow> O(f) \<le> O(g)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
51 |
apply (auto simp add: bigo_alt_def) |
77003 | 52 |
by (metis (no_types, opaque_lifting) mult.assoc mult_le_cancel_iff2 order.trans |
76786 | 53 |
zero_less_mult_iff) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
54 |
|
55821 | 55 |
lemma bigo_refl [intro]: "f \<in> O(f)" |
76786 | 56 |
using bigo_def comm_monoid_mult_class.mult_1 dual_order.eq_iff by blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
57 |
|
55821 | 58 |
lemma bigo_zero: "0 \<in> O(g)" |
76786 | 59 |
using bigo_def mult_le_cancel_left1 by fastforce |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
60 |
|
55821 | 61 |
lemma bigo_zero2: "O(\<lambda>x. 0) = {\<lambda>x. 0}" |
62 |
by (auto simp add: bigo_def) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
63 |
|
55821 | 64 |
lemma bigo_plus_self_subset [intro]: "O(f) + O(f) \<subseteq> O(f)" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25592
diff
changeset
|
65 |
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
|
66 |
apply (rule_tac x = "c + ca" in exI) |
76786 | 67 |
by (smt (verit, best) abs_triangle_ineq add_mono add_pos_pos comm_semiring_class.distrib dual_order.trans) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
68 |
|
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
69 |
lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" |
76786 | 70 |
by (simp add: antisym bigo_plus_self_subset bigo_zero set_zero_plus2) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
71 |
|
55821 | 72 |
lemma bigo_plus_subset [intro]: "O(f + g) \<subseteq> O(f) + O(g)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
73 |
apply (rule subsetI) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25592
diff
changeset
|
74 |
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
|
75 |
apply (subst bigo_pos_const [symmetric])+ |
61945 | 76 |
apply (rule_tac x = "\<lambda>n. if \<bar>g n\<bar> \<le> \<bar>f n\<bar> then x n else 0" in exI) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
77 |
apply (rule conjI) |
63473 | 78 |
apply (rule_tac x = "c + c" in exI) |
79 |
apply (clarsimp) |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
80 |
apply (smt (verit, ccfv_threshold) mult.commute abs_triangle_ineq add_le_cancel_left dual_order.trans mult.left_commute mult_2 mult_le_cancel_iff2) |
76786 | 81 |
apply (simp add: order_less_le) |
61945 | 82 |
apply (rule_tac x = "\<lambda>n. if \<bar>f n\<bar> < \<bar>g n\<bar> then x n else 0" in exI) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
83 |
apply (rule conjI) |
63473 | 84 |
apply (rule_tac x = "c + c" in exI) |
85 |
apply auto |
|
61945 | 86 |
apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> \<le> (c + c) * \<bar>g xa\<bar>") |
76786 | 87 |
apply (metis mult_2 order.trans) |
88 |
apply simp |
|
22665 | 89 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
90 |
|
55821 | 91 |
lemma bigo_plus_subset2 [intro]: "A \<subseteq> O(f) \<Longrightarrow> B \<subseteq> O(f) \<Longrightarrow> A + B \<subseteq> O(f)" |
76786 | 92 |
using bigo_plus_idemp set_plus_mono2 by blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
93 |
|
55821 | 94 |
lemma bigo_plus_eq: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. 0 \<le> g x \<Longrightarrow> O(f + g) = O(f) + O(g)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
95 |
apply (rule equalityI) |
63473 | 96 |
apply (rule bigo_plus_subset) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25592
diff
changeset
|
97 |
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
|
98 |
apply clarify |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
99 |
apply (rule_tac x = "max c ca" in exI) |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
100 |
by (smt (verit, del_insts) add.commute abs_triangle_ineq add_mono_thms_linordered_field(3) distrib_left less_max_iff_disj linorder_not_less max.orderE max_mult_distrib_right order_le_less) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
101 |
|
55821 | 102 |
lemma bigo_bounded_alt: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> c * g x \<Longrightarrow> f \<in> O(g)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
103 |
by (simp add: bigo_def) (metis abs_mult abs_of_nonneg order_trans) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
104 |
|
55821 | 105 |
lemma bigo_bounded: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> g x \<Longrightarrow> f \<in> O(g)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
106 |
by (metis bigo_bounded_alt mult_1) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
107 |
|
55821 | 108 |
lemma bigo_bounded2: "\<forall>x. lb x \<le> f x \<Longrightarrow> \<forall>x. f x \<le> lb x + g x \<Longrightarrow> f \<in> lb +o O(g)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
109 |
by (simp add: add.commute bigo_bounded diff_le_eq set_minus_imp_plus) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
110 |
|
61945 | 111 |
lemma bigo_abs: "(\<lambda>x. \<bar>f x\<bar>) =o O(f)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
112 |
by (smt (verit, del_insts) abs_abs bigo_def bigo_refl mem_Collect_eq) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
113 |
|
61945 | 114 |
lemma bigo_abs2: "f =o O(\<lambda>x. \<bar>f x\<bar>)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
115 |
by (smt (verit, del_insts) abs_abs bigo_def bigo_refl mem_Collect_eq) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
116 |
|
61945 | 117 |
lemma bigo_abs3: "O(f) = O(\<lambda>x. \<bar>f x\<bar>)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
118 |
using bigo_abs bigo_abs2 bigo_elt_subset by blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
119 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
120 |
lemma bigo_abs4: assumes "f =o g +o O(h)" shows "(\<lambda>x. \<bar>f x\<bar>) =o (\<lambda>x. \<bar>g x\<bar>) +o O(h)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
121 |
proof - |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
122 |
{ assume *: "f - g \<in> O(h)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
123 |
have "(\<lambda>x. \<bar>f x\<bar> - \<bar>g x\<bar>) =o O(\<lambda>x. \<bar>\<bar>f x\<bar> - \<bar>g x\<bar>\<bar>)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
124 |
by (rule bigo_abs2) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
125 |
also have "\<dots> \<subseteq> O(\<lambda>x. \<bar>f x - g x\<bar>)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
126 |
by (simp add: abs_triangle_ineq3 bigo_bounded bigo_elt_subset) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
127 |
also have "\<dots> \<subseteq> O(f - g)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
128 |
using bigo_abs3 by fastforce |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
129 |
also from * have "\<dots> \<subseteq> O(h)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
130 |
by (rule bigo_elt_subset) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
131 |
finally have "(\<lambda>x. \<bar>f x\<bar> - \<bar>g x\<bar>) \<in> O(h)" . } |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
132 |
then show ?thesis |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
133 |
by (smt (verit) assms bigo_alt_def fun_diff_def mem_Collect_eq set_minus_imp_plus set_plus_imp_minus) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
134 |
qed |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
135 |
|
61945 | 136 |
lemma bigo_abs5: "f =o O(g) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) =o O(g)" |
63473 | 137 |
by (auto simp: bigo_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
138 |
|
63473 | 139 |
lemma bigo_elt_subset2 [intro]: |
140 |
assumes *: "f \<in> g +o O(h)" |
|
141 |
shows "O(f) \<subseteq> O(g) + O(h)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
142 |
proof - |
63473 | 143 |
note * |
144 |
also have "g +o O(h) \<subseteq> O(g) + O(h)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
145 |
by (auto del: subsetI) |
61945 | 146 |
also have "\<dots> = O(\<lambda>x. \<bar>g x\<bar>) + O(\<lambda>x. \<bar>h x\<bar>)" |
63473 | 147 |
by (subst bigo_abs3 [symmetric])+ (rule refl) |
61945 | 148 |
also have "\<dots> = O((\<lambda>x. \<bar>g x\<bar>) + (\<lambda>x. \<bar>h x\<bar>))" |
55821 | 149 |
by (rule bigo_plus_eq [symmetric]) auto |
150 |
finally have "f \<in> \<dots>" . |
|
151 |
then have "O(f) \<subseteq> \<dots>" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
152 |
by (elim bigo_elt_subset) |
61945 | 153 |
also have "\<dots> = O(\<lambda>x. \<bar>g x\<bar>) + O(\<lambda>x. \<bar>h x\<bar>)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
154 |
by (rule bigo_plus_eq, auto) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
155 |
finally show ?thesis |
68406 | 156 |
by (simp flip: bigo_abs3) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
157 |
qed |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
158 |
|
55821 | 159 |
lemma bigo_mult [intro]: "O(f)*O(g) \<subseteq> O(f * g)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
160 |
apply (rule subsetI) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
161 |
apply (subst bigo_def) |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
162 |
apply (clarsimp 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
|
163 |
apply (rule_tac x = "c * ca" in exI) |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
164 |
by (smt (verit, ccfv_threshold) mult.commute mult.assoc abs_ge_zero abs_mult dual_order.trans mult_mono) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
165 |
|
55821 | 166 |
lemma bigo_mult2 [intro]: "f *o O(g) \<subseteq> O(f * g)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
167 |
by (metis bigo_mult bigo_refl dual_order.trans mult.commute set_times_mono4) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
168 |
|
55821 | 169 |
lemma bigo_mult3: "f \<in> O(h) \<Longrightarrow> g \<in> O(j) \<Longrightarrow> f * g \<in> O(h * j)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
170 |
using bigo_mult mult.commute mult.commute set_times_intro subsetD by blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
171 |
|
55821 | 172 |
lemma bigo_mult4 [intro]: "f \<in> k +o O(h) \<Longrightarrow> g * f \<in> (g * k) +o O(g * h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
173 |
by (metis bigo_mult3 bigo_refl left_diff_distrib' mult.commute set_minus_imp_plus set_plus_imp_minus) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
174 |
|
41528 | 175 |
lemma bigo_mult5: |
55821 | 176 |
fixes f :: "'a \<Rightarrow> 'b::linordered_field" |
177 |
assumes "\<forall>x. f x \<noteq> 0" |
|
178 |
shows "O(f * g) \<subseteq> f *o O(g)" |
|
41528 | 179 |
proof |
180 |
fix h |
|
55821 | 181 |
assume "h \<in> O(f * g)" |
182 |
then have "(\<lambda>x. 1 / (f x)) * h \<in> (\<lambda>x. 1 / f x) *o O(f * g)" |
|
41528 | 183 |
by auto |
55821 | 184 |
also have "\<dots> \<subseteq> O((\<lambda>x. 1 / f x) * (f * g))" |
41528 | 185 |
by (rule bigo_mult2) |
55821 | 186 |
also have "(\<lambda>x. 1 / f x) * (f * g) = g" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
187 |
using assms by auto |
55821 | 188 |
finally have "(\<lambda>x. (1::'b) / f x) * h \<in> O(g)" . |
189 |
then have "f * ((\<lambda>x. (1::'b) / f x) * h) \<in> f *o O(g)" |
|
41528 | 190 |
by auto |
55821 | 191 |
also have "f * ((\<lambda>x. (1::'b) / f x) * h) = h" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
192 |
by (simp add: assms times_fun_def) |
55821 | 193 |
finally show "h \<in> f *o O(g)" . |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
194 |
qed |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
195 |
|
63473 | 196 |
lemma bigo_mult6: "\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = f *o O(g)" |
197 |
for f :: "'a \<Rightarrow> 'b::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
198 |
by (simp add: bigo_mult2 bigo_mult5 subset_antisym) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
199 |
|
63473 | 200 |
lemma bigo_mult7: "\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<subseteq> O(f) * O(g)" |
201 |
for f :: "'a \<Rightarrow> 'b::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
202 |
by (metis bigo_mult6 bigo_refl mult.commute set_times_mono4) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
203 |
|
63473 | 204 |
lemma bigo_mult8: "\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f) * O(g)" |
205 |
for f :: "'a \<Rightarrow> 'b::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
206 |
by (simp add: bigo_mult bigo_mult7 subset_antisym) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
207 |
|
55821 | 208 |
lemma bigo_minus [intro]: "f \<in> O(g) \<Longrightarrow> - f \<in> O(g)" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25592
diff
changeset
|
209 |
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
|
210 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
211 |
lemma bigo_minus2: |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
212 |
assumes "f \<in> g +o O(h)" shows "- f \<in> -g +o O(h)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
213 |
proof - |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
214 |
have "- f + g \<in> O(h)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
215 |
by (metis assms bigo_minus minus_diff_eq set_plus_imp_minus uminus_add_conv_diff) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
216 |
then show ?thesis |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
217 |
by (simp add: set_minus_imp_plus) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
218 |
qed |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
219 |
|
55821 | 220 |
lemma bigo_minus3: "O(- f) = O(f)" |
41528 | 221 |
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
|
222 |
|
63473 | 223 |
lemma bigo_plus_absorb_lemma1: |
224 |
assumes *: "f \<in> O(g)" |
|
225 |
shows "f +o O(g) \<subseteq> O(g)" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
226 |
using assms bigo_plus_idemp set_plus_mono4 by blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
227 |
|
63473 | 228 |
lemma bigo_plus_absorb_lemma2: |
229 |
assumes *: "f \<in> O(g)" |
|
230 |
shows "O(g) \<subseteq> f +o O(g)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
231 |
proof - |
63473 | 232 |
from * have "- f \<in> O(g)" |
233 |
by auto |
|
234 |
then have "- f +o O(g) \<subseteq> O(g)" |
|
235 |
by (elim bigo_plus_absorb_lemma1) |
|
236 |
then have "f +o (- f +o O(g)) \<subseteq> f +o O(g)" |
|
237 |
by auto |
|
238 |
also have "f +o (- f +o O(g)) = O(g)" |
|
239 |
by (simp add: set_plus_rearranges) |
|
240 |
finally show ?thesis . |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
241 |
qed |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
242 |
|
55821 | 243 |
lemma bigo_plus_absorb [simp]: "f \<in> O(g) \<Longrightarrow> f +o O(g) = O(g)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
244 |
by (simp add: bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 subset_antisym) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
245 |
|
55821 | 246 |
lemma bigo_plus_absorb2 [intro]: "f \<in> O(g) \<Longrightarrow> A \<subseteq> O(g) \<Longrightarrow> f +o A \<subseteq> O(g)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
247 |
using bigo_plus_absorb set_plus_mono by blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
248 |
|
55821 | 249 |
lemma bigo_add_commute_imp: "f \<in> g +o O(h) \<Longrightarrow> g \<in> f +o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
250 |
by (metis bigo_minus minus_diff_eq set_minus_imp_plus set_plus_imp_minus) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
251 |
|
55821 | 252 |
lemma bigo_add_commute: "f \<in> g +o O(h) \<longleftrightarrow> g \<in> f +o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
253 |
using bigo_add_commute_imp by blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
254 |
|
55821 | 255 |
lemma bigo_const1: "(\<lambda>x. c) \<in> O(\<lambda>x. 1)" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
256 |
by (auto simp add: bigo_def ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
257 |
|
55821 | 258 |
lemma bigo_const2 [intro]: "O(\<lambda>x. c) \<subseteq> O(\<lambda>x. 1)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
259 |
by (metis bigo_elt_subset bigo_const1) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
260 |
|
63473 | 261 |
lemma bigo_const3: "c \<noteq> 0 \<Longrightarrow> (\<lambda>x. 1) \<in> O(\<lambda>x. c)" |
262 |
for c :: "'a::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
263 |
by (metis bigo_bounded_alt le_numeral_extra(4) nonzero_divide_eq_eq zero_less_one_class.zero_le_one) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
264 |
|
63473 | 265 |
lemma bigo_const4: "c \<noteq> 0 \<Longrightarrow> O(\<lambda>x. 1) \<subseteq> O(\<lambda>x. c)" |
266 |
for c :: "'a::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
267 |
by (metis bigo_elt_subset bigo_const3) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
268 |
|
63473 | 269 |
lemma bigo_const [simp]: "c \<noteq> 0 \<Longrightarrow> O(\<lambda>x. c) = O(\<lambda>x. 1)" |
270 |
for c :: "'a::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
271 |
by (metis equalityI bigo_const2 bigo_const4) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
272 |
|
55821 | 273 |
lemma bigo_const_mult1: "(\<lambda>x. c * f x) \<in> O(f)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
274 |
by (smt (z3) abs_mult bigo_def bigo_refl mem_Collect_eq mult.left_commute mult_commute_abs) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
275 |
|
55821 | 276 |
lemma bigo_const_mult2: "O(\<lambda>x. c * f x) \<subseteq> O(f)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
277 |
by (metis bigo_elt_subset bigo_const_mult1) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
278 |
|
63473 | 279 |
lemma bigo_const_mult3: "c \<noteq> 0 \<Longrightarrow> f \<in> O(\<lambda>x. c * f x)" |
280 |
for c :: "'a::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
281 |
by (simp add: bigo_def) (metis abs_mult field_class.field_divide_inverse mult.commute nonzero_divide_eq_eq order_refl) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
282 |
|
63473 | 283 |
lemma bigo_const_mult4: "c \<noteq> 0 \<Longrightarrow> O(f) \<subseteq> O(\<lambda>x. c * f x)" |
284 |
for c :: "'a::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
285 |
by (simp add: bigo_const_mult3 bigo_elt_subset) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
286 |
|
63473 | 287 |
lemma bigo_const_mult [simp]: "c \<noteq> 0 \<Longrightarrow> O(\<lambda>x. c * f x) = O(f)" |
288 |
for c :: "'a::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
289 |
by (simp add: bigo_const_mult2 bigo_const_mult4 subset_antisym) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
290 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
291 |
lemma bigo_const_mult5 [simp]: "(\<lambda>x. c) *o O(f) = O(f)" if "c \<noteq> 0" |
63473 | 292 |
for c :: "'a::linordered_field" |
77003 | 293 |
proof |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
294 |
show "O(f) \<subseteq> (\<lambda>x. c) *o O(f)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
295 |
using that |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
296 |
apply (clarsimp simp add: bigo_def elt_set_times_def func_times) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
297 |
apply (rule_tac x = "\<lambda>y. inverse c * x y" in exI) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
298 |
apply (simp add: mult.assoc [symmetric] abs_mult) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
299 |
apply (rule_tac x = "\<bar>inverse c\<bar> * ca" in exI) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
300 |
apply auto |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
301 |
done |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
302 |
have "O(\<lambda>x. c * f x) \<subseteq> O(f)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
303 |
by (simp add: bigo_const_mult2) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
304 |
then show "(\<lambda>x. c) *o O(f) \<subseteq> O(f)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
305 |
using order_trans[OF bigo_mult2] by (force simp add: times_fun_def) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
306 |
qed |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
307 |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
308 |
|
55821 | 309 |
lemma bigo_const_mult6 [intro]: "(\<lambda>x. c) *o O(f) \<subseteq> O(f)" |
41528 | 310 |
apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) |
61945 | 311 |
apply (rule_tac x = "ca * \<bar>c\<bar>" in exI) |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
312 |
by (smt (verit, ccfv_SIG) ab_semigroup_mult_class.mult_ac(1) abs_abs abs_le_self_iff abs_mult le_cases3 mult.commute mult_left_mono) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
313 |
|
63473 | 314 |
lemma bigo_const_mult7 [intro]: |
315 |
assumes *: "f =o O(g)" |
|
316 |
shows "(\<lambda>x. c * f x) =o O(g)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
317 |
proof - |
63473 | 318 |
from * have "(\<lambda>x. c) * f =o (\<lambda>x. c) *o O(g)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
319 |
by auto |
55821 | 320 |
also have "(\<lambda>x. c) * f = (\<lambda>x. c * f x)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
321 |
by (simp add: func_times) |
55821 | 322 |
also have "(\<lambda>x. c) *o O(g) \<subseteq> O(g)" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
323 |
by (auto del: subsetI) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
324 |
finally show ?thesis . |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
325 |
qed |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
326 |
|
55821 | 327 |
lemma bigo_compose1: "f =o O(g) \<Longrightarrow> (\<lambda>x. f (k x)) =o O(\<lambda>x. g (k x))" |
63473 | 328 |
by (auto simp: bigo_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
329 |
|
63473 | 330 |
lemma bigo_compose2: "f =o g +o O(h) \<Longrightarrow> (\<lambda>x. f (k x)) =o (\<lambda>x. g (k x)) +o O(\<lambda>x. h(k x))" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
331 |
by (smt (verit, best) set_minus_plus bigo_def fun_diff_def mem_Collect_eq) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
332 |
|
22665 | 333 |
|
64267 | 334 |
subsection \<open>Sum\<close> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
335 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
336 |
lemma bigo_sum_main: |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
337 |
assumes "\<forall>x. \<forall>y \<in> A x. 0 \<le> h x y" and "\<forall>x. \<forall>y \<in> A x. \<bar>f x y\<bar> \<le> c * h x y" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
338 |
shows "(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
339 |
proof - |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
340 |
have "(\<Sum>i\<in>A x. \<bar>f x i\<bar>) \<le> \<bar>c\<bar> * sum (h x) (A x)" for x |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
341 |
by (smt (verit, ccfv_threshold) assms abs_mult_pos abs_of_nonneg abs_of_nonpos dual_order.trans le_cases3 neg_0_le_iff_le sum_distrib_left sum_mono) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
342 |
then show ?thesis |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
343 |
using assms by (fastforce simp add: bigo_def sum_nonneg intro: order_trans [OF sum_abs]) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
344 |
qed |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
345 |
|
64267 | 346 |
lemma bigo_sum1: "\<forall>x y. 0 \<le> h x y \<Longrightarrow> |
61945 | 347 |
\<exists>c. \<forall>x y. \<bar>f x y\<bar> \<le> c * h x y \<Longrightarrow> |
55821 | 348 |
(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
349 |
by (metis (no_types) bigo_sum_main) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
350 |
|
64267 | 351 |
lemma bigo_sum2: "\<forall>y. 0 \<le> h y \<Longrightarrow> |
61945 | 352 |
\<exists>c. \<forall>y. \<bar>f y\<bar> \<le> c * (h y) \<Longrightarrow> |
55821 | 353 |
(\<lambda>x. \<Sum>y \<in> A x. f y) =o O(\<lambda>x. \<Sum>y \<in> A x. h y)" |
64267 | 354 |
by (rule bigo_sum1) auto |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
355 |
|
64267 | 356 |
lemma bigo_sum3: "f =o O(h) \<Longrightarrow> |
61945 | 357 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)" |
64267 | 358 |
apply (rule bigo_sum1) |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
359 |
using abs_ge_zero apply blast |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
360 |
apply (clarsimp simp: bigo_def) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
361 |
by (smt (verit, ccfv_threshold) abs_mult abs_not_less_zero mult.left_commute mult_le_cancel_left) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
362 |
|
64267 | 363 |
lemma bigo_sum4: "f =o g +o O(h) \<Longrightarrow> |
55821 | 364 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o |
365 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * g (k x y)) +o |
|
61945 | 366 |
O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
367 |
using bigo_sum3 [of "f-g" h l k A] |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
368 |
apply (simp add: algebra_simps sum_subtractf) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
369 |
by (smt (verit) bigo_alt_def minus_apply set_minus_imp_plus set_plus_imp_minus mem_Collect_eq) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
370 |
|
64267 | 371 |
lemma bigo_sum5: "f =o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow> |
55821 | 372 |
\<forall>x. 0 \<le> h x \<Longrightarrow> |
373 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o |
|
374 |
O(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y))" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
375 |
using bigo_sum3 [of f h l k A] by simp |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
376 |
|
64267 | 377 |
lemma bigo_sum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow> |
55821 | 378 |
\<forall>x. 0 \<le> h x \<Longrightarrow> |
379 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o |
|
380 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * g (k x y)) +o |
|
381 |
O(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y))" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
382 |
using bigo_sum5 [of "f-g" h l k A] |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
383 |
apply (simp add: algebra_simps sum_subtractf) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
384 |
by (smt (verit, del_insts) bigo_alt_def set_minus_imp_plus minus_apply set_plus_imp_minus mem_Collect_eq) |
22665 | 385 |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
386 |
|
60500 | 387 |
subsection \<open>Misc useful stuff\<close> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
388 |
|
55821 | 389 |
lemma bigo_useful_add: "f =o O(h) \<Longrightarrow> g =o O(h) \<Longrightarrow> f + g =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
390 |
using bigo_plus_idemp set_plus_intro by blast |
55821 | 391 |
|
63473 | 392 |
lemma bigo_useful_const_mult: "c \<noteq> 0 \<Longrightarrow> (\<lambda>x. c) * f =o O(h) \<Longrightarrow> f =o O(h)" |
393 |
for c :: "'a::linordered_field" |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
394 |
using bigo_elt_subset bigo_mult6 by fastforce |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
395 |
|
55821 | 396 |
lemma bigo_fix: "(\<lambda>x::nat. f (x + 1)) =o O(\<lambda>x. h (x + 1)) \<Longrightarrow> f 0 = 0 \<Longrightarrow> f =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
397 |
by (simp add: bigo_alt_def) (metis abs_eq_0_iff abs_ge_zero abs_mult abs_of_pos not0_implies_Suc) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
398 |
|
55821 | 399 |
lemma bigo_fix2: |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
400 |
"(\<lambda>x. f ((x::nat) + 1)) =o (\<lambda>x. g(x + 1)) +o O(\<lambda>x. h(x + 1)) \<Longrightarrow> |
55821 | 401 |
f 0 = g 0 \<Longrightarrow> f =o g +o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
402 |
apply (rule set_minus_imp_plus [OF bigo_fix]) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
403 |
apply (smt (verit, del_insts) bigo_alt_def fun_diff_def set_plus_imp_minus mem_Collect_eq) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
404 |
apply simp |
22665 | 405 |
done |
406 |
||
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
407 |
|
60500 | 408 |
subsection \<open>Less than or equal to\<close> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
409 |
|
55821 | 410 |
definition lesso :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" (infixl "<o" 70) |
411 |
where "f <o g = (\<lambda>x. max (f x - g x) 0)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
412 |
|
61945 | 413 |
lemma bigo_lesseq1: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> \<le> \<bar>f x\<bar> \<Longrightarrow> g =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
414 |
by (smt (verit, del_insts) bigo_def mem_Collect_eq order_trans) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
415 |
|
61945 | 416 |
lemma bigo_lesseq2: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> \<le> f x \<Longrightarrow> g =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
417 |
by (metis (mono_tags, lifting) abs_ge_zero abs_of_nonneg bigo_lesseq1 dual_order.trans) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
418 |
|
55821 | 419 |
lemma bigo_lesseq3: "f =o O(h) \<Longrightarrow> \<forall>x. 0 \<le> g x \<Longrightarrow> \<forall>x. g x \<le> f x \<Longrightarrow> g =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
420 |
by (meson bigo_bounded bigo_elt_subset subsetD) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
421 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
422 |
lemma bigo_lesseq4: "f =o O(h) \<Longrightarrow> \<forall>x. 0 \<le> g x \<Longrightarrow> \<forall>x. g x \<le> \<bar>f x\<bar> \<Longrightarrow> g =o O(h)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
423 |
by (metis abs_of_nonneg bigo_lesseq1) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
424 |
|
55821 | 425 |
lemma bigo_lesso1: "\<forall>x. f x \<le> g x \<Longrightarrow> f <o g =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
426 |
by (smt (verit, del_insts) abs_ge_zero add_0 bigo_abs3 bigo_bounded diff_le_eq lesso_def max_def order_refl) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
427 |
|
63473 | 428 |
lemma bigo_lesso2: "f =o g +o O(h) \<Longrightarrow> \<forall>x. 0 \<le> k x \<Longrightarrow> \<forall>x. k x \<le> f x \<Longrightarrow> k <o g =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
429 |
unfolding lesso_def |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
430 |
apply (rule bigo_lesseq4 [of "f-g"]) |
63473 | 431 |
apply (erule set_plus_imp_minus) |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
432 |
using max.cobounded2 apply blast |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
433 |
by (smt (verit) abs_ge_zero abs_of_nonneg diff_ge_0_iff_ge diff_mono diff_self fun_diff_def order_refl max.coboundedI2 max_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
434 |
|
63473 | 435 |
lemma bigo_lesso3: "f =o g +o O(h) \<Longrightarrow> \<forall>x. 0 \<le> k x \<Longrightarrow> \<forall>x. g x \<le> k x \<Longrightarrow> f <o k =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
436 |
unfolding lesso_def |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
437 |
apply (rule bigo_lesseq4 [of "f-g"]) |
63473 | 438 |
apply (erule set_plus_imp_minus) |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
439 |
using max.cobounded2 apply blast |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
440 |
by (smt (verit) abs_eq_iff abs_ge_zero abs_if abs_minus_le_zero diff_left_mono fun_diff_def le_max_iff_disj order.trans order_eq_refl) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
441 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
442 |
lemma bigo_lesso4: |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
443 |
fixes k :: "'a \<Rightarrow> 'b::linordered_field" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
444 |
assumes f: "f <o g =o O(k)" and g: "g =o h +o O(k)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
445 |
shows "f <o h =o O(k)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
446 |
proof - |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
447 |
have "g - h \<in> O(k)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
448 |
by (simp add: g set_plus_imp_minus) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
449 |
then have "(\<lambda>x. \<bar>g x - h x\<bar>) \<in> O(k)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
450 |
using bigo_abs5 by force |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
451 |
then have \<section>: "(\<lambda>x. max (f x - g x) 0) + (\<lambda>x. \<bar>g x - h x\<bar>) \<in> O(k)" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
452 |
by (metis (mono_tags, lifting) bigo_lesseq1 bigo_useful_add dual_order.eq_iff f lesso_def) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
453 |
have "\<bar>max (f x - h x) 0\<bar> \<le> ((\<lambda>x. max (f x - g x) 0) + (\<lambda>x. \<bar>g x - h x\<bar>)) x" for x |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
454 |
by (auto simp add: func_plus fun_diff_def algebra_simps split: split_max abs_split) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
455 |
then show ?thesis |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
456 |
by (smt (verit, ccfv_SIG) \<section> bigo_lesseq2 lesso_def) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
457 |
qed |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
458 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
459 |
|
77003 | 460 |
lemma bigo_lesso5: |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
461 |
assumes "f <o g =o O(h)" shows "\<exists>C. \<forall>x. f x \<le> g x + C * \<bar>h x\<bar>" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
462 |
proof - |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
463 |
obtain c where "0 < c" and c: "\<And>x. f x - g x \<le> c * \<bar>h x\<bar>" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
464 |
using assms by (auto simp: lesso_def bigo_alt_def) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
465 |
have "\<bar>max (f x - g x) 0\<bar> = max (f x - g x) 0" for x |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
466 |
by (auto simp add: algebra_simps) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
467 |
then show ?thesis |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
468 |
by (metis c add.commute diff_le_eq) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
469 |
qed |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
470 |
|
55821 | 471 |
lemma lesso_add: "f <o g =o O(h) \<Longrightarrow> k <o l =o O(h) \<Longrightarrow> (f + k) <o (g + l) =o O(h)" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
472 |
unfolding lesso_def |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
473 |
using bigo_useful_add by (fastforce split: split_max intro: bigo_lesseq3) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
474 |
|
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
475 |
lemma bigo_LIMSEQ1: "f \<longlonglongrightarrow> 0" if f: "f =o O(g)" and g: "g \<longlonglongrightarrow> 0" |
63473 | 476 |
for f g :: "nat \<Rightarrow> real" |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
477 |
proof - |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
478 |
{ fix r::real |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
479 |
assume "0 < r" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
480 |
obtain c::real where "0 < c" and rc: "\<And>x. \<bar>f x\<bar> \<le> c * \<bar>g x\<bar>" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
481 |
using f by (auto simp: LIMSEQ_iff bigo_alt_def) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
482 |
with g \<open>0 < r\<close> obtain no where "\<forall>n\<ge>no. \<bar>g n\<bar> < r/c" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
483 |
by (fastforce simp: LIMSEQ_iff) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
484 |
then have "\<exists>no. \<forall>n\<ge>no. \<bar>f n\<bar> < r" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
485 |
by (metis \<open>0 < c\<close> mult.commute order_le_less_trans pos_less_divide_eq rc) } |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
486 |
then show ?thesis |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
487 |
by (auto simp: LIMSEQ_iff) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
488 |
qed |
29786 | 489 |
|
77003 | 490 |
lemma bigo_LIMSEQ2: |
77001
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
491 |
fixes f g :: "nat \<Rightarrow> real" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
492 |
assumes "f =o g +o O(h)" "h \<longlonglongrightarrow> 0" and f: "f \<longlonglongrightarrow> a" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
493 |
shows "g \<longlonglongrightarrow> a" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
494 |
proof - |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
495 |
have "f - g \<longlonglongrightarrow> 0" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
496 |
using assms bigo_LIMSEQ1 set_plus_imp_minus by blast |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
497 |
then have "(\<lambda>n. f n - g n) \<longlonglongrightarrow> 0" |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
498 |
by (simp add: fun_diff_def) |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
499 |
then show ?thesis |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
500 |
using Lim_transform_eq f by blast |
68f1fc53c8fd
tidy up of this messy and obsolete theory
paulson <lp15@cam.ac.uk>
parents:
76987
diff
changeset
|
501 |
qed |
29786 | 502 |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
503 |
end |