| author | nipkow | 
| Mon, 29 Jul 2024 15:26:03 +0200 | |
| changeset 80624 | 9f8034d29365 | 
| parent 73296 | 2ac92ba88d6b | 
| permissions | -rw-r--r-- | 
| 38622 | 1 | (* Title: HOL/Library/Function_Algebras.thy | 
| 2 | Author: Jeremy Avigad and Kevin Donnelly; Florian Haftmann, TUM | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 3 | *) | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 4 | |
| 60500 | 5 | section \<open>Pointwise instantiation of functions to algebra type classes\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 6 | |
| 38622 | 7 | theory Function_Algebras | 
| 30738 | 8 | imports Main | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 9 | begin | 
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 10 | |
| 60500 | 11 | text \<open>Pointwise operations\<close> | 
| 25594 | 12 | |
| 13 | instantiation "fun" :: (type, plus) plus | |
| 14 | begin | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 15 | |
| 46575 | 16 | definition "f + g = (\<lambda>x. f x + g x)" | 
| 25594 | 17 | instance .. | 
| 18 | ||
| 19 | end | |
| 20 | ||
| 48173 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 21 | lemma plus_fun_apply [simp]: | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 22 | "(f + g) x = f x + g x" | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 23 | by (simp add: plus_fun_def) | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 24 | |
| 38622 | 25 | instantiation "fun" :: (type, zero) zero | 
| 26 | begin | |
| 27 | ||
| 46575 | 28 | definition "0 = (\<lambda>x. 0)" | 
| 38622 | 29 | instance .. | 
| 30 | ||
| 31 | end | |
| 25594 | 32 | |
| 48173 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 33 | lemma zero_fun_apply [simp]: | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 34 | "0 x = 0" | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 35 | by (simp add: zero_fun_def) | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 36 | |
| 25594 | 37 | instantiation "fun" :: (type, times) times | 
| 38 | begin | |
| 39 | ||
| 46575 | 40 | definition "f * g = (\<lambda>x. f x * g x)" | 
| 25594 | 41 | instance .. | 
| 42 | ||
| 43 | end | |
| 44 | ||
| 48173 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 45 | lemma times_fun_apply [simp]: | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 46 | "(f * g) x = f x * g x" | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 47 | by (simp add: times_fun_def) | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 48 | |
| 25594 | 49 | instantiation "fun" :: (type, one) one | 
| 50 | begin | |
| 51 | ||
| 46575 | 52 | definition "1 = (\<lambda>x. 1)" | 
| 25594 | 53 | instance .. | 
| 54 | ||
| 55 | end | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 56 | |
| 48173 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 57 | lemma one_fun_apply [simp]: | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 58 | "1 x = 1" | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 59 | by (simp add: one_fun_def) | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 60 | |
| 38622 | 61 | |
| 60500 | 62 | text \<open>Additive structures\<close> | 
| 38622 | 63 | |
| 46575 | 64 | instance "fun" :: (type, semigroup_add) semigroup_add | 
| 60679 | 65 | by standard (simp add: fun_eq_iff add.assoc) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 66 | |
| 46575 | 67 | instance "fun" :: (type, cancel_semigroup_add) cancel_semigroup_add | 
| 60679 | 68 | by standard (simp_all add: fun_eq_iff) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 69 | |
| 46575 | 70 | instance "fun" :: (type, ab_semigroup_add) ab_semigroup_add | 
| 60679 | 71 | by standard (simp add: fun_eq_iff add.commute) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 72 | |
| 46575 | 73 | instance "fun" :: (type, cancel_ab_semigroup_add) cancel_ab_semigroup_add | 
| 60679 | 74 | by standard (simp_all add: fun_eq_iff diff_diff_eq) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 75 | |
| 46575 | 76 | instance "fun" :: (type, monoid_add) monoid_add | 
| 60679 | 77 | by standard (simp_all add: fun_eq_iff) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 78 | |
| 46575 | 79 | instance "fun" :: (type, comm_monoid_add) comm_monoid_add | 
| 60679 | 80 | by standard simp | 
| 38622 | 81 | |
| 82 | instance "fun" :: (type, cancel_comm_monoid_add) cancel_comm_monoid_add .. | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 83 | |
| 46575 | 84 | instance "fun" :: (type, group_add) group_add | 
| 60679 | 85 | by standard (simp_all add: fun_eq_iff) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 86 | |
| 46575 | 87 | instance "fun" :: (type, ab_group_add) ab_group_add | 
| 60679 | 88 | by standard simp_all | 
| 38622 | 89 | |
| 90 | ||
| 60500 | 91 | text \<open>Multiplicative structures\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 92 | |
| 46575 | 93 | instance "fun" :: (type, semigroup_mult) semigroup_mult | 
| 60679 | 94 | by standard (simp add: fun_eq_iff mult.assoc) | 
| 38622 | 95 | |
| 46575 | 96 | instance "fun" :: (type, ab_semigroup_mult) ab_semigroup_mult | 
| 60679 | 97 | by standard (simp add: fun_eq_iff mult.commute) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 98 | |
| 46575 | 99 | instance "fun" :: (type, monoid_mult) monoid_mult | 
| 60679 | 100 | by standard (simp_all add: fun_eq_iff) | 
| 38622 | 101 | |
| 46575 | 102 | instance "fun" :: (type, comm_monoid_mult) comm_monoid_mult | 
| 60679 | 103 | by standard simp | 
| 38622 | 104 | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 105 | |
| 60500 | 106 | text \<open>Misc\<close> | 
| 38622 | 107 | |
| 108 | instance "fun" :: (type, "Rings.dvd") "Rings.dvd" .. | |
| 109 | ||
| 46575 | 110 | instance "fun" :: (type, mult_zero) mult_zero | 
| 60679 | 111 | by standard (simp_all add: fun_eq_iff) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 112 | |
| 46575 | 113 | instance "fun" :: (type, zero_neq_one) zero_neq_one | 
| 60679 | 114 | by standard (simp add: fun_eq_iff) | 
| 19736 | 115 | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 116 | |
| 60500 | 117 | text \<open>Ring structures\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 118 | |
| 46575 | 119 | instance "fun" :: (type, semiring) semiring | 
| 60679 | 120 | by standard (simp_all add: fun_eq_iff algebra_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 121 | |
| 46575 | 122 | instance "fun" :: (type, comm_semiring) comm_semiring | 
| 60679 | 123 | by standard (simp add: fun_eq_iff algebra_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 124 | |
| 38622 | 125 | instance "fun" :: (type, semiring_0) semiring_0 .. | 
| 126 | ||
| 127 | instance "fun" :: (type, comm_semiring_0) comm_semiring_0 .. | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 128 | |
| 38622 | 129 | instance "fun" :: (type, semiring_0_cancel) semiring_0_cancel .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 130 | |
| 38622 | 131 | instance "fun" :: (type, comm_semiring_0_cancel) comm_semiring_0_cancel .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 132 | |
| 38622 | 133 | instance "fun" :: (type, semiring_1) semiring_1 .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 134 | |
| 73296 | 135 | lemma numeral_fun: \<^marker>\<open>contributor \<open>Akihisa Yamada\<close>\<close> | 
| 136 | \<open>numeral n = (\<lambda>x::'a. numeral n)\<close> | |
| 137 | by (induction n) (simp_all only: numeral.simps plus_fun_def, simp_all) | |
| 138 | ||
| 139 | lemma numeral_fun_apply [simp]: \<^marker>\<open>contributor \<open>Akihisa Yamada\<close>\<close> | |
| 140 | \<open>numeral n x = numeral n\<close> | |
| 141 | by (simp add: numeral_fun) | |
| 142 | ||
| 46575 | 143 | lemma of_nat_fun: "of_nat n = (\<lambda>x::'a. of_nat n)" | 
| 38622 | 144 | proof - | 
| 145 | have comp: "comp = (\<lambda>f g x. f (g x))" | |
| 146 | by (rule ext)+ simp | |
| 147 | have plus_fun: "plus = (\<lambda>f g x. f x + g x)" | |
| 148 | by (rule ext, rule ext) (fact plus_fun_def) | |
| 149 | have "of_nat n = (comp (plus (1::'b)) ^^ n) (\<lambda>x::'a. 0)" | |
| 150 | by (simp add: of_nat_def plus_fun zero_fun_def one_fun_def comp) | |
| 151 | also have "... = comp ((plus 1) ^^ n) (\<lambda>x::'a. 0)" | |
| 152 | by (simp only: comp_funpow) | |
| 153 | finally show ?thesis by (simp add: of_nat_def comp) | |
| 154 | qed | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 155 | |
| 48173 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 156 | lemma of_nat_fun_apply [simp]: | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 157 | "of_nat n x = of_nat n" | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 158 | by (simp add: of_nat_fun) | 
| 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 159 | |
| 38622 | 160 | instance "fun" :: (type, comm_semiring_1) comm_semiring_1 .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 161 | |
| 38622 | 162 | instance "fun" :: (type, semiring_1_cancel) semiring_1_cancel .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 163 | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 164 | instance "fun" :: (type, comm_semiring_1_cancel) comm_semiring_1_cancel | 
| 60679 | 165 | by standard (auto simp add: times_fun_def algebra_simps) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 166 | |
| 46575 | 167 | instance "fun" :: (type, semiring_char_0) semiring_char_0 | 
| 168 | proof | |
| 38622 | 169 | from inj_of_nat have "inj (\<lambda>n (x::'a). of_nat n :: 'b)" | 
| 170 | by (rule inj_fun) | |
| 171 | then have "inj (\<lambda>n. of_nat n :: 'a \<Rightarrow> 'b)" | |
| 172 | by (simp add: of_nat_fun) | |
| 173 | then show "inj (of_nat :: nat \<Rightarrow> 'a \<Rightarrow> 'b)" . | |
| 174 | qed | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 175 | |
| 38622 | 176 | instance "fun" :: (type, ring) ring .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 177 | |
| 38622 | 178 | instance "fun" :: (type, comm_ring) comm_ring .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 179 | |
| 38622 | 180 | instance "fun" :: (type, ring_1) ring_1 .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 181 | |
| 38622 | 182 | instance "fun" :: (type, comm_ring_1) comm_ring_1 .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 183 | |
| 38622 | 184 | instance "fun" :: (type, ring_char_0) ring_char_0 .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 185 | |
| 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 186 | |
| 60500 | 187 | text \<open>Ordered structures\<close> | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 188 | |
| 46575 | 189 | instance "fun" :: (type, ordered_ab_semigroup_add) ordered_ab_semigroup_add | 
| 60679 | 190 | by standard (auto simp add: le_fun_def intro: add_left_mono) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 191 | |
| 38622 | 192 | instance "fun" :: (type, ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 193 | |
| 46575 | 194 | instance "fun" :: (type, ordered_ab_semigroup_add_imp_le) ordered_ab_semigroup_add_imp_le | 
| 60679 | 195 | by standard (simp add: le_fun_def) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 196 | |
| 38622 | 197 | instance "fun" :: (type, ordered_comm_monoid_add) ordered_comm_monoid_add .. | 
| 198 | ||
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 199 | instance "fun" :: (type, ordered_cancel_comm_monoid_add) ordered_cancel_comm_monoid_add .. | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 200 | |
| 38622 | 201 | instance "fun" :: (type, ordered_ab_group_add) ordered_ab_group_add .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 202 | |
| 46575 | 203 | instance "fun" :: (type, ordered_semiring) ordered_semiring | 
| 60679 | 204 | by standard (auto simp add: le_fun_def intro: mult_left_mono mult_right_mono) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 205 | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 206 | instance "fun" :: (type, dioid) dioid | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 207 | proof standard | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 208 | fix a b :: "'a \<Rightarrow> 'b" | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 209 | show "a \<le> b \<longleftrightarrow> (\<exists>c. b = a + c)" | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 210 | unfolding le_fun_def plus_fun_def fun_eq_iff choice_iff[symmetric, of "\<lambda>x c. b x = a x + c"] | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 211 | by (intro arg_cong[where f=All] ext canonically_ordered_monoid_add_class.le_iff_add) | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 212 | qed | 
| 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
60679diff
changeset | 213 | |
| 46575 | 214 | instance "fun" :: (type, ordered_comm_semiring) ordered_comm_semiring | 
| 60679 | 215 | by standard (fact mult_left_mono) | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 216 | |
| 38622 | 217 | instance "fun" :: (type, ordered_cancel_semiring) ordered_cancel_semiring .. | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 218 | |
| 38622 | 219 | instance "fun" :: (type, ordered_cancel_comm_semiring) ordered_cancel_comm_semiring .. | 
| 220 | ||
| 221 | instance "fun" :: (type, ordered_ring) ordered_ring .. | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 222 | |
| 38622 | 223 | instance "fun" :: (type, ordered_comm_ring) ordered_comm_ring .. | 
| 224 | ||
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 225 | |
| 38622 | 226 | lemmas func_plus = plus_fun_def | 
| 227 | lemmas func_zero = zero_fun_def | |
| 228 | lemmas func_times = times_fun_def | |
| 229 | lemmas func_one = one_fun_def | |
| 19736 | 230 | |
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: diff
changeset | 231 | end | 
| 48173 
c6a5a4336edf
eta-expanded occurences of algebraic functionals are simplified by default
 haftmann parents: 
46575diff
changeset | 232 |