| author | haftmann | 
| Mon, 23 Feb 2009 21:38:45 +0100 | |
| changeset 30077 | c5920259850c | 
| parent 29823 | 0ab754d13ccd | 
| child 32864 | a226f29d4bdc | 
| permissions | -rw-r--r-- | 
| 23449 | 1 | (* Title: HOL/MetisExamples/BigO.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | ||
| 4 | Testing the metis method | |
| 5 | *) | |
| 6 | ||
| 7 | header {* Big O notation *}
 | |
| 8 | ||
| 9 | theory BigO | |
| 29823 
0ab754d13ccd
session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
 haftmann parents: 
29667diff
changeset | 10 | imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Main SetsAndFunctions | 
| 23449 | 11 | begin | 
| 12 | ||
| 13 | subsection {* Definitions *}
 | |
| 14 | ||
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 15 | definition bigo :: "('a => 'b::ordered_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 | ||
| 28592 | 18 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_pos_const"*}
 | 
| 23449 | 19 | lemma bigo_pos_const: "(EX (c::'a::ordered_idom). | 
| 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) | 
| 
7491c00f0915
removed subclass edge ordered_ring < lordered_ring
 haftmann parents: 
25087diff
changeset | 26 | apply (metis abs_ge_minus_self abs_ge_zero abs_minus_cancel abs_of_nonneg equation_minus_iff Orderings.xt1(6) abs_mult) | 
| 23449 | 27 | done | 
| 28 | ||
| 29 | (*** Now various verions with an increasing modulus ***) | |
| 30 | ||
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
26312diff
changeset | 31 | declare [[sledgehammer_modulus = 1]] | 
| 23449 | 32 | |
| 26312 | 33 | lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_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) | |
| 40 | proof (neg_clausify) | |
| 41 | fix c x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 42 | have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * X2\<bar> = \<bar>X2 * X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 43 | by (metis abs_mult mult_commute) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 44 | have 1: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 45 | X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> \<bar>X2\<bar> * X1 = \<bar>X2 * X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 46 | by (metis abs_mult_pos linorder_linear) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 47 | have 2: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 48 | \<not> (0\<Colon>'a\<Colon>ordered_idom) < X1 * X2 \<or> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 49 | \<not> (0\<Colon>'a\<Colon>ordered_idom) \<le> X2 \<or> \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 50 | by (metis linorder_not_less mult_nonneg_nonpos2) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 51 | assume 3: "\<And>x\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 52 | \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 53 | \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 54 | assume 4: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 55 | \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 56 | have 5: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 57 | \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 58 | by (metis 4 abs_mult) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 59 | have 6: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 60 | \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> X1 \<le> \<bar>X2\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 61 | by (metis abs_ge_zero xt1(6)) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 62 | have 7: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 63 | X1 \<le> \<bar>X2\<bar> \<or> (0\<Colon>'a\<Colon>ordered_idom) < X1" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 64 | by (metis not_leE 6) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 65 | have 8: "(0\<Colon>'a\<Colon>ordered_idom) < \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 66 | by (metis 5 7) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 67 | have 9: "\<And>X1\<Colon>'a\<Colon>ordered_idom. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 68 | \<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> \<le> X1 \<or> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 69 | (0\<Colon>'a\<Colon>ordered_idom) < X1" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 70 | by (metis 8 order_less_le_trans) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 71 | have 10: "(0\<Colon>'a\<Colon>ordered_idom) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 72 | < (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 73 | by (metis 3 9) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 74 | have 11: "\<not> (c\<Colon>'a\<Colon>ordered_idom) \<le> (0\<Colon>'a\<Colon>ordered_idom)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 75 | by (metis abs_ge_zero 2 10) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 76 | have 12: "\<And>X1\<Colon>'a\<Colon>ordered_idom. (c\<Colon>'a\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>X1 * c\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 77 | by (metis mult_commute 1 11) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 78 | have 13: "\<And>X1\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 79 | - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1 | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 80 | \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 81 | by (metis 3 abs_le_D2) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 82 | have 14: "\<And>X1\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 83 | - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1 | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 84 | \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 85 | by (metis 0 12 13) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 86 | have 15: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 87 | by (metis abs_mult abs_mult_pos abs_ge_zero) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 88 | have 16: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. X1 \<le> \<bar>X2\<bar> \<or> \<not> X1 \<le> X2" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 89 | by (metis xt1(6) abs_ge_self) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 90 | have 17: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<not> \<bar>X1\<bar> \<le> X2 \<or> X1 \<le> \<bar>X2\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 91 | by (metis 16 abs_le_D1) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 92 | have 18: "\<And>X1\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 93 | (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1 | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 94 | \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 95 | by (metis 17 3 15) | 
| 23449 | 96 | show "False" | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 97 | by (metis abs_le_iff 5 18 14) | 
| 23449 | 98 | qed | 
| 99 | ||
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
26312diff
changeset | 100 | declare [[sledgehammer_modulus = 2]] | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25592diff
changeset | 101 | |
| 23449 | 102 | lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). | 
| 103 | ALL x. (abs (h x)) <= (c * (abs (f x)))) | |
| 104 | = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" | |
| 105 | apply auto | |
| 106 | apply (case_tac "c = 0", simp) | |
| 107 | apply (rule_tac x = "1" in exI, simp) | |
| 108 | apply (rule_tac x = "abs c" in exI, auto); | |
| 109 | proof (neg_clausify) | |
| 110 | fix c x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 111 | have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * X2\<bar> = \<bar>X2 * X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 112 | by (metis abs_mult mult_commute) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 113 | assume 1: "\<And>x\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 114 | \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 115 | \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 116 | assume 2: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 117 | \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 118 | have 3: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 119 | \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 120 | by (metis 2 abs_mult) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 121 | have 4: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 122 | \<not> X1 \<le> (0\<Colon>'a\<Colon>ordered_idom) \<or> X1 \<le> \<bar>X2\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 123 | by (metis abs_ge_zero xt1(6)) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 124 | have 5: "(0\<Colon>'a\<Colon>ordered_idom) < \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 125 | by (metis not_leE 4 3) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 126 | have 6: "(0\<Colon>'a\<Colon>ordered_idom) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 127 | < (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 128 | by (metis 1 order_less_le_trans 5) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 129 | have 7: "\<And>X1\<Colon>'a\<Colon>ordered_idom. (c\<Colon>'a\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>X1 * c\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 130 | by (metis abs_ge_zero linorder_not_less mult_nonneg_nonpos2 6 linorder_linear abs_mult_pos mult_commute) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 131 | have 8: "\<And>X1\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 132 | - (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1 | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 133 | \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 134 | by (metis 0 7 abs_le_D2 1) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 135 | have 9: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<not> \<bar>X1\<bar> \<le> X2 \<or> X1 \<le> \<bar>X2\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 136 | by (metis abs_ge_self xt1(6) abs_le_D1) | 
| 23449 | 137 | show "False" | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 138 | by (metis 8 abs_ge_zero abs_mult_pos abs_mult 1 9 3 abs_le_iff) | 
| 23449 | 139 | qed | 
| 140 | ||
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
26312diff
changeset | 141 | declare [[sledgehammer_modulus = 3]] | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25592diff
changeset | 142 | |
| 23449 | 143 | lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). | 
| 144 | ALL x. (abs (h x)) <= (c * (abs (f x)))) | |
| 145 | = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" | |
| 146 | apply auto | |
| 147 | apply (case_tac "c = 0", simp) | |
| 148 | apply (rule_tac x = "1" in exI, simp) | |
| 149 | apply (rule_tac x = "abs c" in exI, auto); | |
| 150 | proof (neg_clausify) | |
| 151 | fix c x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 152 | assume 0: "\<And>x\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 153 | \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar> | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 154 | \<le> (c\<Colon>'a\<Colon>ordered_idom) * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 155 | assume 1: "\<not> \<bar>(h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) (x\<Colon>'b\<Colon>type)\<bar> | 
| 23449 | 156 | \<le> \<bar>c\<Colon>'a\<Colon>ordered_idom\<bar> * \<bar>(f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) x\<bar>" | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 157 | have 2: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 158 | X1 \<le> \<bar>X2\<bar> \<or> (0\<Colon>'a\<Colon>ordered_idom) < X1" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 159 | by (metis abs_ge_zero xt1(6) not_leE) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 160 | have 3: "\<not> (c\<Colon>'a\<Colon>ordered_idom) \<le> (0\<Colon>'a\<Colon>ordered_idom)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 161 | by (metis abs_ge_zero mult_nonneg_nonpos2 linorder_not_less order_less_le_trans 1 abs_mult 2 0) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 162 | have 4: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2\<Colon>'a\<Colon>ordered_idom. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 163 | by (metis abs_ge_zero abs_mult_pos abs_mult) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 164 | have 5: "\<And>X1\<Colon>'b\<Colon>type. | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 165 | (h\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1 | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 166 | \<le> \<bar>(c\<Colon>'a\<Colon>ordered_idom) * (f\<Colon>'b\<Colon>type \<Rightarrow> 'a\<Colon>ordered_idom) X1\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 167 | by (metis 4 0 xt1(6) abs_ge_self abs_le_D1) | 
| 23449 | 168 | show "False" | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 169 | by (metis abs_mult mult_commute 3 abs_mult_pos linorder_linear 0 abs_le_D2 5 1 abs_le_iff) | 
| 23449 | 170 | qed | 
| 171 | ||
| 172 | ||
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
26312diff
changeset | 173 | declare [[sledgehammer_modulus = 1]] | 
| 24545 | 174 | |
| 175 | lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). | |
| 176 | ALL x. (abs (h x)) <= (c * (abs (f x)))) | |
| 177 | = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" | |
| 178 | apply auto | |
| 179 | apply (case_tac "c = 0", simp) | |
| 180 | apply (rule_tac x = "1" in exI, simp) | |
| 181 | apply (rule_tac x = "abs c" in exI, auto); | |
| 182 | proof (neg_clausify) | |
| 183 | fix c x (*sort/type constraint inserted by hand!*) | |
| 184 | have 0: "\<And>(X1\<Colon>'a\<Colon>ordered_idom) X2. \<bar>X1 * \<bar>X2\<bar>\<bar> = \<bar>X1 * X2\<bar>" | |
| 185 | by (metis abs_ge_zero abs_mult_pos abs_mult) | |
| 186 | assume 1: "\<And>A. \<bar>h A\<bar> \<le> c * \<bar>f A\<bar>" | |
| 187 | have 2: "\<And>X1 X2. \<not> \<bar>X1\<bar> \<le> X2 \<or> (0\<Colon>'a) \<le> X2" | |
| 188 | by (metis abs_ge_zero order_trans) | |
| 189 | have 3: "\<And>X1. (0\<Colon>'a) \<le> c * \<bar>f X1\<bar>" | |
| 190 | by (metis 1 2) | |
| 191 | have 4: "\<And>X1. c * \<bar>f X1\<bar> = \<bar>c * f X1\<bar>" | |
| 192 | by (metis 0 abs_of_nonneg 3) | |
| 193 | have 5: "\<And>X1. - h X1 \<le> c * \<bar>f X1\<bar>" | |
| 194 | by (metis 1 abs_le_D2) | |
| 195 | have 6: "\<And>X1. - h X1 \<le> \<bar>c * f X1\<bar>" | |
| 196 | by (metis 4 5) | |
| 197 | have 7: "\<And>X1. h X1 \<le> c * \<bar>f X1\<bar>" | |
| 198 | by (metis 1 abs_le_D1) | |
| 199 | have 8: "\<And>X1. h X1 \<le> \<bar>c * f X1\<bar>" | |
| 200 | by (metis 4 7) | |
| 201 | assume 9: "\<not> \<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" | |
| 202 | have 10: "\<not> \<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" | |
| 203 | by (metis abs_mult 9) | |
| 204 | show "False" | |
| 205 | by (metis 6 8 10 abs_leI) | |
| 206 | qed | |
| 207 | ||
| 208 | ||
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
26312diff
changeset | 209 | declare [[sledgehammer_sorts = true]] | 
| 24545 | 210 | |
| 23449 | 211 | lemma bigo_alt_def: "O(f) = | 
| 212 |     {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}"
 | |
