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