| author | haftmann | 
| Fri, 06 Mar 2009 20:30:19 +0100 | |
| changeset 30330 | 8291bc63d7c9 | 
| parent 29667 | 53103fc8ffa3 | 
| child 30729 | 461ee3e49ad3 | 
| child 30738 | 0842e906300c | 
| permissions | -rw-r--r-- | 
| 16932 | 1 | (* Title: HOL/Library/SetsAndFunctions.thy | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 2 | Author: Jeremy Avigad and Kevin Donnelly | 
| 
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 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 5 | header {* Operations on sets and functions *}
 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 6 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 7 | theory SetsAndFunctions | 
| 27368 | 8 | imports Plain | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 9 | begin | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 10 | |
| 19736 | 11 | text {*
 | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 12 | This library lifts operations like addition and muliplication to sets and | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 13 | functions of appropriate types. It was designed to support asymptotic | 
| 17161 | 14 | calculations. See the comments at the top of theory @{text BigO}.
 | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 15 | *} | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 16 | |
| 19736 | 17 | subsection {* Basic definitions *}
 | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 18 | |
| 25594 | 19 | definition | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 20 |   set_plus :: "('a::plus) set => 'a set => 'a set"  (infixl "\<oplus>" 65) where
 | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 21 |   "A \<oplus> B == {c. EX a:A. EX b:B. c = a + b}"
 | 
| 25594 | 22 | |
| 23 | instantiation "fun" :: (type, plus) plus | |
| 24 | begin | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 25 | |
| 25594 | 26 | definition | 
| 27 | func_plus: "f + g == (%x. f x + g x)" | |
| 28 | ||
| 29 | instance .. | |
| 30 | ||
| 31 | end | |
| 32 | ||
| 33 | definition | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 34 |   set_times :: "('a::times) set => 'a set => 'a set"  (infixl "\<otimes>" 70) where
 | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 35 |   "A \<otimes> B == {c. EX a:A. EX b:B. c = a * b}"
 | 
| 25594 | 36 | |
| 37 | instantiation "fun" :: (type, times) times | |
| 38 | begin | |
| 39 | ||
| 40 | definition | |
| 41 | func_times: "f * g == (%x. f x * g x)" | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 42 | |
| 25594 | 43 | instance .. | 
| 44 | ||
| 45 | end | |
| 46 | ||
| 47 | ||
| 48 | instantiation "fun" :: (type, zero) zero | |
| 49 | begin | |
| 50 | ||
| 51 | definition | |
| 52 |   func_zero: "0::(('a::type) => ('b::zero)) == %x. 0"
 | |
| 53 | ||
| 54 | instance .. | |
| 55 | ||
| 56 | end | |
| 57 | ||
| 58 | instantiation "fun" :: (type, one) one | |
| 59 | begin | |
| 60 | ||
| 61 | definition | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 62 |   func_one: "1::(('a::type) => ('b::one)) == %x. 1"
 | 
| 25594 | 63 | |
| 64 | instance .. | |
| 65 | ||
| 66 | end | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 67 | |
| 19736 | 68 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20523diff
changeset | 69 | elt_set_plus :: "'a::plus => 'a set => 'a set" (infixl "+o" 70) where | 
| 19736 | 70 |   "a +o B = {c. EX b:B. c = a + b}"
 | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 71 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20523diff
changeset | 72 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20523diff
changeset | 73 | elt_set_times :: "'a::times => 'a set => 'a set" (infixl "*o" 80) where | 
| 19736 | 74 |   "a *o B = {c. EX b:B. c = a * b}"
 | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 75 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 76 | abbreviation (input) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20523diff
changeset | 77 | elt_set_eq :: "'a => 'a set => bool" (infix "=o" 50) where | 
| 19380 | 78 | "x =o A == x : A" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 79 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19736diff
changeset | 80 | instance "fun" :: (type,semigroup_add)semigroup_add | 
| 19380 | 81 | by default (auto simp add: func_plus add_assoc) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 82 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19736diff
changeset | 83 | instance "fun" :: (type,comm_monoid_add)comm_monoid_add | 
| 19380 | 84 | by default (auto simp add: func_zero func_plus add_ac) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 85 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19736diff
changeset | 86 | instance "fun" :: (type,ab_group_add)ab_group_add | 
| 19736 | 87 | apply default | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 88 | apply (simp add: fun_Compl_def func_plus func_zero) | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 89 | apply (simp add: fun_Compl_def func_plus fun_diff_def diff_minus) | 
| 19736 | 90 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 91 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19736diff
changeset | 92 | instance "fun" :: (type,semigroup_mult)semigroup_mult | 
| 19736 | 93 | apply default | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 94 | apply (auto simp add: func_times mult_assoc) | 
| 19736 | 95 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 96 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19736diff
changeset | 97 | instance "fun" :: (type,comm_monoid_mult)comm_monoid_mult | 
| 19736 | 98 | apply default | 
| 99 | apply (auto simp add: func_one func_times mult_ac) | |
| 100 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 101 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
19736diff
changeset | 102 | instance "fun" :: (type,comm_ring_1)comm_ring_1 | 
| 19736 | 103 | apply default | 
| 29667 | 104 | apply (auto simp add: func_plus func_times fun_Compl_def fun_diff_def | 
| 105 | func_one func_zero algebra_simps) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 106 | apply (drule fun_cong) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 107 | apply simp | 
| 19736 | 108 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 109 | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29233diff
changeset | 110 | interpretation set_semigroup_add!: semigroup_add "op \<oplus> :: ('a::semigroup_add) set => 'a set => 'a set"
 | 
| 19736 | 111 | apply default | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 112 | apply (unfold set_plus_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 113 | apply (force simp add: add_assoc) | 
| 19736 | 114 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 115 | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29233diff
changeset | 116 | interpretation set_semigroup_mult!: semigroup_mult "op \<otimes> :: ('a::semigroup_mult) set => 'a set => 'a set"
 | 
| 19736 | 117 | apply default | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 118 | apply (unfold set_times_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 119 | apply (force simp add: mult_assoc) | 
| 19736 | 120 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 121 | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29233diff
changeset | 122 | interpretation set_comm_monoid_add!: comm_monoid_add "{0}" "op \<oplus> :: ('a::comm_monoid_add) set => 'a set => 'a set"
 | 
| 19736 | 123 | apply default | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 124 | apply (unfold set_plus_def) | 
| 19736 | 125 | apply (force simp add: add_ac) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 126 | apply force | 
| 19736 | 127 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 128 | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29233diff
changeset | 129 | interpretation set_comm_monoid_mult!: comm_monoid_mult "{1}" "op \<otimes> :: ('a::comm_monoid_mult) set => 'a set => 'a set"
 | 
| 19736 | 130 | apply default | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 131 | apply (unfold set_times_def) | 
| 19736 | 132 | apply (force simp add: mult_ac) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 133 | apply force | 
| 19736 | 134 | done | 
| 135 | ||
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 136 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 137 | subsection {* Basic properties *}
 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 138 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 139 | lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C \<oplus> D" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 140 | 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 | 141 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 142 | lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C" | 
| 19736 | 143 | 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 | 144 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 145 | lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) \<oplus> | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 146 | (b +o D) = (a + b) +o (C \<oplus> D)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 147 | apply (auto simp add: elt_set_plus_def set_plus_def add_ac) | 
| 19736 | 148 | apply (rule_tac x = "ba + bb" in exI) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 149 | apply (auto simp add: add_ac) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 150 | apply (rule_tac x = "aa + a" in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 151 | apply (auto simp add: add_ac) | 
| 19736 | 152 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 153 | |
| 19736 | 154 | lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = | 
| 155 | (a + b) +o C" | |
| 156 | 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 | 157 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 158 | lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) \<oplus> C = | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 159 | a +o (B \<oplus> C)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 160 | apply (auto simp add: elt_set_plus_def set_plus_def) | 
| 19736 | 161 | apply (blast intro: add_ac) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 162 | 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 | 163 | apply (rule conjI) | 
| 19736 | 164 | apply (rule_tac x = "aa" in bexI) | 
| 165 | apply auto | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 166 | apply (rule_tac x = "ba" in bexI) | 
| 19736 | 167 | apply (auto simp add: add_ac) | 
| 168 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 169 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 170 | theorem set_plus_rearrange4: "C \<oplus> ((a::'a::comm_monoid_add) +o D) = | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 171 | a +o (C \<oplus> D)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 172 | apply (auto intro!: subsetI simp add: elt_set_plus_def set_plus_def add_ac) | 
| 19736 | 173 | apply (rule_tac x = "aa + ba" in exI) | 
| 174 | apply (auto simp add: add_ac) | |
| 175 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 176 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 177 | theorems set_plus_rearranges = set_plus_rearrange set_plus_rearrange2 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 178 | set_plus_rearrange3 set_plus_rearrange4 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 179 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 180 | lemma set_plus_mono [intro!]: "C <= D ==> a +o C <= a +o D" | 
| 19736 | 181 | 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 | 182 | |
| 19736 | 183 | lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==>
 | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 184 | C \<oplus> E <= D \<oplus> F" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 185 | 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 | 186 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 187 | lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C \<oplus> D" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 188 | 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 | 189 | |
| 19736 | 190 | lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==> | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 191 | a +o D <= D \<oplus> C" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 192 | by (auto simp add: elt_set_plus_def set_plus_def add_ac) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 193 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 194 | lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C \<oplus> D" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 195 | apply (subgoal_tac "a +o B <= a +o D") | 
| 19736 | 196 | apply (erule order_trans) | 
| 197 | apply (erule set_plus_mono3) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 198 | apply (erule set_plus_mono) | 
| 19736 | 199 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 200 | |
| 19736 | 201 | lemma set_plus_mono_b: "C <= D ==> x : a +o C | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 202 | ==> x : a +o D" | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 203 | apply (frule set_plus_mono) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 204 | apply auto | 
| 19736 | 205 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 206 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 207 | lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C \<oplus> E ==> | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 208 | x : D \<oplus> F" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 209 | apply (frule set_plus_mono2) | 
| 19736 | 210 | prefer 2 | 
| 211 | apply force | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 212 | apply assumption | 
| 19736 | 213 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 214 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 215 | lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C \<oplus> D" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 216 | apply (frule set_plus_mono3) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 217 | apply auto | 
| 19736 | 218 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 219 | |
| 19736 | 220 | lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==> | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 221 | x : a +o D ==> x : D \<oplus> C" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 222 | apply (frule set_plus_mono4) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 223 | apply auto | 
| 19736 | 224 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 225 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 226 | lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" | 
| 19736 | 227 | 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 | 228 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 229 | lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A \<oplus> B" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 230 | apply (auto intro!: subsetI simp add: set_plus_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 231 | apply (rule_tac x = 0 in bexI) | 
| 19736 | 232 | apply (rule_tac x = x in bexI) | 
| 233 | apply (auto simp add: add_ac) | |
| 234 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 235 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 236 | lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C" | 
| 19736 | 237 | by (auto simp add: elt_set_plus_def add_ac diff_minus) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 238 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 239 | lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C ==> a : b +o C" | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 240 | apply (auto simp add: elt_set_plus_def add_ac diff_minus) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 241 | apply (subgoal_tac "a = (a + - b) + b") | 
| 19736 | 242 | apply (rule bexI, assumption, assumption) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 243 | apply (auto simp add: add_ac) | 
| 19736 | 244 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 245 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 246 | lemma set_minus_plus: "((a::'a::ab_group_add) - b : C) = (a : b +o C)" | 
| 19736 | 247 | 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 | 248 | assumption) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 249 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 250 | lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C \<otimes> D" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 251 | 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 | 252 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 253 | lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C" | 
| 19736 | 254 | 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 | 255 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 256 | lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) \<otimes> | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 257 | (b *o D) = (a * b) *o (C \<otimes> D)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 258 | apply (auto simp add: elt_set_times_def set_times_def) | 
| 19736 | 259 | apply (rule_tac x = "ba * bb" in exI) | 
| 260 | apply (auto simp add: mult_ac) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 261 | apply (rule_tac x = "aa * a" in exI) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 262 | apply (auto simp add: mult_ac) | 
| 19736 | 263 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 264 | |
| 19736 | 265 | lemma set_times_rearrange2: "(a::'a::semigroup_mult) *o (b *o C) = | 
| 266 | (a * b) *o C" | |
| 267 | 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 | 268 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 269 | lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) \<otimes> C = | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 270 | a *o (B \<otimes> C)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 271 | apply (auto simp add: elt_set_times_def set_times_def) | 
| 19736 | 272 | apply (blast intro: mult_ac) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 273 | 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 | 274 | apply (rule conjI) | 
| 19736 | 275 | apply (rule_tac x = "aa" in bexI) | 
| 276 | apply auto | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 277 | apply (rule_tac x = "ba" in bexI) | 
| 19736 | 278 | apply (auto simp add: mult_ac) | 
| 279 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 280 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 281 | theorem set_times_rearrange4: "C \<otimes> ((a::'a::comm_monoid_mult) *o D) = | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 282 | a *o (C \<otimes> D)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 283 | apply (auto intro!: subsetI 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 | 284 | mult_ac) | 
| 19736 | 285 | apply (rule_tac x = "aa * ba" in exI) | 
| 286 | apply (auto simp add: mult_ac) | |
| 287 | done | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 288 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 289 | theorems set_times_rearranges = set_times_rearrange set_times_rearrange2 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 290 | set_times_rearrange3 set_times_rearrange4 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 291 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 292 | lemma set_times_mono [intro]: "C <= D ==> a *o C <= a *o D" | 
| 19736 | 293 | 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 | 294 | |
| 19736 | 295 | lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==>
 | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 296 | C \<otimes> E <= D \<otimes> F" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 297 | 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 | 298 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 299 | lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C \<otimes> D" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 300 | 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 | 301 | |