| 213 | by (auto simp add: bigo_def bigo_pos_const) | |
| 214 | ||
| 28592 | 215 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_elt_subset"*}
 | 
| 23449 | 216 | lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)" | 
| 217 | apply (auto simp add: bigo_alt_def) | |
| 218 | apply (rule_tac x = "ca * c" in exI) | |
| 219 | apply (rule conjI) | |
| 220 | apply (rule mult_pos_pos) | |
| 221 | apply (assumption)+ | |
| 222 | (*sledgehammer*); | |
| 223 | apply (rule allI) | |
| 224 | apply (drule_tac x = "xa" in spec)+ | |
| 225 | apply (subgoal_tac "ca * abs(f xa) <= ca * (c * abs(g xa))"); | |
| 226 | apply (erule order_trans) | |
| 227 | apply (simp add: mult_ac) | |
| 228 | apply (rule mult_left_mono, assumption) | |
| 229 | apply (rule order_less_imp_le, assumption); | |
| 230 | done | |
| 231 | ||
| 232 | ||
| 28592 | 233 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_refl"*}
 | 
| 23449 | 234 | lemma bigo_refl [intro]: "f : O(f)" | 
| 235 | apply(auto simp add: bigo_def) | |
| 236 | proof (neg_clausify) | |
| 237 | fix x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 238 | assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 239 | have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 240 | by (metis mult_le_cancel_right1 order_eq_iff) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 241 | have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 242 | by (metis order_eq_iff 1) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 243 | show "False" | 
| 23449 | 244 | by (metis 0 2) | 
| 245 | qed | |
| 246 | ||
| 28592 | 247 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_zero"*}
 | 
| 23449 | 248 | lemma bigo_zero: "0 : O(g)" | 
| 249 | apply (auto simp add: bigo_def func_zero) | |
| 250 | proof (neg_clausify) | |
| 251 | fix x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 252 | assume 0: "\<And>xa. \<not> (0\<Colon>'b) \<le> xa * \<bar>g (x xa)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 253 | have 1: "\<not> (0\<Colon>'b) \<le> (0\<Colon>'b)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 254 | by (metis 0 mult_eq_0_iff) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 255 | show "False" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 256 | by (metis 1 linorder_neq_iff linorder_antisym_conv1) | 
| 23449 | 257 | qed | 
| 258 | ||
| 259 | lemma bigo_zero2: "O(%x.0) = {%x.0}"
 | |
| 260 | apply (auto simp add: bigo_def) | |
| 261 | apply (rule ext) | |
| 262 | apply auto | |
| 263 | done | |
| 264 | ||
| 265 | lemma bigo_plus_self_subset [intro]: | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 266 | "O(f) \<oplus> O(f) <= O(f)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 267 | apply (auto simp add: bigo_alt_def set_plus_def) | 
| 23449 | 268 | apply (rule_tac x = "c + ca" in exI) | 
| 269 | apply auto | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23464diff
changeset | 270 | apply (simp add: ring_distribs func_plus) | 
| 23449 | 271 | apply (blast intro:order_trans abs_triangle_ineq add_mono elim:) | 
| 272 | done | |
| 273 | ||
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 274 | lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)" | 
| 23449 | 275 | apply (rule equalityI) | 
| 276 | apply (rule bigo_plus_self_subset) | |
| 277 | apply (rule set_zero_plus2) | |
| 278 | apply (rule bigo_zero) | |
| 279 | done | |
| 280 | ||
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 281 | lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)" | 
| 23449 | 282 | apply (rule subsetI) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 283 | apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) | 
| 23449 | 284 | apply (subst bigo_pos_const [symmetric])+ | 
| 285 | apply (rule_tac x = | |
| 286 | "%n. if abs (g n) <= (abs (f n)) then x n else 0" in exI) | |
| 287 | apply (rule conjI) | |
| 288 | apply (rule_tac x = "c + c" in exI) | |
| 289 | apply (clarsimp) | |
| 290 | apply (auto) | |
| 291 | apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (f xa)") | |
| 292 | apply (erule_tac x = xa in allE) | |
| 293 | apply (erule order_trans) | |
| 294 | apply (simp) | |
| 295 | apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") | |
| 296 | apply (erule order_trans) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23464diff
changeset | 297 | apply (simp add: ring_distribs) | 
| 23449 | 298 | apply (rule mult_left_mono) | 
| 299 | apply assumption | |
| 300 | apply (simp add: order_less_le) | |
| 301 | apply (rule mult_left_mono) | |
| 302 | apply (simp add: abs_triangle_ineq) | |
| 303 | apply (simp add: order_less_le) | |
| 304 | apply (rule mult_nonneg_nonneg) | |
| 305 | apply (rule add_nonneg_nonneg) | |
| 306 | apply auto | |
| 307 | apply (rule_tac x = "%n. if (abs (f n)) < abs (g n) then x n else 0" | |
| 308 | in exI) | |
| 309 | apply (rule conjI) | |
| 310 | apply (rule_tac x = "c + c" in exI) | |
| 311 | apply auto | |
| 312 | apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (g xa)") | |
| 313 | apply (erule_tac x = xa in allE) | |
| 314 | apply (erule order_trans) | |
| 315 | apply (simp) | |
| 316 | apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") | |
| 317 | apply (erule order_trans) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23464diff
changeset | 318 | apply (simp add: ring_distribs) | 
| 23449 | 319 | apply (rule mult_left_mono) | 
| 320 | apply (simp add: order_less_le) | |
| 321 | apply (simp add: order_less_le) | |
| 322 | apply (rule mult_left_mono) | |
| 323 | apply (rule abs_triangle_ineq) | |
| 324 | apply (simp add: order_less_le) | |
| 25087 | 325 | apply (metis abs_not_less_zero double_less_0_iff less_not_permute linorder_not_less mult_less_0_iff) | 
| 23449 | 326 | apply (rule ext) | 
| 327 | apply (auto simp add: if_splits linorder_not_le) | |
| 328 | done | |
| 329 | ||
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 330 | 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 | 331 | apply (subgoal_tac "A \<oplus> B <= O(f) \<oplus> O(f)") | 
| 23449 | 332 | apply (erule order_trans) | 
| 333 | apply simp | |
| 334 | apply (auto del: subsetI simp del: bigo_plus_idemp) | |
| 335 | done | |
| 336 | ||
| 28592 | 337 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_plus_eq"*}
 | 
