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