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