| 23449 | 338 | 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 | 339 | O(f + g) = O(f) \<oplus> O(g)" | 
| 23449 | 340 | apply (rule equalityI) | 
| 341 | apply (rule bigo_plus_subset) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 342 | apply (simp add: bigo_alt_def set_plus_def func_plus) | 
| 23449 | 343 | apply clarify | 
| 344 | (*sledgehammer*); | |
| 345 | apply (rule_tac x = "max c ca" in exI) | |
| 346 | apply (rule conjI) | |
| 25087 | 347 | apply (metis Orderings.less_max_iff_disj) | 
| 23449 | 348 | apply clarify | 
| 349 | apply (drule_tac x = "xa" in spec)+ | |
| 350 | apply (subgoal_tac "0 <= f xa + g xa") | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23464diff
changeset | 351 | apply (simp add: ring_distribs) | 
| 23449 | 352 | apply (subgoal_tac "abs(a xa + b xa) <= abs(a xa) + abs(b xa)") | 
| 353 | apply (subgoal_tac "abs(a xa) + abs(b xa) <= | |
| 354 | max c ca * f xa + max c ca * g xa") | |
| 355 | apply (blast intro: order_trans) | |
| 356 | defer 1 | |
| 357 | apply (rule abs_triangle_ineq) | |
| 25087 | 358 | apply (metis add_nonneg_nonneg) | 
| 23449 | 359 | apply (rule add_mono) | 
| 28592 | 360 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_plus_eq_simpler"*} 
 | 
| 24942 | 361 | (*Found by SPASS; SLOW*) | 
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 362 | apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans) | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25592diff
changeset | 363 | apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans) | 
| 23449 | 364 | done | 
| 365 | ||
| 28592 | 366 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded_alt"*}
 | 
| 23449 | 367 | lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> | 
| 368 | f : O(g)" | |
| 369 | apply (auto simp add: bigo_def) | |
| 370 | (*Version 1: one-shot proof*) | |
| 26645 
e114be97befe
Changed naming scheme for theorems generated by interpretations.
 ballarin parents: 
26483diff
changeset | 371 | apply (metis OrderedGroup.abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) Ring_and_Field.abs_mult) | 
| 23449 | 372 | done | 
| 373 | ||
| 26312 | 374 | lemma (*bigo_bounded_alt:*) "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> | 
| 23449 | 375 | f : O(g)" | 
| 376 | apply (auto simp add: bigo_def) | |
| 377 | (*Version 2: single-step proof*) | |
| 378 | proof (neg_clausify) | |
| 379 | fix x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 380 | assume 0: "\<And>x. f x \<le> c * g x" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 381 | assume 1: "\<And>xa. \<not> f (x xa) \<le> xa * \<bar>g (x xa)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 382 | have 2: "\<And>X3. c * g X3 = f X3 \<or> \<not> c * g X3 \<le> f X3" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 383 | by (metis 0 order_antisym_conv) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 384 | have 3: "\<And>X3. \<not> f (x \<bar>X3\<bar>) \<le> \<bar>X3 * g (x \<bar>X3\<bar>)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 385 | by (metis 1 abs_mult) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 386 | have 4: "\<And>X1 X3\<Colon>'b\<Colon>ordered_idom. X3 \<le> X1 \<or> X1 \<le> \<bar>X3\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 387 | by (metis linorder_linear abs_le_D1) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 388 | have 5: "\<And>X3::'b. \<bar>X3\<bar> * \<bar>X3\<bar> = X3 * X3" | 
| 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 | 389 | by (metis abs_mult_self) | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 390 | have 6: "\<And>X3. \<not> X3 * X3 < (0\<Colon>'b\<Colon>ordered_idom)" | 
| 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 | 391 | by (metis not_square_less_zero) | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 392 | have 7: "\<And>X1 X3::'b. \<bar>X1\<bar> * \<bar>X3\<bar> = \<bar>X3 * X1\<bar>" | 
| 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 | 393 | by (metis abs_mult mult_commute) | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 394 | have 8: "\<And>X3::'b. X3 * X3 = \<bar>X3 * X3\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 395 | by (metis abs_mult 5) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 396 | have 9: "\<And>X3. X3 * g (x \<bar>X3\<bar>) \<le> f (x \<bar>X3\<bar>)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 397 | by (metis 3 4) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 398 | have 10: "c * g (x \<bar>c\<bar>) = f (x \<bar>c\<bar>)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 399 | by (metis 2 9) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 400 | have 11: "\<And>X3::'b. \<bar>X3\<bar> * \<bar>\<bar>X3\<bar>\<bar> = \<bar>X3\<bar> * \<bar>X3\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 401 | by (metis abs_idempotent abs_mult 8) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 402 | have 12: "\<And>X3::'b. \<bar>X3 * \<bar>X3\<bar>\<bar> = \<bar>X3\<bar> * \<bar>X3\<bar>" | 
| 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 | 403 | by (metis mult_commute 7 11) | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 404 | have 13: "\<And>X3::'b. \<bar>X3 * \<bar>X3\<bar>\<bar> = X3 * X3" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 405 | by (metis 8 7 12) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 406 | have 14: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> X3 < (0\<Colon>'b)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 407 | by (metis abs_ge_self abs_le_D1 abs_if) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 408 | have 15: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> \<bar>X3\<bar> < (0\<Colon>'b)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 409 | by (metis abs_ge_self abs_le_D1 abs_if) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 410 | have 16: "\<And>X3. X3 * X3 < (0\<Colon>'b) \<or> X3 * \<bar>X3\<bar> \<le> X3 * X3" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 411 | by (metis 15 13) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 412 | have 17: "\<And>X3::'b. X3 * \<bar>X3\<bar> \<le> X3 * X3" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 413 | by (metis 16 6) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 414 | have 18: "\<And>X3. X3 \<le> \<bar>X3\<bar> \<or> \<not> X3 < (0\<Colon>'b)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 415 | by (metis mult_le_cancel_left 17) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 416 | have 19: "\<And>X3::'b. X3 \<le> \<bar>X3\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 417 | by (metis 18 14) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 418 | have 20: "\<not> f (x \<bar>c\<bar>) \<le> \<bar>f (x \<bar>c\<bar>)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 419 | by (metis 3 10) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 420 | show "False" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 421 | by (metis 20 19) | 
| 23449 | 422 | qed | 
| 423 | ||
| 424 | ||
| 425 | text{*So here is the easier (and more natural) problem using transitivity*}
 | |
| 28592 | 426 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded_alt_trans"*}
 | 
| 23449 | 427 | lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" | 
| 428 | apply (auto simp add: bigo_def) | |
| 429 | (*Version 1: one-shot proof*) | |
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25592diff
changeset | 430 | apply (metis Orderings.leD Orderings.leI abs_ge_self abs_le_D1 abs_mult abs_of_nonneg order_le_less) | 
| 23449 | 431 | done | 
| 432 | ||
| 433 | text{*So here is the easier (and more natural) problem using transitivity*}
 | |
| 28592 | 434 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded_alt_trans"*}
 | 
| 23449 | 435 | lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" | 
| 436 | apply (auto simp add: bigo_def) | |
| 437 | (*Version 2: single-step proof*) | |
| 438 | proof (neg_clausify) | |
| 439 | fix x | |
| 23519 | 440 | assume 0: "\<And>A\<Colon>'a\<Colon>type. | 
| 441 | (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) A | |
| 442 | \<le> (c\<Colon>'b\<Colon>ordered_idom) * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) A" | |
| 443 | assume 1: "\<And>A\<Colon>'b\<Colon>ordered_idom. | |
| 444 | \<not> (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) A) | |
| 445 | \<le> A * \<bar>(g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x A)\<bar>" | |
| 446 | have 2: "\<And>X2\<Colon>'a\<Colon>type. | |
| 447 | \<not> (c\<Colon>'b\<Colon>ordered_idom) * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) X2 | |
| 448 | < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) X2" | |
| 449 | by (metis 0 linorder_not_le) | |
| 450 | have 3: "\<And>X2\<Colon>'b\<Colon>ordered_idom. | |
| 451 | \<not> (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>) | |
| 452 | \<le> \<bar>X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)\<bar>" | |
| 453 | by (metis abs_mult 1) | |
| 454 | have 4: "\<And>X2\<Colon>'b\<Colon>ordered_idom. | |
| 455 | \<bar>X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>)\<bar> | |
| 456 | < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)" | |
| 457 | by (metis 3 linorder_not_less) | |
| 458 | have 5: "\<And>X2\<Colon>'b\<Colon>ordered_idom. | |
| 459 | X2 * (g\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) \<bar>X2\<bar>) | |
| 460 | < (f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) (x \<bar>X2\<bar>)" | |
| 461 | by (metis abs_less_iff 4) | |
| 462 | show "False" | |
| 463 | by (metis 2 5) | |
| 23449 | 464 | qed | 
| 465 | ||
| 466 | ||
| 467 | lemma bigo_bounded: "ALL x. 0 <= f x ==> ALL x. f x <= g x ==> | |
| 468 | f : O(g)" | |
| 469 | apply (erule bigo_bounded_alt [of f 1 g]) | |
| 470 | apply simp | |
| 471 | done | |
| 472 | ||
| 28592 | 473 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_bounded2"*}
 | 
