author | chaieb |
Sun, 17 Jun 2007 13:39:27 +0200 | |
changeset 23405 | 8993b3144358 |
parent 21404 | eb85850d3eb7 |
child 23477 | f4b83f03cac9 |
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 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
9 |
imports Main |
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 |
|
17161 | 20 |
instance set :: (plus) plus .. |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
21 |
instance "fun" :: (type, plus) plus .. |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
22 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
23 |
defs (overloaded) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
24 |
func_plus: "f + g == (%x. f x + g x)" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
25 |
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
|
26 |
|
17161 | 27 |
instance set :: (times) times .. |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
28 |
instance "fun" :: (type, times) times .. |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
29 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
30 |
defs (overloaded) |
19736 | 31 |
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
|
32 |
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
|
33 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
34 |
instance "fun" :: (type, minus) minus .. |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
35 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
36 |
defs (overloaded) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
37 |
func_minus: "- f == (%x. - f x)" |
19736 | 38 |
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
|
39 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
40 |
instance "fun" :: (type, zero) zero .. |
17161 | 41 |
instance set :: (zero) zero .. |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
42 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
43 |
defs (overloaded) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
44 |
func_zero: "0::(('a::type) => ('b::zero)) == %x. 0" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
45 |
set_zero: "0::('a::zero)set == {0}" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
46 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
47 |
instance "fun" :: (type, one) one .. |
17161 | 48 |
instance set :: (one) one .. |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
49 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
50 |
defs (overloaded) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
51 |
func_one: "1::(('a::type) => ('b::one)) == %x. 1" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
52 |
set_one: "1::('a::one)set == {1}" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
53 |
|
19736 | 54 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
55 |
elt_set_plus :: "'a::plus => 'a set => 'a set" (infixl "+o" 70) where |
19736 | 56 |
"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
|
57 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
58 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
59 |
elt_set_times :: "'a::times => 'a set => 'a set" (infixl "*o" 80) where |
19736 | 60 |
"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
|
61 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19380
diff
changeset
|
62 |
abbreviation (input) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
63 |
elt_set_eq :: "'a => 'a set => bool" (infix "=o" 50) where |
19380 | 64 |
"x =o A == x : A" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
65 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
66 |
instance "fun" :: (type,semigroup_add)semigroup_add |
19380 | 67 |
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
|
68 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
69 |
instance "fun" :: (type,comm_monoid_add)comm_monoid_add |
19380 | 70 |
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
|
71 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
72 |
instance "fun" :: (type,ab_group_add)ab_group_add |
19736 | 73 |
apply default |
74 |
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
|
75 |
apply (simp add: func_minus func_plus func_diff diff_minus) |
19736 | 76 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
77 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
78 |
instance "fun" :: (type,semigroup_mult)semigroup_mult |
19736 | 79 |
apply default |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
80 |
apply (auto simp add: func_times mult_assoc) |
19736 | 81 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
82 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
83 |
instance "fun" :: (type,comm_monoid_mult)comm_monoid_mult |
19736 | 84 |
apply default |
85 |
apply (auto simp add: func_one func_times mult_ac) |
|
86 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
87 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
88 |
instance "fun" :: (type,comm_ring_1)comm_ring_1 |
19736 | 89 |
apply default |
90 |
apply (auto simp add: func_plus func_times func_minus func_diff ext |
|
91 |
func_one func_zero ring_eq_simps) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
92 |
apply (drule fun_cong) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
93 |
apply simp |
19736 | 94 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
95 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
96 |
instance set :: (semigroup_add)semigroup_add |
19736 | 97 |
apply default |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
98 |
apply (unfold set_plus) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
99 |
apply (force simp add: add_assoc) |
19736 | 100 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
101 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
102 |
instance set :: (semigroup_mult)semigroup_mult |
19736 | 103 |
apply default |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
104 |
apply (unfold set_times) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
105 |
apply (force simp add: mult_assoc) |
19736 | 106 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
107 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
108 |
instance set :: (comm_monoid_add)comm_monoid_add |
19736 | 109 |
apply default |
110 |
apply (unfold set_plus) |
|
111 |
apply (force simp add: add_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
112 |
apply (unfold set_zero) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
113 |
apply force |
19736 | 114 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
115 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
116 |
instance set :: (comm_monoid_mult)comm_monoid_mult |
19736 | 117 |
apply default |
118 |
apply (unfold set_times) |
|
119 |
apply (force simp add: mult_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
120 |
apply (unfold set_one) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
121 |
apply force |
19736 | 122 |
done |
123 |
||
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
124 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
125 |
subsection {* Basic properties *} |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
126 |
|
19736 | 127 |
lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C + D" |
128 |
by (auto simp add: set_plus) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
129 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
130 |
lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C" |
19736 | 131 |
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
|
132 |
|
19736 | 133 |
lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) + |
134 |
(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
|
135 |
apply (auto simp add: elt_set_plus_def set_plus add_ac) |
19736 | 136 |
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
|
137 |
apply (auto simp add: add_ac) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
138 |
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
|
139 |
apply (auto simp add: add_ac) |
19736 | 140 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
141 |
|
19736 | 142 |
lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = |
143 |
(a + b) +o C" |
|
144 |
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
|
145 |
|
19736 | 146 |
lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C = |
147 |
a +o (B + C)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
148 |
apply (auto simp add: elt_set_plus_def set_plus) |
19736 | 149 |
apply (blast intro: add_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
150 |
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
|
151 |
apply (rule conjI) |
19736 | 152 |
apply (rule_tac x = "aa" in bexI) |
153 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
154 |
apply (rule_tac x = "ba" in bexI) |
19736 | 155 |
apply (auto simp add: add_ac) |
156 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
157 |
|
19736 | 158 |
theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) = |
159 |
a +o (C + D)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
160 |
apply (auto intro!: subsetI simp add: elt_set_plus_def set_plus add_ac) |
19736 | 161 |
apply (rule_tac x = "aa + ba" in exI) |
162 |
apply (auto simp add: add_ac) |
|
163 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
164 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
165 |
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
|
166 |
set_plus_rearrange3 set_plus_rearrange4 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
167 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
168 |
lemma set_plus_mono [intro!]: "C <= D ==> a +o C <= a +o D" |
19736 | 169 |
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
|
170 |
|
19736 | 171 |
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
|
172 |
C + E <= D + F" |
19736 | 173 |
by (auto simp add: set_plus) |
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 |
lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C + D" |
19736 | 176 |
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
|
177 |
|
19736 | 178 |
lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==> |
179 |
a +o D <= D + C" |
|
180 |
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
|
181 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
182 |
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
|
183 |
apply (subgoal_tac "a +o B <= a +o D") |
19736 | 184 |
apply (erule order_trans) |
185 |
apply (erule set_plus_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
186 |
apply (erule set_plus_mono) |
19736 | 187 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
188 |
|
19736 | 189 |
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
|
190 |
==> x : a +o D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
191 |
apply (frule set_plus_mono) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
192 |
apply auto |
19736 | 193 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
194 |
|
19736 | 195 |
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
|
196 |
x : D + F" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
197 |
apply (frule set_plus_mono2) |
19736 | 198 |
prefer 2 |
199 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
200 |
apply assumption |
19736 | 201 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
202 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
203 |
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
|
204 |
apply (frule set_plus_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
205 |
apply auto |
19736 | 206 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
207 |
|
19736 | 208 |
lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==> |
209 |
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
|
210 |
apply (frule set_plus_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
211 |
apply auto |
19736 | 212 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
213 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
214 |
lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" |
19736 | 215 |
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
|
216 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
217 |
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
|
218 |
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
|
219 |
apply (rule_tac x = 0 in bexI) |
19736 | 220 |
apply (rule_tac x = x in bexI) |
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 |
lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C" |
19736 | 225 |
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
|
226 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
227 |
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
|
228 |
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
|
229 |
apply (subgoal_tac "a = (a + - b) + b") |
19736 | 230 |
apply (rule bexI, assumption, assumption) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
231 |
apply (auto simp add: add_ac) |
19736 | 232 |
done |
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_minus_plus: "((a::'a::ab_group_add) - b : C) = (a : b +o C)" |
19736 | 235 |
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
|
236 |
assumption) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
237 |
|
19736 | 238 |
lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C * D" |
239 |
by (auto simp add: set_times) |
|
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_times_intro2 [intro!]: "b : C ==> a * b : a *o C" |
19736 | 242 |
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
|
243 |
|
19736 | 244 |
lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) * |
245 |
(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
|
246 |
apply (auto simp add: elt_set_times_def set_times) |
19736 | 247 |
apply (rule_tac x = "ba * bb" in exI) |
248 |
apply (auto simp add: mult_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
249 |
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
|
250 |
apply (auto simp add: mult_ac) |
19736 | 251 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
252 |
|
19736 | 253 |
lemma set_times_rearrange2: "(a::'a::semigroup_mult) *o (b *o C) = |
254 |
(a * b) *o C" |
|
255 |
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
|
256 |
|
19736 | 257 |
lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) * C = |
258 |
a *o (B * C)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
259 |
apply (auto simp add: elt_set_times_def set_times) |
19736 | 260 |
apply (blast intro: mult_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
261 |
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
|
262 |
apply (rule conjI) |
19736 | 263 |
apply (rule_tac x = "aa" in bexI) |
264 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
265 |
apply (rule_tac x = "ba" in bexI) |
19736 | 266 |
apply (auto simp add: mult_ac) |
267 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
268 |
|
19736 | 269 |
theorem set_times_rearrange4: "C * ((a::'a::comm_monoid_mult) *o D) = |
270 |
a *o (C * D)" |
|
271 |
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
|
272 |
mult_ac) |
19736 | 273 |
apply (rule_tac x = "aa * ba" in exI) |
274 |
apply (auto simp add: mult_ac) |
|
275 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
276 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
277 |
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
|
278 |
set_times_rearrange3 set_times_rearrange4 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
279 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
280 |
lemma set_times_mono [intro]: "C <= D ==> a *o C <= a *o D" |
19736 | 281 |
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
|
282 |
|
19736 | 283 |
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
|
284 |
C * E <= D * F" |
19736 | 285 |
by (auto simp add: set_times) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
286 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
287 |
lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C * D" |
19736 | 288 |
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
|
289 |
|
19736 | 290 |
lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==> |
291 |
a *o D <= D * C" |
|
292 |
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
|
293 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
294 |
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
|
295 |
apply (subgoal_tac "a *o B <= a *o D") |
19736 | 296 |
apply (erule order_trans) |
297 |
apply (erule set_times_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
298 |
apply (erule set_times_mono) |
19736 | 299 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
300 |
|
19736 | 301 |
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
|
302 |
==> x : a *o D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
303 |
apply (frule set_times_mono) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
304 |
apply auto |
19736 | 305 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
306 |
|
19736 | 307 |
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
|
308 |
x : D * F" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
309 |
apply (frule set_times_mono2) |
19736 | 310 |
prefer 2 |
311 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
312 |
apply assumption |
19736 | 313 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
314 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
315 |
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
|
316 |
apply (frule set_times_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
317 |
apply auto |
19736 | 318 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
319 |
|
19736 | 320 |
lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==> |
321 |
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
|
322 |
apply (frule set_times_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
323 |
apply auto |
19736 | 324 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
325 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
326 |
lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" |
19736 | 327 |
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
|
328 |
|
19736 | 329 |
lemma set_times_plus_distrib: "(a::'a::semiring) *o (b +o C)= |
330 |
(a * b) +o (a *o C)" |
|
331 |
by (auto simp add: elt_set_plus_def elt_set_times_def ring_distrib) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
332 |
|
19736 | 333 |
lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B + C) = |
334 |
(a *o B) + (a *o C)" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
335 |
apply (auto simp add: set_plus elt_set_times_def ring_distrib) |
19736 | 336 |
apply blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
337 |
apply (rule_tac x = "b + bb" in exI) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
338 |
apply (auto simp add: ring_distrib) |
19736 | 339 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
340 |
|
19736 | 341 |
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
|
342 |
a *o D + C * D" |
19736 | 343 |
apply (auto intro!: subsetI simp add: |
344 |
elt_set_plus_def elt_set_times_def set_times |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
345 |
set_plus ring_distrib) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
346 |
apply auto |
19736 | 347 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
348 |
|
19380 | 349 |
theorems set_times_plus_distribs = |
350 |
set_times_plus_distrib |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
351 |
set_times_plus_distrib2 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
352 |
|
19736 | 353 |
lemma set_neg_intro: "(a::'a::ring_1) : (- 1) *o C ==> |
354 |
- a : C" |
|
355 |
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
|
356 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
357 |
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
|
358 |
- a : (- 1) *o C" |
19736 | 359 |
by (auto simp add: elt_set_times_def) |
360 |
||
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
361 |
end |