author | wenzelm |
Fri, 20 Sep 2024 19:51:08 +0200 | |
changeset 80914 | d97fdabd9e2b |
parent 74690 | 55a4b319b2b9 |
permissions | -rw-r--r-- |
43197 | 1 |
(* Title: HOL/Metis_Examples/Big_O.thy |
2 |
Author: Lawrence C. Paulson, Cambridge University Computer Laboratory |
|
41144 | 3 |
Author: Jasmin Blanchette, TU Muenchen |
23449 | 4 |
|
43197 | 5 |
Metis example featuring the Big O notation. |
23449 | 6 |
*) |
7 |
||
63167 | 8 |
section \<open>Metis Example Featuring the Big O Notation\<close> |
23449 | 9 |
|
43197 | 10 |
theory Big_O |
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41144
diff
changeset
|
11 |
imports |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
64267
diff
changeset
|
12 |
"HOL-Decision_Procs.Dense_Linear_Order" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
64267
diff
changeset
|
13 |
"HOL-Library.Function_Algebras" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
64267
diff
changeset
|
14 |
"HOL-Library.Set_Algebras" |
23449 | 15 |
begin |
16 |
||
63167 | 17 |
subsection \<open>Definitions\<close> |
23449 | 18 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
74690
diff
changeset
|
19 |
definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" (\<open>(1O'(_'))\<close>) where |
61945 | 20 |
"O(f::('a => 'b)) == {h. \<exists>c. \<forall>x. \<bar>h x\<bar> <= c * \<bar>f x\<bar>}" |
23449 | 21 |
|
45575 | 22 |
lemma bigo_pos_const: |
61076 | 23 |
"(\<exists>c::'a::linordered_idom. |
61945 | 24 |
\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) |
25 |
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))" |
|
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
26 |
by (metis (no_types) abs_ge_zero |
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
27 |
algebra_simps mult.comm_neutral |
61824
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61076
diff
changeset
|
28 |
mult_nonpos_nonneg not_le_imp_less order_trans zero_less_one) |
23449 | 29 |
|
36407 | 30 |
(*** Now various verions with an increasing shrink factor ***) |
23449 | 31 |
|
57245 | 32 |
sledgehammer_params [isar_proofs, compress = 1] |
23449 | 33 |
|
45575 | 34 |
lemma |
61076 | 35 |
"(\<exists>c::'a::linordered_idom. |
61945 | 36 |
\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) |
37 |
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))" |
|
23449 | 38 |
apply auto |
39 |
apply (case_tac "c = 0", simp) |
|
40 |
apply (rule_tac x = "1" in exI, simp) |
|
61945 | 41 |
apply (rule_tac x = "\<bar>c\<bar>" in exI, auto) |
36561 | 42 |
proof - |
43 |
fix c :: 'a and x :: 'b |
|
44 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
61076 | 45 |
have F1: "\<forall>x\<^sub>1::'a::linordered_idom. 0 \<le> \<bar>x\<^sub>1\<bar>" by (metis abs_ge_zero) |
46 |
have F2: "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
47 |
have F3: "\<forall>x\<^sub>1 x\<^sub>3. x\<^sub>3 \<le> \<bar>h x\<^sub>1\<bar> \<longrightarrow> x\<^sub>3 \<le> c * \<bar>f x\<^sub>1\<bar>" by (metis A1 order_trans) |
74690 | 48 |
have F4: "\<forall>x\<^sub>2 x\<^sub>3::'a::linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>" by (metis abs_mult) |
49 |
have F5: "\<forall>x\<^sub>3 x\<^sub>1::'a::linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1" by (metis abs_mult_pos) |
|
61076 | 50 |
hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = \<bar>1\<bar> * x\<^sub>1" by (metis F2) |
51 |
hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = x\<^sub>1" by (metis F2 abs_one) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
52 |
hence "\<forall>x\<^sub>3. 0 \<le> \<bar>h x\<^sub>3\<bar> \<longrightarrow> \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis F3) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
53 |
hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis F1) |
61076 | 54 |
hence "\<forall>x\<^sub>3. (0::'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^sub>3\<bar>" by (metis F5) |
55 |
hence "\<forall>x\<^sub>3. (0::'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F4) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
56 |
hence "\<forall>x\<^sub>3. c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F1) |
36561 | 57 |
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1) |
58 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F4) |
|
23449 | 59 |
qed |
60 |
||
57245 | 61 |
sledgehammer_params [isar_proofs, compress = 2] |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25592
diff
changeset
|
62 |
|
45575 | 63 |
lemma |
61076 | 64 |
"(\<exists>c::'a::linordered_idom. |
61945 | 65 |
\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) |
66 |
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))" |
|
23449 | 67 |
apply auto |
68 |
apply (case_tac "c = 0", simp) |
|
69 |
apply (rule_tac x = "1" in exI, simp) |
|
61945 | 70 |
apply (rule_tac x = "\<bar>c\<bar>" in exI, auto) |
36561 | 71 |
proof - |
72 |
fix c :: 'a and x :: 'b |
|
73 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
61076 | 74 |
have F1: "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) |
75 |
have F2: "\<forall>x\<^sub>2 x\<^sub>3::'a::linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>" |
|
36561 | 76 |
by (metis abs_mult) |
61076 | 77 |
have "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_mult_pos abs_one) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
78 |
hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis A1 abs_ge_zero order_trans) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
79 |
hence "\<forall>x\<^sub>3. 0 \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F2 abs_mult_pos) |
36561 | 80 |
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero) |
81 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F2) |
|
23449 | 82 |
qed |
83 |
||
57245 | 84 |
sledgehammer_params [isar_proofs, compress = 3] |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25592
diff
changeset
|
85 |
|
45575 | 86 |
lemma |
61076 | 87 |
"(\<exists>c::'a::linordered_idom. |
61945 | 88 |
\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) |
89 |
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))" |
|
23449 | 90 |
apply auto |
91 |
apply (case_tac "c = 0", simp) |
|
92 |
apply (rule_tac x = "1" in exI, simp) |
|
61945 | 93 |
apply (rule_tac x = "\<bar>c\<bar>" in exI, auto) |
36561 | 94 |
proof - |
95 |
fix c :: 'a and x :: 'b |
|
96 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
61076 | 97 |
have F1: "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) |
98 |
have F2: "\<forall>x\<^sub>3 x\<^sub>1::'a::linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1" by (metis abs_mult_pos) |
|
99 |
hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_one) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
100 |
hence "\<forall>x\<^sub>3. 0 \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^sub>3\<bar>" by (metis F2 A1 abs_ge_zero order_trans) |
46644 | 101 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis A1 abs_ge_zero) |
23449 | 102 |
qed |
103 |
||
57245 | 104 |
sledgehammer_params [isar_proofs, compress = 4] |
24545 | 105 |
|
45575 | 106 |
lemma |
61076 | 107 |
"(\<exists>c::'a::linordered_idom. |
61945 | 108 |
\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>) |
109 |
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>))" |
|
24545 | 110 |
apply auto |
111 |
apply (case_tac "c = 0", simp) |
|
112 |
apply (rule_tac x = "1" in exI, simp) |
|
61945 | 113 |
apply (rule_tac x = "\<bar>c\<bar>" in exI, auto) |
36561 | 114 |
proof - |
115 |
fix c :: 'a and x :: 'b |
|
116 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
61076 | 117 |
have "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51130
diff
changeset
|
118 |
hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" |
36561 | 119 |
by (metis A1 abs_ge_zero order_trans abs_mult_pos abs_one) |
120 |
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero abs_mult_pos abs_mult) |
|
121 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis abs_mult) |
|
24545 | 122 |
qed |
123 |
||
57245 | 124 |
sledgehammer_params [isar_proofs, compress = 1] |
24545 | 125 |
|
61945 | 126 |
lemma bigo_alt_def: "O(f) = {h. \<exists>c. (0 < c \<and> (\<forall>x. \<bar>h x\<bar> <= c * \<bar>f x\<bar>))}" |
23449 | 127 |
by (auto simp add: bigo_def bigo_pos_const) |
128 |
||
67613 | 129 |
lemma bigo_elt_subset [intro]: "f \<in> O(g) \<Longrightarrow> O(f) \<le> O(g)" |
45575 | 130 |
apply (auto simp add: bigo_alt_def) |
131 |
apply (rule_tac x = "ca * c" in exI) |
|
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
132 |
apply (metis algebra_simps mult_le_cancel_left_pos order_trans mult_pos_pos) |
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
133 |
done |
23449 | 134 |
|
67613 | 135 |
lemma bigo_refl [intro]: "f \<in> O(f)" |
46364 | 136 |
unfolding bigo_def mem_Collect_eq |
36844 | 137 |
by (metis mult_1 order_refl) |
23449 | 138 |
|
67613 | 139 |
lemma bigo_zero: "0 \<in> O(g)" |
36561 | 140 |
apply (auto simp add: bigo_def func_zero) |
36844 | 141 |
by (metis mult_zero_left order_refl) |
23449 | 142 |
|
45575 | 143 |
lemma bigo_zero2: "O(\<lambda>x. 0) = {\<lambda>x. 0}" |
144 |
by (auto simp add: bigo_def) |
|
23449 | 145 |
|
43197 | 146 |
lemma bigo_plus_self_subset [intro]: |
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
147 |
"O(f) + O(f) <= O(f)" |
45575 | 148 |
apply (auto simp add: bigo_alt_def set_plus_def) |
149 |
apply (rule_tac x = "c + ca" in exI) |
|
150 |
apply auto |
|
151 |
apply (simp add: ring_distribs func_plus) |
|
152 |
by (metis order_trans abs_triangle_ineq add_mono) |
|
23449 | 153 |
|
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
154 |
lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" |
45575 | 155 |
by (metis bigo_plus_self_subset bigo_zero set_eq_subset set_zero_plus2) |
23449 | 156 |
|
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
157 |
lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)" |
45575 | 158 |
apply (rule subsetI) |
159 |
apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) |
|
160 |
apply (subst bigo_pos_const [symmetric])+ |
|
61945 | 161 |
apply (rule_tac x = "\<lambda>n. if \<bar>g n\<bar> <= \<bar>f n\<bar> then x n else 0" in exI) |
45575 | 162 |
apply (rule conjI) |
163 |
apply (rule_tac x = "c + c" in exI) |
|
164 |
apply clarsimp |
|
61945 | 165 |
apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> <= (c + c) * \<bar>f xa\<bar>") |
56536 | 166 |
apply (metis mult_2 order_trans) |
61945 | 167 |
apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> <= c * (\<bar>f xa\<bar> + \<bar>g xa\<bar>)") |
56536 | 168 |
apply (erule order_trans) |
169 |
apply (simp add: ring_distribs) |
|
170 |
apply (rule mult_left_mono) |
|
171 |
apply (simp add: abs_triangle_ineq) |
|
172 |
apply (simp add: order_less_le) |
|
61945 | 173 |
apply (rule_tac x = "\<lambda>n. if \<bar>f n\<bar> < \<bar>g n\<bar> then x n else 0" in exI) |
45575 | 174 |
apply (rule conjI) |
175 |
apply (rule_tac x = "c + c" in exI) |
|
176 |
apply auto |
|
61945 | 177 |
apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> <= (c + c) * \<bar>g xa\<bar>") |
56536 | 178 |
apply (metis order_trans mult_2) |
61945 | 179 |
apply (subgoal_tac "c * \<bar>f xa + g xa\<bar> <= c * (\<bar>f xa\<bar> + \<bar>g xa\<bar>)") |
56536 | 180 |
apply (erule order_trans) |
181 |
apply (simp add: ring_distribs) |
|
182 |
by (metis abs_triangle_ineq mult_le_cancel_left_pos) |
|
23449 | 183 |
|
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
184 |
lemma bigo_plus_subset2 [intro]: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow> A + B <= O(f)" |
45575 | 185 |
by (metis bigo_plus_idemp set_plus_mono2) |
23449 | 186 |
|
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
187 |
lemma bigo_plus_eq: "\<forall>x. 0 <= f x \<Longrightarrow> \<forall>x. 0 <= g x \<Longrightarrow> O(f + g) = O(f) + O(g)" |
45575 | 188 |
apply (rule equalityI) |
189 |
apply (rule bigo_plus_subset) |
|
190 |
apply (simp add: bigo_alt_def set_plus_def func_plus) |
|
191 |
apply clarify |
|
192 |
(* sledgehammer *) |
|
193 |
apply (rule_tac x = "max c ca" in exI) |
|
46369 | 194 |
|
45575 | 195 |
apply (rule conjI) |
196 |
apply (metis less_max_iff_disj) |
|
197 |
apply clarify |
|
198 |
apply (drule_tac x = "xa" in spec)+ |
|
199 |
apply (subgoal_tac "0 <= f xa + g xa") |
|
200 |
apply (simp add: ring_distribs) |
|
61945 | 201 |
apply (subgoal_tac "\<bar>a xa + b xa\<bar> <= \<bar>a xa\<bar> + \<bar>b xa\<bar>") |
202 |
apply (subgoal_tac "\<bar>a xa\<bar> + \<bar>b xa\<bar> <= max c ca * f xa + max c ca * g xa") |
|
45575 | 203 |
apply (metis order_trans) |
23449 | 204 |
defer 1 |
45575 | 205 |
apply (metis abs_triangle_ineq) |
206 |
apply (metis add_nonneg_nonneg) |
|
46644 | 207 |
apply (rule add_mono) |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
208 |
apply (metis max.cobounded2 linorder_linear max.absorb1 mult_right_mono xt1(6)) |
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
209 |
by (metis max.cobounded2 linorder_not_le mult_le_cancel_right order_trans) |
23449 | 210 |
|
67613 | 211 |
lemma bigo_bounded_alt: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> c * g x \<Longrightarrow> f \<in> O(g)" |
45575 | 212 |
apply (auto simp add: bigo_def) |
36561 | 213 |
(* Version 1: one-line proof *) |
45575 | 214 |
by (metis abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) abs_mult) |
23449 | 215 |
|
67613 | 216 |
lemma "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> c * g x \<Longrightarrow> f \<in> O(g)" |
36561 | 217 |
apply (auto simp add: bigo_def) |
218 |
(* Version 2: structured proof *) |
|
219 |
proof - |
|
220 |
assume "\<forall>x. f x \<le> c * g x" |
|
221 |
thus "\<exists>c. \<forall>x. f x \<le> c * \<bar>g x\<bar>" by (metis abs_mult abs_ge_self order_trans) |
|
23449 | 222 |
qed |
223 |
||
67613 | 224 |
lemma bigo_bounded: "\<forall>x. 0 \<le> f x \<Longrightarrow> \<forall>x. f x \<le> g x \<Longrightarrow> f \<in> O(g)" |
45575 | 225 |
apply (erule bigo_bounded_alt [of f 1 g]) |
226 |
by (metis mult_1) |
|
23449 | 227 |
|
67613 | 228 |
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)" |
36561 | 229 |
apply (rule set_minus_imp_plus) |
230 |
apply (rule bigo_bounded) |
|
46369 | 231 |
apply (metis add_le_cancel_left diff_add_cancel diff_self minus_apply |
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
232 |
algebra_simps) |
46369 | 233 |
by (metis add_le_cancel_left diff_add_cancel func_plus le_fun_def |
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
234 |
algebra_simps) |
23449 | 235 |
|
61945 | 236 |
lemma bigo_abs: "(\<lambda>x. \<bar>f x\<bar>) =o O(f)" |
36561 | 237 |
apply (unfold bigo_def) |
238 |
apply auto |
|
36844 | 239 |
by (metis mult_1 order_refl) |
23449 | 240 |
|
61945 | 241 |
lemma bigo_abs2: "f =o O(\<lambda>x. \<bar>f x\<bar>)" |
36561 | 242 |
apply (unfold bigo_def) |
243 |
apply auto |
|
36844 | 244 |
by (metis mult_1 order_refl) |
43197 | 245 |
|
61945 | 246 |
lemma bigo_abs3: "O(f) = O(\<lambda>x. \<bar>f x\<bar>)" |
36561 | 247 |
proof - |
248 |
have F1: "\<forall>v u. u \<in> O(v) \<longrightarrow> O(u) \<subseteq> O(v)" by (metis bigo_elt_subset) |
|
249 |
have F2: "\<forall>u. (\<lambda>R. \<bar>u R\<bar>) \<in> O(u)" by (metis bigo_abs) |
|
250 |
have "\<forall>u. u \<in> O(\<lambda>R. \<bar>u R\<bar>)" by (metis bigo_abs2) |
|
251 |
thus "O(f) = O(\<lambda>x. \<bar>f x\<bar>)" using F1 F2 by auto |
|
43197 | 252 |
qed |
23449 | 253 |
|
61945 | 254 |
lemma bigo_abs4: "f =o g +o O(h) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) =o (\<lambda>x. \<bar>g x\<bar>) +o O(h)" |
23449 | 255 |
apply (drule set_plus_imp_minus) |
256 |
apply (rule set_minus_imp_plus) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
257 |
apply (subst fun_diff_def) |
23449 | 258 |
proof - |
67613 | 259 |
assume a: "f - g \<in> O(h)" |
61945 | 260 |
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>)" |
23449 | 261 |
by (rule bigo_abs2) |
67613 | 262 |
also have "\<dots> <= O(\<lambda>x. \<bar>f x - g x\<bar>)" |
23449 | 263 |
apply (rule bigo_elt_subset) |
264 |
apply (rule bigo_bounded) |
|
46369 | 265 |
apply (metis abs_ge_zero) |
266 |
by (metis abs_triangle_ineq3) |
|
67613 | 267 |
also have "\<dots> <= O(f - g)" |
23449 | 268 |
apply (rule bigo_elt_subset) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
269 |
apply (subst fun_diff_def) |
23449 | 270 |
apply (rule bigo_abs) |
271 |
done |
|
67613 | 272 |
also have "\<dots> <= O(h)" |
23464 | 273 |
using a by (rule bigo_elt_subset) |
67613 | 274 |
finally show "(\<lambda>x. \<bar>f x\<bar> - \<bar>g x\<bar>) \<in> O(h)" . |
23449 | 275 |
qed |
276 |
||
61945 | 277 |
lemma bigo_abs5: "f =o O(g) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) =o O(g)" |
23449 | 278 |
by (unfold bigo_def, auto) |
279 |
||
67613 | 280 |
lemma bigo_elt_subset2 [intro]: "f \<in> g +o O(h) \<Longrightarrow> O(f) \<le> O(g) + O(h)" |
23449 | 281 |
proof - |
67613 | 282 |
assume "f \<in> g +o O(h)" |
283 |
also have "\<dots> \<le> O(g) + O(h)" |
|
23449 | 284 |
by (auto del: subsetI) |
67613 | 285 |
also have "\<dots> = O(\<lambda>x. \<bar>g x\<bar>) + O(\<lambda>x. \<bar>h x\<bar>)" |
46369 | 286 |
by (metis bigo_abs3) |
61945 | 287 |
also have "... = O((\<lambda>x. \<bar>g x\<bar>) + (\<lambda>x. \<bar>h x\<bar>))" |
23449 | 288 |
by (rule bigo_plus_eq [symmetric], auto) |
67613 | 289 |
finally have "f \<in> \<dots>". |
290 |
then have "O(f) \<le> \<dots>" |
|
23449 | 291 |
by (elim bigo_elt_subset) |
67613 | 292 |
also have "\<dots> = O(\<lambda>x. \<bar>g x\<bar>) + O(\<lambda>x. \<bar>h x\<bar>)" |
23449 | 293 |
by (rule bigo_plus_eq, auto) |
294 |
finally show ?thesis |
|
295 |
by (simp add: bigo_abs3 [symmetric]) |
|
296 |
qed |
|
297 |
||
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
298 |
lemma bigo_mult [intro]: "O(f) * O(g) <= O(f * g)" |
46369 | 299 |
apply (rule subsetI) |
300 |
apply (subst bigo_def) |
|
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
301 |
apply (auto simp del: abs_mult ac_simps |
46369 | 302 |
simp add: bigo_alt_def set_times_def func_times) |
45575 | 303 |
(* sledgehammer *) |
46369 | 304 |
apply (rule_tac x = "c * ca" in exI) |
305 |
apply (rule allI) |
|
306 |
apply (erule_tac x = x in allE)+ |
|
61945 | 307 |
apply (subgoal_tac "c * ca * \<bar>f x * g x\<bar> = (c * \<bar>f x\<bar>) * (ca * \<bar>g x\<bar>)") |
46369 | 308 |
apply (metis (no_types) abs_ge_zero abs_mult mult_mono') |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
309 |
by (metis mult.assoc mult.left_commute abs_of_pos mult.left_commute abs_mult) |
23449 | 310 |
|
311 |
lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" |
|
46369 | 312 |
by (metis bigo_mult bigo_refl set_times_mono3 subset_trans) |
23449 | 313 |
|
67613 | 314 |
lemma bigo_mult3: "f \<in> O(h) \<Longrightarrow> g \<in> O(j) \<Longrightarrow> f * g \<in> O(h * j)" |
69712 | 315 |
by (metis bigo_mult rev_subsetD set_times_intro) |
23449 | 316 |
|
67613 | 317 |
lemma bigo_mult4 [intro]:"f \<in> k +o O(h) \<Longrightarrow> g * f \<in> (g * k) +o O(g * h)" |
23449 | 318 |
by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) |
319 |
||
45575 | 320 |
lemma bigo_mult5: "\<forall>x. f x ~= 0 \<Longrightarrow> |
61076 | 321 |
O(f * g) <= (f::'a => ('b::linordered_field)) *o O(g)" |
23449 | 322 |
proof - |
67613 | 323 |
assume a: "\<forall>x. f x \<noteq> 0" |
23449 | 324 |
show "O(f * g) <= f *o O(g)" |
325 |
proof |
|
326 |
fix h |
|
67613 | 327 |
assume h: "h \<in> O(f * g)" |
328 |
then have "(\<lambda>x. 1 / (f x)) * h \<in> (\<lambda>x. 1 / f x) *o O(f * g)" |
|
23449 | 329 |
by auto |
45575 | 330 |
also have "... <= O((\<lambda>x. 1 / f x) * (f * g))" |
23449 | 331 |
by (rule bigo_mult2) |
45575 | 332 |
also have "(\<lambda>x. 1 / f x) * (f * g) = g" |
59867
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
haftmann
parents:
59557
diff
changeset
|
333 |
by (simp add: fun_eq_iff a) |
67613 | 334 |
finally have "(\<lambda>x. (1::'b) / f x) * h \<in> O(g)". |
335 |
then have "f * ((\<lambda>x. (1::'b) / f x) * h) \<in> f *o O(g)" |
|
23449 | 336 |
by auto |
61076 | 337 |
also have "f * ((\<lambda>x. (1::'b) / f x) * h) = h" |
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
338 |
by (simp add: func_times fun_eq_iff a) |
67613 | 339 |
finally show "h \<in> f *o O(g)". |
23449 | 340 |
qed |
341 |
qed |
|
342 |
||
46369 | 343 |
lemma bigo_mult6: |
61076 | 344 |
"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = (f::'a \<Rightarrow> ('b::linordered_field)) *o O(g)" |
23449 | 345 |
by (metis bigo_mult2 bigo_mult5 order_antisym) |
346 |
||
347 |
(*proof requires relaxing relevance: 2007-01-25*) |
|
45705 | 348 |
declare bigo_mult6 [simp] |
349 |
||
46369 | 350 |
lemma bigo_mult7: |
61076 | 351 |
"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<le> O(f::'a \<Rightarrow> ('b::linordered_field)) * O(g)" |
46369 | 352 |
by (metis bigo_refl bigo_mult6 set_times_mono3) |
23449 | 353 |
|
45575 | 354 |
declare bigo_mult6 [simp del] |
355 |
declare bigo_mult7 [intro!] |
|
356 |
||
46369 | 357 |
lemma bigo_mult8: |
61076 | 358 |
"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f::'a \<Rightarrow> ('b::linordered_field)) * O(g)" |
23449 | 359 |
by (metis bigo_mult bigo_mult7 order_antisym_conv) |
360 |
||
67613 | 361 |
lemma bigo_minus [intro]: "f \<in> O(g) \<Longrightarrow> - f \<in> O(g)" |
46369 | 362 |
by (auto simp add: bigo_def fun_Compl_def) |
23449 | 363 |
|
67613 | 364 |
lemma bigo_minus2: "f \<in> g +o O(h) \<Longrightarrow> -f \<in> -g +o O(h)" |
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
365 |
by (metis (no_types, lifting) bigo_minus diff_minus_eq_add minus_add_distrib |
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
58889
diff
changeset
|
366 |
minus_minus set_minus_imp_plus set_plus_imp_minus) |
23449 | 367 |
|
368 |
lemma bigo_minus3: "O(-f) = O(f)" |
|
46369 | 369 |
by (metis bigo_elt_subset bigo_minus bigo_refl equalityI minus_minus) |
23449 | 370 |
|
67613 | 371 |
lemma bigo_plus_absorb_lemma1: "f \<in> O(g) \<Longrightarrow> f +o O(g) \<le> O(g)" |
46369 | 372 |
by (metis bigo_plus_idemp set_plus_mono3) |
23449 | 373 |
|
67613 | 374 |
lemma bigo_plus_absorb_lemma2: "f \<in> O(g) \<Longrightarrow> O(g) \<le> f +o O(g)" |
46369 | 375 |
by (metis (no_types) bigo_minus bigo_plus_absorb_lemma1 right_minus |
46644 | 376 |
set_plus_mono set_plus_rearrange2 set_zero_plus subsetD subset_refl |
377 |
subset_trans) |
|
23449 | 378 |
|
67613 | 379 |
lemma bigo_plus_absorb [simp]: "f \<in> O(g) \<Longrightarrow> f +o O(g) = O(g)" |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
380 |
by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff) |
23449 | 381 |
|
67613 | 382 |
lemma bigo_plus_absorb2 [intro]: "f \<in> O(g) \<Longrightarrow> A \<subseteq> O(g) \<Longrightarrow> f +o A \<subseteq> O(g)" |
46369 | 383 |
by (metis bigo_plus_absorb set_plus_mono) |
23449 | 384 |
|
67613 | 385 |
lemma bigo_add_commute_imp: "f \<in> g +o O(h) \<Longrightarrow> g \<in> f +o O(h)" |
46369 | 386 |
by (metis bigo_minus minus_diff_eq set_plus_imp_minus set_minus_plus) |
23449 | 387 |
|
67613 | 388 |
lemma bigo_add_commute: "(f \<in> g +o O(h)) = (g \<in> f +o O(h))" |
46369 | 389 |
by (metis bigo_add_commute_imp) |
23449 | 390 |
|
67613 | 391 |
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
|
392 |
by (auto simp add: bigo_def ac_simps) |
23449 | 393 |
|
46369 | 394 |
lemma bigo_const2 [intro]: "O(\<lambda>x. c) \<le> O(\<lambda>x. 1)" |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
395 |
by (metis bigo_const1 bigo_elt_subset) |
23449 | 396 |
|
67613 | 397 |
lemma bigo_const3: "(c::'a::linordered_field) \<noteq> 0 \<Longrightarrow> (\<lambda>x. 1) \<in> O(\<lambda>x. c)" |
23449 | 398 |
apply (simp add: bigo_def) |
36561 | 399 |
by (metis abs_eq_0 left_inverse order_refl) |
23449 | 400 |
|
67613 | 401 |
lemma bigo_const4: "(c::'a::linordered_field) \<noteq> 0 \<Longrightarrow> O(\<lambda>x. 1) \<subseteq> O(\<lambda>x. c)" |
46369 | 402 |
by (metis bigo_elt_subset bigo_const3) |
23449 | 403 |
|
61076 | 404 |
lemma bigo_const [simp]: "(c::'a::linordered_field) ~= 0 \<Longrightarrow> |
45575 | 405 |
O(\<lambda>x. c) = O(\<lambda>x. 1)" |
46369 | 406 |
by (metis bigo_const2 bigo_const4 equalityI) |
23449 | 407 |
|
67613 | 408 |
lemma bigo_const_mult1: "(\<lambda>x. c * f x) \<in> O(f)" |
46369 | 409 |
apply (simp add: bigo_def abs_mult) |
36561 | 410 |
by (metis le_less) |
23449 | 411 |
|
46369 | 412 |
lemma bigo_const_mult2: "O(\<lambda>x. c * f x) \<le> O(f)" |
23449 | 413 |
by (rule bigo_elt_subset, rule bigo_const_mult1) |
414 |
||
67613 | 415 |
lemma bigo_const_mult3: "(c::'a::linordered_field) \<noteq> 0 \<Longrightarrow> f \<in> O(\<lambda>x. c * f x)" |
45575 | 416 |
apply (simp add: bigo_def) |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
417 |
by (metis (no_types) abs_mult mult.assoc mult_1 order_refl left_inverse) |
23449 | 418 |
|
46369 | 419 |
lemma bigo_const_mult4: |
61076 | 420 |
"(c::'a::linordered_field) \<noteq> 0 \<Longrightarrow> O(f) \<le> O(\<lambda>x. c * f x)" |
46369 | 421 |
by (metis bigo_elt_subset bigo_const_mult3) |
23449 | 422 |
|
67613 | 423 |
lemma bigo_const_mult [simp]: "(c::'a::linordered_field) \<noteq> 0 \<Longrightarrow> |
45575 | 424 |
O(\<lambda>x. c * f x) = O(f)" |
46369 | 425 |
by (metis equalityI bigo_const_mult2 bigo_const_mult4) |
23449 | 426 |
|
67613 | 427 |
lemma bigo_const_mult5 [simp]: "(c::'a::linordered_field) \<noteq> 0 \<Longrightarrow> |
45575 | 428 |
(\<lambda>x. c) *o O(f) = O(f)" |
23449 | 429 |
apply (auto del: subsetI) |
430 |
apply (rule order_trans) |
|
431 |
apply (rule bigo_mult2) |
|
432 |
apply (simp add: func_times) |
|
433 |
apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times) |
|
45575 | 434 |
apply (rule_tac x = "\<lambda>y. inverse c * x y" in exI) |
43197 | 435 |
apply (rename_tac g d) |
24942 | 436 |
apply safe |
43197 | 437 |
apply (rule_tac [2] ext) |
438 |
prefer 2 |
|
26041
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents:
25710
diff
changeset
|
439 |
apply simp |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
440 |
apply (simp add: mult.assoc [symmetric] abs_mult) |
39259
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
441 |
(* couldn't get this proof without the step above *) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
442 |
proof - |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
443 |
fix g :: "'b \<Rightarrow> 'a" and d :: 'a |
61076 | 444 |
assume A1: "c \<noteq> (0::'a)" |
445 |
assume A2: "\<forall>x::'b. \<bar>g x\<bar> \<le> d * \<bar>f x\<bar>" |
|
39259
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
446 |
have F1: "inverse \<bar>c\<bar> = \<bar>inverse c\<bar>" using A1 by (metis nonzero_abs_inverse) |
61076 | 447 |
have F2: "(0::'a) < \<bar>c\<bar>" using A1 by (metis zero_less_abs_iff) |
448 |
have "(0::'a) < \<bar>c\<bar> \<longrightarrow> (0::'a) < \<bar>inverse c\<bar>" using F1 by (metis positive_imp_inverse_positive) |
|
449 |
hence "(0::'a) < \<bar>inverse c\<bar>" using F2 by metis |
|
450 |
hence F3: "(0::'a) \<le> \<bar>inverse c\<bar>" by (metis order_le_less) |
|
451 |
have "\<exists>(u::'a) SKF\<^sub>7::'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar>" |
|
39259
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
452 |
using A2 by metis |
61076 | 453 |
hence F4: "\<exists>(u::'a) SKF\<^sub>7::'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<and> (0::'a) \<le> \<bar>inverse c\<bar>" |
39259
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
454 |
using F3 by metis |
61076 | 455 |
hence "\<exists>(v::'a) (u::'a) SKF\<^sub>7::'a \<Rightarrow> 'b. \<bar>inverse c\<bar> * \<bar>g (SKF\<^sub>7 (u * v))\<bar> \<le> u * (v * \<bar>f (SKF\<^sub>7 (u * v))\<bar>)" |
59557 | 456 |
by (metis mult_left_mono) |
61076 | 457 |
then show "\<exists>ca::'a. \<forall>x::'b. inverse \<bar>c\<bar> * \<bar>g x\<bar> \<le> ca * \<bar>f x\<bar>" |
68536 | 458 |
using A2 F4 by (metis F1 \<open>0 < \<bar>inverse c\<bar>\<close> mult.assoc mult_le_cancel_left_pos) |
39259
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
459 |
qed |
23449 | 460 |
|
45575 | 461 |
lemma bigo_const_mult6 [intro]: "(\<lambda>x. c) *o O(f) <= O(f)" |
23449 | 462 |
apply (auto intro!: subsetI |
463 |
simp add: bigo_def elt_set_times_def func_times |
|
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
464 |
simp del: abs_mult ac_simps) |
45575 | 465 |
(* sledgehammer *) |
61945 | 466 |
apply (rule_tac x = "ca * \<bar>c\<bar>" in exI) |
23449 | 467 |
apply (rule allI) |
61945 | 468 |
apply (subgoal_tac "ca * \<bar>c\<bar> * \<bar>f x\<bar> = \<bar>c\<bar> * (ca * \<bar>f x\<bar>)") |
23449 | 469 |
apply (erule ssubst) |
470 |
apply (subst abs_mult) |
|
471 |
apply (rule mult_left_mono) |
|
472 |
apply (erule spec) |
|
473 |
apply simp |
|
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
474 |
apply (simp add: ac_simps) |
23449 | 475 |
done |
476 |
||
45575 | 477 |
lemma bigo_const_mult7 [intro]: "f =o O(g) \<Longrightarrow> (\<lambda>x. c * f x) =o O(g)" |
46369 | 478 |
by (metis bigo_const_mult1 bigo_elt_subset order_less_le psubsetD) |
23449 | 479 |
|
45575 | 480 |
lemma bigo_compose1: "f =o O(g) \<Longrightarrow> (\<lambda>x. f(k x)) =o O(\<lambda>x. g(k x))" |
23449 | 481 |
by (unfold bigo_def, auto) |
482 |
||
46369 | 483 |
lemma bigo_compose2: |
484 |
"f =o g +o O(h) \<Longrightarrow> (\<lambda>x. f(k x)) =o (\<lambda>x. g(k x)) +o O(\<lambda>x. h(k x))" |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53015
diff
changeset
|
485 |
apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus) |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53015
diff
changeset
|
486 |
apply (drule bigo_compose1 [of "f - g" h k]) |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53015
diff
changeset
|
487 |
apply (simp add: fun_diff_def) |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53015
diff
changeset
|
488 |
done |
23449 | 489 |
|
64267 | 490 |
subsection \<open>Sum\<close> |
23449 | 491 |
|
67613 | 492 |
lemma bigo_sum_main: "\<forall>x. \<forall>y \<in> A x. 0 \<le> h x y \<Longrightarrow> |
61945 | 493 |
\<exists>c. \<forall>x. \<forall>y \<in> A x. \<bar>f x y\<bar> <= c * (h x y) \<Longrightarrow> |
61954 | 494 |
(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)" |
46369 | 495 |
apply (auto simp add: bigo_def) |
61945 | 496 |
apply (rule_tac x = "\<bar>c\<bar>" in exI) |
46369 | 497 |
apply (subst abs_of_nonneg) back back |
64267 | 498 |
apply (rule sum_nonneg) |
46369 | 499 |
apply force |
64267 | 500 |
apply (subst sum_distrib_left) |
46369 | 501 |
apply (rule allI) |
502 |
apply (rule order_trans) |
|
64267 | 503 |
apply (rule sum_abs) |
504 |
apply (rule sum_mono) |
|
46369 | 505 |
by (metis abs_ge_self abs_mult_pos order_trans) |
23449 | 506 |
|
64267 | 507 |
lemma bigo_sum1: "\<forall>x y. 0 <= h x y \<Longrightarrow> |
61945 | 508 |
\<exists>c. \<forall>x y. \<bar>f x y\<bar> <= c * (h x y) \<Longrightarrow> |
61954 | 509 |
(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)" |
64267 | 510 |
by (metis (no_types) bigo_sum_main) |
23449 | 511 |
|
64267 | 512 |
lemma bigo_sum2: "\<forall>y. 0 <= h y \<Longrightarrow> |
61945 | 513 |
\<exists>c. \<forall>y. \<bar>f y\<bar> <= c * (h y) \<Longrightarrow> |
61954 | 514 |
(\<lambda>x. \<Sum>y \<in> A x. f y) =o O(\<lambda>x. \<Sum>y \<in> A x. h y)" |
64267 | 515 |
apply (rule bigo_sum1) |
46369 | 516 |
by metis+ |
23449 | 517 |
|
64267 | 518 |
lemma bigo_sum3: "f =o O(h) \<Longrightarrow> |
61954 | 519 |
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * f(k x y)) =o |
520 |
O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h(k x y)\<bar>)" |
|
64267 | 521 |
apply (rule bigo_sum1) |
45575 | 522 |
apply (rule allI)+ |
523 |
apply (rule abs_ge_zero) |
|
524 |
apply (unfold bigo_def) |
|
525 |
apply (auto simp add: abs_mult) |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
526 |
by (metis abs_ge_zero mult.left_commute mult_left_mono) |
23449 | 527 |
|
64267 | 528 |
lemma bigo_sum4: "f =o g +o O(h) \<Longrightarrow> |
61954 | 529 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * f(k x y)) =o |
530 |
(\<lambda>x. \<Sum>y \<in> A x. l x y * g(k x y)) +o |
|
531 |
O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h(k x y)\<bar>)" |
|
45575 | 532 |
apply (rule set_minus_imp_plus) |
533 |
apply (subst fun_diff_def) |
|
64267 | 534 |
apply (subst sum_subtractf [symmetric]) |
45575 | 535 |
apply (subst right_diff_distrib [symmetric]) |
64267 | 536 |
apply (rule bigo_sum3) |
46369 | 537 |
by (metis (lifting, no_types) fun_diff_def set_plus_imp_minus ext) |
23449 | 538 |
|
64267 | 539 |
lemma bigo_sum5: "f =o O(h) \<Longrightarrow> \<forall>x y. 0 <= l x y \<Longrightarrow> |
45575 | 540 |
\<forall>x. 0 <= h x \<Longrightarrow> |
61954 | 541 |
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * f(k x y)) =o |
542 |
O(\<lambda>x. \<Sum>y \<in> A x. (l x y) * h(k x y))" |
|
543 |
apply (subgoal_tac "(\<lambda>x. \<Sum>y \<in> A x. (l x y) * h(k x y)) = |
|
544 |
(\<lambda>x. \<Sum>y \<in> A x. \<bar>(l x y) * h(k x y)\<bar>)") |
|
46369 | 545 |
apply (erule ssubst) |
64267 | 546 |
apply (erule bigo_sum3) |
46369 | 547 |
apply (rule ext) |
64267 | 548 |
apply (rule sum.cong) |
57418 | 549 |
apply (rule refl) |
46369 | 550 |
by (metis abs_of_nonneg zero_le_mult_iff) |
23449 | 551 |
|
64267 | 552 |
lemma bigo_sum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 <= l x y \<Longrightarrow> |
45575 | 553 |
\<forall>x. 0 <= h x \<Longrightarrow> |
61954 | 554 |
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * f(k x y)) =o |
555 |
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * g(k x y)) +o |
|
556 |
O(\<lambda>x. \<Sum>y \<in> A x. (l x y) * h(k x y))" |
|
23449 | 557 |
apply (rule set_minus_imp_plus) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
558 |
apply (subst fun_diff_def) |
64267 | 559 |
apply (subst sum_subtractf [symmetric]) |
23449 | 560 |
apply (subst right_diff_distrib [symmetric]) |
64267 | 561 |
apply (rule bigo_sum5) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
562 |
apply (subst fun_diff_def [symmetric]) |
23449 | 563 |
apply (drule set_plus_imp_minus) |
564 |
apply auto |
|
565 |
done |
|
566 |
||
63167 | 567 |
subsection \<open>Misc useful stuff\<close> |
23449 | 568 |
|
45575 | 569 |
lemma bigo_useful_intro: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow> |
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47108
diff
changeset
|
570 |
A + B <= O(f)" |
23449 | 571 |
apply (subst bigo_plus_idemp [symmetric]) |
572 |
apply (rule set_plus_mono2) |
|
573 |
apply assumption+ |
|
574 |
done |
|
575 |
||
45575 | 576 |
lemma bigo_useful_add: "f =o O(h) \<Longrightarrow> g =o O(h) \<Longrightarrow> f + g =o O(h)" |
23449 | 577 |
apply (subst bigo_plus_idemp [symmetric]) |
578 |
apply (rule set_plus_intro) |
|
579 |
apply assumption+ |
|
580 |
done |
|
43197 | 581 |
|
61076 | 582 |
lemma bigo_useful_const_mult: "(c::'a::linordered_field) ~= 0 \<Longrightarrow> |
45575 | 583 |
(\<lambda>x. c) * f =o O(h) \<Longrightarrow> f =o O(h)" |
23449 | 584 |
apply (rule subsetD) |
45575 | 585 |
apply (subgoal_tac "(\<lambda>x. 1 / c) *o O(h) <= O(h)") |
23449 | 586 |
apply assumption |
587 |
apply (rule bigo_const_mult6) |
|
45575 | 588 |
apply (subgoal_tac "f = (\<lambda>x. 1 / c) * ((\<lambda>x. c) * f)") |
23449 | 589 |
apply (erule ssubst) |
590 |
apply (erule set_times_intro2) |
|
43197 | 591 |
apply (simp add: func_times) |
23449 | 592 |
done |
593 |
||
61076 | 594 |
lemma bigo_fix: "(\<lambda>x. f ((x::nat) + 1)) =o O(\<lambda>x. h(x + 1)) \<Longrightarrow> f 0 = 0 \<Longrightarrow> |
23449 | 595 |
f =o O(h)" |
45575 | 596 |
apply (simp add: bigo_alt_def) |
597 |
by (metis abs_ge_zero abs_mult abs_of_pos abs_zero not0_implies_Suc) |
|
23449 | 598 |
|
43197 | 599 |
lemma bigo_fix2: |
61076 | 600 |
"(\<lambda>x. f ((x::nat) + 1)) =o (\<lambda>x. g(x + 1)) +o O(\<lambda>x. h(x + 1)) \<Longrightarrow> |
45575 | 601 |
f 0 = g 0 \<Longrightarrow> f =o g +o O(h)" |
23449 | 602 |
apply (rule set_minus_imp_plus) |
603 |
apply (rule bigo_fix) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
604 |
apply (subst fun_diff_def) |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
605 |
apply (subst fun_diff_def [symmetric]) |
23449 | 606 |
apply (rule set_plus_imp_minus) |
607 |
apply simp |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
608 |
apply (simp add: fun_diff_def) |
23449 | 609 |
done |
610 |
||
63167 | 611 |
subsection \<open>Less than or equal to\<close> |
23449 | 612 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
74690
diff
changeset
|
613 |
definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl \<open><o\<close> 70) where |
45575 | 614 |
"f <o g == (\<lambda>x. max (f x - g x) 0)" |
23449 | 615 |
|
61945 | 616 |
lemma bigo_lesseq1: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> <= \<bar>f x\<bar> \<Longrightarrow> |
23449 | 617 |
g =o O(h)" |
618 |
apply (unfold bigo_def) |
|
619 |
apply clarsimp |
|
43197 | 620 |
apply (blast intro: order_trans) |
23449 | 621 |
done |
622 |
||
61945 | 623 |
lemma bigo_lesseq2: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> <= f x \<Longrightarrow> |
23449 | 624 |
g =o O(h)" |
625 |
apply (erule bigo_lesseq1) |
|
43197 | 626 |
apply (blast intro: abs_ge_self order_trans) |
23449 | 627 |
done |
628 |
||
45575 | 629 |
lemma bigo_lesseq3: "f =o O(h) \<Longrightarrow> \<forall>x. 0 <= g x \<Longrightarrow> \<forall>x. g x <= f x \<Longrightarrow> |
23449 | 630 |
g =o O(h)" |
631 |
apply (erule bigo_lesseq2) |
|
632 |
apply (rule allI) |
|
633 |
apply (subst abs_of_nonneg) |
|
634 |
apply (erule spec)+ |
|
635 |
done |
|
636 |
||
45575 | 637 |
lemma bigo_lesseq4: "f =o O(h) \<Longrightarrow> |
61945 | 638 |
\<forall>x. 0 <= g x \<Longrightarrow> \<forall>x. g x <= \<bar>f x\<bar> \<Longrightarrow> |
23449 | 639 |
g =o O(h)" |
640 |
apply (erule bigo_lesseq1) |
|
641 |
apply (rule allI) |
|
642 |
apply (subst abs_of_nonneg) |
|
643 |
apply (erule spec)+ |
|
644 |
done |
|
645 |
||
45575 | 646 |
lemma bigo_lesso1: "\<forall>x. f x <= g x \<Longrightarrow> f <o g =o O(h)" |
36561 | 647 |
apply (unfold lesso_def) |
45575 | 648 |
apply (subgoal_tac "(\<lambda>x. max (f x - g x) 0) = 0") |
649 |
apply (metis bigo_zero) |
|
46364 | 650 |
by (metis (lifting, no_types) func_zero le_fun_def le_iff_diff_le_0 |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
651 |
max.absorb2 order_eq_iff) |
23449 | 652 |
|
45575 | 653 |
lemma bigo_lesso2: "f =o g +o O(h) \<Longrightarrow> |
654 |
\<forall>x. 0 <= k x \<Longrightarrow> \<forall>x. k x <= f x \<Longrightarrow> |
|
23449 | 655 |
k <o g =o O(h)" |
656 |
apply (unfold lesso_def) |
|
657 |
apply (rule bigo_lesseq4) |
|
658 |
apply (erule set_plus_imp_minus) |
|
659 |
apply (rule allI) |
|
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
660 |
apply (rule max.cobounded2) |
23449 | 661 |
apply (rule allI) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
662 |
apply (subst fun_diff_def) |
23449 | 663 |
apply (erule thin_rl) |
45575 | 664 |
(* sledgehammer *) |
665 |
apply (case_tac "0 <= k x - g x") |
|
46644 | 666 |
apply (metis (lifting) abs_le_D1 linorder_linear min_diff_distrib_left |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
667 |
min.absorb1 min.absorb2 max.absorb1) |
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
668 |
by (metis abs_ge_zero le_cases max.absorb2) |
23449 | 669 |
|
45575 | 670 |
lemma bigo_lesso3: "f =o g +o O(h) \<Longrightarrow> |
671 |
\<forall>x. 0 <= k x \<Longrightarrow> \<forall>x. g x <= k x \<Longrightarrow> |
|
23449 | 672 |
f <o k =o O(h)" |
46644 | 673 |
apply (unfold lesso_def) |
674 |
apply (rule bigo_lesseq4) |
|
23449 | 675 |
apply (erule set_plus_imp_minus) |
46644 | 676 |
apply (rule allI) |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
677 |
apply (rule max.cobounded2) |
46644 | 678 |
apply (rule allI) |
679 |
apply (subst fun_diff_def) |
|
680 |
apply (erule thin_rl) |
|
681 |
(* sledgehammer *) |
|
682 |
apply (case_tac "0 <= f x - k x") |
|
683 |
apply simp |
|
684 |
apply (subst abs_of_nonneg) |
|
23449 | 685 |
apply (drule_tac x = x in spec) back |
61824
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61076
diff
changeset
|
686 |
apply (metis diff_less_0_iff_less linorder_not_le not_le_imp_less xt1(12) xt1(6)) |
45575 | 687 |
apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff) |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54230
diff
changeset
|
688 |
by (metis abs_ge_zero linorder_linear max.absorb1 max.commute) |
23449 | 689 |
|
45705 | 690 |
lemma bigo_lesso4: |
61076 | 691 |
"f <o g =o O(k::'a=>'b::{linordered_field}) \<Longrightarrow> |
45705 | 692 |
g =o h +o O(k) \<Longrightarrow> f <o h =o O(k)" |
693 |
apply (unfold lesso_def) |
|
694 |
apply (drule set_plus_imp_minus) |
|
695 |
apply (drule bigo_abs5) back |
|
696 |
apply (simp add: fun_diff_def) |
|
697 |
apply (drule bigo_useful_add, assumption) |
|
698 |
apply (erule bigo_lesseq2) back |
|
699 |
apply (rule allI) |
|
700 |
by (auto simp add: func_plus fun_diff_def algebra_simps |
|
23449 | 701 |
split: split_max abs_split) |
702 |
||
61945 | 703 |
lemma bigo_lesso5: "f <o g =o O(h) \<Longrightarrow> \<exists>C. \<forall>x. f x <= g x + C * \<bar>h x\<bar>" |
45705 | 704 |
apply (simp only: lesso_def bigo_alt_def) |
705 |
apply clarsimp |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
706 |
by (metis add.commute diff_le_eq) |
23449 | 707 |
|
708 |
end |