| author | blanchet | 
| Wed, 15 Mar 2023 15:28:44 +0100 | |
| changeset 77670 | b9e9b818d7b0 | 
| parent 77003 | ab905b5bb206 | 
| child 80914 | d97fdabd9e2b | 
| 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 | 
| 77003 | 15 | was designed to support asymptotic calculations for the now-obsolete BigO theory, | 
| 16 | but has other uses. | |
| 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 | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 90 | lemma sumset_empty [simp]: "A + {} = {}" "{} + A = {}"
 | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 91 | by (auto simp: set_plus_def) | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 92 | |
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 93 | lemma Un_set_plus: "(A \<union> B) + C = (A+C) \<union> (B+C)" and set_plus_Un: "C + (A \<union> B) = (C+A) \<union> (C+B)" | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 94 | by (auto simp: set_plus_def) | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 95 | |
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 96 | lemma | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 97 | fixes A :: "'a::comm_monoid_add set" | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 98 | shows insert_set_plus: "(insert a A) + B = (A+B) \<union> (((+)a) ` B)" and set_plus_insert: "B + (insert a A) = (B+A) \<union> (((+)a) ` B)" | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 99 | using add.commute by (auto simp: set_plus_def) | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 100 | |
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 101 | lemma set_add_0 [simp]: | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 102 | fixes A :: "'a::comm_monoid_add set" | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 103 |   shows "{0} + A = A"
 | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 104 | by (metis comm_monoid_add_class.add_0 set_zero) | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 105 | |
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 106 | lemma set_add_0_right [simp]: | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 107 | fixes A :: "'a::comm_monoid_add set" | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 108 |   shows "A + {0} = A"
 | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 109 | by (metis add.comm_neutral set_zero) | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 110 | |
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 111 | lemma card_plus_sing: | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 112 | fixes A :: "'a::ab_group_add set" | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 113 |   shows "card (A + {a}) = card A"
 | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 114 | proof (rule bij_betw_same_card) | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 115 |   show "bij_betw ((+) (-a)) (A + {a}) A"
 | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 116 | by (fastforce simp: set_plus_def bij_betw_def image_iff) | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 117 | qed | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 118 | |
