author | blanchet |
Fri, 18 Nov 2011 11:47:12 +0100 | |
changeset 45574 | 7a39df11bcf6 |
parent 45532 | 74b17a0881b3 |
child 45575 | 3a865fc42bbf |
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 |
||
43197 | 8 |
header {* Metis Example Featuring the Big O Notation *} |
23449 | 9 |
|
43197 | 10 |
theory Big_O |
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41144
diff
changeset
|
11 |
imports |
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41144
diff
changeset
|
12 |
"~~/src/HOL/Decision_Procs/Dense_Linear_Order" |
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41144
diff
changeset
|
13 |
Main |
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41144
diff
changeset
|
14 |
"~~/src/HOL/Library/Function_Algebras" |
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41144
diff
changeset
|
15 |
"~~/src/HOL/Library/Set_Algebras" |
23449 | 16 |
begin |
17 |
||
42103
6066a35f6678
Metis examples use the new Skolemizer to test it
blanchet
parents:
41865
diff
changeset
|
18 |
declare [[metis_new_skolemizer]] |
6066a35f6678
Metis examples use the new Skolemizer to test it
blanchet
parents:
41865
diff
changeset
|
19 |
|
23449 | 20 |
subsection {* Definitions *} |
21 |
||
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
22 |
definition bigo :: "('a => 'b::{linordered_idom,number_ring}) => ('a => 'b) set" ("(1O'(_'))") where |
23449 | 23 |
"O(f::('a => 'b)) == {h. EX c. ALL x. abs (h x) <= c * abs (f x)}" |
24 |
||
38991 | 25 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_pos_const" ]] |
43197 | 26 |
lemma bigo_pos_const: "(EX (c::'a::linordered_idom). |
23449 | 27 |
ALL x. (abs (h x)) <= (c * (abs (f x)))) |
28 |
= (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" |
|
29 |
apply auto |
|
30 |
apply (case_tac "c = 0", simp) |
|
31 |
apply (rule_tac x = "1" in exI, simp) |
|
25304
7491c00f0915
removed subclass edge ordered_ring < lordered_ring
haftmann
parents:
25087
diff
changeset
|
32 |
apply (rule_tac x = "abs c" in exI, auto) |
36561 | 33 |
apply (metis abs_ge_zero abs_of_nonneg Orderings.xt1(6) abs_mult) |
23449 | 34 |
done |
35 |
||
36407 | 36 |
(*** Now various verions with an increasing shrink factor ***) |
23449 | 37 |
|
36925 | 38 |
sledgehammer_params [isar_proof, isar_shrink_factor = 1] |
23449 | 39 |
|
43197 | 40 |
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom). |
23449 | 41 |
ALL x. (abs (h x)) <= (c * (abs (f x)))) |
42 |
= (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" |
|
43 |
apply auto |
|
44 |
apply (case_tac "c = 0", simp) |
|
45 |
apply (rule_tac x = "1" in exI, simp) |
|
46 |
apply (rule_tac x = "abs c" in exI, auto) |
|
36561 | 47 |
proof - |
48 |
fix c :: 'a and x :: 'b |
|
49 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
50 |
have F1: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> \<bar>x\<^isub>1\<bar>" by (metis abs_ge_zero) |
|
36844 | 51 |
have F2: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1) |
36561 | 52 |
have F3: "\<forall>x\<^isub>1 x\<^isub>3. x\<^isub>3 \<le> \<bar>h x\<^isub>1\<bar> \<longrightarrow> x\<^isub>3 \<le> c * \<bar>f x\<^isub>1\<bar>" by (metis A1 order_trans) |
53 |
have F4: "\<forall>x\<^isub>2 x\<^isub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^isub>3\<bar> * \<bar>x\<^isub>2\<bar> = \<bar>x\<^isub>3 * x\<^isub>2\<bar>" |
|
54 |
by (metis abs_mult) |
|
55 |
have F5: "\<forall>x\<^isub>3 x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^isub>1 \<longrightarrow> \<bar>x\<^isub>3 * x\<^isub>1\<bar> = \<bar>x\<^isub>3\<bar> * x\<^isub>1" |
|
56 |
by (metis abs_mult_pos) |
|
57 |
hence "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = \<bar>1\<bar> * x\<^isub>1" by (metis F2) |
|
58 |
hence "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^isub>1" by (metis F2 abs_one) |
|
59 |
hence "\<forall>x\<^isub>3. 0 \<le> \<bar>h x\<^isub>3\<bar> \<longrightarrow> \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>" by (metis F3) |
|
60 |
hence "\<forall>x\<^isub>3. \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>" by (metis F1) |
|
61 |
hence "\<forall>x\<^isub>3. (0\<Colon>'a) \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^isub>3\<bar>" by (metis F5) |
|
62 |
hence "\<forall>x\<^isub>3. (0\<Colon>'a) \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c * f x\<^isub>3\<bar>" by (metis F4) |
|
63 |
hence "\<forall>x\<^isub>3. c * \<bar>f x\<^isub>3\<bar> = \<bar>c * f x\<^isub>3\<bar>" by (metis F1) |
|
64 |
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1) |
|
65 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F4) |
|
23449 | 66 |
qed |
67 |
||
36925 | 68 |
sledgehammer_params [isar_proof, isar_shrink_factor = 2] |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25592
diff
changeset
|
69 |
|
43197 | 70 |
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom). |
23449 | 71 |
ALL x. (abs (h x)) <= (c * (abs (f x)))) |
72 |
= (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" |
|
73 |
apply auto |
|
74 |
apply (case_tac "c = 0", simp) |
|
75 |
apply (rule_tac x = "1" in exI, simp) |
|
36844 | 76 |
apply (rule_tac x = "abs c" in exI, auto) |
36561 | 77 |
proof - |
78 |
fix c :: 'a and x :: 'b |
|
79 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
36844 | 80 |
have F1: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1) |
36561 | 81 |
have F2: "\<forall>x\<^isub>2 x\<^isub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^isub>3\<bar> * \<bar>x\<^isub>2\<bar> = \<bar>x\<^isub>3 * x\<^isub>2\<bar>" |
82 |
by (metis abs_mult) |
|
83 |
have "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^isub>1" by (metis F1 abs_mult_pos abs_one) |
|
84 |
hence "\<forall>x\<^isub>3. \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>" by (metis A1 abs_ge_zero order_trans) |
|
85 |
hence "\<forall>x\<^isub>3. 0 \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c * f x\<^isub>3\<bar>" by (metis F2 abs_mult_pos) |
|
86 |
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero) |
|
87 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F2) |
|
23449 | 88 |
qed |
89 |
||
36925 | 90 |
sledgehammer_params [isar_proof, isar_shrink_factor = 3] |
25710
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
paulson
parents:
25592
diff
changeset
|
91 |
|
43197 | 92 |
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom). |
23449 | 93 |
ALL x. (abs (h x)) <= (c * (abs (f x)))) |
94 |
= (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" |
|
95 |
apply auto |
|
96 |
apply (case_tac "c = 0", simp) |
|
97 |
apply (rule_tac x = "1" in exI, simp) |
|
36561 | 98 |
apply (rule_tac x = "abs c" in exI, auto) |
99 |
proof - |
|
100 |
fix c :: 'a and x :: 'b |
|
101 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
36844 | 102 |
have F1: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1) |
36561 | 103 |
have F2: "\<forall>x\<^isub>3 x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^isub>1 \<longrightarrow> \<bar>x\<^isub>3 * x\<^isub>1\<bar> = \<bar>x\<^isub>3\<bar> * x\<^isub>1" by (metis abs_mult_pos) |
104 |
hence "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^isub>1" by (metis F1 abs_one) |
|
105 |
hence "\<forall>x\<^isub>3. 0 \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^isub>3\<bar>" by (metis F2 A1 abs_ge_zero order_trans) |
|
106 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis A1 abs_mult abs_ge_zero) |
|
23449 | 107 |
qed |
108 |
||
36925 | 109 |
sledgehammer_params [isar_proof, isar_shrink_factor = 4] |
24545 | 110 |
|
43197 | 111 |
lemma (*bigo_pos_const:*) "(EX (c::'a::linordered_idom). |
24545 | 112 |
ALL x. (abs (h x)) <= (c * (abs (f x)))) |
113 |
= (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" |
|
114 |
apply auto |
|
115 |
apply (case_tac "c = 0", simp) |
|
116 |
apply (rule_tac x = "1" in exI, simp) |
|
36561 | 117 |
apply (rule_tac x = "abs c" in exI, auto) |
118 |
proof - |
|
119 |
fix c :: 'a and x :: 'b |
|
120 |
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>" |
|
36844 | 121 |
have "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1) |
36561 | 122 |
hence "\<forall>x\<^isub>3. \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>" |
123 |
by (metis A1 abs_ge_zero order_trans abs_mult_pos abs_one) |
|
124 |
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero abs_mult_pos abs_mult) |
|
125 |
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis abs_mult) |
|
24545 | 126 |
qed |
127 |
||
36925 | 128 |
sledgehammer_params [isar_proof, isar_shrink_factor = 1] |
24545 | 129 |
|
43197 | 130 |
lemma bigo_alt_def: "O(f) = |
23449 | 131 |
{h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}" |
132 |
by (auto simp add: bigo_def bigo_pos_const) |
|
133 |
||
38991 | 134 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_elt_subset" ]] |
23449 | 135 |
lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)" |
136 |
apply (auto simp add: bigo_alt_def) |
|
137 |
apply (rule_tac x = "ca * c" in exI) |
|
138 |
apply (rule conjI) |
|
139 |
apply (rule mult_pos_pos) |
|
43197 | 140 |
apply (assumption)+ |
36844 | 141 |
(*sledgehammer*) |
23449 | 142 |
apply (rule allI) |
143 |
apply (drule_tac x = "xa" in spec)+ |
|
36844 | 144 |
apply (subgoal_tac "ca * abs(f xa) <= ca * (c * abs(g xa))") |
23449 | 145 |
apply (erule order_trans) |
146 |
apply (simp add: mult_ac) |
|
147 |
apply (rule mult_left_mono, assumption) |
|
36844 | 148 |
apply (rule order_less_imp_le, assumption) |
23449 | 149 |
done |
150 |
||
151 |
||
38991 | 152 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_refl" ]] |
23449 | 153 |
lemma bigo_refl [intro]: "f : O(f)" |
36561 | 154 |
apply (auto simp add: bigo_def) |
36844 | 155 |
by (metis mult_1 order_refl) |
23449 | 156 |
|
38991 | 157 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_zero" ]] |
23449 | 158 |
lemma bigo_zero: "0 : O(g)" |
36561 | 159 |
apply (auto simp add: bigo_def func_zero) |
36844 | 160 |
by (metis mult_zero_left order_refl) |
23449 | 161 |
|
162 |
lemma bigo_zero2: "O(%x.0) = {%x.0}" |
|
43197 | 163 |
by (auto simp add: bigo_def) |
23449 | 164 |
|
43197 | 165 |
lemma bigo_plus_self_subset [intro]: |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
166 |
"O(f) \<oplus> O(f) <= O(f)" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
167 |
apply (auto simp add: bigo_alt_def set_plus_def) |
23449 | 168 |
apply (rule_tac x = "c + ca" in exI) |
169 |
apply auto |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
23464
diff
changeset
|
170 |
apply (simp add: ring_distribs func_plus) |
43197 | 171 |
apply (blast intro:order_trans abs_triangle_ineq add_mono elim:) |
23449 | 172 |
done |
173 |
||
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
174 |
lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)" |
23449 | 175 |
apply (rule equalityI) |
176 |
apply (rule bigo_plus_self_subset) |
|
43197 | 177 |
apply (rule set_zero_plus2) |
23449 | 178 |
apply (rule bigo_zero) |
179 |
done |
|
180 |
||
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
181 |
lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)" |
23449 | 182 |
apply (rule subsetI) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
183 |
apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) |
23449 | 184 |
apply (subst bigo_pos_const [symmetric])+ |
43197 | 185 |
apply (rule_tac x = |
23449 | 186 |
"%n. if abs (g n) <= (abs (f n)) then x n else 0" in exI) |
187 |
apply (rule conjI) |
|
188 |
apply (rule_tac x = "c + c" in exI) |
|
189 |
apply (clarsimp) |
|
190 |
apply (auto) |
|
191 |
apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (f xa)") |
|
192 |
apply (erule_tac x = xa in allE) |
|
193 |
apply (erule order_trans) |
|
194 |
apply (simp) |
|
195 |
apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") |
|
196 |
apply (erule order_trans) |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
23464
diff
changeset
|
197 |
apply (simp add: ring_distribs) |
23449 | 198 |
apply (rule mult_left_mono) |
199 |
apply (simp add: abs_triangle_ineq) |
|
200 |
apply (simp add: order_less_le) |
|
201 |
apply (rule mult_nonneg_nonneg) |
|
202 |
apply auto |
|
43197 | 203 |
apply (rule_tac x = "%n. if (abs (f n)) < abs (g n) then x n else 0" |
23449 | 204 |
in exI) |
205 |
apply (rule conjI) |
|
206 |
apply (rule_tac x = "c + c" in exI) |
|
207 |
apply auto |
|
208 |
apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (g xa)") |
|
209 |
apply (erule_tac x = xa in allE) |
|
210 |
apply (erule order_trans) |
|
211 |
apply (simp) |
|
212 |
apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") |
|
213 |
apply (erule order_trans) |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
23464
diff
changeset
|
214 |
apply (simp add: ring_distribs) |
23449 | 215 |
apply (rule mult_left_mono) |
216 |
apply (rule abs_triangle_ineq) |
|
217 |
apply (simp add: order_less_le) |
|
45532
74b17a0881b3
Int.thy: remove duplicate lemmas double_less_0_iff and odd_less_0, use {even,odd}_less_0_iff instead
huffman
parents:
45504
diff
changeset
|
218 |
apply (metis abs_not_less_zero even_less_0_iff less_not_permute linorder_not_less mult_less_0_iff) |
23449 | 219 |
done |
220 |
||
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
221 |
lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \<oplus> B <= O(f)" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
222 |
apply (subgoal_tac "A \<oplus> B <= O(f) \<oplus> O(f)") |
23449 | 223 |
apply (erule order_trans) |
224 |
apply simp |
|
225 |
apply (auto del: subsetI simp del: bigo_plus_idemp) |
|
226 |
done |
|
227 |
||
38991 | 228 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_plus_eq" ]] |
43197 | 229 |
lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
230 |
O(f + g) = O(f) \<oplus> O(g)" |
23449 | 231 |
apply (rule equalityI) |
232 |
apply (rule bigo_plus_subset) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
233 |
apply (simp add: bigo_alt_def set_plus_def func_plus) |
43197 | 234 |
apply clarify |
235 |
(*sledgehammer*) |
|
23449 | 236 |
apply (rule_tac x = "max c ca" in exI) |
237 |
apply (rule conjI) |
|
25087 | 238 |
apply (metis Orderings.less_max_iff_disj) |
23449 | 239 |
apply clarify |
240 |
apply (drule_tac x = "xa" in spec)+ |
|
241 |
apply (subgoal_tac "0 <= f xa + g xa") |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
23464
diff
changeset
|
242 |
apply (simp add: ring_distribs) |
23449 | 243 |
apply (subgoal_tac "abs(a xa + b xa) <= abs(a xa) + abs(b xa)") |
43197 | 244 |
apply (subgoal_tac "abs(a xa) + abs(b xa) <= |
23449 | 245 |
max c ca * f xa + max c ca * g xa") |
246 |
apply (blast intro: order_trans) |
|
247 |
defer 1 |
|
248 |
apply (rule abs_triangle_ineq) |
|
25087 | 249 |
apply (metis add_nonneg_nonneg) |
23449 | 250 |
apply (rule add_mono) |
39259
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
251 |
using [[ sledgehammer_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
252 |
apply (metis le_maxI2 linorder_linear min_max.sup_absorb1 mult_right_mono xt1(6)) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
253 |
apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans) |
23449 | 254 |
done |
255 |
||
38991 | 256 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded_alt" ]] |
43197 | 257 |
lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> |
258 |
f : O(g)" |
|
23449 | 259 |
apply (auto simp add: bigo_def) |
36561 | 260 |
(* Version 1: one-line proof *) |
35050
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents:
35028
diff
changeset
|
261 |
apply (metis abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) abs_mult) |
23449 | 262 |
done |
263 |
||
43197 | 264 |
lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> |
36561 | 265 |
f : O(g)" |
266 |
apply (auto simp add: bigo_def) |
|
267 |
(* Version 2: structured proof *) |
|
268 |
proof - |
|
269 |
assume "\<forall>x. f x \<le> c * g x" |
|
270 |
thus "\<exists>c. \<forall>x. f x \<le> c * \<bar>g x\<bar>" by (metis abs_mult abs_ge_self order_trans) |
|
23449 | 271 |
qed |
272 |
||
36561 | 273 |
text{*So here is the easier (and more natural) problem using transitivity*} |
38991 | 274 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] |
43197 | 275 |
lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" |
36561 | 276 |
apply (auto simp add: bigo_def) |
277 |
(* Version 1: one-line proof *) |
|
278 |
by (metis abs_ge_self abs_mult order_trans) |
|
23449 | 279 |
|
280 |
text{*So here is the easier (and more natural) problem using transitivity*} |
|
38991 | 281 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] |
43197 | 282 |
lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" |
23449 | 283 |
apply (auto simp add: bigo_def) |
36561 | 284 |
(* Version 2: structured proof *) |
285 |
proof - |
|
286 |
assume "\<forall>x. f x \<le> c * g x" |
|
287 |
thus "\<exists>c. \<forall>x. f x \<le> c * \<bar>g x\<bar>" by (metis abs_mult abs_ge_self order_trans) |
|
23449 | 288 |
qed |
289 |
||
43197 | 290 |
lemma bigo_bounded: "ALL x. 0 <= f x ==> ALL x. f x <= g x ==> |
291 |
f : O(g)" |
|
23449 | 292 |
apply (erule bigo_bounded_alt [of f 1 g]) |
293 |
apply simp |
|
294 |
done |
|
295 |
||
38991 | 296 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded2" ]] |
23449 | 297 |
lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==> |
298 |
f : lb +o O(g)" |
|
36561 | 299 |
apply (rule set_minus_imp_plus) |
300 |
apply (rule bigo_bounded) |
|
301 |
apply (auto simp add: diff_minus fun_Compl_def func_plus) |
|
302 |
prefer 2 |
|
303 |
apply (drule_tac x = x in spec)+ |
|
36844 | 304 |
apply (metis add_right_mono add_commute diff_add_cancel diff_minus_eq_add le_less order_trans) |
36561 | 305 |
proof - |
306 |
fix x :: 'a |
|
307 |
assume "\<forall>x. lb x \<le> f x" |
|
308 |
thus "(0\<Colon>'b) \<le> f x + - lb x" by (metis not_leE diff_minus less_iff_diff_less_0 less_le_not_le) |
|
23449 | 309 |
qed |
310 |
||
38991 | 311 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_abs" ]] |
43197 | 312 |
lemma bigo_abs: "(%x. abs(f x)) =o O(f)" |
36561 | 313 |
apply (unfold bigo_def) |
314 |
apply auto |
|
36844 | 315 |
by (metis mult_1 order_refl) |
23449 | 316 |
|
38991 | 317 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_abs2" ]] |
23449 | 318 |
lemma bigo_abs2: "f =o O(%x. abs(f x))" |
36561 | 319 |
apply (unfold bigo_def) |
320 |
apply auto |
|
36844 | 321 |
by (metis mult_1 order_refl) |
43197 | 322 |
|
23449 | 323 |
lemma bigo_abs3: "O(f) = O(%x. abs(f x))" |
36561 | 324 |
proof - |
325 |
have F1: "\<forall>v u. u \<in> O(v) \<longrightarrow> O(u) \<subseteq> O(v)" by (metis bigo_elt_subset) |
|
326 |
have F2: "\<forall>u. (\<lambda>R. \<bar>u R\<bar>) \<in> O(u)" by (metis bigo_abs) |
|
327 |
have "\<forall>u. u \<in> O(\<lambda>R. \<bar>u R\<bar>)" by (metis bigo_abs2) |
|
328 |
thus "O(f) = O(\<lambda>x. \<bar>f x\<bar>)" using F1 F2 by auto |
|
43197 | 329 |
qed |
23449 | 330 |
|
43197 | 331 |
lemma bigo_abs4: "f =o g +o O(h) ==> |
23449 | 332 |
(%x. abs (f x)) =o (%x. abs (g x)) +o O(h)" |
333 |
apply (drule set_plus_imp_minus) |
|
334 |
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
|
335 |
apply (subst fun_diff_def) |
23449 | 336 |
proof - |
337 |
assume a: "f - g : O(h)" |
|
338 |
have "(%x. abs (f x) - abs (g x)) =o O(%x. abs(abs (f x) - abs (g x)))" |
|
339 |
by (rule bigo_abs2) |
|
340 |
also have "... <= O(%x. abs (f x - g x))" |
|
341 |
apply (rule bigo_elt_subset) |
|
342 |
apply (rule bigo_bounded) |
|
343 |
apply force |
|
344 |
apply (rule allI) |
|
345 |
apply (rule abs_triangle_ineq3) |
|
346 |
done |
|
347 |
also have "... <= O(f - g)" |
|
348 |
apply (rule bigo_elt_subset) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
349 |
apply (subst fun_diff_def) |
23449 | 350 |
apply (rule bigo_abs) |
351 |
done |
|
352 |
also have "... <= O(h)" |
|
23464 | 353 |
using a by (rule bigo_elt_subset) |
23449 | 354 |
finally show "(%x. abs (f x) - abs (g x)) : O(h)". |
355 |
qed |
|
356 |
||
43197 | 357 |
lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" |
23449 | 358 |
by (unfold bigo_def, auto) |
359 |
||
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
360 |
lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \<oplus> O(h)" |
23449 | 361 |
proof - |
362 |
assume "f : g +o O(h)" |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
363 |
also have "... <= O(g) \<oplus> O(h)" |
23449 | 364 |
by (auto del: subsetI) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
365 |
also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))" |
23449 | 366 |
apply (subst bigo_abs3 [symmetric])+ |
367 |
apply (rule refl) |
|
368 |
done |
|
369 |
also have "... = O((%x. abs(g x)) + (%x. abs(h x)))" |
|
370 |
by (rule bigo_plus_eq [symmetric], auto) |
|
371 |
finally have "f : ...". |
|
372 |
then have "O(f) <= ..." |
|
373 |
by (elim bigo_elt_subset) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
374 |
also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))" |
23449 | 375 |
by (rule bigo_plus_eq, auto) |
376 |
finally show ?thesis |
|
377 |
by (simp add: bigo_abs3 [symmetric]) |
|
378 |
qed |
|
379 |
||
38991 | 380 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult" ]] |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
381 |
lemma bigo_mult [intro]: "O(f)\<otimes>O(g) <= O(f * g)" |
23449 | 382 |
apply (rule subsetI) |
383 |
apply (subst bigo_def) |
|
384 |
apply (auto simp del: abs_mult mult_ac |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
385 |
simp add: bigo_alt_def set_times_def func_times) |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
386 |
(*sledgehammer*) |
23449 | 387 |
apply (rule_tac x = "c * ca" in exI) |
388 |
apply(rule allI) |
|
389 |
apply(erule_tac x = x in allE)+ |
|
43197 | 390 |
apply(subgoal_tac "c * ca * abs(f x * g x) = |
23449 | 391 |
(c * abs(f x)) * (ca * abs(g x))") |
38991 | 392 |
using [[ sledgehammer_problem_prefix = "BigO__bigo_mult_simpler" ]] |
43197 | 393 |
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
|
394 |
apply (metis mult_assoc mult_left_commute |
35050
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents:
35028
diff
changeset
|
395 |
abs_of_pos mult_left_commute |
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents:
35028
diff
changeset
|
396 |
abs_mult mult_pos_pos) |
43197 | 397 |
apply (erule ssubst) |
23449 | 398 |
apply (subst abs_mult) |
36561 | 399 |
(* not quite as hard as BigO__bigo_mult_simpler_1 (a hard problem!) since |
400 |
abs_mult has just been done *) |
|
401 |
by (metis abs_ge_zero mult_mono') |
|
23449 | 402 |
|
38991 | 403 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult2" ]] |
23449 | 404 |
lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" |
405 |
apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) |
|
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
406 |
(*sledgehammer*) |
23449 | 407 |
apply (rule_tac x = c in exI) |
408 |
apply clarify |
|
409 |
apply (drule_tac x = x in spec) |
|
38991 | 410 |
using [[ sledgehammer_problem_prefix = "BigO__bigo_mult2_simpler" ]] |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
411 |
(*sledgehammer [no luck]*) |
23449 | 412 |
apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))") |
413 |
apply (simp add: mult_ac) |
|
414 |
apply (rule mult_left_mono, assumption) |
|
415 |
apply (rule abs_ge_zero) |
|
416 |
done |
|
417 |
||
38991 | 418 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult3" ]] |
23449 | 419 |
lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)" |
36561 | 420 |
by (metis bigo_mult set_rev_mp set_times_intro) |
23449 | 421 |
|
38991 | 422 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult4" ]] |
23449 | 423 |
lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)" |
424 |
by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) |
|
425 |
||
426 |
||
427 |
lemma bigo_mult5: "ALL x. f x ~= 0 ==> |
|
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
428 |
O(f * g) <= (f::'a => ('b::{linordered_field,number_ring})) *o O(g)" |
23449 | 429 |
proof - |
41541 | 430 |
assume a: "ALL x. f x ~= 0" |
23449 | 431 |
show "O(f * g) <= f *o O(g)" |
432 |
proof |
|
433 |
fix h |
|
41541 | 434 |
assume h: "h : O(f * g)" |
23449 | 435 |
then have "(%x. 1 / (f x)) * h : (%x. 1 / f x) *o O(f * g)" |
436 |
by auto |
|
437 |
also have "... <= O((%x. 1 / f x) * (f * g))" |
|
438 |
by (rule bigo_mult2) |
|
439 |
also have "(%x. 1 / f x) * (f * g) = g" |
|
43197 | 440 |
apply (simp add: func_times) |
23449 | 441 |
apply (rule ext) |
41541 | 442 |
apply (simp add: a h nonzero_divide_eq_eq mult_ac) |
23449 | 443 |
done |
444 |
finally have "(%x. (1::'b) / f x) * h : O(g)". |
|
445 |
then have "f * ((%x. (1::'b) / f x) * h) : f *o O(g)" |
|
446 |
by auto |
|
447 |
also have "f * ((%x. (1::'b) / f x) * h) = h" |
|
43197 | 448 |
apply (simp add: func_times) |
23449 | 449 |
apply (rule ext) |
41541 | 450 |
apply (simp add: a h nonzero_divide_eq_eq mult_ac) |
23449 | 451 |
done |
452 |
finally show "h : f *o O(g)". |
|
453 |
qed |
|
454 |
qed |
|
455 |
||
38991 | 456 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult6" ]] |
23449 | 457 |
lemma bigo_mult6: "ALL x. f x ~= 0 ==> |
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
458 |
O(f * g) = (f::'a => ('b::{linordered_field,number_ring})) *o O(g)" |
23449 | 459 |
by (metis bigo_mult2 bigo_mult5 order_antisym) |
460 |
||
461 |
(*proof requires relaxing relevance: 2007-01-25*) |
|
38991 | 462 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult7" ]] |
23449 | 463 |
declare bigo_mult6 [simp] |
464 |
lemma bigo_mult7: "ALL x. f x ~= 0 ==> |
|
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
465 |
O(f * g) <= O(f::'a => ('b::{linordered_field,number_ring})) \<otimes> O(g)" |
23449 | 466 |
(*sledgehammer*) |
467 |
apply (subst bigo_mult6) |
|
468 |
apply assumption |
|
43197 | 469 |
apply (rule set_times_mono3) |
23449 | 470 |
apply (rule bigo_refl) |
471 |
done |
|
472 |
declare bigo_mult6 [simp del] |
|
473 |
||
38991 | 474 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult8" ]] |
23449 | 475 |
declare bigo_mult7[intro!] |
476 |
lemma bigo_mult8: "ALL x. f x ~= 0 ==> |
|
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
477 |
O(f * g) = O(f::'a => ('b::{linordered_field,number_ring})) \<otimes> O(g)" |
23449 | 478 |
by (metis bigo_mult bigo_mult7 order_antisym_conv) |
479 |
||
480 |
lemma bigo_minus [intro]: "f : O(g) ==> - f : O(g)" |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
481 |
by (auto simp add: bigo_def fun_Compl_def) |
23449 | 482 |
|
483 |
lemma bigo_minus2: "f : g +o O(h) ==> -f : -g +o O(h)" |
|
484 |
apply (rule set_minus_imp_plus) |
|
485 |
apply (drule set_plus_imp_minus) |
|
486 |
apply (drule bigo_minus) |
|
487 |
apply (simp add: diff_minus) |
|
488 |
done |
|
489 |
||
490 |
lemma bigo_minus3: "O(-f) = O(f)" |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
491 |
by (auto simp add: bigo_def fun_Compl_def abs_minus_cancel) |
23449 | 492 |
|
493 |
lemma bigo_plus_absorb_lemma1: "f : O(g) ==> f +o O(g) <= O(g)" |
|
494 |
proof - |
|
495 |
assume a: "f : O(g)" |
|
496 |
show "f +o O(g) <= O(g)" |
|
497 |
proof - |
|
498 |
have "f : O(f)" by auto |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
499 |
then have "f +o O(g) <= O(f) \<oplus> O(g)" |
23449 | 500 |
by (auto del: subsetI) |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
501 |
also have "... <= O(g) \<oplus> O(g)" |
23449 | 502 |
proof - |
503 |
from a have "O(f) <= O(g)" by (auto del: subsetI) |
|
504 |
thus ?thesis by (auto del: subsetI) |
|
505 |
qed |
|
506 |
also have "... <= O(g)" by (simp add: bigo_plus_idemp) |
|
507 |
finally show ?thesis . |
|
508 |
qed |
|
509 |
qed |
|
510 |
||
511 |
lemma bigo_plus_absorb_lemma2: "f : O(g) ==> O(g) <= f +o O(g)" |
|
512 |
proof - |
|
513 |
assume a: "f : O(g)" |
|
514 |
show "O(g) <= f +o O(g)" |
|
515 |
proof - |
|
516 |
from a have "-f : O(g)" by auto |
|
517 |
then have "-f +o O(g) <= O(g)" by (elim bigo_plus_absorb_lemma1) |
|
518 |
then have "f +o (-f +o O(g)) <= f +o O(g)" by auto |
|
519 |
also have "f +o (-f +o O(g)) = O(g)" |
|
520 |
by (simp add: set_plus_rearranges) |
|
521 |
finally show ?thesis . |
|
522 |
qed |
|
523 |
qed |
|
524 |
||
38991 | 525 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_plus_absorb" ]] |
23449 | 526 |
lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)" |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
527 |
by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff) |
23449 | 528 |
|
529 |
lemma bigo_plus_absorb2 [intro]: "f : O(g) ==> A <= O(g) ==> f +o A <= O(g)" |
|
530 |
apply (subgoal_tac "f +o A <= f +o O(g)") |
|
531 |
apply force+ |
|
532 |
done |
|
533 |
||
534 |
lemma bigo_add_commute_imp: "f : g +o O(h) ==> g : f +o O(h)" |
|
535 |
apply (subst set_minus_plus [symmetric]) |
|
536 |
apply (subgoal_tac "g - f = - (f - g)") |
|
537 |
apply (erule ssubst) |
|
538 |
apply (rule bigo_minus) |
|
539 |
apply (subst set_minus_plus) |
|
540 |
apply assumption |
|
541 |
apply (simp add: diff_minus add_ac) |
|
542 |
done |
|
543 |
||
544 |
lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))" |
|
545 |
apply (rule iffI) |
|
546 |
apply (erule bigo_add_commute_imp)+ |
|
547 |
done |
|
548 |
||
549 |
lemma bigo_const1: "(%x. c) : O(%x. 1)" |
|
550 |
by (auto simp add: bigo_def mult_ac) |
|
551 |
||
38991 | 552 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_const2" ]] |
23449 | 553 |
lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)" |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
554 |
by (metis bigo_const1 bigo_elt_subset) |
23449 | 555 |
|
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
556 |
lemma bigo_const2 [intro]: "O(%x. c::'b::{linordered_idom,number_ring}) <= O(%x. 1)" |
36561 | 557 |
(* "thus" had to be replaced by "show" with an explicit reference to "F1" *) |
558 |
proof - |
|
559 |
have F1: "\<forall>u. (\<lambda>Q. u) \<in> O(\<lambda>Q. 1)" by (metis bigo_const1) |
|
560 |
show "O(\<lambda>x. c) \<subseteq> O(\<lambda>x. 1)" by (metis F1 bigo_elt_subset) |
|
23449 | 561 |
qed |
562 |
||
38991 | 563 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_const3" ]] |
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
564 |
lemma bigo_const3: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> (%x. 1) : O(%x. c)" |
23449 | 565 |
apply (simp add: bigo_def) |
36561 | 566 |
by (metis abs_eq_0 left_inverse order_refl) |
23449 | 567 |
|
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
568 |
lemma bigo_const4: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> O(%x. 1) <= O(%x. c)" |
23449 | 569 |
by (rule bigo_elt_subset, rule bigo_const3, assumption) |
570 |
||
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
571 |
lemma bigo_const [simp]: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> |
23449 | 572 |
O(%x. c) = O(%x. 1)" |
573 |
by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption) |
|
574 |
||
38991 | 575 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult1" ]] |
23449 | 576 |
lemma bigo_const_mult1: "(%x. c * f x) : O(f)" |
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
577 |
apply (simp add: bigo_def abs_mult) |
36561 | 578 |
by (metis le_less) |
23449 | 579 |
|
580 |
lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)" |
|
581 |
by (rule bigo_elt_subset, rule bigo_const_mult1) |
|
582 |
||
38991 | 583 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult3" ]] |
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
584 |
lemma bigo_const_mult3: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> f : O(%x. c * f x)" |
23449 | 585 |
apply (simp add: bigo_def) |
36561 | 586 |
(*sledgehammer [no luck]*) |
23449 | 587 |
apply (rule_tac x = "abs(inverse c)" in exI) |
588 |
apply (simp only: abs_mult [symmetric] mult_assoc [symmetric]) |
|
43197 | 589 |
apply (subst left_inverse) |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
590 |
apply (auto ) |
23449 | 591 |
done |
592 |
||
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
593 |
lemma bigo_const_mult4: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> |
23449 | 594 |
O(f) <= O(%x. c * f x)" |
595 |
by (rule bigo_elt_subset, rule bigo_const_mult3, assumption) |
|
596 |
||
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
597 |
lemma bigo_const_mult [simp]: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> |
23449 | 598 |
O(%x. c * f x) = O(f)" |
599 |
by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4) |
|
600 |
||
38991 | 601 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult5" ]] |
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
602 |
lemma bigo_const_mult5 [simp]: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> |
23449 | 603 |
(%x. c) *o O(f) = O(f)" |
604 |
apply (auto del: subsetI) |
|
605 |
apply (rule order_trans) |
|
606 |
apply (rule bigo_mult2) |
|
607 |
apply (simp add: func_times) |
|
608 |
apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times) |
|
609 |
apply (rule_tac x = "%y. inverse c * x y" in exI) |
|
43197 | 610 |
apply (rename_tac g d) |
24942 | 611 |
apply safe |
43197 | 612 |
apply (rule_tac [2] ext) |
613 |
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
|
614 |
apply simp |
24942 | 615 |
apply (simp add: mult_assoc [symmetric] abs_mult) |
39259
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
616 |
(* couldn't get this proof without the step above *) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
617 |
proof - |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
618 |
fix g :: "'b \<Rightarrow> 'a" and d :: 'a |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
619 |
assume A1: "c \<noteq> (0\<Colon>'a)" |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
620 |
assume A2: "\<forall>x\<Colon>'b. \<bar>g x\<bar> \<le> d * \<bar>f x\<bar>" |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
621 |
have F1: "inverse \<bar>c\<bar> = \<bar>inverse c\<bar>" using A1 by (metis nonzero_abs_inverse) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
622 |
have F2: "(0\<Colon>'a) < \<bar>c\<bar>" using A1 by (metis zero_less_abs_iff) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
623 |
have "(0\<Colon>'a) < \<bar>c\<bar> \<longrightarrow> (0\<Colon>'a) < \<bar>inverse c\<bar>" using F1 by (metis positive_imp_inverse_positive) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
624 |
hence "(0\<Colon>'a) < \<bar>inverse c\<bar>" using F2 by metis |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
625 |
hence F3: "(0\<Colon>'a) \<le> \<bar>inverse c\<bar>" by (metis order_le_less) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
626 |
have "\<exists>(u\<Colon>'a) SKF\<^isub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar>" |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
627 |
using A2 by metis |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
628 |
hence F4: "\<exists>(u\<Colon>'a) SKF\<^isub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar> \<and> (0\<Colon>'a) \<le> \<bar>inverse c\<bar>" |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
629 |
using F3 by metis |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
630 |
hence "\<exists>(v\<Colon>'a) (u\<Colon>'a) SKF\<^isub>7\<Colon>'a \<Rightarrow> 'b. \<bar>inverse c\<bar> * \<bar>g (SKF\<^isub>7 (u * v))\<bar> \<le> u * (v * \<bar>f (SKF\<^isub>7 (u * v))\<bar>)" |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
631 |
by (metis comm_mult_left_mono) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
632 |
thus "\<exists>ca\<Colon>'a. \<forall>x\<Colon>'b. \<bar>inverse c\<bar> * \<bar>g x\<bar> \<le> ca * \<bar>f x\<bar>" |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
633 |
using A2 F4 by (metis ab_semigroup_mult_class.mult_ac(1) comm_mult_left_mono) |
194014eb4f9f
replace two slow "metis" proofs with faster proofs
blanchet
parents:
38991
diff
changeset
|
634 |
qed |
23449 | 635 |
|
636 |
||
38991 | 637 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult6" ]] |
23449 | 638 |
lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)" |
639 |
apply (auto intro!: subsetI |
|
640 |
simp add: bigo_def elt_set_times_def func_times |
|
641 |
simp del: abs_mult mult_ac) |
|
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
642 |
(*sledgehammer*) |
23449 | 643 |
apply (rule_tac x = "ca * (abs c)" in exI) |
644 |
apply (rule allI) |
|
645 |
apply (subgoal_tac "ca * abs(c) * abs(f x) = abs(c) * (ca * abs(f x))") |
|
646 |
apply (erule ssubst) |
|
647 |
apply (subst abs_mult) |
|
648 |
apply (rule mult_left_mono) |
|
649 |
apply (erule spec) |
|
650 |
apply simp |
|
651 |
apply(simp add: mult_ac) |
|
652 |
done |
|
653 |
||
654 |
lemma bigo_const_mult7 [intro]: "f =o O(g) ==> (%x. c * f x) =o O(g)" |
|
655 |
proof - |
|
656 |
assume "f =o O(g)" |
|
657 |
then have "(%x. c) * f =o (%x. c) *o O(g)" |
|
658 |
by auto |
|
659 |
also have "(%x. c) * f = (%x. c * f x)" |
|
660 |
by (simp add: func_times) |
|
661 |
also have "(%x. c) *o O(g) <= O(g)" |
|
662 |
by (auto del: subsetI) |
|
663 |
finally show ?thesis . |
|
664 |
qed |
|
665 |
||
666 |
lemma bigo_compose1: "f =o O(g) ==> (%x. f(k x)) =o O(%x. g(k x))" |
|
667 |
by (unfold bigo_def, auto) |
|
668 |
||
43197 | 669 |
lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o |
23449 | 670 |
O(%x. h(k x))" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
671 |
apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def |
23449 | 672 |
func_plus) |
673 |
apply (erule bigo_compose1) |
|
674 |
done |
|
675 |
||
676 |
subsection {* Setsum *} |
|
677 |
||
43197 | 678 |
lemma bigo_setsum_main: "ALL x. ALL y : A x. 0 <= h x y ==> |
23449 | 679 |
EX c. ALL x. ALL y : A x. abs(f x y) <= c * (h x y) ==> |
43197 | 680 |
(%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" |
23449 | 681 |
apply (auto simp add: bigo_def) |
682 |
apply (rule_tac x = "abs c" in exI) |
|
683 |
apply (subst abs_of_nonneg) back back |
|
684 |
apply (rule setsum_nonneg) |
|
685 |
apply force |
|
686 |
apply (subst setsum_right_distrib) |
|
687 |
apply (rule allI) |
|
688 |
apply (rule order_trans) |
|
689 |
apply (rule setsum_abs) |
|
690 |
apply (rule setsum_mono) |
|
43197 | 691 |
apply (blast intro: order_trans mult_right_mono abs_ge_self) |
23449 | 692 |
done |
693 |
||
38991 | 694 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_setsum1" ]] |
43197 | 695 |
lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> |
23449 | 696 |
EX c. ALL x y. abs(f x y) <= c * (h x y) ==> |
697 |
(%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" |
|
698 |
apply (rule bigo_setsum_main) |
|
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
699 |
(*sledgehammer*) |
23449 | 700 |
apply force |
701 |
apply clarsimp |
|
702 |
apply (rule_tac x = c in exI) |
|
703 |
apply force |
|
704 |
done |
|
705 |
||
43197 | 706 |
lemma bigo_setsum2: "ALL y. 0 <= h y ==> |
23449 | 707 |
EX c. ALL y. abs(f y) <= c * (h y) ==> |
708 |
(%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)" |
|
43197 | 709 |
by (rule bigo_setsum1, auto) |
23449 | 710 |
|
38991 | 711 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_setsum3" ]] |
23449 | 712 |
lemma bigo_setsum3: "f =o O(h) ==> |
713 |
(%x. SUM y : A x. (l x y) * f(k x y)) =o |
|
714 |
O(%x. SUM y : A x. abs(l x y * h(k x y)))" |
|
715 |
apply (rule bigo_setsum1) |
|
716 |
apply (rule allI)+ |
|
717 |
apply (rule abs_ge_zero) |
|
718 |
apply (unfold bigo_def) |
|
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
719 |
apply (auto simp add: abs_mult) |
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
720 |
(*sledgehammer*) |
23449 | 721 |
apply (rule_tac x = c in exI) |
722 |
apply (rule allI)+ |
|
723 |
apply (subst mult_left_commute) |
|
724 |
apply (rule mult_left_mono) |
|
725 |
apply (erule spec) |
|
726 |
apply (rule abs_ge_zero) |
|
727 |
done |
|
728 |
||
729 |
lemma bigo_setsum4: "f =o g +o O(h) ==> |
|
730 |
(%x. SUM y : A x. l x y * f(k x y)) =o |
|
731 |
(%x. SUM y : A x. l x y * g(k x y)) +o |
|
732 |
O(%x. SUM y : A x. abs(l x y * h(k x y)))" |
|
733 |
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
|
734 |
apply (subst fun_diff_def) |
23449 | 735 |
apply (subst setsum_subtractf [symmetric]) |
736 |
apply (subst right_diff_distrib [symmetric]) |
|
737 |
apply (rule bigo_setsum3) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
738 |
apply (subst fun_diff_def [symmetric]) |
23449 | 739 |
apply (erule set_plus_imp_minus) |
740 |
done |
|
741 |
||
38991 | 742 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_setsum5" ]] |
43197 | 743 |
lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> |
23449 | 744 |
ALL x. 0 <= h x ==> |
745 |
(%x. SUM y : A x. (l x y) * f(k x y)) =o |
|
43197 | 746 |
O(%x. SUM y : A x. (l x y) * h(k x y))" |
747 |
apply (subgoal_tac "(%x. SUM y : A x. (l x y) * h(k x y)) = |
|
23449 | 748 |
(%x. SUM y : A x. abs((l x y) * h(k x y)))") |
749 |
apply (erule ssubst) |
|
750 |
apply (erule bigo_setsum3) |
|
751 |
apply (rule ext) |
|
752 |
apply (rule setsum_cong2) |
|
43197 | 753 |
apply (thin_tac "f \<in> O(h)") |
24942 | 754 |
apply (metis abs_of_nonneg zero_le_mult_iff) |
23449 | 755 |
done |
756 |
||
757 |
lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==> |
|
758 |
ALL x. 0 <= h x ==> |
|
759 |
(%x. SUM y : A x. (l x y) * f(k x y)) =o |
|
760 |
(%x. SUM y : A x. (l x y) * g(k x y)) +o |
|
43197 | 761 |
O(%x. SUM y : A x. (l x y) * h(k x y))" |
23449 | 762 |
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
|
763 |
apply (subst fun_diff_def) |
23449 | 764 |
apply (subst setsum_subtractf [symmetric]) |
765 |
apply (subst right_diff_distrib [symmetric]) |
|
766 |
apply (rule bigo_setsum5) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
767 |
apply (subst fun_diff_def [symmetric]) |
23449 | 768 |
apply (drule set_plus_imp_minus) |
769 |
apply auto |
|
770 |
done |
|
771 |
||
772 |
subsection {* Misc useful stuff *} |
|
773 |
||
774 |
lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==> |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
775 |
A \<oplus> B <= O(f)" |
23449 | 776 |
apply (subst bigo_plus_idemp [symmetric]) |
777 |
apply (rule set_plus_mono2) |
|
778 |
apply assumption+ |
|
779 |
done |
|
780 |
||
781 |
lemma bigo_useful_add: "f =o O(h) ==> g =o O(h) ==> f + g =o O(h)" |
|
782 |
apply (subst bigo_plus_idemp [symmetric]) |
|
783 |
apply (rule set_plus_intro) |
|
784 |
apply assumption+ |
|
785 |
done |
|
43197 | 786 |
|
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
787 |
lemma bigo_useful_const_mult: "(c::'a::{linordered_field,number_ring}) ~= 0 ==> |
23449 | 788 |
(%x. c) * f =o O(h) ==> f =o O(h)" |
789 |
apply (rule subsetD) |
|
790 |
apply (subgoal_tac "(%x. 1 / c) *o O(h) <= O(h)") |
|
791 |
apply assumption |
|
792 |
apply (rule bigo_const_mult6) |
|
793 |
apply (subgoal_tac "f = (%x. 1 / c) * ((%x. c) * f)") |
|
794 |
apply (erule ssubst) |
|
795 |
apply (erule set_times_intro2) |
|
43197 | 796 |
apply (simp add: func_times) |
23449 | 797 |
done |
798 |
||
38991 | 799 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_fix" ]] |
23449 | 800 |
lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==> |
801 |
f =o O(h)" |
|
802 |
apply (simp add: bigo_alt_def) |
|
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
803 |
(*sledgehammer*) |
23449 | 804 |
apply clarify |
805 |
apply (rule_tac x = c in exI) |
|
806 |
apply safe |
|
807 |
apply (case_tac "x = 0") |
|
43197 | 808 |
apply (metis abs_ge_zero abs_zero order_less_le split_mult_pos_le) |
23449 | 809 |
apply (subgoal_tac "x = Suc (x - 1)") |
23816 | 810 |
apply metis |
23449 | 811 |
apply simp |
812 |
done |
|
813 |
||
814 |
||
43197 | 815 |
lemma bigo_fix2: |
816 |
"(%x. f ((x::nat) + 1)) =o (%x. g(x + 1)) +o O(%x. h(x + 1)) ==> |
|
23449 | 817 |
f 0 = g 0 ==> f =o g +o O(h)" |
818 |
apply (rule set_minus_imp_plus) |
|
819 |
apply (rule bigo_fix) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
820 |
apply (subst fun_diff_def) |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
821 |
apply (subst fun_diff_def [symmetric]) |
23449 | 822 |
apply (rule set_plus_imp_minus) |
823 |
apply simp |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
824 |
apply (simp add: fun_diff_def) |
23449 | 825 |
done |
826 |
||
827 |
subsection {* Less than or equal to *} |
|
828 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35050
diff
changeset
|
829 |
definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) where |
23449 | 830 |
"f <o g == (%x. max (f x - g x) 0)" |
831 |
||
832 |
lemma bigo_lesseq1: "f =o O(h) ==> ALL x. abs (g x) <= abs (f x) ==> |
|
833 |
g =o O(h)" |
|
834 |
apply (unfold bigo_def) |
|
835 |
apply clarsimp |
|
43197 | 836 |
apply (blast intro: order_trans) |
23449 | 837 |
done |
838 |
||
839 |
lemma bigo_lesseq2: "f =o O(h) ==> ALL x. abs (g x) <= f x ==> |
|
840 |
g =o O(h)" |
|
841 |
apply (erule bigo_lesseq1) |
|
43197 | 842 |
apply (blast intro: abs_ge_self order_trans) |
23449 | 843 |
done |
844 |
||
845 |
lemma bigo_lesseq3: "f =o O(h) ==> ALL x. 0 <= g x ==> ALL x. g x <= f x ==> |
|
846 |
g =o O(h)" |
|
847 |
apply (erule bigo_lesseq2) |
|
848 |
apply (rule allI) |
|
849 |
apply (subst abs_of_nonneg) |
|
850 |
apply (erule spec)+ |
|
851 |
done |
|
852 |
||
853 |
lemma bigo_lesseq4: "f =o O(h) ==> |
|
854 |
ALL x. 0 <= g x ==> ALL x. g x <= abs (f x) ==> |
|
855 |
g =o O(h)" |
|
856 |
apply (erule bigo_lesseq1) |
|
857 |
apply (rule allI) |
|
858 |
apply (subst abs_of_nonneg) |
|
859 |
apply (erule spec)+ |
|
860 |
done |
|
861 |
||
38991 | 862 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso1" ]] |
23449 | 863 |
lemma bigo_lesso1: "ALL x. f x <= g x ==> f <o g =o O(h)" |
36561 | 864 |
apply (unfold lesso_def) |
865 |
apply (subgoal_tac "(%x. max (f x - g x) 0) = 0") |
|
866 |
proof - |
|
867 |
assume "(\<lambda>x. max (f x - g x) 0) = 0" |
|
868 |
thus "(\<lambda>x. max (f x - g x) 0) \<in> O(h)" by (metis bigo_zero) |
|
869 |
next |
|
870 |
show "\<forall>x\<Colon>'a. f x \<le> g x \<Longrightarrow> (\<lambda>x\<Colon>'a. max (f x - g x) (0\<Colon>'b)) = (0\<Colon>'a \<Rightarrow> 'b)" |
|
23449 | 871 |
apply (unfold func_zero) |
872 |
apply (rule ext) |
|
36561 | 873 |
by (simp split: split_max) |
874 |
qed |
|
23449 | 875 |
|
38991 | 876 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso2" ]] |
23449 | 877 |
lemma bigo_lesso2: "f =o g +o O(h) ==> |
878 |
ALL x. 0 <= k x ==> ALL x. k x <= f x ==> |
|
879 |
k <o g =o O(h)" |
|
880 |
apply (unfold lesso_def) |
|
881 |
apply (rule bigo_lesseq4) |
|
882 |
apply (erule set_plus_imp_minus) |
|
883 |
apply (rule allI) |
|
884 |
apply (rule le_maxI2) |
|
885 |
apply (rule allI) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
886 |
apply (subst fun_diff_def) |
23449 | 887 |
apply (erule thin_rl) |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
888 |
(*sledgehammer*) |
23449 | 889 |
apply (case_tac "0 <= k x - g x") |
36561 | 890 |
(* apply (metis abs_le_iff add_le_imp_le_right diff_minus le_less |
891 |
le_max_iff_disj min_max.le_supE min_max.sup_absorb2 |
|
892 |
min_max.sup_commute) *) |
|
37320 | 893 |
proof - |
894 |
fix x :: 'a |
|
895 |
assume "\<forall>x\<Colon>'a. k x \<le> f x" |
|
896 |
hence F1: "\<forall>x\<^isub>1\<Colon>'a. max (k x\<^isub>1) (f x\<^isub>1) = f x\<^isub>1" by (metis min_max.sup_absorb2) |
|
897 |
assume "(0\<Colon>'b) \<le> k x - g x" |
|
898 |
hence F2: "max (0\<Colon>'b) (k x - g x) = k x - g x" by (metis min_max.sup_absorb2) |
|
899 |
have F3: "\<forall>x\<^isub>1\<Colon>'b. x\<^isub>1 \<le> \<bar>x\<^isub>1\<bar>" by (metis abs_le_iff le_less) |
|
900 |
have "\<forall>(x\<^isub>2\<Colon>'b) x\<^isub>1\<Colon>'b. max x\<^isub>1 x\<^isub>2 \<le> x\<^isub>2 \<or> max x\<^isub>1 x\<^isub>2 \<le> x\<^isub>1" by (metis le_less le_max_iff_disj) |
|
901 |
hence "\<forall>(x\<^isub>3\<Colon>'b) (x\<^isub>2\<Colon>'b) x\<^isub>1\<Colon>'b. x\<^isub>1 - x\<^isub>2 \<le> x\<^isub>3 - x\<^isub>2 \<or> x\<^isub>3 \<le> x\<^isub>1" by (metis add_le_imp_le_right diff_minus min_max.le_supE) |
|
902 |
hence "k x - g x \<le> f x - g x" by (metis F1 le_less min_max.sup_absorb2 min_max.sup_commute) |
|
903 |
hence "k x - g x \<le> \<bar>f x - g x\<bar>" by (metis F3 le_max_iff_disj min_max.sup_absorb2) |
|
904 |
thus "max (k x - g x) (0\<Colon>'b) \<le> \<bar>f x - g x\<bar>" by (metis F2 min_max.sup_commute) |
|
36561 | 905 |
next |
906 |
show "\<And>x\<Colon>'a. |
|
907 |
\<lbrakk>\<forall>x\<Colon>'a. (0\<Colon>'b) \<le> k x; \<forall>x\<Colon>'a. k x \<le> f x; \<not> (0\<Colon>'b) \<le> k x - g x\<rbrakk> |
|
908 |
\<Longrightarrow> max (k x - g x) (0\<Colon>'b) \<le> \<bar>f x - g x\<bar>" |
|
909 |
by (metis abs_ge_zero le_cases min_max.sup_absorb2) |
|
24545 | 910 |
qed |
23449 | 911 |
|
38991 | 912 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso3" ]] |
23449 | 913 |
lemma bigo_lesso3: "f =o g +o O(h) ==> |
914 |
ALL x. 0 <= k x ==> ALL x. g x <= k x ==> |
|
915 |
f <o k =o O(h)" |
|
916 |
apply (unfold lesso_def) |
|
917 |
apply (rule bigo_lesseq4) |
|
918 |
apply (erule set_plus_imp_minus) |
|
919 |
apply (rule allI) |
|
920 |
apply (rule le_maxI2) |
|
921 |
apply (rule allI) |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
922 |
apply (subst fun_diff_def) |
43197 | 923 |
apply (erule thin_rl) |
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41541
diff
changeset
|
924 |
(*sledgehammer*) |
23449 | 925 |
apply (case_tac "0 <= f x - k x") |
29667 | 926 |
apply (simp) |
23449 | 927 |
apply (subst abs_of_nonneg) |
928 |
apply (drule_tac x = x in spec) back |
|
38991 | 929 |
using [[ sledgehammer_problem_prefix = "BigO__bigo_lesso3_simpler" ]] |
24545 | 930 |
apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6)) |
931 |
apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff) |
|
29511
7071b017cb35
migrated class package to new locale implementation
haftmann
parents:
28592
diff
changeset
|
932 |
apply (metis abs_ge_zero linorder_linear min_max.sup_absorb1 min_max.sup_commute) |
23449 | 933 |
done |
934 |
||
45504
cad35ed6effa
Metis_Examples/Big_O.thy: add number_ring class constraints, adapt proofs to use cancellation simprocs
huffman
parents:
45270
diff
changeset
|
935 |
lemma bigo_lesso4: "f <o g =o O(k::'a=>'b::{linordered_field,number_ring}) ==> |
23449 | 936 |
g =o h +o O(k) ==> f <o h =o O(k)" |
937 |
apply (unfold lesso_def) |
|
938 |
apply (drule set_plus_imp_minus) |
|
939 |
apply (drule bigo_abs5) back |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
26645
diff
changeset
|
940 |
apply (simp add: fun_diff_def) |
23449 | 941 |
apply (drule bigo_useful_add) |
942 |
apply assumption |
|
943 |
apply (erule bigo_lesseq2) back |
|
944 |
apply (rule allI) |
|
29667 | 945 |
apply (auto simp add: func_plus fun_diff_def algebra_simps |
23449 | 946 |
split: split_max abs_split) |
947 |
done |
|
948 |
||
38991 | 949 |
declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso5" ]] |
23449 | 950 |
lemma bigo_lesso5: "f <o g =o O(h) ==> |
951 |
EX C. ALL x. f x <= g x + C * abs(h x)" |
|
952 |
apply (simp only: lesso_def bigo_alt_def) |
|
953 |
apply clarsimp |
|
43197 | 954 |
apply (metis abs_if abs_mult add_commute diff_le_eq less_not_permute) |
23449 | 955 |
done |
956 |
||
957 |
end |