| 23449 | 474 | lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==> | 
| 475 | f : lb +o O(g)" | |
| 476 | apply (rule set_minus_imp_plus) | |
| 477 | apply (rule bigo_bounded) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 478 | apply (auto simp add: diff_minus fun_Compl_def func_plus) | 
| 23449 | 479 | prefer 2 | 
| 480 | apply (drule_tac x = x in spec)+ | |
| 481 | apply arith (*not clear that it's provable otherwise*) | |
| 482 | proof (neg_clausify) | |
| 483 | fix x | |
| 484 | assume 0: "\<And>y. lb y \<le> f y" | |
| 485 | assume 1: "\<not> (0\<Colon>'b) \<le> f x + - lb x" | |
| 486 | have 2: "\<And>X3. (0\<Colon>'b) + X3 = X3" | |
| 487 | by (metis diff_eq_eq right_minus_eq) | |
| 488 | have 3: "\<not> (0\<Colon>'b) \<le> f x - lb x" | |
| 29667 | 489 | by (metis 1 diff_minus) | 
| 23449 | 490 | have 4: "\<not> (0\<Colon>'b) + lb x \<le> f x" | 
| 491 | by (metis 3 le_diff_eq) | |
| 492 | show "False" | |
| 493 | by (metis 4 2 0) | |
| 494 | qed | |
| 495 | ||
| 28592 | 496 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_abs"*}
 | 
| 23449 | 497 | lemma bigo_abs: "(%x. abs(f x)) =o O(f)" | 
| 498 | apply (unfold bigo_def) | |
| 499 | apply auto | |
| 500 | proof (neg_clausify) | |
| 501 | fix x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 502 | assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 503 | have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 504 | by (metis mult_le_cancel_right1 order_eq_iff) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 505 | have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 506 | by (metis order_eq_iff 1) | 
| 23449 | 507 | show "False" | 
| 508 | by (metis 0 2) | |
| 509 | qed | |
| 510 | ||
| 28592 | 511 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_abs2"*}
 | 
| 23449 | 512 | lemma bigo_abs2: "f =o O(%x. abs(f x))" | 
| 513 | apply (unfold bigo_def) | |
| 514 | apply auto | |
| 515 | proof (neg_clausify) | |
| 516 | fix x | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 517 | assume 0: "\<And>xa. \<not> \<bar>f (x xa)\<bar> \<le> xa * \<bar>f (x xa)\<bar>" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 518 | have 1: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2 \<or> \<not> (1\<Colon>'b) \<le> (1\<Colon>'b)" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 519 | by (metis mult_le_cancel_right1 order_eq_iff) | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 520 | have 2: "\<And>X2. X2 \<le> (1\<Colon>'b) * X2" | 
| 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 521 | by (metis order_eq_iff 1) | 
| 23449 | 522 | show "False" | 
| 523 | by (metis 0 2) | |
| 524 | qed | |
| 525 | ||
| 526 | lemma bigo_abs3: "O(f) = O(%x. abs(f x))" | |
| 527 | apply (rule equalityI) | |
| 528 | apply (rule bigo_elt_subset) | |
| 529 | apply (rule bigo_abs2) | |
| 530 | apply (rule bigo_elt_subset) | |
| 531 | apply (rule bigo_abs) | |
| 532 | done | |
| 533 | ||
| 534 | lemma bigo_abs4: "f =o g +o O(h) ==> | |
| 535 | (%x. abs (f x)) =o (%x. abs (g x)) +o O(h)" | |
| 536 | apply (drule set_plus_imp_minus) | |
| 537 | apply (rule set_minus_imp_plus) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 538 | apply (subst fun_diff_def) | 
| 23449 | 539 | proof - | 
| 540 | assume a: "f - g : O(h)" | |
| 541 | have "(%x. abs (f x) - abs (g x)) =o O(%x. abs(abs (f x) - abs (g x)))" | |
| 542 | by (rule bigo_abs2) | |
| 543 | also have "... <= O(%x. abs (f x - g x))" | |
| 544 | apply (rule bigo_elt_subset) | |
| 545 | apply (rule bigo_bounded) | |
| 546 | apply force | |
| 547 | apply (rule allI) | |
| 548 | apply (rule abs_triangle_ineq3) | |
| 549 | done | |
| 550 | also have "... <= O(f - g)" | |
| 551 | apply (rule bigo_elt_subset) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 552 | apply (subst fun_diff_def) | 
| 23449 | 553 | apply (rule bigo_abs) | 
| 554 | done | |
| 555 | also have "... <= O(h)" | |
| 23464 | 556 | using a by (rule bigo_elt_subset) | 
| 23449 | 557 | finally show "(%x. abs (f x) - abs (g x)) : O(h)". | 
| 558 | qed | |
| 559 | ||
| 560 | lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" | |
| 561 | by (unfold bigo_def, auto) | |
| 562 | ||
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 563 | lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \<oplus> O(h)" | 
| 23449 | 564 | proof - | 
| 565 | assume "f : g +o O(h)" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 566 | also have "... <= O(g) \<oplus> O(h)" | 
| 23449 | 567 | by (auto del: subsetI) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 568 | also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))" | 
| 23449 | 569 | apply (subst bigo_abs3 [symmetric])+ | 
| 570 | apply (rule refl) | |
| 571 | done | |
| 572 | also have "... = O((%x. abs(g x)) + (%x. abs(h x)))" | |
| 573 | by (rule bigo_plus_eq [symmetric], auto) | |
| 574 | finally have "f : ...". | |
| 575 | then have "O(f) <= ..." | |
| 576 | by (elim bigo_elt_subset) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 577 | also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))" | 
| 23449 | 578 | by (rule bigo_plus_eq, auto) | 
| 579 | finally show ?thesis | |
| 580 | by (simp add: bigo_abs3 [symmetric]) | |
| 581 | qed | |
| 582 | ||
| 28592 | 583 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult"*}
 | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 584 | lemma bigo_mult [intro]: "O(f)\<otimes>O(g) <= O(f * g)" | 
