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