| 56899 | 119 | 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 | 120 | 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 | 121 | |
| 53596 | 122 | lemma set_plus_elim: | 
| 123 | assumes "x \<in> A + B" | |
| 124 | obtains a b where "x = a + b" and "a \<in> A" and "b \<in> B" | |
| 125 | using assms unfolding set_plus_def by fast | |
| 126 | ||
| 56899 | 127 | lemma set_plus_intro2 [intro]: "b \<in> C \<Longrightarrow> a + b \<in> a +o C" | 
| 19736 | 128 | 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 | 129 | |
| 63473 | 130 | lemma set_plus_rearrange: "(a +o C) + (b +o D) = (a + b) +o (C + D)" | 
| 131 | for a b :: "'a::comm_monoid_add" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 132 | by (auto simp: elt_set_plus_def set_plus_def; metis group_cancel.add1 group_cancel.add2) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 133 | |
| 63473 | 134 | lemma set_plus_rearrange2: "a +o (b +o C) = (a + b) +o C" | 
| 135 | for a b :: "'a::semigroup_add" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56899diff
changeset | 136 | 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 | 137 | |
| 63473 | 138 | lemma set_plus_rearrange3: "(a +o B) + C = a +o (B + C)" | 
| 139 | for a :: "'a::semigroup_add" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 140 | by (auto simp add: elt_set_plus_def set_plus_def; metis add.assoc) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 141 | |
| 63473 | 142 | theorem set_plus_rearrange4: "C + (a +o D) = a +o (C + D)" | 
| 143 | for a :: "'a::comm_monoid_add" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 144 | by (metis add.commute set_plus_rearrange3) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 145 | |
| 61337 | 146 | 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 | 147 | set_plus_rearrange3 set_plus_rearrange4 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 148 | |
| 56899 | 149 | lemma set_plus_mono [intro!]: "C \<subseteq> D \<Longrightarrow> a +o C \<subseteq> a +o D" | 
| 19736 | 150 | 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 | 151 | |
| 63473 | 152 | lemma set_plus_mono2 [intro]: "C \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> C + E \<subseteq> D + F" | 
| 153 | 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 | 154 | 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 | 155 | |
| 56899 | 156 | 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 | 157 | 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 | 158 | |
| 63473 | 159 | lemma set_plus_mono4 [intro]: "a \<in> C \<Longrightarrow> a +o D \<subseteq> D + C" | 
| 160 | for a :: "'a::comm_monoid_add" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 161 | 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 | 162 | |
| 56899 | 163 | lemma set_plus_mono5: "a \<in> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> a +o B \<subseteq> C + D" | 
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 164 | using order_subst2 by blast | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 165 | |
| 56899 | 166 | lemma set_plus_mono_b: "C \<subseteq> D \<Longrightarrow> x \<in> a +o C \<Longrightarrow> x \<in> a +o D" | 
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 167 | using set_plus_mono by blast | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 168 | |
| 63473 | 169 | lemma set_zero_plus [simp]: "0 +o C = C" | 
| 170 | for C :: "'a::comm_monoid_add set" | |
| 19736 | 171 | 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 | 172 | |
| 63473 | 173 | lemma set_zero_plus2: "0 \<in> A \<Longrightarrow> B \<subseteq> A + B" | 
| 174 | for A B :: "'a::comm_monoid_add set" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 175 | using set_plus_intro by fastforce | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 176 | |
| 63473 | 177 | lemma set_plus_imp_minus: "a \<in> b +o C \<Longrightarrow> a - b \<in> C" | 
| 178 | 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 | 179 | 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 | 180 | |
| 63473 | 181 | lemma set_minus_imp_plus: "a - b \<in> C \<Longrightarrow> a \<in> b +o C" | 
| 182 | for a b :: "'a::ab_group_add" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 183 | by (metis add.commute diff_add_cancel set_plus_intro2) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 184 | |
| 63473 | 185 | lemma set_minus_plus: "a - b \<in> C \<longleftrightarrow> a \<in> b +o C" | 
| 186 | for a b :: "'a::ab_group_add" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 187 | by (meson set_minus_imp_plus set_plus_imp_minus) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 188 | |
| 56899 | 189 | 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 | 190 | 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 | 191 | |
| 53596 | 192 | lemma set_times_elim: | 
| 193 | assumes "x \<in> A * B" | |
| 194 | obtains a b where "x = a * b" and "a \<in> A" and "b \<in> B" | |
| 195 | using assms unfolding set_times_def by fast | |
| 196 | ||
| 56899 | 197 | lemma set_times_intro2 [intro!]: "b \<in> C \<Longrightarrow> a * b \<in> a *o C" | 
| 19736 | 198 | 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 | 199 | |
| 63473 | 200 | lemma set_times_rearrange: "(a *o C) * (b *o D) = (a * b) *o (C * D)" | 
| 201 | for a b :: "'a::comm_monoid_mult" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 202 | by (auto simp add: elt_set_times_def set_times_def; metis mult.assoc mult.left_commute) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 203 | |
| 63473 | 204 | lemma set_times_rearrange2: "a *o (b *o C) = (a * b) *o C" | 
| 205 | for a b :: "'a::semigroup_mult" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56899diff
changeset | 206 | 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 | 207 | |
| 63473 | 208 | lemma set_times_rearrange3: "(a *o B) * C = a *o (B * C)" | 
| 209 | for a :: "'a::semigroup_mult" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 210 | by (auto simp add: elt_set_times_def set_times_def; metis mult.assoc) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 211 | |
| 63473 | 212 | theorem set_times_rearrange4: "C * (a *o D) = a *o (C * D)" | 
| 213 | for a :: "'a::comm_monoid_mult" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 214 | by (metis mult.commute set_times_rearrange3) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 215 | |
| 61337 | 216 | 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 | 217 | set_times_rearrange3 set_times_rearrange4 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 218 | |
| 56899 | 219 | lemma set_times_mono [intro]: "C \<subseteq> D \<Longrightarrow> a *o C \<subseteq> a *o D" | 
| 19736 | 220 | 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 | 221 | |
| 63473 | 222 | lemma set_times_mono2 [intro]: "C \<subseteq> D \<Longrightarrow> E \<subseteq> F \<Longrightarrow> C * E \<subseteq> D * F" | 
| 223 | 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 | 224 | 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 | 225 | |
| 56899 | 226 | 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 | 227 | 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 | 228 | |
| 63473 | 229 | lemma set_times_mono4 [intro]: "a \<in> C \<Longrightarrow> a *o D \<subseteq> D * C" | 
| 230 | for a :: "'a::comm_monoid_mult" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 231 | 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 | 232 | |
| 56899 | 233 | lemma set_times_mono5: "a \<in> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> a *o B \<subseteq> C * D" | 
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 234 | by (meson dual_order.trans set_times_mono set_times_mono3) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 235 | |
| 63473 | 236 | lemma set_one_times [simp]: "1 *o C = C" | 
| 237 | for C :: "'a::comm_monoid_mult set" | |
| 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 | |
| 63473 | 240 | lemma set_times_plus_distrib: "a *o (b +o C) = (a * b) +o (a *o C)" | 
| 241 | for a b :: "'a::semiring" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
21404diff
changeset | 242 | 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 | 243 | |
| 63473 | 244 | lemma set_times_plus_distrib2: "a *o (B + C) = (a *o B) + (a *o C)" | 
| 245 | for a :: "'a::semiring" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 246 | by (auto simp: set_plus_def elt_set_times_def; metis distrib_left) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 247 | |
| 63473 | 248 | lemma set_times_plus_distrib3: "(a +o C) * D \<subseteq> a *o D + C * D" | 
| 249 | for a :: "'a::semiring" | |
| 75450 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 250 | using distrib_right | 
| 
f16d83de3e4a
Tidied up some super-messy proofs
 paulson <lp15@cam.ac.uk> parents: 
