author | huffman |
Wed, 10 Aug 2011 18:02:16 -0700 | |
changeset 44142 | 8e27e0177518 |
parent 40887 | ee8d0548c148 |
child 44890 | 22f665a2e91c |
permissions | -rw-r--r-- |
38622 | 1 |
(* Title: HOL/Library/Set_Algebras.thy |
2 |
Author: Jeremy Avigad and Kevin Donnelly; Florian Haftmann, TUM |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
3 |
*) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
4 |
|
38622 | 5 |
header {* Algebraic operations on sets *} |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
6 |
|
38622 | 7 |
theory Set_Algebras |
30738 | 8 |
imports Main |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
9 |
begin |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
10 |
|
19736 | 11 |
text {* |
38622 | 12 |
This library lifts operations like addition and muliplication to |
13 |
sets. It was designed to support asymptotic calculations. See the |
|
14 |
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
|
15 |
*} |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
16 |
|
38622 | 17 |
definition set_plus :: "'a::plus set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "\<oplus>" 65) where |
18 |
"A \<oplus> B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a + b}" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
19 |
|
38622 | 20 |
definition set_times :: "'a::times set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "\<otimes>" 70) where |
21 |
"A \<otimes> B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a * b}" |
|
25594 | 22 |
|
38622 | 23 |
definition elt_set_plus :: "'a::plus \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "+o" 70) where |
24 |
"a +o B = {c. \<exists>b\<in>B. c = a + b}" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
25 |
|
38622 | 26 |
definition elt_set_times :: "'a::times \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "*o" 80) where |
27 |
"a *o B = {c. \<exists>b\<in>B. c = a * b}" |
|
25594 | 28 |
|
38622 | 29 |
abbreviation (input) elt_set_eq :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infix "=o" 50) where |
30 |
"x =o A \<equiv> x \<in> A" |
|
25594 | 31 |
|
38622 | 32 |
interpretation set_add!: semigroup "set_plus :: 'a::semigroup_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof |
33 |
qed (force simp add: set_plus_def add.assoc) |
|
25594 | 34 |
|
38622 | 35 |
interpretation set_add!: abel_semigroup "set_plus :: 'a::ab_semigroup_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof |
36 |
qed (force simp add: set_plus_def add.commute) |
|
25594 | 37 |
|
38622 | 38 |
interpretation set_add!: monoid "set_plus :: 'a::monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" proof |
39 |
qed (simp_all add: set_plus_def) |
|
25594 | 40 |
|
38622 | 41 |
interpretation set_add!: comm_monoid "set_plus :: 'a::comm_monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" proof |
42 |
qed (simp add: set_plus_def) |
|
25594 | 43 |
|
38622 | 44 |
definition listsum_set :: "('a::monoid_add set) list \<Rightarrow> 'a set" where |
45 |
"listsum_set = monoid_add.listsum set_plus {0}" |
|
25594 | 46 |
|
38622 | 47 |
interpretation set_add!: monoid_add "set_plus :: 'a::monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" where |
48 |
"monoid_add.listsum set_plus {0::'a} = listsum_set" |
|
49 |
proof - |
|
50 |
show "class.monoid_add set_plus {0::'a}" proof |
|
51 |
qed (simp_all add: set_add.assoc) |
|
52 |
then interpret set_add!: monoid_add "set_plus :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" . |
|
53 |
show "monoid_add.listsum set_plus {0::'a} = listsum_set" |
|
54 |
by (simp only: listsum_set_def) |
|
55 |
qed |
|
25594 | 56 |
|
38622 | 57 |
definition setsum_set :: "('b \<Rightarrow> ('a::comm_monoid_add) set) \<Rightarrow> 'b set \<Rightarrow> 'a set" where |
58 |
"setsum_set f A = (if finite A then fold_image set_plus f {0} A else {0})" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
59 |
|
38622 | 60 |
interpretation set_add!: |
61 |
comm_monoid_big "set_plus :: 'a::comm_monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" setsum_set |
|
62 |
proof |
|
63 |
qed (fact setsum_set_def) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
64 |
|
38622 | 65 |
interpretation set_add!: comm_monoid_add "set_plus :: 'a::comm_monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" where |
66 |
"monoid_add.listsum set_plus {0::'a} = listsum_set" |
|
67 |
and "comm_monoid_add.setsum set_plus {0::'a} = setsum_set" |
|
68 |
proof - |
|
69 |
show "class.comm_monoid_add (set_plus :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set) {0}" proof |
|
70 |
qed (simp_all add: set_add.commute) |
|
71 |
then interpret set_add!: comm_monoid_add "set_plus :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" . |
|
72 |
show "monoid_add.listsum set_plus {0::'a} = listsum_set" |
|
73 |
by (simp only: listsum_set_def) |
|
74 |
show "comm_monoid_add.setsum set_plus {0::'a} = setsum_set" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
75 |
by (simp add: set_add.setsum_def setsum_set_def fun_eq_iff) |
38622 | 76 |
qed |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
77 |
|
38622 | 78 |
interpretation set_mult!: semigroup "set_times :: 'a::semigroup_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof |
79 |
qed (force simp add: set_times_def mult.assoc) |
|
80 |
||
81 |
interpretation set_mult!: abel_semigroup "set_times :: 'a::ab_semigroup_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof |
|
82 |
qed (force simp add: set_times_def mult.commute) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
83 |
|
38622 | 84 |
interpretation set_mult!: monoid "set_times :: 'a::monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" proof |
85 |
qed (simp_all add: set_times_def) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
86 |
|
38622 | 87 |
interpretation set_mult!: comm_monoid "set_times :: 'a::comm_monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" proof |
88 |
qed (simp add: set_times_def) |
|
89 |
||
90 |
definition power_set :: "nat \<Rightarrow> ('a::monoid_mult set) \<Rightarrow> 'a set" where |
|
91 |
"power_set n A = power.power {1} set_times A n" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
92 |
|
38622 | 93 |
interpretation set_mult!: monoid_mult "{1}" "set_times :: 'a::monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" where |
94 |
"power.power {1} set_times = (\<lambda>A n. power_set n A)" |
|
95 |
proof - |
|
96 |
show "class.monoid_mult {1} (set_times :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set)" proof |
|
97 |
qed (simp_all add: set_mult.assoc) |
|
98 |
show "power.power {1} set_times = (\<lambda>A n. power_set n A)" |
|
99 |
by (simp add: power_set_def) |
|
100 |
qed |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
101 |
|
38622 | 102 |
definition setprod_set :: "('b \<Rightarrow> ('a::comm_monoid_mult) set) \<Rightarrow> 'b set \<Rightarrow> 'a set" where |
103 |
"setprod_set f A = (if finite A then fold_image set_times f {1} A else {1})" |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
104 |
|
38622 | 105 |
interpretation set_mult!: |
106 |
comm_monoid_big "set_times :: 'a::comm_monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" setprod_set |
|
107 |
proof |
|
108 |
qed (fact setprod_set_def) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
109 |
|
38622 | 110 |
interpretation set_mult!: comm_monoid_mult "set_times :: 'a::comm_monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" where |
111 |
"power.power {1} set_times = (\<lambda>A n. power_set n A)" |
|
112 |
and "comm_monoid_mult.setprod set_times {1::'a} = setprod_set" |
|
113 |
proof - |
|
114 |
show "class.comm_monoid_mult (set_times :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set) {1}" proof |
|
115 |
qed (simp_all add: set_mult.commute) |
|
116 |
then interpret set_mult!: comm_monoid_mult "set_times :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" . |
|
117 |
show "power.power {1} set_times = (\<lambda>A n. power_set n A)" |
|
118 |
by (simp add: power_set_def) |
|
119 |
show "comm_monoid_mult.setprod set_times {1::'a} = setprod_set" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
120 |
by (simp add: set_mult.setprod_def setprod_set_def fun_eq_iff) |
38622 | 121 |
qed |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
122 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
123 |
lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C \<oplus> D" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
124 |
by (auto simp add: set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
125 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
126 |
lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C" |
19736 | 127 |
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
|
128 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
129 |
lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) \<oplus> |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
130 |
(b +o D) = (a + b) +o (C \<oplus> D)" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
131 |
apply (auto simp add: elt_set_plus_def set_plus_def add_ac) |
19736 | 132 |
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
|
133 |
apply (auto simp add: add_ac) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
134 |
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
|
135 |
apply (auto simp add: add_ac) |
19736 | 136 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
137 |
|
19736 | 138 |
lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = |
139 |
(a + b) +o C" |
|
140 |
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
|
141 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
142 |
lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) \<oplus> C = |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
143 |
a +o (B \<oplus> C)" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
144 |
apply (auto simp add: elt_set_plus_def set_plus_def) |
19736 | 145 |
apply (blast intro: add_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
146 |
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
|
147 |
apply (rule conjI) |
19736 | 148 |
apply (rule_tac x = "aa" in bexI) |
149 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
150 |
apply (rule_tac x = "ba" in bexI) |
19736 | 151 |
apply (auto simp add: add_ac) |
152 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
153 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
154 |
theorem set_plus_rearrange4: "C \<oplus> ((a::'a::comm_monoid_add) +o D) = |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
155 |
a +o (C \<oplus> D)" |
44142 | 156 |
apply (auto simp add: elt_set_plus_def set_plus_def add_ac) |
19736 | 157 |
apply (rule_tac x = "aa + ba" in exI) |
158 |
apply (auto simp add: add_ac) |
|
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 |
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
|
162 |
set_plus_rearrange3 set_plus_rearrange4 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
163 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
164 |
lemma set_plus_mono [intro!]: "C <= D ==> a +o C <= a +o D" |
19736 | 165 |
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
|
166 |
|
19736 | 167 |
lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==> |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
168 |
C \<oplus> E <= D \<oplus> F" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
169 |
by (auto simp add: set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
170 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
171 |
lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C \<oplus> D" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
172 |
by (auto simp add: elt_set_plus_def set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
173 |
|
19736 | 174 |
lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==> |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
175 |
a +o D <= D \<oplus> C" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
176 |
by (auto simp add: elt_set_plus_def set_plus_def add_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
177 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
178 |
lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C \<oplus> D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
179 |
apply (subgoal_tac "a +o B <= a +o D") |
19736 | 180 |
apply (erule order_trans) |
181 |
apply (erule set_plus_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
182 |
apply (erule set_plus_mono) |
19736 | 183 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
184 |
|
19736 | 185 |
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
|
186 |
==> x : a +o D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
187 |
apply (frule set_plus_mono) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
188 |
apply auto |
19736 | 189 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
190 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
191 |
lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C \<oplus> E ==> |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
192 |
x : D \<oplus> F" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
193 |
apply (frule set_plus_mono2) |
19736 | 194 |
prefer 2 |
195 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
196 |
apply assumption |
19736 | 197 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
198 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
199 |
lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C \<oplus> D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
200 |
apply (frule set_plus_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
201 |
apply auto |
19736 | 202 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
203 |
|
19736 | 204 |
lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==> |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
205 |
x : a +o D ==> x : D \<oplus> C" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
206 |
apply (frule set_plus_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
207 |
apply auto |
19736 | 208 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
209 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
210 |
lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" |
19736 | 211 |
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
|
212 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
213 |
lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A \<oplus> B" |
44142 | 214 |
apply (auto simp add: set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
215 |
apply (rule_tac x = 0 in bexI) |
19736 | 216 |
apply (rule_tac x = x in bexI) |
217 |
apply (auto simp add: add_ac) |
|
218 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
219 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
220 |
lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C" |
19736 | 221 |
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
|
222 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
223 |
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
|
224 |
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
|
225 |
apply (subgoal_tac "a = (a + - b) + b") |
19736 | 226 |
apply (rule bexI, assumption, assumption) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
227 |
apply (auto simp add: add_ac) |
19736 | 228 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
229 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
230 |
lemma set_minus_plus: "((a::'a::ab_group_add) - b : C) = (a : b +o C)" |
19736 | 231 |
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
|
232 |
assumption) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
233 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
234 |
lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C \<otimes> D" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
235 |
by (auto simp add: set_times_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
236 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
237 |
lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C" |
19736 | 238 |
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
|
239 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
240 |
lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) \<otimes> |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
241 |
(b *o D) = (a * b) *o (C \<otimes> D)" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
242 |
apply (auto simp add: elt_set_times_def set_times_def) |
19736 | 243 |
apply (rule_tac x = "ba * bb" in exI) |
244 |
apply (auto simp add: mult_ac) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
245 |
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
|
246 |
apply (auto simp add: mult_ac) |
19736 | 247 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
248 |
|
19736 | 249 |
lemma set_times_rearrange2: "(a::'a::semigroup_mult) *o (b *o C) = |
250 |
(a * b) *o C" |
|
251 |
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
|
252 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
253 |
lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) \<otimes> C = |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
254 |
a *o (B \<otimes> C)" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
255 |
apply (auto simp add: elt_set_times_def set_times_def) |
19736 | 256 |
apply (blast intro: mult_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
257 |
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
|
258 |
apply (rule conjI) |
19736 | 259 |
apply (rule_tac x = "aa" in bexI) |
260 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
261 |
apply (rule_tac x = "ba" in bexI) |
19736 | 262 |
apply (auto simp add: mult_ac) |
263 |
done |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
264 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
265 |
theorem set_times_rearrange4: "C \<otimes> ((a::'a::comm_monoid_mult) *o D) = |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
266 |
a *o (C \<otimes> D)" |
44142 | 267 |
apply (auto simp add: elt_set_times_def set_times_def |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
268 |
mult_ac) |
19736 | 269 |
apply (rule_tac x = "aa * ba" in exI) |
270 |
apply (auto simp add: mult_ac) |
|
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 |
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
|
274 |
set_times_rearrange3 set_times_rearrange4 |
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_times_mono [intro]: "C <= D ==> a *o C <= a *o D" |
19736 | 277 |
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
|
278 |
|
19736 | 279 |
lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==> |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
280 |
C \<otimes> E <= D \<otimes> F" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
281 |
by (auto simp add: set_times_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
282 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
283 |
lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C \<otimes> D" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
284 |
by (auto simp add: elt_set_times_def set_times_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
285 |
|
19736 | 286 |
lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==> |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
287 |
a *o D <= D \<otimes> C" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
288 |
by (auto simp add: elt_set_times_def set_times_def mult_ac) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
289 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
290 |
lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C \<otimes> D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
291 |
apply (subgoal_tac "a *o B <= a *o D") |
19736 | 292 |
apply (erule order_trans) |
293 |
apply (erule set_times_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
294 |
apply (erule set_times_mono) |
19736 | 295 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
296 |
|
19736 | 297 |
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
|
298 |
==> x : a *o D" |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
299 |
apply (frule set_times_mono) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
300 |
apply auto |
19736 | 301 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
302 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
303 |
lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C \<otimes> E ==> |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
304 |
x : D \<otimes> F" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
305 |
apply (frule set_times_mono2) |
19736 | 306 |
prefer 2 |
307 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
308 |
apply assumption |
19736 | 309 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
310 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
311 |
lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C \<otimes> D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
312 |
apply (frule set_times_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
313 |
apply auto |
19736 | 314 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
315 |
|
19736 | 316 |
lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==> |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
317 |
x : a *o D ==> x : D \<otimes> C" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
318 |
apply (frule set_times_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
319 |
apply auto |
19736 | 320 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
321 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
322 |
lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" |
19736 | 323 |
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
|
324 |
|
19736 | 325 |
lemma set_times_plus_distrib: "(a::'a::semiring) *o (b +o C)= |
326 |
(a * b) +o (a *o C)" |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
21404
diff
changeset
|
327 |
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
|
328 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
329 |
lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B \<oplus> C) = |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
330 |
(a *o B) \<oplus> (a *o C)" |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
331 |
apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) |
19736 | 332 |
apply blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
333 |
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
|
334 |
apply (auto simp add: ring_distribs) |
19736 | 335 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
336 |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
337 |
lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) \<otimes> D <= |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
338 |
a *o D \<oplus> C \<otimes> D" |
44142 | 339 |
apply (auto simp add: |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
340 |
elt_set_plus_def elt_set_times_def set_times_def |
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
341 |
set_plus_def ring_distribs) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
342 |
apply auto |
19736 | 343 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
344 |
|
19380 | 345 |
theorems set_times_plus_distribs = |
346 |
set_times_plus_distrib |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
347 |
set_times_plus_distrib2 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
348 |
|
19736 | 349 |
lemma set_neg_intro: "(a::'a::ring_1) : (- 1) *o C ==> |
350 |
- a : C" |
|
351 |
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
|
352 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
353 |
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
|
354 |
- a : (- 1) *o C" |
19736 | 355 |
by (auto simp add: elt_set_times_def) |
356 |
||
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
357 |
lemma set_plus_image: |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
358 |
fixes S T :: "'n::semigroup_add set" shows "S \<oplus> T = (\<lambda>(x, y). x + y) ` (S \<times> T)" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
359 |
unfolding set_plus_def by (fastsimp simp: image_iff) |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
360 |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
361 |
lemma set_setsum_alt: |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
362 |
assumes fin: "finite I" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
363 |
shows "setsum_set S I = {setsum s I |s. \<forall>i\<in>I. s i \<in> S i}" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
364 |
(is "_ = ?setsum I") |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
365 |
using fin proof induct |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
366 |
case (insert x F) |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
367 |
have "setsum_set S (insert x F) = S x \<oplus> ?setsum F" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
368 |
using insert.hyps by auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
369 |
also have "...= {s x + setsum s F |s. \<forall> i\<in>insert x F. s i \<in> S i}" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
370 |
unfolding set_plus_def |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
371 |
proof safe |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
372 |
fix y s assume "y \<in> S x" "\<forall>i\<in>F. s i \<in> S i" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
373 |
then show "\<exists>s'. y + setsum s F = s' x + setsum s' F \<and> (\<forall>i\<in>insert x F. s' i \<in> S i)" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
374 |
using insert.hyps |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
375 |
by (intro exI[of _ "\<lambda>i. if i \<in> F then s i else y"]) (auto simp add: set_plus_def) |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
376 |
qed auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
377 |
finally show ?case |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
378 |
using insert.hyps by auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
379 |
qed auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
380 |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
381 |
lemma setsum_set_cond_linear: |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
382 |
fixes f :: "('a::comm_monoid_add) set \<Rightarrow> ('b::comm_monoid_add) set" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
383 |
assumes [intro!]: "\<And>A B. P A \<Longrightarrow> P B \<Longrightarrow> P (A \<oplus> B)" "P {0}" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
384 |
and f: "\<And>A B. P A \<Longrightarrow> P B \<Longrightarrow> f (A \<oplus> B) = f A \<oplus> f B" "f {0} = {0}" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
385 |
assumes all: "\<And>i. i \<in> I \<Longrightarrow> P (S i)" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
386 |
shows "f (setsum_set S I) = setsum_set (f \<circ> S) I" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
387 |
proof cases |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
388 |
assume "finite I" from this all show ?thesis |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
389 |
proof induct |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
390 |
case (insert x F) |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
391 |
from `finite F` `\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)` have "P (setsum_set S F)" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
392 |
by induct auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
393 |
with insert show ?case |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
394 |
by (simp, subst f) auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
395 |
qed (auto intro!: f) |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
396 |
qed (auto intro!: f) |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
397 |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
398 |
lemma setsum_set_linear: |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
399 |
fixes f :: "('a::comm_monoid_add) set => ('b::comm_monoid_add) set" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
400 |
assumes "\<And>A B. f(A) \<oplus> f(B) = f(A \<oplus> B)" "f {0} = {0}" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
401 |
shows "f (setsum_set S I) = setsum_set (f \<circ> S) I" |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
402 |
using setsum_set_cond_linear[of "\<lambda>x. True" f I S] assms by auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
403 |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
404 |
end |