| 23449 | 585 | apply (rule subsetI) | 
| 586 | apply (subst bigo_def) | |
| 587 | 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 | 588 | simp add: bigo_alt_def set_times_def func_times) | 
| 23449 | 589 | (*sledgehammer*); | 
| 590 | apply (rule_tac x = "c * ca" in exI) | |
| 591 | apply(rule allI) | |
| 592 | apply(erule_tac x = x in allE)+ | |
| 593 | apply(subgoal_tac "c * ca * abs(f x * g x) = | |
| 594 | (c * abs(f x)) * (ca * abs(g x))") | |
| 28592 | 595 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult_simpler"*}
 | 
| 23449 | 596 | 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 | 597 | apply (metis mult_assoc mult_left_commute | 
| 
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 | 598 | OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute | 
| 
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 | 599 | Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos) | 
| 
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 | 600 | apply (erule ssubst) | 
| 23449 | 601 | apply (subst abs_mult) | 
| 602 | (*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has | |
| 603 | just been done*) | |
| 604 | proof (neg_clausify) | |
| 605 | fix a c b ca x | |
| 606 | assume 0: "(0\<Colon>'b\<Colon>ordered_idom) < (c\<Colon>'b\<Colon>ordered_idom)" | |
| 607 | assume 1: "\<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> | |
| 608 | \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>" | |
| 609 | assume 2: "\<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> | |
| 610 | \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>" | |
| 611 | assume 3: "\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> * | |
| 612 | \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> | |
| 613 | \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> * | |
| 614 | ((ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>)" | |
| 615 | have 4: "\<bar>c\<Colon>'b\<Colon>ordered_idom\<bar> = c" | |
| 616 | by (metis OrderedGroup.abs_of_pos 0) | |
| 617 | have 5: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (c\<Colon>'b\<Colon>ordered_idom) * \<bar>X1\<bar> = \<bar>c * X1\<bar>" | |
| 618 | by (metis Ring_and_Field.abs_mult 4) | |
| 619 | have 6: "(0\<Colon>'b\<Colon>ordered_idom) = (1\<Colon>'b\<Colon>ordered_idom) \<or> | |
| 620 | (0\<Colon>'b\<Colon>ordered_idom) < (1\<Colon>'b\<Colon>ordered_idom)" | |
| 621 | by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_ordered_idom) | |
| 622 | have 7: "(0\<Colon>'b\<Colon>ordered_idom) < (1\<Colon>'b\<Colon>ordered_idom)" | |
| 623 | by (metis 6 Ring_and_Field.one_neq_zero) | |
| 624 | have 8: "\<bar>1\<Colon>'b\<Colon>ordered_idom\<bar> = (1\<Colon>'b\<Colon>ordered_idom)" | |
| 625 | by (metis OrderedGroup.abs_of_pos 7) | |
| 626 | have 9: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (0\<Colon>'b\<Colon>ordered_idom) \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>X1\<bar>" | |
| 627 | by (metis OrderedGroup.abs_ge_zero 5) | |
| 628 | have 10: "\<And>X1\<Colon>'b\<Colon>ordered_idom. X1 * (1\<Colon>'b\<Colon>ordered_idom) = X1" | |
| 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 | 629 | by (metis Ring_and_Field.mult_cancel_right2 mult_commute) | 
| 23449 | 630 | have 11: "\<And>X1\<Colon>'b\<Colon>ordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar> * \<bar>1\<Colon>'b\<Colon>ordered_idom\<bar>" | 
| 631 | by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10) | |
| 632 | have 12: "\<And>X1\<Colon>'b\<Colon>ordered_idom. \<bar>\<bar>X1\<bar>\<bar> = \<bar>X1\<bar>" | |
| 633 | by (metis 11 8 10) | |
| 634 | have 13: "\<And>X1\<Colon>'b\<Colon>ordered_idom. (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>X1\<bar>" | |
| 635 | by (metis OrderedGroup.abs_ge_zero 12) | |
| 636 | have 14: "\<not> (0\<Colon>'b\<Colon>ordered_idom) | |
| 637 | \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> \<or> | |
| 638 | \<not> (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or> | |
| 639 | \<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or> | |
| 640 | \<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<le> c * \<bar>f x\<bar>" | |
| 641 | by (metis 3 Ring_and_Field.mult_mono) | |
| 642 | have 15: "\<not> (0\<Colon>'b\<Colon>ordered_idom) \<le> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> \<or> | |
| 643 | \<not> \<bar>b x\<bar> \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or> | |
| 644 | \<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> | |
| 645 | \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>" | |
| 646 | by (metis 14 9) | |
| 647 | have 16: "\<not> \<bar>(b\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> | |
| 648 | \<le> (ca\<Colon>'b\<Colon>ordered_idom) * \<bar>(g\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> \<or> | |
| 649 | \<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar> | |
| 650 | \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>" | |
| 651 | by (metis 15 13) | |
| 652 | have 17: "\<not> \<bar>(a\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) (x\<Colon>'a)\<bar> | |
| 653 | \<le> (c\<Colon>'b\<Colon>ordered_idom) * \<bar>(f\<Colon>'a \<Rightarrow> 'b\<Colon>ordered_idom) x\<bar>" | |
| 654 | by (metis 16 2) | |
| 655 | show 18: "False" | |
| 656 | by (metis 17 1) | |
| 657 | qed | |
| 658 | ||
| 659 | ||
| 28592 | 660 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult2"*}
 | 
| 23449 | 661 | lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" | 
| 662 | apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) | |
| 663 | (*sledgehammer*); | |
| 664 | apply (rule_tac x = c in exI) | |
| 665 | apply clarify | |
| 666 | apply (drule_tac x = x in spec) | |
| 28592 | 667 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult2_simpler"*}
 | 
| 24942 | 668 | (*sledgehammer [no luck]*); | 
| 23449 | 669 | apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))") | 
| 670 | apply (simp add: mult_ac) | |
| 671 | apply (rule mult_left_mono, assumption) | |
| 672 | apply (rule abs_ge_zero) | |
| 673 | done | |
| 674 | ||
| 28592 | 675 | ML_command{*AtpWrapper.problem_name:="BigO__bigo_mult3"*}
 | 
| 23449 | 676 | lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)" | 
| 677 | by (metis bigo_mult set_times_intro subset_iff) | |
| 678 | ||
| 28592 | 679 | ML_command{*AtpWrapper.problem_name:="BigO__bigo_mult4"*}
 | 
| 23449 | 680 | lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)" | 
| 681 | by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) | |
| 682 | ||
| 683 | ||
| 684 | lemma bigo_mult5: "ALL x. f x ~= 0 ==> | |
| 685 |     O(f * g) <= (f::'a => ('b::ordered_field)) *o O(g)"
 | |
| 686 | proof - | |
| 687 | assume "ALL x. f x ~= 0" | |
| 688 | show "O(f * g) <= f *o O(g)" | |
| 689 | proof | |
| 690 | fix h | |
| 691 | assume "h : O(f * g)" | |
| 692 | then have "(%x. 1 / (f x)) * h : (%x. 1 / f x) *o O(f * g)" | |
| 693 | by auto | |
| 694 | also have "... <= O((%x. 1 / f x) * (f * g))" | |
| 695 | by (rule bigo_mult2) | |
| 696 | also have "(%x. 1 / f x) * (f * g) = g" | |
| 697 | apply (simp add: func_times) | |
| 698 | apply (rule ext) | |
| 699 | apply (simp add: prems nonzero_divide_eq_eq mult_ac) | |
| 700 | done | |
| 701 | finally have "(%x. (1::'b) / f x) * h : O(g)". | |
| 702 | then have "f * ((%x. (1::'b) / f x) * h) : f *o O(g)" | |
| 703 | by auto | |
| 704 | also have "f * ((%x. (1::'b) / f x) * h) = h" | |
| 705 | apply (simp add: func_times) | |
| 706 | apply (rule ext) | |
| 707 | apply (simp add: prems nonzero_divide_eq_eq mult_ac) | |
| 708 | done | |
| 709 | finally show "h : f *o O(g)". | |
| 710 | qed | |
| 711 | qed | |
| 712 | ||
| 28592 | 713 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult6"*}
 | 
| 23449 | 714 | lemma bigo_mult6: "ALL x. f x ~= 0 ==> | 
| 715 |     O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)"
 | |
| 716 | by (metis bigo_mult2 bigo_mult5 order_antisym) | |
| 717 | ||
| 718 | (*proof requires relaxing relevance: 2007-01-25*) | |
| 28592 | 719 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult7"*}
 | 
| 23449 | 720 | declare bigo_mult6 [simp] | 
| 721 | lemma bigo_mult7: "ALL x. f x ~= 0 ==> | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 722 |     O(f * g) <= O(f::'a => ('b::ordered_field)) \<otimes> O(g)"
 | 
| 23449 | 723 | (*sledgehammer*) | 
| 724 | apply (subst bigo_mult6) | |
| 725 | apply assumption | |
| 726 | apply (rule set_times_mono3) | |
| 727 | apply (rule bigo_refl) | |
| 728 | done | |
| 729 | declare bigo_mult6 [simp del] | |
| 730 | ||
| 28592 | 731 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_mult8"*}
 | 
| 23449 | 732 | declare bigo_mult7[intro!] | 
| 733 | lemma bigo_mult8: "ALL x. f x ~= 0 ==> | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 734 |     O(f * g) = O(f::'a => ('b::ordered_field)) \<otimes> O(g)"
 | 
| 23449 | 735 | by (metis bigo_mult bigo_mult7 order_antisym_conv) | 
| 736 | ||
| 737 | 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 | 738 | by (auto simp add: bigo_def fun_Compl_def) | 
| 23449 | 739 | |
| 740 | lemma bigo_minus2: "f : g +o O(h) ==> -f : -g +o O(h)" | |
| 741 | apply (rule set_minus_imp_plus) | |
| 742 | apply (drule set_plus_imp_minus) | |
| 743 | apply (drule bigo_minus) | |
| 744 | apply (simp add: diff_minus) | |
| 745 | done | |
| 746 | ||
| 747 | 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 | 748 | by (auto simp add: bigo_def fun_Compl_def abs_minus_cancel) | 
| 23449 | 749 | |
| 750 | lemma bigo_plus_absorb_lemma1: "f : O(g) ==> f +o O(g) <= O(g)" | |
| 751 | proof - | |
| 752 | assume a: "f : O(g)" | |
| 753 | show "f +o O(g) <= O(g)" | |
| 754 | proof - | |
| 755 | have "f : O(f)" by auto | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 756 | then have "f +o O(g) <= O(f) \<oplus> O(g)" | 
| 23449 | 757 | by (auto del: subsetI) | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 758 | also have "... <= O(g) \<oplus> O(g)" | 
| 23449 | 759 | proof - | 
| 760 | from a have "O(f) <= O(g)" by (auto del: subsetI) | |
| 761 | thus ?thesis by (auto del: subsetI) | |
| 762 | qed | |
| 763 | also have "... <= O(g)" by (simp add: bigo_plus_idemp) | |
| 764 | finally show ?thesis . | |
| 765 | qed | |
| 766 | qed | |
| 767 | ||
| 768 | lemma bigo_plus_absorb_lemma2: "f : O(g) ==> O(g) <= f +o O(g)" | |
| 769 | proof - | |
| 770 | assume a: "f : O(g)" | |
| 771 | show "O(g) <= f +o O(g)" | |
| 772 | proof - | |
| 773 | from a have "-f : O(g)" by auto | |
| 774 | then have "-f +o O(g) <= O(g)" by (elim bigo_plus_absorb_lemma1) | |
| 775 | then have "f +o (-f +o O(g)) <= f +o O(g)" by auto | |
| 776 | also have "f +o (-f +o O(g)) = O(g)" | |
| 777 | by (simp add: set_plus_rearranges) | |
| 778 | finally show ?thesis . | |
| 779 | qed | |
| 780 | qed | |
| 781 | ||
| 28592 | 782 | ML_command{*AtpWrapper.problem_name:="BigO__bigo_plus_absorb"*}
 | 