| 19736 | 302 | lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==> | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 303 | a *o D <= D \<otimes> C" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 304 | by (auto simp add: elt_set_times_def set_times_def mult_ac) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 305 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 306 | lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C \<otimes> D" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 307 | apply (subgoal_tac "a *o B <= a *o D") | 
| 19736 | 308 | apply (erule order_trans) | 
| 309 | apply (erule set_times_mono3) | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 310 | apply (erule set_times_mono) | 
| 19736 | 311 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 312 | |
| 19736 | 313 | lemma set_times_mono_b: "C <= D ==> x : a *o C | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 314 | ==> x : a *o D" | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 315 | apply (frule set_times_mono) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 316 | apply auto | 
| 19736 | 317 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 318 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 319 | lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C \<otimes> E ==> | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 320 | x : D \<otimes> F" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 321 | apply (frule set_times_mono2) | 
| 19736 | 322 | prefer 2 | 
| 323 | apply force | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 324 | apply assumption | 
| 19736 | 325 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 326 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 327 | lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C \<otimes> D" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 328 | apply (frule set_times_mono3) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 329 | apply auto | 
| 19736 | 330 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 331 | |
| 19736 | 332 | lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==> | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 333 | x : a *o D ==> x : D \<otimes> C" | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 334 | apply (frule set_times_mono4) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 335 | apply auto | 
| 19736 | 336 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 337 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 338 | lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" | 
| 19736 | 339 | 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 | 340 | |
| 19736 | 341 | lemma set_times_plus_distrib: "(a::'a::semiring) *o (b +o C)= | 
| 342 | (a * b) +o (a *o C)" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
21404diff
changeset | 343 | 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 | 344 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 345 | lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B \<oplus> C) = | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 346 | (a *o B) \<oplus> (a *o C)" | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 347 | apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) | 
| 19736 | 348 | apply blast | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 349 | apply (rule_tac x = "b + bb" in exI) | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
21404diff
changeset | 350 | apply (auto simp add: ring_distribs) | 
| 19736 | 351 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 352 | |
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 353 | lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) \<otimes> D <= | 
| 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 354 | a *o D \<oplus> C \<otimes> D" | 
| 19736 | 355 | apply (auto intro!: subsetI simp add: | 
| 26814 
b3e8d5ec721d
Replaced + and * on sets by \<oplus> and \<otimes>, to avoid clash with
 berghofe parents: 
25764diff
changeset | 356 | 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: 
25764diff
changeset | 357 | set_plus_def ring_distribs) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 358 | apply auto | 
| 19736 | 359 | done | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 360 | |
| 19380 | 361 | theorems set_times_plus_distribs = | 
| 362 | set_times_plus_distrib | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 363 | set_times_plus_distrib2 | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 364 | |
| 19736 | 365 | lemma set_neg_intro: "(a::'a::ring_1) : (- 1) *o C ==> | 
| 366 | - a : C" | |
| 367 | 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 | 368 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 369 | lemma set_neg_intro2: "(a::'a::ring_1) : C ==> | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 370 | - a : (- 1) *o C" | 
| 19736 | 371 | by (auto simp add: elt_set_times_def) | 
| 372 | ||
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 373 | end |