| author | haftmann | 
| Tue, 20 Jun 2017 13:07:49 +0200 | |
| changeset 66149 | 4bf16fb7c14d | 
| parent 64267 | b9a1486e79be | 
| child 69313 | b021008c5397 | 
| permissions | -rw-r--r-- | 
| 38622 | 1 | (* Title: HOL/Library/Set_Algebras.thy | 
| 63473 | 2 | Author: Jeremy Avigad | 
| 3 | Author: Kevin Donnelly | |
| 4 | Author: Florian Haftmann, TUM | |
| 16908 
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 | |
| 60500 | 7 | 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 | 8 | |
| 38622 | 9 | theory Set_Algebras | 
| 63485 | 10 | imports Main | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 11 | begin | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 12 | |
| 60500 | 13 | text \<open> | 
| 63485 | 14 | This library lifts operations like addition and multiplication to sets. It | 
| 15 | was designed to support asymptotic calculations. See the comments at the top | |
| 63680 | 16 | of \<^file>\<open>BigO.thy\<close>. | 
| 60500 | 17 | \<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 18 | |
| 47443 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 19 | instantiation set :: (plus) plus | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 20 | begin | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 21 | |
| 63473 | 22 | definition plus_set :: "'a::plus set \<Rightarrow> 'a set \<Rightarrow> 'a set" | 
| 23 |   where set_plus_def: "A + B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a + b}"
 | |
| 47443 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 24 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 25 | instance .. | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 26 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 27 | end | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 28 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 29 | instantiation set :: (times) times | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 30 | begin | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 31 | |
| 63473 | 32 | definition times_set :: "'a::times set \<Rightarrow> 'a set \<Rightarrow> 'a set" | 
| 33 |   where set_times_def: "A * B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a * b}"
 | |
| 47443 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 34 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 35 | instance .. | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 36 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 37 | end | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 38 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 39 | instantiation set :: (zero) zero | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 40 | begin | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 41 | |
| 63473 | 42 | definition set_zero[simp]: "(0::'a::zero set) = {0}"
 | 
| 47443 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 43 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 44 | instance .. | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 45 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 46 | end | 
| 56899 | 47 | |
| 47443 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 48 | instantiation set :: (one) one | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 49 | begin | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 50 | |
| 63473 | 51 | definition set_one[simp]: "(1::'a::one set) = {1}"
 | 
| 47443 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 52 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 53 | instance .. | 
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 54 | |
| 
aeff49a3369b
backported Set_Algebras to use type classes (basically reverting b3e8d5ec721d from 2008)
 krauss parents: 
44890diff
changeset | 55 | end | 
| 25594 | 56 | |
| 63473 | 57 | definition elt_set_plus :: "'a::plus \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "+o" 70) | 
| 58 |   where "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 | |
| 63473 | 60 | definition elt_set_times :: "'a::times \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "*o" 80) | 
| 61 |   where "a *o B = {c. \<exists>b\<in>B. c = a * b}"
 | |
| 25594 | 62 | |
| 63473 | 63 | abbreviation (input) elt_set_eq :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infix "=o" 50) | 
| 64 | where "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: 
44890diff
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: 
44890diff
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: 
44890diff
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: 
44890diff
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: 
44890diff
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: 
44890diff
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: 
44890diff
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: 
44890diff
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: 
25764diff
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 | |
| 63473 | 101 | lemma set_plus_rearrange: "(a +o C) + (b +o D) = (a + b) +o (C + D)" | 
| 102 | for a b :: "'a::comm_monoid_add" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
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) | 
| 63473 | 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: 
57512diff
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 | |
| 63473 | 110 | lemma set_plus_rearrange2: "a +o (b +o C) = (a + b) +o C" | 
| 111 | for a b :: "'a::semigroup_add" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56899diff
changeset | 112 | 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 | 113 | |
| 63473 | 114 | lemma set_plus_rearrange3: "(a +o B) + C = a +o (B + C)" | 
| 115 | for a :: "'a::semigroup_add" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 116 | 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: 
57512diff
changeset | 117 | apply (blast intro: ac_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 118 | 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 | 119 | apply (rule conjI) | 
| 19736 | 120 | apply (rule_tac x = "aa" in bexI) | 
| 121 | apply auto | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 122 | apply (rule_tac x = "ba" in bexI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 123 | apply (auto simp add: ac_simps) | 
| 19736 | 124 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 125 | |
| 63473 | 126 | theorem set_plus_rearrange4: "C + (a +o D) = a +o (C + D)" | 
| 127 | for a :: "'a::comm_monoid_add" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 128 | apply (auto simp add: elt_set_plus_def set_plus_def ac_simps) | 
| 19736 | 129 | apply (rule_tac x = "aa + ba" in exI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 130 | apply (auto simp add: ac_simps) | 
| 19736 | 131 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 132 | |
| 61337 | 133 | 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 | 134 | set_plus_rearrange3 set_plus_rearrange4 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 135 | |
| 56899 | 136 | lemma set_plus_mono [intro!]: "C \<subseteq> D \<Longrightarrow> a +o C \<subseteq> a +o D" | 
| 19736 | 137 | 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 | 138 | |
| 63473 | 139 | lemma set_plus_mono2 [intro]: "C \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> C + E \<subseteq> D + F" | 
| 140 | for C D E F :: "'a::plus set" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 141 | 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 | 142 | |
| 56899 | 143 | 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: 
25764diff
changeset | 144 | 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 | 145 | |
| 63473 | 146 | lemma set_plus_mono4 [intro]: "a \<in> C \<Longrightarrow> a +o D \<subseteq> D + C" | 
| 147 | for a :: "'a::comm_monoid_add" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 148 | 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 | 149 | |
| 56899 | 150 | lemma set_plus_mono5: "a \<in> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> a +o B \<subseteq> C + D" | 
| 151 | apply (subgoal_tac "a +o B \<subseteq> a +o D") | |
| 19736 | 152 | apply (erule order_trans) | 
| 153 | apply (erule set_plus_mono3) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 154 | apply (erule set_plus_mono) | 
| 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_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 | 158 | apply (frule set_plus_mono) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 159 | apply auto | 
| 19736 | 160 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 161 | |
| 56899 | 162 | 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 | 163 | apply (frule set_plus_mono2) | 
| 19736 | 164 | prefer 2 | 
| 165 | apply force | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 166 | apply assumption | 
| 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_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 | 170 | apply (frule set_plus_mono3) | 
| 
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 | |
| 63473 | 174 | lemma set_plus_mono4_b: "a \<in> C \<Longrightarrow> x \<in> a +o D \<Longrightarrow> x \<in> D + C" | 
| 175 | for a x :: "'a::comm_monoid_add" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 176 | apply (frule set_plus_mono4) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 177 | apply auto | 
| 19736 | 178 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 179 | |
| 63473 | 180 | lemma set_zero_plus [simp]: "0 +o C = C" | 
| 181 | for C :: "'a::comm_monoid_add set" | |
| 19736 | 182 | 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 | 183 | |
| 63473 | 184 | lemma set_zero_plus2: "0 \<in> A \<Longrightarrow> B \<subseteq> A + B" | 
| 185 | for A B :: "'a::comm_monoid_add set" | |
| 44142 | 186 | 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 | 187 | apply (rule_tac x = 0 in bexI) | 
| 19736 | 188 | apply (rule_tac x = x in bexI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 189 | apply (auto simp add: ac_simps) | 
| 19736 | 190 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 191 | |
| 63473 | 192 | lemma set_plus_imp_minus: "a \<in> b +o C \<Longrightarrow> a - b \<in> C" | 
| 193 | for a b :: "'a::ab_group_add" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 194 | 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 | 195 | |
| 63473 | 196 | lemma set_minus_imp_plus: "a - b \<in> C \<Longrightarrow> a \<in> b +o C" | 
| 197 | for a b :: "'a::ab_group_add" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 198 | 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 | 199 | apply (subgoal_tac "a = (a + - b) + b") | 
| 63473 | 200 | apply (rule bexI) | 
| 201 | apply assumption | |
| 202 | apply (auto simp add: ac_simps) | |
| 19736 | 203 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 204 | |
| 63473 | 205 | lemma set_minus_plus: "a - b \<in> C \<longleftrightarrow> a \<in> b +o C" | 
| 206 | for a b :: "'a::ab_group_add" | |
| 207 | apply (rule iffI) | |
| 208 | apply (rule set_minus_imp_plus) | |
| 209 | apply assumption | |
| 210 | apply (rule set_plus_imp_minus) | |
| 211 | apply assumption | |
| 212 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 213 | |
| 56899 | 214 | 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: 
25764diff
changeset | 215 | 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 | 216 | |
| 53596 | 217 | lemma set_times_elim: | 
| 218 | assumes "x \<in> A * B" | |
| 219 | obtains a b where "x = a * b" and "a \<in> A" and "b \<in> B" | |
| 220 | using assms unfolding set_times_def by fast | |
| 221 | ||
| 56899 | 222 | lemma set_times_intro2 [intro!]: "b \<in> C \<Longrightarrow> a * b \<in> a *o C" | 
| 19736 | 223 | 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 | 224 | |
| 63473 | 225 | lemma set_times_rearrange: "(a *o C) * (b *o D) = (a * b) *o (C * D)" | 
| 226 | for a b :: "'a::comm_monoid_mult" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 227 | apply (auto simp add: elt_set_times_def set_times_def) | 
| 19736 | 228 | apply (rule_tac x = "ba * bb" in exI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 229 | apply (auto simp add: ac_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 230 | apply (rule_tac x = "aa * a" in exI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 231 | apply (auto simp add: ac_simps) | 
| 19736 | 232 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 233 | |
| 63473 | 234 | lemma set_times_rearrange2: "a *o (b *o C) = (a * b) *o C" | 
| 235 | for a b :: "'a::semigroup_mult" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56899diff
changeset | 236 | 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 | 237 | |
| 63473 | 238 | lemma set_times_rearrange3: "(a *o B) * C = a *o (B * C)" | 
| 239 | for a :: "'a::semigroup_mult" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 240 | 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: 
57512diff
changeset | 241 | apply (blast intro: ac_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 242 | 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 | 243 | apply (rule conjI) | 
| 19736 | 244 | apply (rule_tac x = "aa" in bexI) | 
| 245 | apply auto | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 246 | apply (rule_tac x = "ba" in bexI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 247 | apply (auto simp add: ac_simps) | 
| 19736 | 248 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 249 | |
| 63473 | 250 | theorem set_times_rearrange4: "C * (a *o D) = a *o (C * D)" | 
| 251 | for a :: "'a::comm_monoid_mult" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 252 | apply (auto simp add: elt_set_times_def set_times_def ac_simps) | 
| 19736 | 253 | apply (rule_tac x = "aa * ba" in exI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 254 | apply (auto simp add: ac_simps) | 
| 19736 | 255 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 256 | |
| 61337 | 257 | 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 | 258 | set_times_rearrange3 set_times_rearrange4 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 259 | |
| 56899 | 260 | lemma set_times_mono [intro]: "C \<subseteq> D \<Longrightarrow> a *o C \<subseteq> a *o D" | 
| 19736 | 261 | 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 | 262 | |
| 63473 | 263 | lemma set_times_mono2 [intro]: "C \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> C * E \<subseteq> D * F" | 
| 264 | for C D E F :: "'a::times set" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 265 | 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 | 266 | |
| 56899 | 267 | 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: 
25764diff
changeset | 268 | 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 | 269 | |
| 63473 | 270 | lemma set_times_mono4 [intro]: "a \<in> C \<Longrightarrow> a *o D \<subseteq> D * C" | 
| 271 | for a :: "'a::comm_monoid_mult" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 272 | 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 | 273 | |
| 56899 | 274 | lemma set_times_mono5: "a \<in> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> a *o B \<subseteq> C * D" | 
| 275 | apply (subgoal_tac "a *o B \<subseteq> a *o D") | |
| 19736 | 276 | apply (erule order_trans) | 
| 277 | apply (erule set_times_mono3) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 278 | apply (erule set_times_mono) | 
| 19736 | 279 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 280 | |
| 56899 | 281 | 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 | 282 | apply (frule set_times_mono) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 283 | apply auto | 
| 19736 | 284 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 285 | |
| 56899 | 286 | 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 | 287 | apply (frule set_times_mono2) | 
| 19736 | 288 | prefer 2 | 
| 289 | apply force | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 290 | apply assumption | 
| 19736 | 291 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 292 | |
| 56899 | 293 | 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 | 294 | apply (frule set_times_mono3) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 295 | apply auto | 
| 19736 | 296 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 297 | |
| 63473 | 298 | lemma set_times_mono4_b: "a \<in> C \<Longrightarrow> x \<in> a *o D \<Longrightarrow> x \<in> D * C" | 
| 299 | for a x :: "'a::comm_monoid_mult" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 300 | apply (frule set_times_mono4) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 301 | apply auto | 
| 19736 | 302 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 303 | |
| 63473 | 304 | lemma set_one_times [simp]: "1 *o C = C" | 
| 305 | for C :: "'a::comm_monoid_mult set" | |
| 19736 | 306 | 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 | 307 | |
| 63473 | 308 | lemma set_times_plus_distrib: "a *o (b +o C) = (a * b) +o (a *o C)" | 
| 309 | for a b :: "'a::semiring" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
21404diff
changeset | 310 | 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 | 311 | |
| 63473 | 312 | lemma set_times_plus_distrib2: "a *o (B + C) = (a *o B) + (a *o C)" | 
| 313 | for a :: "'a::semiring" | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 314 | apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) | 
| 19736 | 315 | apply blast | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 316 | apply (rule_tac x = "b + bb" in exI) | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
21404diff
changeset | 317 | apply (auto simp add: ring_distribs) | 
| 19736 | 318 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 319 | |
| 63473 | 320 | lemma set_times_plus_distrib3: "(a +o C) * D \<subseteq> a *o D + C * D" | 
| 321 | for a :: "'a::semiring" | |
| 322 | apply (auto simp: elt_set_plus_def elt_set_times_def set_times_def set_plus_def ring_distribs) | |
| 16908 
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 | |
| 61337 | 326 | lemmas set_times_plus_distribs = | 
| 19380 | 327 | set_times_plus_distrib | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 328 | set_times_plus_distrib2 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 329 | |
| 63473 | 330 | lemma set_neg_intro: "a \<in> (- 1) *o C \<Longrightarrow> - a \<in> C" | 
| 331 | for a :: "'a::ring_1" | |
| 19736 | 332 | 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 | 333 | |
| 63473 | 334 | lemma set_neg_intro2: "a \<in> C \<Longrightarrow> - a \<in> (- 1) *o C" | 
| 335 | for a :: "'a::ring_1" | |
| 19736 | 336 | by (auto simp add: elt_set_times_def) | 
| 337 | ||
| 53596 | 338 | lemma set_plus_image: "S + T = (\<lambda>(x, y). x + y) ` (S \<times> T)" | 
| 63473 | 339 | by (fastforce simp: set_plus_def image_iff) | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 340 | |
| 53596 | 341 | lemma set_times_image: "S * T = (\<lambda>(x, y). x * y) ` (S \<times> T)" | 
| 63473 | 342 | by (fastforce simp: set_times_def image_iff) | 
| 53596 | 343 | |
| 56899 | 344 | lemma finite_set_plus: "finite s \<Longrightarrow> finite t \<Longrightarrow> finite (s + t)" | 
| 63473 | 345 | by (simp add: set_plus_image) | 
| 53596 | 346 | |
| 56899 | 347 | lemma finite_set_times: "finite s \<Longrightarrow> finite t \<Longrightarrow> finite (s * t)" | 
| 63473 | 348 | by (simp add: set_times_image) | 
| 53596 | 349 | |
| 64267 | 350 | lemma set_sum_alt: | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 351 | assumes fin: "finite I" | 
| 64267 | 352 |   shows "sum S I = {sum s I |s. \<forall>i\<in>I. s i \<in> S i}"
 | 
| 353 | (is "_ = ?sum I") | |
| 56899 | 354 | using fin | 
| 355 | proof induct | |
| 356 | case empty | |
| 357 | then show ?case by simp | |
| 358 | next | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 359 | case (insert x F) | 
| 64267 | 360 | have "sum S (insert x F) = S x + ?sum F" | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 361 | using insert.hyps by auto | 
| 64267 | 362 |   also have "\<dots> = {s x + sum 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: 
39302diff
changeset | 363 | unfolding set_plus_def | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 364 | proof safe | 
| 56899 | 365 | fix y s | 
| 366 | assume "y \<in> S x" "\<forall>i\<in>F. s i \<in> S i" | |
| 64267 | 367 | then show "\<exists>s'. y + sum s F = s' x + sum s' F \<and> (\<forall>i\<in>insert x F. s' i \<in> S i)" | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 368 | using insert.hyps | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 369 | 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: 
39302diff
changeset | 370 | qed auto | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 371 | finally show ?case | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 372 | using insert.hyps by auto | 
| 56899 | 373 | qed | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 374 | |
| 64267 | 375 | lemma sum_set_cond_linear: | 
| 56899 | 376 | 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: 
47444diff
changeset | 377 |   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: 
47444diff
changeset | 378 |     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: 
39302diff
changeset | 379 | assumes all: "\<And>i. i \<in> I \<Longrightarrow> P (S i)" | 
| 64267 | 380 | shows "f (sum S I) = sum (f \<circ> S) I" | 
| 56899 | 381 | proof (cases "finite I") | 
| 382 | case True | |
| 383 | from this all show ?thesis | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 384 | proof induct | 
| 56899 | 385 | case empty | 
| 386 | then show ?case by (auto intro!: f) | |
| 387 | next | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 388 | case (insert x F) | 
| 64267 | 389 | from \<open>finite F\<close> \<open>\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)\<close> have "P (sum S F)" | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 390 | by induct auto | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 391 | with insert show ?case | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 392 | by (simp, subst f) auto | 
| 56899 | 393 | qed | 
| 394 | next | |
| 395 | case False | |
| 396 | then show ?thesis by (auto intro!: f) | |
| 397 | qed | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 398 | |
| 64267 | 399 | lemma sum_set_linear: | 
| 56899 | 400 | 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: 
47444diff
changeset | 401 |   assumes "\<And>A B. f(A) + f(B) = f(A + B)" "f {0} = {0}"
 | 
| 64267 | 402 | shows "f (sum S I) = sum (f \<circ> S) I" | 
| 403 | using sum_set_cond_linear[of "\<lambda>x. True" f I S] assms by auto | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 404 | |
| 47446 | 405 | lemma set_times_Un_distrib: | 
| 406 | "A * (B \<union> C) = A * B \<union> A * C" | |
| 407 | "(A \<union> B) * C = A * C \<union> B * C" | |
| 56899 | 408 | by (auto simp: set_times_def) | 
| 47446 | 409 | |
| 410 | lemma set_times_UNION_distrib: | |
| 56899 | 411 | "A * UNION I M = (\<Union>i\<in>I. A * M i)" | 
| 412 | "UNION I M * A = (\<Union>i\<in>I. M i * A)" | |
| 413 | by (auto simp: set_times_def) | |
| 47446 | 414 | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 415 | end |