| 23449 | 783 | lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)" | 
| 784 | by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff); | |
| 785 | ||
| 786 | lemma bigo_plus_absorb2 [intro]: "f : O(g) ==> A <= O(g) ==> f +o A <= O(g)" | |
| 787 | apply (subgoal_tac "f +o A <= f +o O(g)") | |
| 788 | apply force+ | |
| 789 | done | |
| 790 | ||
| 791 | lemma bigo_add_commute_imp: "f : g +o O(h) ==> g : f +o O(h)" | |
| 792 | apply (subst set_minus_plus [symmetric]) | |
| 793 | apply (subgoal_tac "g - f = - (f - g)") | |
| 794 | apply (erule ssubst) | |
| 795 | apply (rule bigo_minus) | |
| 796 | apply (subst set_minus_plus) | |
| 797 | apply assumption | |
| 798 | apply (simp add: diff_minus add_ac) | |
| 799 | done | |
| 800 | ||
| 801 | lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))" | |
| 802 | apply (rule iffI) | |
| 803 | apply (erule bigo_add_commute_imp)+ | |
| 804 | done | |
| 805 | ||
| 806 | lemma bigo_const1: "(%x. c) : O(%x. 1)" | |
| 807 | by (auto simp add: bigo_def mult_ac) | |
| 808 | ||
| 28592 | 809 | ML_command{*AtpWrapper.problem_name:="BigO__bigo_const2"*}
 | 
| 23449 | 810 | lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)" | 
| 811 | by (metis bigo_const1 bigo_elt_subset); | |
| 812 | ||
| 24855 | 813 | lemma bigo_const2 [intro]: "O(%x. c::'b::ordered_idom) <= O(%x. 1)"; | 
| 23449 | 814 | (*??FAILS because the two occurrences of COMBK have different polymorphic types | 
| 815 | proof (neg_clausify) | |
| 816 | assume 0: "\<not> O(COMBK (c\<Colon>'b\<Colon>ordered_idom)) \<subseteq> O(COMBK (1\<Colon>'b\<Colon>ordered_idom))" | |
| 817 | have 1: "COMBK (c\<Colon>'b\<Colon>ordered_idom) \<notin> O(COMBK (1\<Colon>'b\<Colon>ordered_idom))" | |
| 818 | apply (rule notI) | |
| 819 | apply (rule 0 [THEN notE]) | |
| 820 | apply (rule bigo_elt_subset) | |
| 821 | apply assumption; | |
| 822 | sorry | |
| 823 | by (metis 0 bigo_elt_subset) loops?? | |
| 824 | show "False" | |
| 825 | by (metis 1 bigo_const1) | |
| 826 | qed | |
| 827 | *) | |
| 828 | apply (rule bigo_elt_subset) | |
| 829 | apply (rule bigo_const1) | |
| 830 | done | |
| 831 | ||
| 28592 | 832 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_const3"*}
 | 
| 23449 | 833 | lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)" | 
| 834 | apply (simp add: bigo_def) | |
| 835 | proof (neg_clausify) | |
| 836 | assume 0: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> (0\<Colon>'a\<Colon>ordered_field)" | |
| 23519 | 837 | assume 1: "\<And>A\<Colon>'a\<Colon>ordered_field. \<not> (1\<Colon>'a\<Colon>ordered_field) \<le> A * \<bar>c\<Colon>'a\<Colon>ordered_field\<bar>" | 
| 23449 | 838 | have 2: "(0\<Colon>'a\<Colon>ordered_field) = \<bar>c\<Colon>'a\<Colon>ordered_field\<bar> \<or> | 
| 839 | \<not> (1\<Colon>'a\<Colon>ordered_field) \<le> (1\<Colon>'a\<Colon>ordered_field)" | |
| 840 | by (metis 1 field_inverse) | |
| 841 | have 3: "\<bar>c\<Colon>'a\<Colon>ordered_field\<bar> = (0\<Colon>'a\<Colon>ordered_field)" | |
| 23519 | 842 | by (metis linorder_neq_iff linorder_antisym_conv1 2) | 
| 23449 | 843 | have 4: "(0\<Colon>'a\<Colon>ordered_field) = (c\<Colon>'a\<Colon>ordered_field)" | 
| 23519 | 844 | by (metis 3 abs_eq_0) | 
| 845 | show "False" | |
| 846 | by (metis 0 4) | |
| 23449 | 847 | qed | 
| 848 | ||
| 849 | lemma bigo_const4: "(c::'a::ordered_field) ~= 0 ==> O(%x. 1) <= O(%x. c)" | |
| 850 | by (rule bigo_elt_subset, rule bigo_const3, assumption) | |
| 851 | ||
| 852 | lemma bigo_const [simp]: "(c::'a::ordered_field) ~= 0 ==> | |
| 853 | O(%x. c) = O(%x. 1)" | |
| 854 | by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption) | |
| 855 | ||
| 28592 | 856 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult1"*}
 | 
| 23449 | 857 | 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 | 858 | apply (simp add: bigo_def abs_mult) | 
| 23449 | 859 | proof (neg_clausify) | 
| 860 | fix x | |
| 25304 
7491c00f0915
removed subclass edge ordered_ring < lordered_ring
 haftmann parents: 
25087diff
changeset | 861 | assume 0: "\<And>xa\<Colon>'b\<Colon>ordered_idom. | 
| 
7491c00f0915
removed subclass edge ordered_ring < lordered_ring
 haftmann parents: 
25087diff
changeset | 862 | \<not> \<bar>c\<Colon>'b\<Colon>ordered_idom\<bar> * | 
| 
7491c00f0915
removed subclass edge ordered_ring < lordered_ring
 haftmann parents: 
25087diff
changeset | 863 | \<bar>(f\<Colon>'a\<Colon>type \<Rightarrow> 'b\<Colon>ordered_idom) ((x\<Colon>'b\<Colon>ordered_idom \<Rightarrow> 'a\<Colon>type) xa)\<bar> | 
| 
7491c00f0915
removed subclass edge ordered_ring < lordered_ring
 haftmann parents: 
25087diff
changeset | 864 | \<le> xa * \<bar>f (x xa)\<bar>" | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 865 | show "False" | 
| 25304 
7491c00f0915
removed subclass edge ordered_ring < lordered_ring
 haftmann parents: 
25087diff
changeset | 866 | by (metis linorder_neq_iff linorder_antisym_conv1 0) | 
| 23449 | 867 | qed | 
| 868 | ||
| 869 | lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)" | |
| 870 | by (rule bigo_elt_subset, rule bigo_const_mult1) | |
| 871 | ||
| 28592 | 872 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult3"*}
 | 
| 23449 | 873 | lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)" | 
| 874 | apply (simp add: bigo_def) | |
| 24942 | 875 | (*sledgehammer [no luck]*); | 
| 23449 | 876 | apply (rule_tac x = "abs(inverse c)" in exI) | 
| 877 | apply (simp only: abs_mult [symmetric] mult_assoc [symmetric]) | |
| 878 | apply (subst left_inverse) | |
| 879 | apply (auto ); | |
| 880 | done | |
| 881 | ||
| 882 | lemma bigo_const_mult4: "(c::'a::ordered_field) ~= 0 ==> | |
| 883 | O(f) <= O(%x. c * f x)" | |
| 884 | by (rule bigo_elt_subset, rule bigo_const_mult3, assumption) | |
| 885 | ||
| 886 | lemma bigo_const_mult [simp]: "(c::'a::ordered_field) ~= 0 ==> | |
| 887 | O(%x. c * f x) = O(f)" | |
| 888 | by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4) | |
| 889 | ||
| 28592 | 890 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult5"*}
 | 
| 23449 | 891 | lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==> | 
| 892 | (%x. c) *o O(f) = O(f)" | |
| 893 | apply (auto del: subsetI) | |
| 894 | apply (rule order_trans) | |
| 895 | apply (rule bigo_mult2) | |
| 896 | apply (simp add: func_times) | |
| 897 | apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times) | |
| 898 | apply (rule_tac x = "%y. inverse c * x y" in exI) | |
| 24942 | 899 | apply (rename_tac g d) | 
| 900 | apply safe | |
| 901 | apply (rule_tac [2] ext) | |
| 902 | 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 | 903 | apply simp | 
| 24942 | 904 | apply (simp add: mult_assoc [symmetric] abs_mult) | 
| 905 | (*couldn't get this proof without the step above; SLOW*) | |
| 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 | 906 | apply (metis mult_assoc abs_ge_zero mult_left_mono) | 
| 23449 | 907 | done | 
| 908 | ||
| 909 | ||
| 28592 | 910 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_const_mult6"*}
 | 
| 23449 | 911 | lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)" | 
| 912 | apply (auto intro!: subsetI | |
| 913 | simp add: bigo_def elt_set_times_def func_times | |
| 914 | simp del: abs_mult mult_ac) | |
| 915 | (*sledgehammer*); | |
| 916 | apply (rule_tac x = "ca * (abs c)" in exI) | |
| 917 | apply (rule allI) | |
| 918 | apply (subgoal_tac "ca * abs(c) * abs(f x) = abs(c) * (ca * abs(f x))") | |
| 919 | apply (erule ssubst) | |
| 920 | apply (subst abs_mult) | |
| 921 | apply (rule mult_left_mono) | |
| 922 | apply (erule spec) | |
| 923 | apply simp | |
| 924 | apply(simp add: mult_ac) | |
| 925 | done | |
| 926 | ||
| 927 | lemma bigo_const_mult7 [intro]: "f =o O(g) ==> (%x. c * f x) =o O(g)" | |
| 928 | proof - | |
| 929 | assume "f =o O(g)" | |
| 930 | then have "(%x. c) * f =o (%x. c) *o O(g)" | |
| 931 | by auto | |
| 932 | also have "(%x. c) * f = (%x. c * f x)" | |
| 933 | by (simp add: func_times) | |
| 934 | also have "(%x. c) *o O(g) <= O(g)" | |
| 935 | by (auto del: subsetI) | |
| 936 | finally show ?thesis . | |
| 937 | qed | |
| 938 | ||
| 939 | lemma bigo_compose1: "f =o O(g) ==> (%x. f(k x)) =o O(%x. g(k x))" | |
| 940 | by (unfold bigo_def, auto) | |
| 941 | ||
| 942 | lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o | |
| 943 | O(%x. h(k x))" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 944 | apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def | 
| 23449 | 945 | func_plus) | 
| 946 | apply (erule bigo_compose1) | |
| 947 | done | |
| 948 | ||
| 949 | subsection {* Setsum *}
 | |