69313diff
changeset | 251 | by (fastforce simp add: elt_set_plus_def elt_set_times_def set_times_def set_plus_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 252 | |
| 61337 | 253 | lemmas set_times_plus_distribs = | 
| 19380 | 254 | set_times_plus_distrib | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 255 | set_times_plus_distrib2 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 256 | |
| 63473 | 257 | lemma set_neg_intro: "a \<in> (- 1) *o C \<Longrightarrow> - a \<in> C" | 
| 258 | for a :: "'a::ring_1" | |
| 19736 | 259 | 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 | 260 | |
| 63473 | 261 | lemma set_neg_intro2: "a \<in> C \<Longrightarrow> - a \<in> (- 1) *o C" | 
| 262 | for a :: "'a::ring_1" | |
| 19736 | 263 | by (auto simp add: elt_set_times_def) | 
| 264 | ||
| 53596 | 265 | lemma set_plus_image: "S + T = (\<lambda>(x, y). x + y) ` (S \<times> T)" | 
| 63473 | 266 | by (fastforce simp: set_plus_def image_iff) | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 267 | |
| 53596 | 268 | lemma set_times_image: "S * T = (\<lambda>(x, y). x * y) ` (S \<times> T)" | 
| 63473 | 269 | by (fastforce simp: set_times_def image_iff) | 
| 53596 | 270 | |
| 56899 | 271 | lemma finite_set_plus: "finite s \<Longrightarrow> finite t \<Longrightarrow> finite (s + t)" | 
| 63473 | 272 | by (simp add: set_plus_image) | 
| 53596 | 273 | |
| 56899 | 274 | lemma finite_set_times: "finite s \<Longrightarrow> finite t \<Longrightarrow> finite (s * t)" | 
| 63473 | 275 | by (simp add: set_times_image) | 
| 53596 | 276 | |
| 64267 | 277 | lemma set_sum_alt: | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 278 | assumes fin: "finite I" | 
| 64267 | 279 |   shows "sum S I = {sum s I |s. \<forall>i\<in>I. s i \<in> S i}"
 | 
| 280 | (is "_ = ?sum I") | |
| 56899 | 281 | using fin | 
| 282 | proof induct | |
| 283 | case empty | |
| 284 | then show ?case by simp | |
| 285 | next | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 286 | case (insert x F) | 
| 64267 | 287 | 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 | 288 | using insert.hyps by auto | 
| 64267 | 289 |   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 | 290 | unfolding set_plus_def | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 291 | proof safe | 
| 56899 | 292 | fix y s | 
| 293 | assume "y \<in> S x" "\<forall>i\<in>F. s i \<in> S i" | |
| 64267 | 294 | 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 | 295 | using insert.hyps | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 296 | 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 | 297 | qed auto | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 298 | finally show ?case | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 299 | using insert.hyps by auto | 
| 56899 | 300 | qed | 
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 301 | |
| 64267 | 302 | lemma sum_set_cond_linear: | 
| 56899 | 303 | 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 | 304 |   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 | 305 |     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 | 306 | assumes all: "\<And>i. i \<in> I \<Longrightarrow> P (S i)" | 
| 64267 | 307 | shows "f (sum S I) = sum (f \<circ> S) I" | 
| 56899 | 308 | proof (cases "finite I") | 
| 309 | case True | |
| 310 | from this all show ?thesis | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 311 | proof induct | 
| 56899 | 312 | case empty | 
| 313 | then show ?case by (auto intro!: f) | |
| 314 | next | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 315 | case (insert x F) | 
| 64267 | 316 | 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 | 317 | by induct auto | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 318 | with insert show ?case | 
| 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 319 | by (simp, subst f) auto | 
| 56899 | 320 | qed | 
| 321 | next | |
| 322 | case False | |
| 323 | then show ?thesis by (auto intro!: f) | |
| 324 | qed | |
| 40887 
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
 hoelzl parents: 
39302diff
changeset | 325 | |
| 64267 | 326 | lemma sum_set_linear: | 
| 56899 | 327 | 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 | 328 |   assumes "\<And>A B. f(A) + f(B) = f(A + B)" "f {0} = {0}"
 | 
| 64267 | 329 | shows "f (sum S I) = sum (f \<circ> S) I" | 
| 330 | 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 | 331 | |
| 47446 | 332 | lemma set_times_Un_distrib: | 
| 333 | "A * (B \<union> C) = A * B \<union> A * C" | |
| 334 | "(A \<union> B) * C = A * C \<union> B * C" | |
| 56899 | 335 | by (auto simp: set_times_def) | 
| 47446 | 336 | |
| 337 | lemma set_times_UNION_distrib: | |
| 69313 | 338 | "A * \<Union>(M ` I) = (\<Union>i\<in>I. A * M i)" | 
| 339 | "\<Union>(M ` I) * A = (\<Union>i\<in>I. M i * A)" | |
| 56899 | 340 | by (auto simp: set_times_def) | 
| 47446 | 341 | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 342 | end |