author | haftmann |
Mon, 10 Dec 2007 11:24:09 +0100 | |
changeset 25594 | 43c718438f9f |
parent 23477 | f4b83f03cac9 |
child 25691 | 8f8d83af100a |
permissions | -rwxr-xr-x |
16932 | 1 |
(* Title: HOL/Library/SetsAndFunctions.thy |
19736 | 2 |
ID: $Id$ |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
3 |
Author: Jeremy Avigad and Kevin Donnelly |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
4 |
*) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
5 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
6 |
header {* Operations on sets and functions *} |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
7 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
8 |
theory SetsAndFunctions |
25594 | 9 |
imports PreList |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
10 |
begin |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
11 |
|
19736 | 12 |
text {* |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
13 |
This library lifts operations like addition and muliplication to sets and |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
14 |
functions of appropriate types. It was designed to support asymptotic |
17161 | 15 |
calculations. See the comments at the top of theory @{text BigO}. |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
16 |
*} |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
17 |
|
19736 | 18 |
subsection {* Basic definitions *} |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
19 |
|
25594 | 20 |
instantiation set :: (plus) plus |
21 |
begin |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
22 |
|
25594 | 23 |
definition |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
24 |
set_plus: "A + B == {c. EX a:A. EX b:B. c = a + b}" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
25 |
|
25594 | 26 |
instance .. |
27 |
||
28 |
end |
|
29 |
||
30 |
instantiation "fun" :: (type, plus) plus |
|
31 |
begin |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
32 |
|
25594 | 33 |
definition |
34 |
func_plus: "f + g == (%x. f x + g x)" |
|
35 |
||
36 |
instance .. |
|
37 |
||
38 |
end |
|
39 |
||
40 |
instantiation set :: (times) times |
|
41 |
begin |
|
42 |
||
43 |
definition |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
44 |
set_times:"A * B == {c. EX a:A. EX b:B. c = a * b}" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
45 |
|
25594 | 46 |
instance .. |
47 |
||
48 |
end |
|
49 |
||
50 |
instantiation "fun" :: (type, times) times |
|
51 |
begin |
|
52 |
||
53 |
definition |
|
54 |
func_times: "f * g == (%x. f x * g x)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
55 |
|
25594 | 56 |
instance .. |
57 |
||
58 |
end |
|
59 |
||
60 |
instantiation "fun" :: (type, minus) minus |
|
61 |
begin |
|
62 |
||
63 |
definition |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
64 |
func_minus: "- f == (%x. - f x)" |
25594 | 65 |
|
66 |
definition |
|
19736 | 67 |
func_diff: "f - g == %x. f x - g x" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
68 |
|
25594 | 69 |
instance .. |
70 |
||
71 |
end |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
72 |
|
25594 | 73 |
instantiation set :: (zero) zero |
74 |
begin |
|
75 |
||
76 |
definition |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
77 |
set_zero: "0::('a::zero)set == {0}" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
78 |
|
25594 | 79 |
instance .. |
80 |
||
81 |
end |
|
82 |
||
83 |
instantiation "fun" :: (type, zero) zero |
|
84 |
begin |
|
85 |
||
86 |
definition |
|
87 |
func_zero: "0::(('a::type) => ('b::zero)) == %x. 0" |
|
88 |
||
89 |
instance .. |
|
90 |
||
91 |
end |
|
92 |
||
93 |
instantiation set :: (one) one |
|
94 |
begin |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
95 |
|
25594 | 96 |
definition |
97 |
set_one: "1::('a::one)set == {1}" |
|
98 |
||
99 |
instance .. |
|
100 |
||
101 |
end |
|
102 |
||
103 |
instantiation "fun" :: (type, one) one |
|
104 |
begin |
|
105 |
||
106 |
definition |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
107 |
func_one: "1::(('a::type) => ('b::one)) == %x. 1" |
25594 | 108 |
|
109 |
instance .. |
|
110 |
||
111 |
end |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
112 |
|
19736 | 113 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
114 |
elt_set_plus :: "'a::plus => 'a set => 'a set" (infixl "+o" 70) where |
19736 | 115 |
"a +o B = {c. EX b:B. c = a + b}" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
116 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
117 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
118 |
elt_set_times :: "'a::times => 'a set => 'a set" (infixl "*o" 80) where |
19736 | 119 |
"a *o B = {c. EX b:B. c = a * b}" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
120 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19380
diff
changeset
|
121 |
abbreviation (input) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
122 |
elt_set_eq :: "'a => 'a set => bool" (infix "=o" 50) where |
19380 | 123 |
"x =o A == x : A" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
124 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
125 |
instance "fun" :: (type,semigroup_add)semigroup_add |
19380 | 126 |
by default (auto simp add: func_plus add_assoc) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
127 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
128 |
instance "fun" :: (type,comm_monoid_add)comm_monoid_add |
19380 | 129 |
by default (auto simp add: func_zero func_plus add_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
130 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
131 |
instance "fun" :: (type,ab_group_add)ab_group_add |
19736 | 132 |
apply default |
133 |
apply (simp add: func_minus func_plus func_zero) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
134 |
apply (simp add: func_minus func_plus func_diff diff_minus) |
19736 | 135 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
136 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
137 |
instance "fun" :: (type,semigroup_mult)semigroup_mult |
19736 | 138 |
apply default |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
139 |
apply (auto simp add: func_times mult_assoc) |
19736 | 140 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
141 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
142 |
instance "fun" :: (type,comm_monoid_mult)comm_monoid_mult |
19736 | 143 |
apply default |
144 |
apply (auto simp add: func_one func_times mult_ac) |
|
145 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
146 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
147 |
instance "fun" :: (type,comm_ring_1)comm_ring_1 |
19736 | 148 |
apply default |
149 |
apply (auto simp add: func_plus func_times func_minus func_diff ext |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
21404
diff
changeset
|
150 |
func_one func_zero ring_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
151 |
apply (drule fun_cong) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
152 |
apply simp |
19736 | 153 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
154 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
155 |
instance set :: (semigroup_add)semigroup_add |
19736 | 156 |
apply default |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
157 |
apply (unfold set_plus) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
158 |
apply (force simp add: add_assoc) |
19736 | 159 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
160 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
161 |
instance set :: (semigroup_mult)semigroup_mult |
19736 | 162 |
apply default |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
163 |
apply (unfold set_times) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
164 |
apply (force simp add: mult_assoc) |
19736 | 165 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
166 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
167 |
instance set :: (comm_monoid_add)comm_monoid_add |
19736 | 168 |
apply default |
169 |
apply (unfold set_plus) |
|
170 |
apply (force simp add: add_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
171 |
apply (unfold set_zero) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
172 |
apply force |
19736 | 173 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
174 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
175 |
instance set :: (comm_monoid_mult)comm_monoid_mult |
19736 | 176 |
apply default |
177 |
apply (unfold set_times) |
|
178 |
apply (force simp add: mult_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
179 |
apply (unfold set_one) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
180 |
apply force |
19736 | 181 |
done |
182 |
||
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
183 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
184 |
subsection {* Basic properties *} |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
185 |
|
19736 | 186 |
lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C + D" |
187 |
by (auto simp add: set_plus) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
188 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
189 |
lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C" |
19736 | 190 |
by (auto simp add: elt_set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
191 |
|
19736 | 192 |
lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) + |
193 |
(b +o D) = (a + b) +o (C + D)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
194 |
apply (auto simp add: elt_set_plus_def set_plus add_ac) |
19736 | 195 |
apply (rule_tac x = "ba + bb" in exI) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
196 |
apply (auto simp add: add_ac) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
197 |
apply (rule_tac x = "aa + a" in exI) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
198 |
apply (auto simp add: add_ac) |
19736 | 199 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
200 |
|
19736 | 201 |
lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = |
202 |
(a + b) +o C" |
|
203 |
by (auto simp add: elt_set_plus_def add_assoc) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
204 |
|
19736 | 205 |
lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C = |
206 |
a +o (B + C)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
207 |
apply (auto simp add: elt_set_plus_def set_plus) |
19736 | 208 |
apply (blast intro: add_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
209 |
apply (rule_tac x = "a + aa" in exI) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
210 |
apply (rule conjI) |
19736 | 211 |
apply (rule_tac x = "aa" in bexI) |
212 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
213 |
apply (rule_tac x = "ba" in bexI) |
19736 | 214 |
apply (auto simp add: add_ac) |
215 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
216 |
|
19736 | 217 |
theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) = |
218 |
a +o (C + D)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
219 |
apply (auto intro!: subsetI simp add: elt_set_plus_def set_plus add_ac) |
19736 | 220 |
apply (rule_tac x = "aa + ba" in exI) |
221 |
apply (auto simp add: add_ac) |
|
222 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
223 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
224 |
theorems set_plus_rearranges = set_plus_rearrange set_plus_rearrange2 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
225 |
set_plus_rearrange3 set_plus_rearrange4 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
226 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
227 |
lemma set_plus_mono [intro!]: "C <= D ==> a +o C <= a +o D" |
19736 | 228 |
by (auto simp add: elt_set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
229 |
|
19736 | 230 |
lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
231 |
C + E <= D + F" |
19736 | 232 |
by (auto simp add: set_plus) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
233 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
234 |
lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C + D" |
19736 | 235 |
by (auto simp add: elt_set_plus_def set_plus) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
236 |
|
19736 | 237 |
lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==> |
238 |
a +o D <= D + C" |
|
239 |
by (auto simp add: elt_set_plus_def set_plus add_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
240 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
241 |
lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C + D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
242 |
apply (subgoal_tac "a +o B <= a +o D") |
19736 | 243 |
apply (erule order_trans) |
244 |
apply (erule set_plus_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
245 |
apply (erule set_plus_mono) |
19736 | 246 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
247 |
|
19736 | 248 |
lemma set_plus_mono_b: "C <= D ==> x : a +o C |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
249 |
==> x : a +o D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
250 |
apply (frule set_plus_mono) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
251 |
apply auto |
19736 | 252 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
253 |
|
19736 | 254 |
lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C + E ==> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
255 |
x : D + F" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
256 |
apply (frule set_plus_mono2) |
19736 | 257 |
prefer 2 |
258 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
259 |
apply assumption |
19736 | 260 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
261 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
262 |
lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C + D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
263 |
apply (frule set_plus_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
264 |
apply auto |
19736 | 265 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
266 |
|
19736 | 267 |
lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==> |
268 |
x : a +o D ==> x : D + C" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
269 |
apply (frule set_plus_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
270 |
apply auto |
19736 | 271 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
272 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
273 |
lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" |
19736 | 274 |
by (auto simp add: elt_set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
275 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
276 |
lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A + B" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
277 |
apply (auto intro!: subsetI simp add: set_plus) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
278 |
apply (rule_tac x = 0 in bexI) |
19736 | 279 |
apply (rule_tac x = x in bexI) |
280 |
apply (auto simp add: add_ac) |
|
281 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
282 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
283 |
lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C" |
19736 | 284 |
by (auto simp add: elt_set_plus_def add_ac diff_minus) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
285 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
286 |
lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C ==> a : b +o C" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
287 |
apply (auto simp add: elt_set_plus_def add_ac diff_minus) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
288 |
apply (subgoal_tac "a = (a + - b) + b") |
19736 | 289 |
apply (rule bexI, assumption, assumption) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
290 |
apply (auto simp add: add_ac) |
19736 | 291 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
292 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
293 |
lemma set_minus_plus: "((a::'a::ab_group_add) - b : C) = (a : b +o C)" |
19736 | 294 |
by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus, |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
295 |
assumption) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
296 |
|
19736 | 297 |
lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C * D" |
298 |
by (auto simp add: set_times) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
299 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
300 |
lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C" |
19736 | 301 |
by (auto simp add: elt_set_times_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
302 |
|
19736 | 303 |
lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) * |
304 |
(b *o D) = (a * b) *o (C * D)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
305 |
apply (auto simp add: elt_set_times_def set_times) |
19736 | 306 |
apply (rule_tac x = "ba * bb" in exI) |
307 |
apply (auto simp add: mult_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
308 |
apply (rule_tac x = "aa * a" in exI) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
309 |
apply (auto simp add: mult_ac) |
19736 | 310 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
311 |
|
19736 | 312 |
lemma set_times_rearrange2: "(a::'a::semigroup_mult) *o (b *o C) = |
313 |
(a * b) *o C" |
|
314 |
by (auto simp add: elt_set_times_def mult_assoc) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
315 |
|
19736 | 316 |
lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) * C = |
317 |
a *o (B * C)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
318 |
apply (auto simp add: elt_set_times_def set_times) |
19736 | 319 |
apply (blast intro: mult_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
320 |
apply (rule_tac x = "a * aa" in exI) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
321 |
apply (rule conjI) |
19736 | 322 |
apply (rule_tac x = "aa" in bexI) |
323 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
324 |
apply (rule_tac x = "ba" in bexI) |
19736 | 325 |
apply (auto simp add: mult_ac) |
326 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
327 |
|
19736 | 328 |
theorem set_times_rearrange4: "C * ((a::'a::comm_monoid_mult) *o D) = |
329 |
a *o (C * D)" |
|
330 |
apply (auto intro!: subsetI simp add: elt_set_times_def set_times |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
331 |
mult_ac) |
19736 | 332 |
apply (rule_tac x = "aa * ba" in exI) |
333 |
apply (auto simp add: mult_ac) |
|
334 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
335 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
336 |
theorems set_times_rearranges = set_times_rearrange set_times_rearrange2 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
337 |
set_times_rearrange3 set_times_rearrange4 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
338 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
339 |
lemma set_times_mono [intro]: "C <= D ==> a *o C <= a *o D" |
19736 | 340 |
by (auto simp add: elt_set_times_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
341 |
|
19736 | 342 |
lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
343 |
C * E <= D * F" |
19736 | 344 |
by (auto simp add: set_times) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
345 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
346 |
lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C * D" |
19736 | 347 |
by (auto simp add: elt_set_times_def set_times) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
348 |
|
19736 | 349 |
lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==> |
350 |
a *o D <= D * C" |
|
351 |
by (auto simp add: elt_set_times_def set_times mult_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
352 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
353 |
lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C * D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
354 |
apply (subgoal_tac "a *o B <= a *o D") |
19736 | 355 |
apply (erule order_trans) |
356 |
apply (erule set_times_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
357 |
apply (erule set_times_mono) |
19736 | 358 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
359 |
|
19736 | 360 |
lemma set_times_mono_b: "C <= D ==> x : a *o C |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
361 |
==> x : a *o D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
362 |
apply (frule set_times_mono) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
363 |
apply auto |
19736 | 364 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
365 |
|
19736 | 366 |
lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C * E ==> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
367 |
x : D * F" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
368 |
apply (frule set_times_mono2) |
19736 | 369 |
prefer 2 |
370 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
371 |
apply assumption |
19736 | 372 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
373 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
374 |
lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C * D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
375 |
apply (frule set_times_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
376 |
apply auto |
19736 | 377 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
378 |
|
19736 | 379 |
lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==> |
380 |
x : a *o D ==> x : D * C" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
381 |
apply (frule set_times_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
382 |
apply auto |
19736 | 383 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
384 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
385 |
lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" |
19736 | 386 |
by (auto simp add: elt_set_times_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
387 |
|
19736 | 388 |
lemma set_times_plus_distrib: "(a::'a::semiring) *o (b +o C)= |
389 |
(a * b) +o (a *o C)" |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
21404
diff
changeset
|
390 |
by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
391 |
|
19736 | 392 |
lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B + C) = |
393 |
(a *o B) + (a *o C)" |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
21404
diff
changeset
|
394 |
apply (auto simp add: set_plus elt_set_times_def ring_distribs) |
19736 | 395 |
apply blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
396 |
apply (rule_tac x = "b + bb" in exI) |
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
21404
diff
changeset
|
397 |
apply (auto simp add: ring_distribs) |
19736 | 398 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
399 |
|
19736 | 400 |
lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) * D <= |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
401 |
a *o D + C * D" |
19736 | 402 |
apply (auto intro!: subsetI simp add: |
403 |
elt_set_plus_def elt_set_times_def set_times |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
21404
diff
changeset
|
404 |
set_plus ring_distribs) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
405 |
apply auto |
19736 | 406 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
407 |
|
19380 | 408 |
theorems set_times_plus_distribs = |
409 |
set_times_plus_distrib |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
410 |
set_times_plus_distrib2 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
411 |
|
19736 | 412 |
lemma set_neg_intro: "(a::'a::ring_1) : (- 1) *o C ==> |
413 |
- a : C" |
|
414 |
by (auto simp add: elt_set_times_def) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
415 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
416 |
lemma set_neg_intro2: "(a::'a::ring_1) : C ==> |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
417 |
- a : (- 1) *o C" |
19736 | 418 |
by (auto simp add: elt_set_times_def) |
419 |
||
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
420 |
end |