author | wenzelm |
Tue, 06 Oct 2015 15:14:28 +0200 | |
changeset 61337 | 4645502c3c64 |
parent 60679 | ade12ef2773c |
child 61585 | a9599d3d7610 |
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 |
|
60500 | 5 |
section \<open>Algebraic operations on sets\<close> |
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 |
|
60500 | 11 |
text \<open> |
56899 | 12 |
This library lifts operations like addition and multiplication to |
38622 | 13 |
sets. It was designed to support asymptotic calculations. See the |
14 |
comments at the top of theory @{text BigO}. |
|
60500 | 15 |
\<close> |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
16 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
17 |
instantiation set :: (plus) plus |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
18 |
begin |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
19 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
20 |
definition plus_set :: "'a::plus set \<Rightarrow> 'a set \<Rightarrow> 'a set" where |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
21 |
set_plus_def: "A + B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a + b}" |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
22 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
23 |
instance .. |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
24 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
25 |
end |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
26 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
27 |
instantiation set :: (times) times |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
28 |
begin |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
29 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
30 |
definition times_set :: "'a::times set \<Rightarrow> 'a set \<Rightarrow> 'a set" where |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
31 |
set_times_def: "A * B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a * b}" |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
32 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
33 |
instance .. |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
34 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
35 |
end |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
36 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
37 |
instantiation set :: (zero) zero |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
38 |
begin |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
39 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
40 |
definition |
56899 | 41 |
set_zero[simp]: "(0::'a::zero set) = {0}" |
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
42 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
43 |
instance .. |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
44 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
45 |
end |
56899 | 46 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
47 |
instantiation set :: (one) one |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
48 |
begin |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
49 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
50 |
definition |
56899 | 51 |
set_one[simp]: "(1::'a::one set) = {1}" |
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
52 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
53 |
instance .. |
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
54 |
|
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
55 |
end |
25594 | 56 |
|
38622 | 57 |
definition elt_set_plus :: "'a::plus \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "+o" 70) where |
58 |
"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
|
59 |
|
38622 | 60 |
definition elt_set_times :: "'a::times \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "*o" 80) where |
61 |
"a *o B = {c. \<exists>b\<in>B. c = a * b}" |
|
25594 | 62 |
|
38622 | 63 |
abbreviation (input) elt_set_eq :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infix "=o" 50) where |
64 |
"x =o A \<equiv> x \<in> A" |
|
25594 | 65 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
66 |
instance set :: (semigroup_add) semigroup_add |
60679 | 67 |
by standard (force simp add: set_plus_def add.assoc) |
25594 | 68 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
69 |
instance set :: (ab_semigroup_add) ab_semigroup_add |
60679 | 70 |
by standard (force simp add: set_plus_def add.commute) |
25594 | 71 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
72 |
instance set :: (monoid_add) monoid_add |
60679 | 73 |
by standard (simp_all add: set_plus_def) |
25594 | 74 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
75 |
instance set :: (comm_monoid_add) comm_monoid_add |
60679 | 76 |
by standard (simp_all add: set_plus_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
77 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
78 |
instance set :: (semigroup_mult) semigroup_mult |
60679 | 79 |
by standard (force simp add: set_times_def mult.assoc) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
80 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
81 |
instance set :: (ab_semigroup_mult) ab_semigroup_mult |
60679 | 82 |
by standard (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 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
84 |
instance set :: (monoid_mult) monoid_mult |
60679 | 85 |
by standard (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 |
|
47443
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
krauss
parents:
44890
diff
changeset
|
87 |
instance set :: (comm_monoid_mult) comm_monoid_mult |
60679 | 88 |
by standard (simp_all add: set_times_def) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
89 |
|
56899 | 90 |
lemma set_plus_intro [intro]: "a \<in> C \<Longrightarrow> b \<in> D \<Longrightarrow> a + b \<in> C + D" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
91 |
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
|
92 |
|
53596 | 93 |
lemma set_plus_elim: |
94 |
assumes "x \<in> A + B" |
|
95 |
obtains a b where "x = a + b" and "a \<in> A" and "b \<in> B" |
|
96 |
using assms unfolding set_plus_def by fast |
|
97 |
||
56899 | 98 |
lemma set_plus_intro2 [intro]: "b \<in> C \<Longrightarrow> a + b \<in> a +o C" |
19736 | 99 |
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
|
100 |
|
56899 | 101 |
lemma set_plus_rearrange: |
102 |
"((a::'a::comm_monoid_add) +o C) + (b +o D) = (a + b) +o (C + D)" |
|
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
103 |
apply (auto simp add: elt_set_plus_def set_plus_def ac_simps) |
19736 | 104 |
apply (rule_tac x = "ba + bb" in exI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
105 |
apply (auto simp add: ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
106 |
apply (rule_tac x = "aa + a" in exI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
107 |
apply (auto simp add: ac_simps) |
19736 | 108 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
109 |
|
56899 | 110 |
lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = (a + b) +o C" |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56899
diff
changeset
|
111 |
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
|
112 |
|
56899 | 113 |
lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C = a +o (B + C)" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
114 |
apply (auto simp add: elt_set_plus_def set_plus_def) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
115 |
apply (blast intro: ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
116 |
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
|
117 |
apply (rule conjI) |
19736 | 118 |
apply (rule_tac x = "aa" in bexI) |
119 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
120 |
apply (rule_tac x = "ba" in bexI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
121 |
apply (auto simp add: ac_simps) |
19736 | 122 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
123 |
|
56899 | 124 |
theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) = a +o (C + D)" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
125 |
apply (auto simp add: elt_set_plus_def set_plus_def ac_simps) |
19736 | 126 |
apply (rule_tac x = "aa + ba" in exI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
127 |
apply (auto simp add: ac_simps) |
19736 | 128 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
129 |
|
61337 | 130 |
lemmas set_plus_rearranges = set_plus_rearrange set_plus_rearrange2 |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
131 |
set_plus_rearrange3 set_plus_rearrange4 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
132 |
|
56899 | 133 |
lemma set_plus_mono [intro!]: "C \<subseteq> D \<Longrightarrow> a +o C \<subseteq> a +o D" |
19736 | 134 |
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
|
135 |
|
56899 | 136 |
lemma set_plus_mono2 [intro]: "(C::'a::plus set) \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> C + E \<subseteq> D + F" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
137 |
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
|
138 |
|
56899 | 139 |
lemma set_plus_mono3 [intro]: "a \<in> C \<Longrightarrow> a +o D \<subseteq> C + D" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
140 |
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
|
141 |
|
56899 | 142 |
lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) \<in> C \<Longrightarrow> a +o D \<subseteq> D + C" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
143 |
by (auto simp add: elt_set_plus_def set_plus_def ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
144 |
|
56899 | 145 |
lemma set_plus_mono5: "a \<in> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> a +o B \<subseteq> C + D" |
146 |
apply (subgoal_tac "a +o B \<subseteq> a +o D") |
|
19736 | 147 |
apply (erule order_trans) |
148 |
apply (erule set_plus_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
149 |
apply (erule set_plus_mono) |
19736 | 150 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
151 |
|
56899 | 152 |
lemma set_plus_mono_b: "C \<subseteq> D \<Longrightarrow> x \<in> a +o C \<Longrightarrow> x \<in> a +o D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
153 |
apply (frule set_plus_mono) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
154 |
apply auto |
19736 | 155 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
156 |
|
56899 | 157 |
lemma set_plus_mono2_b: "C \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> x \<in> C + E \<Longrightarrow> x \<in> D + F" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
158 |
apply (frule set_plus_mono2) |
19736 | 159 |
prefer 2 |
160 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
161 |
apply assumption |
19736 | 162 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
163 |
|
56899 | 164 |
lemma set_plus_mono3_b: "a \<in> C \<Longrightarrow> x \<in> a +o D \<Longrightarrow> x \<in> C + D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
165 |
apply (frule set_plus_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
166 |
apply auto |
19736 | 167 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
168 |
|
56899 | 169 |
lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C \<Longrightarrow> x \<in> a +o D \<Longrightarrow> x \<in> D + C" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
170 |
apply (frule set_plus_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
171 |
apply auto |
19736 | 172 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
173 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
174 |
lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" |
19736 | 175 |
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
|
176 |
|
56899 | 177 |
lemma set_zero_plus2: "(0::'a::comm_monoid_add) \<in> A \<Longrightarrow> B \<subseteq> A + B" |
44142 | 178 |
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
|
179 |
apply (rule_tac x = 0 in bexI) |
19736 | 180 |
apply (rule_tac x = x in bexI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
181 |
apply (auto simp add: ac_simps) |
19736 | 182 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
183 |
|
56899 | 184 |
lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C \<Longrightarrow> (a - b) \<in> C" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
185 |
by (auto simp add: elt_set_plus_def ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
186 |
|
56899 | 187 |
lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C \<Longrightarrow> a \<in> b +o C" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
188 |
apply (auto simp add: elt_set_plus_def ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
189 |
apply (subgoal_tac "a = (a + - b) + b") |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53596
diff
changeset
|
190 |
apply (rule bexI, assumption) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
191 |
apply (auto simp add: ac_simps) |
19736 | 192 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
193 |
|
56899 | 194 |
lemma set_minus_plus: "(a::'a::ab_group_add) - b \<in> C \<longleftrightarrow> a \<in> b +o C" |
195 |
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
|
196 |
|
56899 | 197 |
lemma set_times_intro [intro]: "a \<in> C \<Longrightarrow> b \<in> D \<Longrightarrow> a * b \<in> C * D" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
198 |
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
|
199 |
|
53596 | 200 |
lemma set_times_elim: |
201 |
assumes "x \<in> A * B" |
|
202 |
obtains a b where "x = a * b" and "a \<in> A" and "b \<in> B" |
|
203 |
using assms unfolding set_times_def by fast |
|
204 |
||
56899 | 205 |
lemma set_times_intro2 [intro!]: "b \<in> C \<Longrightarrow> a * b \<in> a *o C" |
19736 | 206 |
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
|
207 |
|
56899 | 208 |
lemma set_times_rearrange: |
209 |
"((a::'a::comm_monoid_mult) *o C) * (b *o D) = (a * b) *o (C * D)" |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
210 |
apply (auto simp add: elt_set_times_def set_times_def) |
19736 | 211 |
apply (rule_tac x = "ba * bb" in exI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
212 |
apply (auto simp add: ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
213 |
apply (rule_tac x = "aa * a" in exI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
214 |
apply (auto simp add: ac_simps) |
19736 | 215 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
216 |
|
56899 | 217 |
lemma set_times_rearrange2: |
218 |
"(a::'a::semigroup_mult) *o (b *o C) = (a * b) *o C" |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56899
diff
changeset
|
219 |
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
|
220 |
|
56899 | 221 |
lemma set_times_rearrange3: |
222 |
"((a::'a::semigroup_mult) *o B) * C = a *o (B * C)" |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
223 |
apply (auto simp add: elt_set_times_def set_times_def) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
224 |
apply (blast intro: ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
225 |
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
|
226 |
apply (rule conjI) |
19736 | 227 |
apply (rule_tac x = "aa" in bexI) |
228 |
apply auto |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
229 |
apply (rule_tac x = "ba" in bexI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
230 |
apply (auto simp add: ac_simps) |
19736 | 231 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
232 |
|
56899 | 233 |
theorem set_times_rearrange4: |
234 |
"C * ((a::'a::comm_monoid_mult) *o D) = a *o (C * D)" |
|
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
235 |
apply (auto simp add: elt_set_times_def set_times_def ac_simps) |
19736 | 236 |
apply (rule_tac x = "aa * ba" in exI) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
237 |
apply (auto simp add: ac_simps) |
19736 | 238 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
239 |
|
61337 | 240 |
lemmas set_times_rearranges = set_times_rearrange set_times_rearrange2 |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
241 |
set_times_rearrange3 set_times_rearrange4 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
242 |
|
56899 | 243 |
lemma set_times_mono [intro]: "C \<subseteq> D \<Longrightarrow> a *o C \<subseteq> a *o D" |
19736 | 244 |
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
|
245 |
|
56899 | 246 |
lemma set_times_mono2 [intro]: "(C::'a::times set) \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> C * E \<subseteq> D * F" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
247 |
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
|
248 |
|
56899 | 249 |
lemma set_times_mono3 [intro]: "a \<in> C \<Longrightarrow> a *o D \<subseteq> C * D" |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
250 |
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
|
251 |
|
56899 | 252 |
lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C \<Longrightarrow> a *o D \<subseteq> D * C" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
253 |
by (auto simp add: elt_set_times_def set_times_def ac_simps) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
254 |
|
56899 | 255 |
lemma set_times_mono5: "a \<in> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> a *o B \<subseteq> C * D" |
256 |
apply (subgoal_tac "a *o B \<subseteq> a *o D") |
|
19736 | 257 |
apply (erule order_trans) |
258 |
apply (erule set_times_mono3) |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
259 |
apply (erule set_times_mono) |
19736 | 260 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
261 |
|
56899 | 262 |
lemma set_times_mono_b: "C \<subseteq> D \<Longrightarrow> x \<in> a *o C \<Longrightarrow> x \<in> a *o D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
263 |
apply (frule set_times_mono) |
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 |
|
56899 | 267 |
lemma set_times_mono2_b: "C \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> x \<in> C * E \<Longrightarrow> x \<in> D * F" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
268 |
apply (frule set_times_mono2) |
19736 | 269 |
prefer 2 |
270 |
apply force |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
271 |
apply assumption |
19736 | 272 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
273 |
|
56899 | 274 |
lemma set_times_mono3_b: "a \<in> C \<Longrightarrow> x \<in> a *o D \<Longrightarrow> x \<in> C * D" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
275 |
apply (frule set_times_mono3) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
276 |
apply auto |
19736 | 277 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
278 |
|
56899 | 279 |
lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) \<in> C \<Longrightarrow> x \<in> a *o D \<Longrightarrow> x \<in> D * C" |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
280 |
apply (frule set_times_mono4) |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
281 |
apply auto |
19736 | 282 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
283 |
|
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
284 |
lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" |
19736 | 285 |
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
|
286 |
|
56899 | 287 |
lemma set_times_plus_distrib: |
288 |
"(a::'a::semiring) *o (b +o C) = (a * b) +o (a *o C)" |
|
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
21404
diff
changeset
|
289 |
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
|
290 |
|
56899 | 291 |
lemma set_times_plus_distrib2: |
292 |
"(a::'a::semiring) *o (B + C) = (a *o B) + (a *o C)" |
|
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
293 |
apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) |
19736 | 294 |
apply blast |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
295 |
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
|
296 |
apply (auto simp add: ring_distribs) |
19736 | 297 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
298 |
|
56899 | 299 |
lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) * D \<subseteq> a *o D + C * D" |
44142 | 300 |
apply (auto simp add: |
26814
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
berghofe
parents:
25764
diff
changeset
|
301 |
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
|
302 |
set_plus_def ring_distribs) |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
303 |
apply auto |
19736 | 304 |
done |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
305 |
|
61337 | 306 |
lemmas set_times_plus_distribs = |
19380 | 307 |
set_times_plus_distrib |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
308 |
set_times_plus_distrib2 |
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
309 |
|
56899 | 310 |
lemma set_neg_intro: "(a::'a::ring_1) \<in> (- 1) *o C \<Longrightarrow> - a \<in> C" |
19736 | 311 |
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
|
312 |
|
56899 | 313 |
lemma set_neg_intro2: "(a::'a::ring_1) \<in> C \<Longrightarrow> - a \<in> (- 1) *o C" |
19736 | 314 |
by (auto simp add: elt_set_times_def) |
315 |
||
53596 | 316 |
lemma set_plus_image: "S + T = (\<lambda>(x, y). x + y) ` (S \<times> T)" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44142
diff
changeset
|
317 |
unfolding set_plus_def by (fastforce simp: image_iff) |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
318 |
|
53596 | 319 |
lemma set_times_image: "S * T = (\<lambda>(x, y). x * y) ` (S \<times> T)" |
320 |
unfolding set_times_def by (fastforce simp: image_iff) |
|
321 |
||
56899 | 322 |
lemma finite_set_plus: "finite s \<Longrightarrow> finite t \<Longrightarrow> finite (s + t)" |
323 |
unfolding set_plus_image by simp |
|
53596 | 324 |
|
56899 | 325 |
lemma finite_set_times: "finite s \<Longrightarrow> finite t \<Longrightarrow> finite (s * t)" |
326 |
unfolding set_times_image by simp |
|
53596 | 327 |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
328 |
lemma set_setsum_alt: |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
329 |
assumes fin: "finite I" |
47444
d21c95af2df7
removed "setsum_set", now subsumed by generic setsum
krauss
parents:
47443
diff
changeset
|
330 |
shows "setsum S I = {setsum s I |s. \<forall>i\<in>I. s i \<in> S i}" |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
331 |
(is "_ = ?setsum I") |
56899 | 332 |
using fin |
333 |
proof induct |
|
334 |
case empty |
|
335 |
then show ?case by simp |
|
336 |
next |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
337 |
case (insert x F) |
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47444
diff
changeset
|
338 |
have "setsum S (insert x F) = S x + ?setsum F" |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
339 |
using insert.hyps by auto |
56899 | 340 |
also have "\<dots> = {s x + setsum s F |s. \<forall> i\<in>insert x F. s i \<in> S i}" |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
341 |
unfolding set_plus_def |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
342 |
proof safe |
56899 | 343 |
fix y s |
344 |
assume "y \<in> S x" "\<forall>i\<in>F. s i \<in> S i" |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
345 |
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
|
346 |
using insert.hyps |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
347 |
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
|
348 |
qed auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
349 |
finally show ?case |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
350 |
using insert.hyps by auto |
56899 | 351 |
qed |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
352 |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
353 |
lemma setsum_set_cond_linear: |
56899 | 354 |
fixes f :: "'a::comm_monoid_add set \<Rightarrow> 'b::comm_monoid_add set" |
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47444
diff
changeset
|
355 |
assumes [intro!]: "\<And>A B. P A \<Longrightarrow> P B \<Longrightarrow> P (A + B)" "P {0}" |
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47444
diff
changeset
|
356 |
and f: "\<And>A B. P A \<Longrightarrow> P B \<Longrightarrow> f (A + B) = f A + f B" "f {0} = {0}" |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
357 |
assumes all: "\<And>i. i \<in> I \<Longrightarrow> P (S i)" |
47444
d21c95af2df7
removed "setsum_set", now subsumed by generic setsum
krauss
parents:
47443
diff
changeset
|
358 |
shows "f (setsum S I) = setsum (f \<circ> S) I" |
56899 | 359 |
proof (cases "finite I") |
360 |
case True |
|
361 |
from this all show ?thesis |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
362 |
proof induct |
56899 | 363 |
case empty |
364 |
then show ?case by (auto intro!: f) |
|
365 |
next |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
366 |
case (insert x F) |
60500 | 367 |
from \<open>finite F\<close> \<open>\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)\<close> have "P (setsum S F)" |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
368 |
by induct auto |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
369 |
with insert show ?case |
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
370 |
by (simp, subst f) auto |
56899 | 371 |
qed |
372 |
next |
|
373 |
case False |
|
374 |
then show ?thesis by (auto intro!: f) |
|
375 |
qed |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
376 |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
377 |
lemma setsum_set_linear: |
56899 | 378 |
fixes f :: "'a::comm_monoid_add set \<Rightarrow> 'b::comm_monoid_add set" |
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47444
diff
changeset
|
379 |
assumes "\<And>A B. f(A) + f(B) = f(A + B)" "f {0} = {0}" |
47444
d21c95af2df7
removed "setsum_set", now subsumed by generic setsum
krauss
parents:
47443
diff
changeset
|
380 |
shows "f (setsum S I) = setsum (f \<circ> S) I" |
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
39302
diff
changeset
|
381 |
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
|
382 |
|
47446 | 383 |
lemma set_times_Un_distrib: |
384 |
"A * (B \<union> C) = A * B \<union> A * C" |
|
385 |
"(A \<union> B) * C = A * C \<union> B * C" |
|
56899 | 386 |
by (auto simp: set_times_def) |
47446 | 387 |
|
388 |
lemma set_times_UNION_distrib: |
|
56899 | 389 |
"A * UNION I M = (\<Union>i\<in>I. A * M i)" |
390 |
"UNION I M * A = (\<Union>i\<in>I. M i * A)" |
|
391 |
by (auto simp: set_times_def) |
|
47446 | 392 |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
diff
changeset
|
393 |
end |