| 950 | ||
| 951 | lemma bigo_setsum_main: "ALL x. ALL y : A x. 0 <= h x y ==> | |
| 952 | EX c. ALL x. ALL y : A x. abs(f x y) <= c * (h x y) ==> | |
| 953 | (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" | |
| 954 | apply (auto simp add: bigo_def) | |
| 955 | apply (rule_tac x = "abs c" in exI) | |
| 956 | apply (subst abs_of_nonneg) back back | |
| 957 | apply (rule setsum_nonneg) | |
| 958 | apply force | |
| 959 | apply (subst setsum_right_distrib) | |
| 960 | apply (rule allI) | |
| 961 | apply (rule order_trans) | |
| 962 | apply (rule setsum_abs) | |
| 963 | apply (rule setsum_mono) | |
| 964 | apply (blast intro: order_trans mult_right_mono abs_ge_self) | |
| 965 | done | |
| 966 | ||
| 28592 | 967 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_setsum1"*}
 | 
| 23449 | 968 | lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> | 
| 969 | EX c. ALL x y. abs(f x y) <= c * (h x y) ==> | |
| 970 | (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" | |
| 971 | apply (rule bigo_setsum_main) | |
| 972 | (*sledgehammer*); | |
| 973 | apply force | |
| 974 | apply clarsimp | |
| 975 | apply (rule_tac x = c in exI) | |
| 976 | apply force | |
| 977 | done | |
| 978 | ||
| 979 | lemma bigo_setsum2: "ALL y. 0 <= h y ==> | |
| 980 | EX c. ALL y. abs(f y) <= c * (h y) ==> | |
| 981 | (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)" | |
| 982 | by (rule bigo_setsum1, auto) | |
| 983 | ||
| 28592 | 984 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_setsum3"*}
 | 
| 23449 | 985 | lemma bigo_setsum3: "f =o O(h) ==> | 
| 986 | (%x. SUM y : A x. (l x y) * f(k x y)) =o | |
| 987 | O(%x. SUM y : A x. abs(l x y * h(k x y)))" | |
| 988 | apply (rule bigo_setsum1) | |
| 989 | apply (rule allI)+ | |
| 990 | apply (rule abs_ge_zero) | |
| 991 | apply (unfold bigo_def) | |
| 992 | apply (auto simp add: abs_mult); | |
| 993 | (*sledgehammer*); | |
| 994 | apply (rule_tac x = c in exI) | |
| 995 | apply (rule allI)+ | |
| 996 | apply (subst mult_left_commute) | |
| 997 | apply (rule mult_left_mono) | |
| 998 | apply (erule spec) | |
| 999 | apply (rule abs_ge_zero) | |
| 1000 | done | |
| 1001 | ||
| 1002 | lemma bigo_setsum4: "f =o g +o O(h) ==> | |
| 1003 | (%x. SUM y : A x. l x y * f(k x y)) =o | |
| 1004 | (%x. SUM y : A x. l x y * g(k x y)) +o | |
| 1005 | O(%x. SUM y : A x. abs(l x y * h(k x y)))" | |
| 1006 | apply (rule set_minus_imp_plus) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1007 | apply (subst fun_diff_def) | 
| 23449 | 1008 | apply (subst setsum_subtractf [symmetric]) | 
| 1009 | apply (subst right_diff_distrib [symmetric]) | |
| 1010 | apply (rule bigo_setsum3) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1011 | apply (subst fun_diff_def [symmetric]) | 
| 23449 | 1012 | apply (erule set_plus_imp_minus) | 
| 1013 | done | |
| 1014 | ||
| 28592 | 1015 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_setsum5"*}
 | 
| 23449 | 1016 | lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> | 
| 1017 | ALL x. 0 <= h x ==> | |
| 1018 | (%x. SUM y : A x. (l x y) * f(k x y)) =o | |
| 1019 | O(%x. SUM y : A x. (l x y) * h(k x y))" | |
| 1020 | apply (subgoal_tac "(%x. SUM y : A x. (l x y) * h(k x y)) = | |
| 1021 | (%x. SUM y : A x. abs((l x y) * h(k x y)))") | |
| 1022 | apply (erule ssubst) | |
| 1023 | apply (erule bigo_setsum3) | |
| 1024 | apply (rule ext) | |
| 1025 | apply (rule setsum_cong2) | |
| 1026 | apply (thin_tac "f \<in> O(h)") | |
| 24942 | 1027 | apply (metis abs_of_nonneg zero_le_mult_iff) | 
| 23449 | 1028 | done | 
| 1029 | ||
| 1030 | lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==> | |
| 1031 | ALL x. 0 <= h x ==> | |
| 1032 | (%x. SUM y : A x. (l x y) * f(k x y)) =o | |
| 1033 | (%x. SUM y : A x. (l x y) * g(k x y)) +o | |
| 1034 | O(%x. SUM y : A x. (l x y) * h(k x y))" | |
| 1035 | apply (rule set_minus_imp_plus) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1036 | apply (subst fun_diff_def) | 
| 23449 | 1037 | apply (subst setsum_subtractf [symmetric]) | 
| 1038 | apply (subst right_diff_distrib [symmetric]) | |
| 1039 | apply (rule bigo_setsum5) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1040 | apply (subst fun_diff_def [symmetric]) | 
| 23449 | 1041 | apply (drule set_plus_imp_minus) | 
| 1042 | apply auto | |
| 1043 | done | |
| 1044 | ||
| 1045 | subsection {* Misc useful stuff *}
 | |
| 1046 | ||
| 1047 | 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 | 1048 | A \<oplus> B <= O(f)" | 
| 23449 | 1049 | apply (subst bigo_plus_idemp [symmetric]) | 
| 1050 | apply (rule set_plus_mono2) | |
| 1051 | apply assumption+ | |
| 1052 | done | |
| 1053 | ||
| 1054 | lemma bigo_useful_add: "f =o O(h) ==> g =o O(h) ==> f + g =o O(h)" | |
| 1055 | apply (subst bigo_plus_idemp [symmetric]) | |
| 1056 | apply (rule set_plus_intro) | |
| 1057 | apply assumption+ | |
| 1058 | done | |
| 1059 | ||
| 1060 | lemma bigo_useful_const_mult: "(c::'a::ordered_field) ~= 0 ==> | |
| 1061 | (%x. c) * f =o O(h) ==> f =o O(h)" | |
| 1062 | apply (rule subsetD) | |
| 1063 | apply (subgoal_tac "(%x. 1 / c) *o O(h) <= O(h)") | |
| 1064 | apply assumption | |
| 1065 | apply (rule bigo_const_mult6) | |
| 1066 | apply (subgoal_tac "f = (%x. 1 / c) * ((%x. c) * f)") | |
| 1067 | apply (erule ssubst) | |
| 1068 | apply (erule set_times_intro2) | |
| 1069 | apply (simp add: func_times) | |
| 1070 | done | |
| 1071 | ||
| 28592 | 1072 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_fix"*}
 | 
| 23449 | 1073 | lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==> | 
| 1074 | f =o O(h)" | |
| 1075 | apply (simp add: bigo_alt_def) | |
| 1076 | (*sledgehammer*); | |
| 1077 | apply clarify | |
| 1078 | apply (rule_tac x = c in exI) | |
| 1079 | apply safe | |
| 1080 | apply (case_tac "x = 0") | |
| 23816 | 1081 | apply (metis OrderedGroup.abs_ge_zero OrderedGroup.abs_zero order_less_le Ring_and_Field.split_mult_pos_le) | 
| 23449 | 1082 | apply (subgoal_tac "x = Suc (x - 1)") | 
| 23816 | 1083 | apply metis | 
| 23449 | 1084 | apply simp | 
| 1085 | done | |
| 1086 | ||
| 1087 | ||
| 1088 | lemma bigo_fix2: | |
| 1089 | "(%x. f ((x::nat) + 1)) =o (%x. g(x + 1)) +o O(%x. h(x + 1)) ==> | |
| 1090 | f 0 = g 0 ==> f =o g +o O(h)" | |
| 1091 | apply (rule set_minus_imp_plus) | |
| 1092 | apply (rule bigo_fix) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1093 | apply (subst fun_diff_def) | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1094 | apply (subst fun_diff_def [symmetric]) | 
| 23449 | 1095 | apply (rule set_plus_imp_minus) | 
| 1096 | apply simp | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1097 | apply (simp add: fun_diff_def) | 
| 23449 | 1098 | done | 
| 1099 | ||
| 1100 | subsection {* Less than or equal to *}
 | |
| 1101 | ||
| 1102 | constdefs | |
| 1103 |   lesso :: "('a => 'b::ordered_idom) => ('a => 'b) => ('a => 'b)"
 | |
| 1104 | (infixl "<o" 70) | |
| 1105 | "f <o g == (%x. max (f x - g x) 0)" | |
| 1106 | ||
| 1107 | lemma bigo_lesseq1: "f =o O(h) ==> ALL x. abs (g x) <= abs (f x) ==> | |
| 1108 | g =o O(h)" | |
| 1109 | apply (unfold bigo_def) | |
| 1110 | apply clarsimp | |
| 1111 | apply (blast intro: order_trans) | |
| 1112 | done | |
| 1113 | ||
| 1114 | lemma bigo_lesseq2: "f =o O(h) ==> ALL x. abs (g x) <= f x ==> | |
| 1115 | g =o O(h)" | |
| 1116 | apply (erule bigo_lesseq1) | |
| 1117 | apply (blast intro: abs_ge_self order_trans) | |
| 1118 | done | |
| 1119 | ||
| 1120 | lemma bigo_lesseq3: "f =o O(h) ==> ALL x. 0 <= g x ==> ALL x. g x <= f x ==> | |
| 1121 | g =o O(h)" | |
| 1122 | apply (erule bigo_lesseq2) | |
| 1123 | apply (rule allI) | |
| 1124 | apply (subst abs_of_nonneg) | |
| 1125 | apply (erule spec)+ | |
| 1126 | done | |
| 1127 | ||
| 1128 | lemma bigo_lesseq4: "f =o O(h) ==> | |
| 1129 | ALL x. 0 <= g x ==> ALL x. g x <= abs (f x) ==> | |
| 1130 | g =o O(h)" | |
| 1131 | apply (erule bigo_lesseq1) | |
| 1132 | apply (rule allI) | |
| 1133 | apply (subst abs_of_nonneg) | |
| 1134 | apply (erule spec)+ | |
| 1135 | done | |
| 1136 | ||
| 28592 | 1137 | ML_command{*AtpWrapper.problem_name:="BigO__bigo_lesso1"*}
 | 
| 23449 | 1138 | lemma bigo_lesso1: "ALL x. f x <= g x ==> f <o g =o O(h)" | 
| 1139 | apply (unfold lesso_def) | |
| 1140 | apply (subgoal_tac "(%x. max (f x - g x) 0) = 0") | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24855diff
changeset | 1141 | (*??Translation of TSTP raised an exception: Type unification failed: Variable ?'X2.0::type not of sort ord*) | 
| 25082 | 1142 | apply (metis bigo_zero) | 
| 23449 | 1143 | apply (unfold func_zero) | 
| 1144 | apply (rule ext) | |
| 1145 | apply (simp split: split_max) | |
| 1146 | done | |
| 1147 | ||
| 1148 | ||
| 28592 | 1149 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso2"*}
 | 
| 23449 | 1150 | lemma bigo_lesso2: "f =o g +o O(h) ==> | 
| 1151 | ALL x. 0 <= k x ==> ALL x. k x <= f x ==> | |
| 1152 | k <o g =o O(h)" | |
| 1153 | apply (unfold lesso_def) | |
| 1154 | apply (rule bigo_lesseq4) | |
| 1155 | apply (erule set_plus_imp_minus) | |
| 1156 | apply (rule allI) | |
| 1157 | apply (rule le_maxI2) | |
| 1158 | apply (rule allI) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1159 | apply (subst fun_diff_def) | 
| 23449 | 1160 | apply (erule thin_rl) | 
| 1161 | (*sledgehammer*); | |
| 1162 | apply (case_tac "0 <= k x - g x") | |
| 24545 | 1163 | prefer 2 (*re-order subgoals because I don't know what to put after a structured proof*) | 
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 1164 | apply (metis abs_ge_zero abs_minus_commute linorder_linear min_max.sup_absorb1 min_max.sup_commute) | 
| 24545 | 1165 | proof (neg_clausify) | 
| 1166 | fix x | |
| 1167 | assume 0: "\<And>A. k A \<le> f A" | |
| 1168 | have 1: "\<And>(X1\<Colon>'b\<Colon>ordered_idom) X2. \<not> max X1 X2 < X1" | |
| 1169 | by (metis linorder_not_less le_maxI1) (*sort inserted by hand*) | |
| 1170 | assume 2: "(0\<Colon>'b) \<le> k x - g x" | |
| 1171 | have 3: "\<not> k x - g x < (0\<Colon>'b)" | |
| 1172 | by (metis 2 linorder_not_less) | |
| 1173 | have 4: "\<And>X1 X2. min X1 (k X2) \<le> f X2" | |
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 1174 | by (metis min_max.inf_le2 min_max.le_inf_iff min_max.le_iff_inf 0) | 
| 24545 | 1175 | have 5: "\<bar>g x - f x\<bar> = f x - g x" | 
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 1176 | by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.inf_commute 4 linorder_not_le min_max.le_iff_inf 3 diff_less_0_iff_less linorder_not_less) | 
| 24545 | 1177 | have 6: "max (0\<Colon>'b) (k x - g x) = k x - g x" | 
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 1178 | by (metis min_max.le_iff_sup 2) | 
| 24545 | 1179 | assume 7: "\<not> max (k x - g x) (0\<Colon>'b) \<le> \<bar>f x - g x\<bar>" | 
| 1180 | have 8: "\<not> k x - g x \<le> f x - g x" | |
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 1181 | by (metis 5 abs_minus_commute 7 min_max.sup_commute 6) | 
| 24545 | 1182 | show "False" | 
| 29511 
7071b017cb35
migrated class package to new locale implementation
 haftmann parents: 
28592diff
changeset | 1183 | by (metis min_max.sup_commute min_max.inf_commute min_max.sup_inf_absorb min_max.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8) | 
| 24545 | 1184 | qed | 
| 23449 | 1185 | |
| 28592 | 1186 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso3"*}
 | 
| 23449 | 1187 | lemma bigo_lesso3: "f =o g +o O(h) ==> | 
| 1188 | ALL x. 0 <= k x ==> ALL x. g x <= k x ==> | |
| 1189 | f <o k =o O(h)" | |
| 1190 | apply (unfold lesso_def) | |
| 1191 | apply (rule bigo_lesseq4) | |
| 1192 | apply (erule set_plus_imp_minus) | |
| 1193 | apply (rule allI) | |
| 1194 | apply (rule le_maxI2) | |
| 1195 | apply (rule allI) | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1196 | apply (subst fun_diff_def) | 
| 23449 | 1197 | apply (erule thin_rl) | 
| 1198 | (*sledgehammer*); | |
| 1199 | apply (case_tac "0 <= f x - k x") | |
| 29667 | 1200 | apply (simp) | 
| 23449 | 1201 | apply (subst abs_of_nonneg) | 
| 1202 | apply (drule_tac x = x in spec) back | |
| 28592 | 1203 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso3_simpler"*}
 | 
| 24545 | 1204 | apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6)) | 
| 1205 | 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 | 1206 | apply (metis abs_ge_zero linorder_linear min_max.sup_absorb1 min_max.sup_commute) | 
| 23449 | 1207 | done | 
| 1208 | ||
| 1209 | lemma bigo_lesso4: "f <o g =o O(k::'a=>'b::ordered_field) ==> | |
| 1210 | g =o h +o O(k) ==> f <o h =o O(k)" | |
| 1211 | apply (unfold lesso_def) | |
| 1212 | apply (drule set_plus_imp_minus) | |
| 1213 | apply (drule bigo_abs5) back | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
26645diff
changeset | 1214 | apply (simp add: fun_diff_def) | 
| 23449 | 1215 | apply (drule bigo_useful_add) | 
| 1216 | apply assumption | |
| 1217 | apply (erule bigo_lesseq2) back | |
| 1218 | apply (rule allI) | |
| 29667 | 1219 | apply (auto simp add: func_plus fun_diff_def algebra_simps | 
| 23449 | 1220 | split: split_max abs_split) | 
| 1221 | done | |
| 1222 | ||
| 28592 | 1223 | ML_command{*AtpWrapper.problem_name := "BigO__bigo_lesso5"*}
 | 
| 23449 | 1224 | lemma bigo_lesso5: "f <o g =o O(h) ==> | 
| 1225 | EX C. ALL x. f x <= g x + C * abs(h x)" | |
| 1226 | apply (simp only: lesso_def bigo_alt_def) | |
| 1227 | apply clarsimp | |
| 24855 | 1228 | apply (metis abs_if abs_mult add_commute diff_le_eq less_not_permute) | 
| 23449 | 1229 | done | 
| 1230 | ||
| 1231 | end |