| author | wenzelm | 
| Fri, 20 Oct 2023 22:19:05 +0200 | |
| changeset 78805 | 62616d8422c5 | 
| parent 70219 | b21efbf64292 | 
| child 79541 | 4f40225936d1 | 
| permissions | -rw-r--r-- | 
| 62479 | 1 | (* Title: HOL/Nonstandard_Analysis/StarDef.thy | 
| 2 | Author: Jacques D. Fleuriot and Brian Huffman | |
| 27468 | 3 | *) | 
| 4 | ||
| 61975 | 5 | section \<open>Construction of Star Types Using Ultrafilters\<close> | 
| 27468 | 6 | |
| 7 | theory StarDef | |
| 64435 | 8 | imports Free_Ultrafilter | 
| 27468 | 9 | begin | 
| 10 | ||
| 61975 | 11 | subsection \<open>A Free Ultrafilter over the Naturals\<close> | 
| 27468 | 12 | |
| 69597 | 13 | definition FreeUltrafilterNat :: "nat filter" (\<open>\<U>\<close>) | 
| 64435 | 14 | where "\<U> = (SOME U. freeultrafilter U)" | 
| 27468 | 15 | |
| 16 | lemma freeultrafilter_FreeUltrafilterNat: "freeultrafilter \<U>" | |
| 70219 | 17 | unfolding FreeUltrafilterNat_def | 
| 18 | by (simp add: freeultrafilter_Ex someI_ex) | |
| 27468 | 19 | |
| 64438 | 20 | interpretation FreeUltrafilterNat: freeultrafilter \<U> | 
| 64435 | 21 | by (rule freeultrafilter_FreeUltrafilterNat) | 
| 22 | ||
| 27468 | 23 | |
| 61975 | 24 | subsection \<open>Definition of \<open>star\<close> type constructor\<close> | 
| 27468 | 25 | |
| 64435 | 26 | definition starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set" | 
| 27 |   where "starrel = {(X, Y). eventually (\<lambda>n. X n = Y n) \<U>}"
 | |
| 27468 | 28 | |
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 29 | definition "star = (UNIV :: (nat \<Rightarrow> 'a) set) // starrel" | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 30 | |
| 49834 | 31 | typedef 'a star = "star :: (nat \<Rightarrow> 'a) set set" | 
| 64435 | 32 | by (auto simp: star_def intro: quotientI) | 
| 27468 | 33 | |
| 64435 | 34 | definition star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star" | 
| 35 |   where "star_n X = Abs_star (starrel `` {X})"
 | |
| 27468 | 36 | |
| 37 | theorem star_cases [case_names star_n, cases type: star]: | |
| 64435 | 38 | obtains X where "x = star_n X" | 
| 39 | by (cases x) (auto simp: star_n_def star_def elim: quotientE) | |
| 27468 | 40 | |
| 64435 | 41 | lemma all_star_eq: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>X. P (star_n X))" | 
| 70219 | 42 | by (metis star_cases) | 
| 27468 | 43 | |
| 64435 | 44 | lemma ex_star_eq: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>X. P (star_n X))" | 
| 70219 | 45 | by (metis star_cases) | 
| 27468 | 46 | |
| 69597 | 47 | text \<open>Proving that \<^term>\<open>starrel\<close> is an equivalence relation.\<close> | 
| 27468 | 48 | |
| 64435 | 49 | lemma starrel_iff [iff]: "(X, Y) \<in> starrel \<longleftrightarrow> eventually (\<lambda>n. X n = Y n) \<U>" | 
| 50 | by (simp add: starrel_def) | |
| 27468 | 51 | |
| 52 | lemma equiv_starrel: "equiv UNIV starrel" | |
| 40815 | 53 | proof (rule equivI) | 
| 30198 | 54 | show "refl starrel" by (simp add: refl_on_def) | 
| 27468 | 55 | show "sym starrel" by (simp add: sym_def eq_commute) | 
| 60041 | 56 | show "trans starrel" by (intro transI) (auto elim: eventually_elim2) | 
| 27468 | 57 | qed | 
| 58 | ||
| 64435 | 59 | lemmas equiv_starrel_iff = eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I] | 
| 27468 | 60 | |
| 61 | lemma starrel_in_star: "starrel``{x} \<in> star"
 | |
| 64435 | 62 | by (simp add: star_def quotientI) | 
| 27468 | 63 | |
| 64435 | 64 | lemma star_n_eq_iff: "star_n X = star_n Y \<longleftrightarrow> eventually (\<lambda>n. X n = Y n) \<U>" | 
| 65 | by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff) | |
| 27468 | 66 | |
| 67 | ||
| 61975 | 68 | subsection \<open>Transfer principle\<close> | 
| 27468 | 69 | |
| 61975 | 70 | text \<open>This introduction rule starts each transfer proof.\<close> | 
| 64435 | 71 | lemma transfer_start: "P \<equiv> eventually (\<lambda>n. Q) \<U> \<Longrightarrow> Trueprop P \<equiv> Trueprop Q" | 
| 60041 | 72 | by (simp add: FreeUltrafilterNat.proper) | 
| 27468 | 73 | |
| 64270 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 Simon Wimmer <wimmers@in.tum.de> parents: 
64242diff
changeset | 74 | text \<open>Standard principles that play a central role in the transfer tactic.\<close> | 
| 69597 | 75 | definition Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" (\<open>(_ \<star>/ _)\<close> [300, 301] 300)
 | 
| 64435 | 76 | where "Ifun f \<equiv> | 
| 77 |     \<lambda>x. Abs_star (\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})"
 | |
| 64270 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 Simon Wimmer <wimmers@in.tum.de> parents: 
64242diff
changeset | 78 | |
| 64435 | 79 | lemma Ifun_congruent2: "congruent2 starrel starrel (\<lambda>F X. starrel``{\<lambda>n. F n (X n)})"
 | 
| 80 | by (auto simp add: congruent2_def equiv_starrel_iff elim!: eventually_rev_mp) | |
| 64270 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 Simon Wimmer <wimmers@in.tum.de> parents: 
64242diff
changeset | 81 | |
| 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 Simon Wimmer <wimmers@in.tum.de> parents: 
64242diff
changeset | 82 | lemma Ifun_star_n: "star_n F \<star> star_n X = star_n (\<lambda>n. F n (X n))" | 
| 64435 | 83 | by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star | 
| 84 | UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2]) | |
| 64270 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 Simon Wimmer <wimmers@in.tum.de> parents: 
64242diff
changeset | 85 | |
| 64435 | 86 | lemma transfer_Ifun: "f \<equiv> star_n F \<Longrightarrow> x \<equiv> star_n X \<Longrightarrow> f \<star> x \<equiv> star_n (\<lambda>n. F n (X n))" | 
| 87 | by (simp only: Ifun_star_n) | |
| 64270 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 Simon Wimmer <wimmers@in.tum.de> parents: 
64242diff
changeset | 88 | |
| 64435 | 89 | definition star_of :: "'a \<Rightarrow> 'a star" | 
| 90 | where "star_of x \<equiv> star_n (\<lambda>n. x)" | |
| 64270 
bf474d719011
Modified transfer principle in HOL/NSA to cause less ho-unficiation
 Simon Wimmer <wimmers@in.tum.de> parents: 
64242diff
changeset | 91 | |
| 61975 | 92 | text \<open>Initialize transfer tactic.\<close> | 
| 69605 | 93 | ML_file \<open>transfer_principle.ML\<close> | 
| 27468 | 94 | |
| 64435 | 95 | method_setup transfer = | 
| 96 | \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (Transfer_Principle.transfer_tac ctxt ths))\<close> | |
| 97 | "transfer principle" | |
| 47432 | 98 | |
| 99 | ||
| 61975 | 100 | text \<open>Transfer introduction rules.\<close> | 
| 27468 | 101 | |
| 102 | lemma transfer_ex [transfer_intro]: | |
| 64435 | 103 | "(\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow> | 
| 104 | \<exists>x::'a star. p x \<equiv> eventually (\<lambda>n. \<exists>x. P n x) \<U>" | |
| 105 | by (simp only: ex_star_eq eventually_ex) | |
| 27468 | 106 | |
| 107 | lemma transfer_all [transfer_intro]: | |
| 64435 | 108 | "(\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow> | 
| 109 | \<forall>x::'a star. p x \<equiv> eventually (\<lambda>n. \<forall>x. P n x) \<U>" | |
| 110 | by (simp only: all_star_eq FreeUltrafilterNat.eventually_all_iff) | |
| 27468 | 111 | |
| 64435 | 112 | lemma transfer_not [transfer_intro]: "p \<equiv> eventually P \<U> \<Longrightarrow> \<not> p \<equiv> eventually (\<lambda>n. \<not> P n) \<U>" | 
| 113 | by (simp only: FreeUltrafilterNat.eventually_not_iff) | |
| 27468 | 114 | |
| 115 | lemma transfer_conj [transfer_intro]: | |
| 64435 | 116 | "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p \<and> q \<equiv> eventually (\<lambda>n. P n \<and> Q n) \<U>" | 
| 117 | by (simp only: eventually_conj_iff) | |
| 27468 | 118 | |
| 119 | lemma transfer_disj [transfer_intro]: | |
| 64435 | 120 | "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p \<or> q \<equiv> eventually (\<lambda>n. P n \<or> Q n) \<U>" | 
| 121 | by (simp only: FreeUltrafilterNat.eventually_disj_iff) | |
| 27468 | 122 | |
| 123 | lemma transfer_imp [transfer_intro]: | |
| 64435 | 124 | "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p \<longrightarrow> q \<equiv> eventually (\<lambda>n. P n \<longrightarrow> Q n) \<U>" | 
| 125 | by (simp only: FreeUltrafilterNat.eventually_imp_iff) | |
| 27468 | 126 | |
| 127 | lemma transfer_iff [transfer_intro]: | |
| 64435 | 128 | "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p = q \<equiv> eventually (\<lambda>n. P n = Q n) \<U>" | 
| 129 | by (simp only: FreeUltrafilterNat.eventually_iff_iff) | |
| 27468 | 130 | |
| 131 | lemma transfer_if_bool [transfer_intro]: | |
| 64435 | 132 | "p \<equiv> eventually P \<U> \<Longrightarrow> x \<equiv> eventually X \<U> \<Longrightarrow> y \<equiv> eventually Y \<U> \<Longrightarrow> | 
| 133 | (if p then x else y) \<equiv> eventually (\<lambda>n. if P n then X n else Y n) \<U>" | |
| 134 | by (simp only: if_bool_eq_conj transfer_conj transfer_imp transfer_not) | |
| 27468 | 135 | |
| 136 | lemma transfer_eq [transfer_intro]: | |
| 64435 | 137 | "x \<equiv> star_n X \<Longrightarrow> y \<equiv> star_n Y \<Longrightarrow> x = y \<equiv> eventually (\<lambda>n. X n = Y n) \<U>" | 
| 138 | by (simp only: star_n_eq_iff) | |
| 27468 | 139 | |
| 140 | lemma transfer_if [transfer_intro]: | |
| 64435 | 141 | "p \<equiv> eventually (\<lambda>n. P n) \<U> \<Longrightarrow> x \<equiv> star_n X \<Longrightarrow> y \<equiv> star_n Y \<Longrightarrow> | 
| 142 | (if p then x else y) \<equiv> star_n (\<lambda>n. if P n then X n else Y n)" | |
| 143 | by (rule eq_reflection) (auto simp: star_n_eq_iff transfer_not elim!: eventually_mono) | |
| 27468 | 144 | |
| 145 | lemma transfer_fun_eq [transfer_intro]: | |
| 64435 | 146 | "(\<And>X. f (star_n X) = g (star_n X) \<equiv> eventually (\<lambda>n. F n (X n) = G n (X n)) \<U>) \<Longrightarrow> | 
| 147 | f = g \<equiv> eventually (\<lambda>n. F n = G n) \<U>" | |
| 148 | by (simp only: fun_eq_iff transfer_all) | |
| 27468 | 149 | |
| 150 | lemma transfer_star_n [transfer_intro]: "star_n X \<equiv> star_n (\<lambda>n. X n)" | |
| 64435 | 151 | by (rule reflexive) | 
| 27468 | 152 | |
| 60041 | 153 | lemma transfer_bool [transfer_intro]: "p \<equiv> eventually (\<lambda>n. p) \<U>" | 
| 64435 | 154 | by (simp add: FreeUltrafilterNat.proper) | 
| 27468 | 155 | |
| 156 | ||
| 61975 | 157 | subsection \<open>Standard elements\<close> | 
| 27468 | 158 | |
| 64435 | 159 | definition Standard :: "'a star set" | 
| 160 | where "Standard = range star_of" | |
| 27468 | 161 | |
| 69597 | 162 | text \<open>Transfer tactic should remove occurrences of \<^term>\<open>star_of\<close>.\<close> | 
| 163 | setup \<open>Transfer_Principle.add_const \<^const_name>\<open>star_of\<close>\<close> | |
| 27468 | 164 | |
| 64435 | 165 | lemma star_of_inject: "star_of x = star_of y \<longleftrightarrow> x = y" | 
| 166 | by transfer (rule refl) | |
| 27468 | 167 | |
| 168 | lemma Standard_star_of [simp]: "star_of x \<in> Standard" | |
| 64435 | 169 | by (simp add: Standard_def) | 
| 170 | ||
| 27468 | 171 | |
| 61975 | 172 | subsection \<open>Internal functions\<close> | 
| 27468 | 173 | |
| 69597 | 174 | text \<open>Transfer tactic should remove occurrences of \<^term>\<open>Ifun\<close>.\<close> | 
| 175 | setup \<open>Transfer_Principle.add_const \<^const_name>\<open>Ifun\<close>\<close> | |
| 27468 | 176 | |
| 177 | lemma Ifun_star_of [simp]: "star_of f \<star> star_of x = star_of (f x)" | |
| 64435 | 178 | by transfer (rule refl) | 
| 27468 | 179 | |
| 64435 | 180 | lemma Standard_Ifun [simp]: "f \<in> Standard \<Longrightarrow> x \<in> Standard \<Longrightarrow> f \<star> x \<in> Standard" | 
| 181 | by (auto simp add: Standard_def) | |
| 27468 | 182 | |
| 183 | ||
| 64435 | 184 | text \<open>Nonstandard extensions of functions.\<close> | 
| 27468 | 185 | |
| 69597 | 186 | definition starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a star \<Rightarrow> 'b star"  (\<open>*f* _\<close> [80] 80)
 | 
| 64435 | 187 | where "starfun f \<equiv> \<lambda>x. star_of f \<star> x" | 
| 188 | ||
| 69597 | 189 | definition starfun2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> 'c star"  (\<open>*f2* _\<close> [80] 80)
 | 
| 64435 | 190 | where "starfun2 f \<equiv> \<lambda>x y. star_of f \<star> x \<star> y" | 
| 27468 | 191 | |
| 192 | declare starfun_def [transfer_unfold] | |
| 193 | declare starfun2_def [transfer_unfold] | |
| 194 | ||
| 195 | lemma starfun_star_n: "( *f* f) (star_n X) = star_n (\<lambda>n. f (X n))" | |
| 64435 | 196 | by (simp only: starfun_def star_of_def Ifun_star_n) | 
| 27468 | 197 | |
| 64435 | 198 | lemma starfun2_star_n: "( *f2* f) (star_n X) (star_n Y) = star_n (\<lambda>n. f (X n) (Y n))" | 
| 199 | by (simp only: starfun2_def star_of_def Ifun_star_n) | |
| 27468 | 200 | |
| 201 | lemma starfun_star_of [simp]: "( *f* f) (star_of x) = star_of (f x)" | |
| 64435 | 202 | by transfer (rule refl) | 
| 27468 | 203 | |
| 204 | lemma starfun2_star_of [simp]: "( *f2* f) (star_of x) = *f* f x" | |
| 64435 | 205 | by transfer (rule refl) | 
| 27468 | 206 | |
| 207 | lemma Standard_starfun [simp]: "x \<in> Standard \<Longrightarrow> starfun f x \<in> Standard" | |
| 64435 | 208 | by (simp add: starfun_def) | 
| 27468 | 209 | |
| 64435 | 210 | lemma Standard_starfun2 [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> starfun2 f x y \<in> Standard" | 
| 211 | by (simp add: starfun2_def) | |
| 27468 | 212 | |
| 213 | lemma Standard_starfun_iff: | |
| 214 | assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y" | |
| 64435 | 215 | shows "starfun f x \<in> Standard \<longleftrightarrow> x \<in> Standard" | 
| 27468 | 216 | proof | 
| 217 | assume "x \<in> Standard" | |
| 64435 | 218 | then show "starfun f x \<in> Standard" by simp | 
| 27468 | 219 | next | 
| 64435 | 220 | from inj have inj': "\<And>x y. starfun f x = starfun f y \<Longrightarrow> x = y" | 
| 221 | by transfer | |
| 27468 | 222 | assume "starfun f x \<in> Standard" | 
| 223 | then obtain b where b: "starfun f x = star_of b" | |
| 224 | unfolding Standard_def .. | |
| 64435 | 225 | then have "\<exists>x. starfun f x = star_of b" .. | 
| 226 | then have "\<exists>a. f a = b" by transfer | |
| 27468 | 227 | then obtain a where "f a = b" .. | 
| 64435 | 228 | then have "starfun f (star_of a) = star_of b" by transfer | 
| 27468 | 229 | with b have "starfun f x = starfun f (star_of a)" by simp | 
| 64435 | 230 | then have "x = star_of a" by (rule inj') | 
| 231 | then show "x \<in> Standard" by (simp add: Standard_def) | |
| 27468 | 232 | qed | 
| 233 | ||
| 234 | lemma Standard_starfun2_iff: | |
| 235 | assumes inj: "\<And>a b a' b'. f a b = f a' b' \<Longrightarrow> a = a' \<and> b = b'" | |
| 64435 | 236 | shows "starfun2 f x y \<in> Standard \<longleftrightarrow> x \<in> Standard \<and> y \<in> Standard" | 
| 27468 | 237 | proof | 
| 238 | assume "x \<in> Standard \<and> y \<in> Standard" | |
| 64435 | 239 | then show "starfun2 f x y \<in> Standard" by simp | 
| 27468 | 240 | next | 
| 241 | have inj': "\<And>x y z w. starfun2 f x y = starfun2 f z w \<Longrightarrow> x = z \<and> y = w" | |
| 242 | using inj by transfer | |
| 243 | assume "starfun2 f x y \<in> Standard" | |
| 244 | then obtain c where c: "starfun2 f x y = star_of c" | |
| 245 | unfolding Standard_def .. | |
| 64435 | 246 | then have "\<exists>x y. starfun2 f x y = star_of c" by auto | 
| 247 | then have "\<exists>a b. f a b = c" by transfer | |
| 27468 | 248 | then obtain a b where "f a b = c" by auto | 
| 64435 | 249 | then have "starfun2 f (star_of a) (star_of b) = star_of c" by transfer | 
| 250 | with c have "starfun2 f x y = starfun2 f (star_of a) (star_of b)" by simp | |
| 251 | then have "x = star_of a \<and> y = star_of b" by (rule inj') | |
| 252 | then show "x \<in> Standard \<and> y \<in> Standard" by (simp add: Standard_def) | |
| 27468 | 253 | qed | 
| 254 | ||
| 255 | ||
| 61975 | 256 | subsection \<open>Internal predicates\<close> | 
| 27468 | 257 | |
| 64435 | 258 | definition unstar :: "bool star \<Rightarrow> bool" | 
| 259 | where "unstar b \<longleftrightarrow> b = star_of True" | |
| 27468 | 260 | |
| 64435 | 261 | lemma unstar_star_n: "unstar (star_n P) \<longleftrightarrow> eventually P \<U>" | 
| 262 | by (simp add: unstar_def star_of_def star_n_eq_iff) | |
| 27468 | 263 | |
| 264 | lemma unstar_star_of [simp]: "unstar (star_of p) = p" | |
| 64435 | 265 | by (simp add: unstar_def star_of_inject) | 
| 27468 | 266 | |
| 69597 | 267 | text \<open>Transfer tactic should remove occurrences of \<^term>\<open>unstar\<close>.\<close> | 
| 268 | setup \<open>Transfer_Principle.add_const \<^const_name>\<open>unstar\<close>\<close> | |
| 27468 | 269 | |
| 64435 | 270 | lemma transfer_unstar [transfer_intro]: "p \<equiv> star_n P \<Longrightarrow> unstar p \<equiv> eventually P \<U>" | 
| 271 | by (simp only: unstar_star_n) | |
| 27468 | 272 | |
| 69597 | 273 | definition starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool"  (\<open>*p* _\<close> [80] 80)
 | 
| 64435 | 274 | where "*p* P = (\<lambda>x. unstar (star_of P \<star> x))" | 
| 27468 | 275 | |
| 69597 | 276 | definition starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool"  (\<open>*p2* _\<close> [80] 80)
 | 
| 64435 | 277 | where "*p2* P = (\<lambda>x y. unstar (star_of P \<star> x \<star> y))" | 
| 27468 | 278 | |
| 279 | declare starP_def [transfer_unfold] | |
| 280 | declare starP2_def [transfer_unfold] | |
| 281 | ||
| 64435 | 282 | lemma starP_star_n: "( *p* P) (star_n X) = eventually (\<lambda>n. P (X n)) \<U>" | 
| 283 | by (simp only: starP_def star_of_def Ifun_star_n unstar_star_n) | |
| 27468 | 284 | |
| 64435 | 285 | lemma starP2_star_n: "( *p2* P) (star_n X) (star_n Y) = (eventually (\<lambda>n. P (X n) (Y n)) \<U>)" | 
| 286 | by (simp only: starP2_def star_of_def Ifun_star_n unstar_star_n) | |
| 27468 | 287 | |
| 288 | lemma starP_star_of [simp]: "( *p* P) (star_of x) = P x" | |
| 64435 | 289 | by transfer (rule refl) | 
| 27468 | 290 | |
| 291 | lemma starP2_star_of [simp]: "( *p2* P) (star_of x) = *p* P x" | |
| 64435 | 292 | by transfer (rule refl) | 
| 27468 | 293 | |
| 294 | ||
| 61975 | 295 | subsection \<open>Internal sets\<close> | 
| 27468 | 296 | |
| 64435 | 297 | definition Iset :: "'a set star \<Rightarrow> 'a star set" | 
| 67399 | 298 |   where "Iset A = {x. ( *p2* (\<in>)) x A}"
 | 
| 27468 | 299 | |
| 64435 | 300 | lemma Iset_star_n: "(star_n X \<in> Iset (star_n A)) = (eventually (\<lambda>n. X n \<in> A n) \<U>)" | 
| 301 | by (simp add: Iset_def starP2_star_n) | |
| 27468 | 302 | |
| 69597 | 303 | text \<open>Transfer tactic should remove occurrences of \<^term>\<open>Iset\<close>.\<close> | 
| 304 | setup \<open>Transfer_Principle.add_const \<^const_name>\<open>Iset\<close>\<close> | |
| 27468 | 305 | |
| 306 | lemma transfer_mem [transfer_intro]: | |
| 64435 | 307 | "x \<equiv> star_n X \<Longrightarrow> a \<equiv> Iset (star_n A) \<Longrightarrow> x \<in> a \<equiv> eventually (\<lambda>n. X n \<in> A n) \<U>" | 
| 308 | by (simp only: Iset_star_n) | |
| 27468 | 309 | |
| 310 | lemma transfer_Collect [transfer_intro]: | |
| 64435 | 311 | "(\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow> | 
| 312 | Collect p \<equiv> Iset (star_n (\<lambda>n. Collect (P n)))" | |
| 313 | by (simp add: atomize_eq set_eq_iff all_star_eq Iset_star_n) | |
| 27468 | 314 | |
| 315 | lemma transfer_set_eq [transfer_intro]: | |
| 64435 | 316 | "a \<equiv> Iset (star_n A) \<Longrightarrow> b \<equiv> Iset (star_n B) \<Longrightarrow> a = b \<equiv> eventually (\<lambda>n. A n = B n) \<U>" | 
| 317 | by (simp only: set_eq_iff transfer_all transfer_iff transfer_mem) | |
| 27468 | 318 | |
| 319 | lemma transfer_ball [transfer_intro]: | |
| 64435 | 320 | "a \<equiv> Iset (star_n A) \<Longrightarrow> (\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow> | 
| 321 | \<forall>x\<in>a. p x \<equiv> eventually (\<lambda>n. \<forall>x\<in>A n. P n x) \<U>" | |
| 322 | by (simp only: Ball_def transfer_all transfer_imp transfer_mem) | |
| 27468 | 323 | |
| 324 | lemma transfer_bex [transfer_intro]: | |
| 64435 | 325 | "a \<equiv> Iset (star_n A) \<Longrightarrow> (\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow> | 
| 326 | \<exists>x\<in>a. p x \<equiv> eventually (\<lambda>n. \<exists>x\<in>A n. P n x) \<U>" | |
| 327 | by (simp only: Bex_def transfer_ex transfer_conj transfer_mem) | |
| 27468 | 328 | |
| 64435 | 329 | lemma transfer_Iset [transfer_intro]: "a \<equiv> star_n A \<Longrightarrow> Iset a \<equiv> Iset (star_n (\<lambda>n. A n))" | 
| 330 | by simp | |
| 331 | ||
| 27468 | 332 | |
| 61975 | 333 | text \<open>Nonstandard extensions of sets.\<close> | 
| 27468 | 334 | |
| 69597 | 335 | definition starset :: "'a set \<Rightarrow> 'a star set" (\<open>*s* _\<close> [80] 80) | 
| 64435 | 336 | where "starset A = Iset (star_of A)" | 
| 27468 | 337 | |
| 338 | declare starset_def [transfer_unfold] | |
| 339 | ||
| 64435 | 340 | lemma starset_mem: "star_of x \<in> *s* A \<longleftrightarrow> x \<in> A" | 
| 341 | by transfer (rule refl) | |
| 27468 | 342 | |
| 343 | lemma starset_UNIV: "*s* (UNIV::'a set) = (UNIV::'a star set)" | |
| 64435 | 344 | by (transfer UNIV_def) (rule refl) | 
| 27468 | 345 | |
| 346 | lemma starset_empty: "*s* {} = {}"
 | |
| 64435 | 347 | by (transfer empty_def) (rule refl) | 
| 27468 | 348 | |
| 349 | lemma starset_insert: "*s* (insert x A) = insert (star_of x) ( *s* A)" | |
| 64435 | 350 | by (transfer insert_def Un_def) (rule refl) | 
| 27468 | 351 | |
| 352 | lemma starset_Un: "*s* (A \<union> B) = *s* A \<union> *s* B" | |
| 64435 | 353 | by (transfer Un_def) (rule refl) | 
| 27468 | 354 | |
| 355 | lemma starset_Int: "*s* (A \<inter> B) = *s* A \<inter> *s* B" | |
| 64435 | 356 | by (transfer Int_def) (rule refl) | 
| 27468 | 357 | |
| 358 | lemma starset_Compl: "*s* -A = -( *s* A)" | |
| 64435 | 359 | by (transfer Compl_eq) (rule refl) | 
| 27468 | 360 | |
| 361 | lemma starset_diff: "*s* (A - B) = *s* A - *s* B" | |
| 64435 | 362 | by (transfer set_diff_eq) (rule refl) | 
| 27468 | 363 | |
| 364 | lemma starset_image: "*s* (f ` A) = ( *f* f) ` ( *s* A)" | |
| 64435 | 365 | by (transfer image_def) (rule refl) | 
| 27468 | 366 | |
| 367 | lemma starset_vimage: "*s* (f -` A) = ( *f* f) -` ( *s* A)" | |
| 64435 | 368 | by (transfer vimage_def) (rule refl) | 
| 27468 | 369 | |
| 64435 | 370 | lemma starset_subset: "( *s* A \<subseteq> *s* B) \<longleftrightarrow> A \<subseteq> B" | 
| 371 | by (transfer subset_eq) (rule refl) | |
| 27468 | 372 | |
| 64435 | 373 | lemma starset_eq: "( *s* A = *s* B) \<longleftrightarrow> A = B" | 
| 374 | by transfer (rule refl) | |
| 27468 | 375 | |
| 376 | lemmas starset_simps [simp] = | |
| 377 | starset_mem starset_UNIV | |
| 378 | starset_empty starset_insert | |
| 379 | starset_Un starset_Int | |
| 380 | starset_Compl starset_diff | |
| 381 | starset_image starset_vimage | |
| 382 | starset_subset starset_eq | |
| 383 | ||
| 384 | ||
| 61975 | 385 | subsection \<open>Syntactic classes\<close> | 
| 27468 | 386 | |
| 387 | instantiation star :: (zero) zero | |
| 388 | begin | |
| 64435 | 389 | definition star_zero_def: "0 \<equiv> star_of 0" | 
| 390 | instance .. | |
| 27468 | 391 | end | 
| 392 | ||
| 393 | instantiation star :: (one) one | |
| 394 | begin | |
| 64435 | 395 | definition star_one_def: "1 \<equiv> star_of 1" | 
| 396 | instance .. | |
| 27468 | 397 | end | 
| 398 | ||
| 399 | instantiation star :: (plus) plus | |
| 400 | begin | |
| 67399 | 401 | definition star_add_def: "(+) \<equiv> *f2* (+)" | 
| 64435 | 402 | instance .. | 
| 27468 | 403 | end | 
| 404 | ||
| 405 | instantiation star :: (times) times | |
| 406 | begin | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67635diff
changeset | 407 | definition star_mult_def: "((*)) \<equiv> *f2* ((*))" | 
| 64435 | 408 | instance .. | 
| 27468 | 409 | end | 
| 410 | ||
| 411 | instantiation star :: (uminus) uminus | |
| 412 | begin | |
| 64435 | 413 | definition star_minus_def: "uminus \<equiv> *f* uminus" | 
| 414 | instance .. | |
| 27468 | 415 | end | 
| 416 | ||
| 417 | instantiation star :: (minus) minus | |
| 418 | begin | |
| 67399 | 419 | definition star_diff_def: "(-) \<equiv> *f2* (-)" | 
| 64435 | 420 | instance .. | 
| 27468 | 421 | end | 
| 422 | ||
| 423 | instantiation star :: (abs) abs | |
| 424 | begin | |
| 64435 | 425 | definition star_abs_def: "abs \<equiv> *f* abs" | 
| 426 | instance .. | |
| 27468 | 427 | end | 
| 428 | ||
| 429 | instantiation star :: (sgn) sgn | |
| 430 | begin | |
| 64435 | 431 | definition star_sgn_def: "sgn \<equiv> *f* sgn" | 
| 432 | instance .. | |
| 27468 | 433 | end | 
| 434 | ||
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 435 | instantiation star :: (divide) divide | 
| 27468 | 436 | begin | 
| 64435 | 437 | definition star_divide_def: "divide \<equiv> *f2* divide" | 
| 438 | instance .. | |
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 439 | end | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 440 | |
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 441 | instantiation star :: (inverse) inverse | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 442 | begin | 
| 64435 | 443 | definition star_inverse_def: "inverse \<equiv> *f* inverse" | 
| 444 | instance .. | |
| 27468 | 445 | end | 
| 446 | ||
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
35043diff
changeset | 447 | instance star :: (Rings.dvd) Rings.dvd .. | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27468diff
changeset | 448 | |
| 63950 
cdc1e59aa513
syntactic type class for operation mod named after mod;
 haftmann parents: 
63456diff
changeset | 449 | instantiation star :: (modulo) modulo | 
| 27468 | 450 | begin | 
| 67399 | 451 | definition star_mod_def: "(mod) \<equiv> *f2* (mod)" | 
| 64435 | 452 | instance .. | 
| 27468 | 453 | end | 
| 454 | ||
| 455 | instantiation star :: (ord) ord | |
| 456 | begin | |
| 67399 | 457 | definition star_le_def: "(\<le>) \<equiv> *p2* (\<le>)" | 
| 458 | definition star_less_def: "(<) \<equiv> *p2* (<)" | |
| 64435 | 459 | instance .. | 
| 27468 | 460 | end | 
| 461 | ||
| 462 | lemmas star_class_defs [transfer_unfold] = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 463 | star_zero_def star_one_def | 
| 27468 | 464 | star_add_def star_diff_def star_minus_def | 
| 465 | star_mult_def star_divide_def star_inverse_def | |
| 466 | star_le_def star_less_def star_abs_def star_sgn_def | |
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 467 | star_mod_def | 
| 27468 | 468 | |
| 64435 | 469 | |
| 470 | text \<open>Class operations preserve standard elements.\<close> | |
| 27468 | 471 | |
| 472 | lemma Standard_zero: "0 \<in> Standard" | |
| 64435 | 473 | by (simp add: star_zero_def) | 
| 27468 | 474 | |
| 475 | lemma Standard_one: "1 \<in> Standard" | |
| 64435 | 476 | by (simp add: star_one_def) | 
| 27468 | 477 | |
| 64435 | 478 | lemma Standard_add: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x + y \<in> Standard" | 
| 479 | by (simp add: star_add_def) | |
| 27468 | 480 | |
| 64435 | 481 | lemma Standard_diff: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x - y \<in> Standard" | 
| 482 | by (simp add: star_diff_def) | |
| 27468 | 483 | |
| 484 | lemma Standard_minus: "x \<in> Standard \<Longrightarrow> - x \<in> Standard" | |
| 64435 | 485 | by (simp add: star_minus_def) | 
| 27468 | 486 | |
| 64435 | 487 | lemma Standard_mult: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x * y \<in> Standard" | 
| 488 | by (simp add: star_mult_def) | |
| 27468 | 489 | |
| 64435 | 490 | lemma Standard_divide: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x / y \<in> Standard" | 
| 491 | by (simp add: star_divide_def) | |
| 27468 | 492 | |
| 493 | lemma Standard_inverse: "x \<in> Standard \<Longrightarrow> inverse x \<in> Standard" | |
| 64435 | 494 | by (simp add: star_inverse_def) | 
| 27468 | 495 | |
| 61945 | 496 | lemma Standard_abs: "x \<in> Standard \<Longrightarrow> \<bar>x\<bar> \<in> Standard" | 
| 64435 | 497 | by (simp add: star_abs_def) | 
| 27468 | 498 | |
| 64435 | 499 | lemma Standard_mod: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x mod y \<in> Standard" | 
| 500 | by (simp add: star_mod_def) | |
| 27468 | 501 | |
| 502 | lemmas Standard_simps [simp] = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 503 | Standard_zero Standard_one | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 504 | Standard_add Standard_diff Standard_minus | 
| 27468 | 505 | Standard_mult Standard_divide Standard_inverse | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 506 | Standard_abs Standard_mod | 
| 27468 | 507 | |
| 64435 | 508 | |
| 69597 | 509 | text \<open>\<^term>\<open>star_of\<close> preserves class operations.\<close> | 
| 27468 | 510 | |
| 511 | lemma star_of_add: "star_of (x + y) = star_of x + star_of y" | |
| 64435 | 512 | by transfer (rule refl) | 
| 27468 | 513 | |
| 514 | lemma star_of_diff: "star_of (x - y) = star_of x - star_of y" | |
| 64435 | 515 | by transfer (rule refl) | 
| 27468 | 516 | |
| 517 | lemma star_of_minus: "star_of (-x) = - star_of x" | |
| 64435 | 518 | by transfer (rule refl) | 
| 27468 | 519 | |
| 520 | lemma star_of_mult: "star_of (x * y) = star_of x * star_of y" | |
| 64435 | 521 | by transfer (rule refl) | 
| 27468 | 522 | |
| 523 | lemma star_of_divide: "star_of (x / y) = star_of x / star_of y" | |
| 64435 | 524 | by transfer (rule refl) | 
| 27468 | 525 | |
| 526 | lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)" | |
| 64435 | 527 | by transfer (rule refl) | 
| 27468 | 528 | |
| 529 | lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y" | |
| 64435 | 530 | by transfer (rule refl) | 
| 27468 | 531 | |
| 61945 | 532 | lemma star_of_abs: "star_of \<bar>x\<bar> = \<bar>star_of x\<bar>" | 
| 64435 | 533 | by transfer (rule refl) | 
| 27468 | 534 | |
| 64435 | 535 | |
| 69597 | 536 | text \<open>\<^term>\<open>star_of\<close> preserves numerals.\<close> | 
| 27468 | 537 | |
| 538 | lemma star_of_zero: "star_of 0 = 0" | |
| 64435 | 539 | by transfer (rule refl) | 
| 27468 | 540 | |
| 541 | lemma star_of_one: "star_of 1 = 1" | |
| 64435 | 542 | by transfer (rule refl) | 
| 27468 | 543 | |
| 64435 | 544 | |
| 69597 | 545 | text \<open>\<^term>\<open>star_of\<close> preserves orderings.\<close> | 
| 27468 | 546 | |
| 547 | lemma star_of_less: "(star_of x < star_of y) = (x < y)" | |
| 548 | by transfer (rule refl) | |
| 549 | ||
| 550 | lemma star_of_le: "(star_of x \<le> star_of y) = (x \<le> y)" | |
| 551 | by transfer (rule refl) | |
| 552 | ||
| 553 | lemma star_of_eq: "(star_of x = star_of y) = (x = y)" | |
| 554 | by transfer (rule refl) | |
| 555 | ||
| 64435 | 556 | |
| 557 | text \<open>As above, for \<open>0\<close>.\<close> | |
| 27468 | 558 | |
| 559 | lemmas star_of_0_less = star_of_less [of 0, simplified star_of_zero] | |
| 560 | lemmas star_of_0_le = star_of_le [of 0, simplified star_of_zero] | |
| 561 | lemmas star_of_0_eq = star_of_eq [of 0, simplified star_of_zero] | |
| 562 | ||
| 563 | lemmas star_of_less_0 = star_of_less [of _ 0, simplified star_of_zero] | |
| 564 | lemmas star_of_le_0 = star_of_le [of _ 0, simplified star_of_zero] | |
| 565 | lemmas star_of_eq_0 = star_of_eq [of _ 0, simplified star_of_zero] | |
| 566 | ||
| 64435 | 567 | |
| 568 | text \<open>As above, for \<open>1\<close>.\<close> | |
| 27468 | 569 | |
| 570 | lemmas star_of_1_less = star_of_less [of 1, simplified star_of_one] | |
| 571 | lemmas star_of_1_le = star_of_le [of 1, simplified star_of_one] | |
| 572 | lemmas star_of_1_eq = star_of_eq [of 1, simplified star_of_one] | |
| 573 | ||
| 574 | lemmas star_of_less_1 = star_of_less [of _ 1, simplified star_of_one] | |
| 575 | lemmas star_of_le_1 = star_of_le [of _ 1, simplified star_of_one] | |
| 576 | lemmas star_of_eq_1 = star_of_eq [of _ 1, simplified star_of_one] | |
| 577 | ||
| 578 | lemmas star_of_simps [simp] = | |
| 579 | star_of_add star_of_diff star_of_minus | |
| 580 | star_of_mult star_of_divide star_of_inverse | |
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60041diff
changeset | 581 | star_of_mod star_of_abs | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 582 | star_of_zero star_of_one | 
| 27468 | 583 | star_of_less star_of_le star_of_eq | 
| 584 | star_of_0_less star_of_0_le star_of_0_eq | |
| 585 | star_of_less_0 star_of_le_0 star_of_eq_0 | |
| 586 | star_of_1_less star_of_1_le star_of_1_eq | |
| 587 | star_of_less_1 star_of_le_1 star_of_eq_1 | |
| 588 | ||
| 64435 | 589 | |
| 61975 | 590 | subsection \<open>Ordering and lattice classes\<close> | 
| 27468 | 591 | |
| 592 | instance star :: (order) order | |
| 70219 | 593 | proof | 
| 594 | show "\<And>x y::'a star. (x < y) = (x \<le> y \<and> \<not> y \<le> x)" | |
| 595 | by transfer (rule less_le_not_le) | |
| 596 | show "\<And>x::'a star. x \<le> x" | |
| 597 | by transfer (rule order_refl) | |
| 598 | show "\<And>x y z::'a star. \<lbrakk>x \<le> y; y \<le> z\<rbrakk> \<Longrightarrow> x \<le> z" | |
| 599 | by transfer (rule order_trans) | |
| 600 | show "\<And>x y::'a star. \<lbrakk>x \<le> y; y \<le> x\<rbrakk> \<Longrightarrow> x = y" | |
| 601 | by transfer (rule order_antisym) | |
| 602 | qed | |
| 27468 | 603 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 604 | instantiation star :: (semilattice_inf) semilattice_inf | 
| 27468 | 605 | begin | 
| 64435 | 606 | definition star_inf_def [transfer_unfold]: "inf \<equiv> *f2* inf" | 
| 607 | instance by (standard; transfer) auto | |
| 27468 | 608 | end | 
| 609 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 610 | instantiation star :: (semilattice_sup) semilattice_sup | 
| 27468 | 611 | begin | 
| 64435 | 612 | definition star_sup_def [transfer_unfold]: "sup \<equiv> *f2* sup" | 
| 613 | instance by (standard; transfer) auto | |
| 27468 | 614 | end | 
| 615 | ||
| 616 | instance star :: (lattice) lattice .. | |
| 617 | ||
| 618 | instance star :: (distrib_lattice) distrib_lattice | |
| 60867 | 619 | by (standard; transfer) (auto simp add: sup_inf_distrib1) | 
| 27468 | 620 | |
| 64435 | 621 | lemma Standard_inf [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> inf x y \<in> Standard" | 
| 622 | by (simp add: star_inf_def) | |
| 27468 | 623 | |
| 64435 | 624 | lemma Standard_sup [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> sup x y \<in> Standard" | 
| 625 | by (simp add: star_sup_def) | |
| 27468 | 626 | |
| 627 | lemma star_of_inf [simp]: "star_of (inf x y) = inf (star_of x) (star_of y)" | |
| 64435 | 628 | by transfer (rule refl) | 
| 27468 | 629 | |
| 630 | lemma star_of_sup [simp]: "star_of (sup x y) = sup (star_of x) (star_of y)" | |
| 64435 | 631 | by transfer (rule refl) | 
| 27468 | 632 | |
| 633 | instance star :: (linorder) linorder | |
| 64435 | 634 | by (intro_classes, transfer, rule linorder_linear) | 
| 27468 | 635 | |
| 636 | lemma star_max_def [transfer_unfold]: "max = *f2* max" | |
| 70219 | 637 | unfolding max_def | 
| 638 | by (intro ext, transfer, simp) | |
| 27468 | 639 | |
| 640 | lemma star_min_def [transfer_unfold]: "min = *f2* min" | |
| 70219 | 641 | unfolding min_def | 
| 642 | by (intro ext, transfer, simp) | |
| 27468 | 643 | |
| 64435 | 644 | lemma Standard_max [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> max x y \<in> Standard" | 
| 645 | by (simp add: star_max_def) | |
| 27468 | 646 | |
| 64435 | 647 | lemma Standard_min [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> min x y \<in> Standard" | 
| 648 | by (simp add: star_min_def) | |
| 27468 | 649 | |
| 650 | lemma star_of_max [simp]: "star_of (max x y) = max (star_of x) (star_of y)" | |
| 64435 | 651 | by transfer (rule refl) | 
| 27468 | 652 | |
| 653 | lemma star_of_min [simp]: "star_of (min x y) = min (star_of x) (star_of y)" | |
| 64435 | 654 | by transfer (rule refl) | 
| 27468 | 655 | |
| 656 | ||
| 61975 | 657 | subsection \<open>Ordered group classes\<close> | 
| 27468 | 658 | |
| 659 | instance star :: (semigroup_add) semigroup_add | |
| 64435 | 660 | by (intro_classes, transfer, rule add.assoc) | 
| 27468 | 661 | |
| 662 | instance star :: (ab_semigroup_add) ab_semigroup_add | |
| 64435 | 663 | by (intro_classes, transfer, rule add.commute) | 
| 27468 | 664 | |
| 665 | instance star :: (semigroup_mult) semigroup_mult | |
| 64435 | 666 | by (intro_classes, transfer, rule mult.assoc) | 
| 27468 | 667 | |
| 668 | instance star :: (ab_semigroup_mult) ab_semigroup_mult | |
| 64435 | 669 | by (intro_classes, transfer, rule mult.commute) | 
| 27468 | 670 | |
| 671 | instance star :: (comm_monoid_add) comm_monoid_add | |
| 64435 | 672 | by (intro_classes, transfer, rule comm_monoid_add_class.add_0) | 
| 27468 | 673 | |
| 674 | instance star :: (monoid_mult) monoid_mult | |
| 64435 | 675 | apply intro_classes | 
| 676 | apply (transfer, rule mult_1_left) | |
| 677 | apply (transfer, rule mult_1_right) | |
| 678 | done | |
| 27468 | 679 | |
| 60867 | 680 | instance star :: (power) power .. | 
| 681 | ||
| 27468 | 682 | instance star :: (comm_monoid_mult) comm_monoid_mult | 
| 64435 | 683 | by (intro_classes, transfer, rule mult_1) | 
| 27468 | 684 | |
| 685 | instance star :: (cancel_semigroup_add) cancel_semigroup_add | |
| 64435 | 686 | apply intro_classes | 
| 687 | apply (transfer, erule add_left_imp_eq) | |
| 688 | apply (transfer, erule add_right_imp_eq) | |
| 689 | done | |
| 27468 | 690 | |
| 691 | instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add | |
| 64435 | 692 | by intro_classes (transfer, simp add: diff_diff_eq)+ | 
| 27468 | 693 | |
| 29904 | 694 | instance star :: (cancel_comm_monoid_add) cancel_comm_monoid_add .. | 
| 695 | ||
| 27468 | 696 | instance star :: (ab_group_add) ab_group_add | 
| 64435 | 697 | apply intro_classes | 
| 698 | apply (transfer, rule left_minus) | |
| 699 | apply (transfer, rule diff_conv_add_uminus) | |
| 700 | done | |
| 27468 | 701 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 702 | instance star :: (ordered_ab_semigroup_add) ordered_ab_semigroup_add | 
| 64435 | 703 | by (intro_classes, transfer, rule add_left_mono) | 
| 27468 | 704 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 705 | instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add .. | 
| 27468 | 706 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 707 | instance star :: (ordered_ab_semigroup_add_imp_le) ordered_ab_semigroup_add_imp_le | 
| 64435 | 708 | by (intro_classes, transfer, rule add_le_imp_le_left) | 
| 27468 | 709 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 710 | instance star :: (ordered_comm_monoid_add) ordered_comm_monoid_add .. | 
| 63456 
3365c8ec67bd
sharing simp rules between ordered monoids and rings
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
62479diff
changeset | 711 | instance star :: (ordered_ab_semigroup_monoid_add_imp_le) ordered_ab_semigroup_monoid_add_imp_le .. | 
| 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: 
61975diff
changeset | 712 | instance star :: (ordered_cancel_comm_monoid_add) ordered_cancel_comm_monoid_add .. | 
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 713 | instance star :: (ordered_ab_group_add) ordered_ab_group_add .. | 
| 27468 | 714 | |
| 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: 
61975diff
changeset | 715 | instance star :: (ordered_ab_group_add_abs) ordered_ab_group_add_abs | 
| 64435 | 716 | by intro_classes (transfer, simp add: abs_ge_self abs_leI abs_triangle_ineq)+ | 
| 27468 | 717 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 718 | instance star :: (linordered_cancel_ab_semigroup_add) linordered_cancel_ab_semigroup_add .. | 
| 27468 | 719 | |
| 720 | ||
| 61975 | 721 | subsection \<open>Ring and field classes\<close> | 
| 27468 | 722 | |
| 723 | instance star :: (semiring) semiring | |
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 724 | by (intro_classes; transfer) (fact distrib_right distrib_left)+ | 
| 27468 | 725 | |
| 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: 
61975diff
changeset | 726 | instance star :: (semiring_0) semiring_0 | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 727 | by (intro_classes; transfer) simp_all | 
| 27468 | 728 | |
| 729 | instance star :: (semiring_0_cancel) semiring_0_cancel .. | |
| 730 | ||
| 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: 
61975diff
changeset | 731 | instance star :: (comm_semiring) comm_semiring | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 732 | by (intro_classes; transfer) (fact distrib_right) | 
| 27468 | 733 | |
| 734 | instance star :: (comm_semiring_0) comm_semiring_0 .. | |
| 735 | instance star :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. | |
| 736 | ||
| 737 | instance star :: (zero_neq_one) zero_neq_one | |
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 738 | by (intro_classes; transfer) (fact zero_neq_one) | 
| 27468 | 739 | |
| 740 | instance star :: (semiring_1) semiring_1 .. | |
| 741 | instance star :: (comm_semiring_1) comm_semiring_1 .. | |
| 742 | ||
| 59680 | 743 | declare dvd_def [transfer_refold] | 
| 59676 | 744 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60516diff
changeset | 745 | instance star :: (comm_semiring_1_cancel) comm_semiring_1_cancel | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 746 | by (intro_classes; transfer) (fact right_diff_distrib') | 
| 59676 | 747 | |
| 59833 
ab828c2c5d67
clarified no_zero_devisors: makes only sense in a semiring;
 haftmann parents: 
59816diff
changeset | 748 | instance star :: (semiring_no_zero_divisors) semiring_no_zero_divisors | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 749 | by (intro_classes; transfer) (fact no_zero_divisors) | 
| 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 750 | |
| 60867 | 751 | instance star :: (semiring_1_no_zero_divisors) semiring_1_no_zero_divisors .. | 
| 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: 
61975diff
changeset | 752 | |
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 753 | instance star :: (semiring_no_zero_divisors_cancel) semiring_no_zero_divisors_cancel | 
| 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 754 | by (intro_classes; transfer) simp_all | 
| 27468 | 755 | |
| 756 | instance star :: (semiring_1_cancel) semiring_1_cancel .. | |
| 757 | instance star :: (ring) ring .. | |
| 758 | instance star :: (comm_ring) comm_ring .. | |
| 759 | instance star :: (ring_1) ring_1 .. | |
| 760 | instance star :: (comm_ring_1) comm_ring_1 .. | |
| 59833 
ab828c2c5d67
clarified no_zero_devisors: makes only sense in a semiring;
 haftmann parents: 
59816diff
changeset | 761 | instance star :: (semidom) semidom .. | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 762 | |
| 60353 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 763 | instance star :: (semidom_divide) semidom_divide | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 764 | by (intro_classes; transfer) simp_all | 
| 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 765 | |
| 27468 | 766 | instance star :: (ring_no_zero_divisors) ring_no_zero_divisors .. | 
| 767 | instance star :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors .. | |
| 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: 
61975diff
changeset | 768 | instance star :: (idom) idom .. | 
| 60353 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 769 | instance star :: (idom_divide) idom_divide .. | 
| 27468 | 770 | |
| 771 | instance star :: (division_ring) division_ring | |
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 772 | by (intro_classes; transfer) (simp_all add: divide_inverse) | 
| 27468 | 773 | |
| 774 | instance star :: (field) field | |
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 775 | by (intro_classes; transfer) (simp_all add: divide_inverse) | 
| 27468 | 776 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 777 | instance star :: (ordered_semiring) ordered_semiring | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 778 | by (intro_classes; transfer) (fact mult_left_mono mult_right_mono)+ | 
| 27468 | 779 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 780 | instance star :: (ordered_cancel_semiring) ordered_cancel_semiring .. | 
| 27468 | 781 | |
| 35043 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 haftmann parents: 
35035diff
changeset | 782 | instance star :: (linordered_semiring_strict) linordered_semiring_strict | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 783 | by (intro_classes; transfer) (fact mult_strict_left_mono mult_strict_right_mono)+ | 
| 27468 | 784 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 785 | instance star :: (ordered_comm_semiring) ordered_comm_semiring | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 786 | by (intro_classes; transfer) (fact mult_left_mono) | 
| 27468 | 787 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 788 | instance star :: (ordered_cancel_comm_semiring) ordered_cancel_comm_semiring .. | 
| 27468 | 789 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 790 | instance star :: (linordered_comm_semiring_strict) linordered_comm_semiring_strict | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 791 | by (intro_classes; transfer) (fact mult_strict_left_mono) | 
| 27468 | 792 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 793 | instance star :: (ordered_ring) ordered_ring .. | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 794 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 795 | instance star :: (ordered_ring_abs) ordered_ring_abs | 
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 796 | by (intro_classes; transfer) (fact abs_eq_mult) | 
| 27468 | 797 | |
| 798 | instance star :: (abs_if) abs_if | |
| 60516 
0826b7025d07
generalized some theorems about integral domains and moved to HOL theories
 haftmann parents: 
60429diff
changeset | 799 | by (intro_classes; transfer) (fact abs_if) | 
| 27468 | 800 | |
| 35043 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 haftmann parents: 
35035diff
changeset | 801 | instance star :: (linordered_ring_strict) linordered_ring_strict .. | 
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 802 | instance star :: (ordered_comm_ring) ordered_comm_ring .. | 
| 27468 | 803 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 804 | instance star :: (linordered_semidom) linordered_semidom | 
| 64290 | 805 | by (intro_classes; transfer) (fact zero_less_one le_add_diff_inverse2)+ | 
| 27468 | 806 | |
| 64290 | 807 | instance star :: (linordered_idom) linordered_idom | 
| 808 | by (intro_classes; transfer) (fact sgn_if) | |
| 809 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 810 | instance star :: (linordered_field) linordered_field .. | 
| 27468 | 811 | |
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 812 | instance star :: (algebraic_semidom) algebraic_semidom .. | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 813 | |
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 814 | instantiation star :: (normalization_semidom) normalization_semidom | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 815 | begin | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 816 | |
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 817 | definition unit_factor_star :: "'a star \<Rightarrow> 'a star" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 818 | where [transfer_unfold]: "unit_factor_star = *f* unit_factor" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 819 | |
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 820 | definition normalize_star :: "'a star \<Rightarrow> 'a star" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 821 | where [transfer_unfold]: "normalize_star = *f* normalize" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 822 | |
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 823 | instance | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 824 | by standard (transfer; simp add: is_unit_unit_factor unit_factor_mult)+ | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 825 | |
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 826 | end | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
64600diff
changeset | 827 | |
| 66815 | 828 | instance star :: (semidom_modulo) semidom_modulo | 
| 829 | by standard (transfer; simp) | |
| 830 | ||
| 831 | ||
| 64435 | 832 | |
| 61975 | 833 | subsection \<open>Power\<close> | 
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 834 | |
| 67399 | 835 | lemma star_power_def [transfer_unfold]: "(^) \<equiv> \<lambda>x n. ( *f* (\<lambda>x. x ^ n)) x" | 
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 836 | proof (rule eq_reflection, rule ext, rule ext) | 
| 64435 | 837 | show "x ^ n = ( *f* (\<lambda>x. x ^ n)) x" for n :: nat and x :: "'a star" | 
| 838 | proof (induct n arbitrary: x) | |
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 839 | case 0 | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 840 | have "\<And>x::'a star. ( *f* (\<lambda>x. 1)) x = 1" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 841 | by transfer simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 842 | then show ?case by simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 843 | next | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 844 | case (Suc n) | 
| 61076 | 845 | have "\<And>x::'a star. x * ( *f* (\<lambda>x::'a. x ^ n)) x = ( *f* (\<lambda>x::'a. x * x ^ n)) x" | 
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 846 | by transfer simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 847 | with Suc show ?case by simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 848 | qed | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 849 | qed | 
| 27468 | 850 | |
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 851 | lemma Standard_power [simp]: "x \<in> Standard \<Longrightarrow> x ^ n \<in> Standard" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 852 | by (simp add: star_power_def) | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 853 | |
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 854 | lemma star_of_power [simp]: "star_of (x ^ n) = star_of x ^ n" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 855 | by transfer (rule refl) | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 856 | |
| 27468 | 857 | |
| 61975 | 858 | subsection \<open>Number classes\<close> | 
| 27468 | 859 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 860 | instance star :: (numeral) numeral .. | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 861 | |
| 64435 | 862 | lemma star_numeral_def [transfer_unfold]: "numeral k = star_of (numeral k)" | 
| 863 | by (induct k) (simp_all only: numeral.simps star_of_one star_of_add) | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 864 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 865 | lemma Standard_numeral [simp]: "numeral k \<in> Standard" | 
| 64435 | 866 | by (simp add: star_numeral_def) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 867 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 868 | lemma star_of_numeral [simp]: "star_of (numeral k) = numeral k" | 
| 64435 | 869 | by transfer (rule refl) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 870 | |
| 27468 | 871 | lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)" | 
| 64435 | 872 | by (induct n) simp_all | 
| 27468 | 873 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 874 | lemmas star_of_compare_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 875 | star_of_less [of "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 876 | star_of_le [of "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 877 | star_of_eq [of "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 878 | star_of_less [of _ "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 879 | star_of_le [of _ "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 880 | star_of_eq [of _ "numeral k", simplified star_of_numeral] | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 881 | star_of_less [of "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 882 | star_of_le [of "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 883 | star_of_eq [of "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 884 | star_of_less [of _ "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 885 | star_of_le [of _ "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 886 | star_of_eq [of _ "- numeral k", simplified star_of_numeral] for k | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 887 | |
| 27468 | 888 | lemma Standard_of_nat [simp]: "of_nat n \<in> Standard" | 
| 64435 | 889 | by (simp add: star_of_nat_def) | 
| 27468 | 890 | |
| 891 | lemma star_of_of_nat [simp]: "star_of (of_nat n) = of_nat n" | |
| 64435 | 892 | by transfer (rule refl) | 
| 27468 | 893 | |
| 894 | lemma star_of_int_def [transfer_unfold]: "of_int z = star_of (of_int z)" | |
| 64435 | 895 | by (rule int_diff_cases [of z]) simp | 
| 27468 | 896 | |
| 897 | lemma Standard_of_int [simp]: "of_int z \<in> Standard" | |
| 64435 | 898 | by (simp add: star_of_int_def) | 
| 27468 | 899 | |
| 900 | lemma star_of_of_int [simp]: "star_of (of_int z) = of_int z" | |
| 64435 | 901 | by transfer (rule refl) | 
| 27468 | 902 | |
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62376diff
changeset | 903 | instance star :: (semiring_char_0) semiring_char_0 | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62376diff
changeset | 904 | proof | 
| 64435 | 905 | have "inj (star_of :: 'a \<Rightarrow> 'a star)" | 
| 906 | by (rule injI) simp | |
| 907 | then have "inj (star_of \<circ> of_nat :: nat \<Rightarrow> 'a star)" | |
| 69700 
7a92cbec7030
new material about summations and powers, along with some tweaks
 paulson <lp15@cam.ac.uk> parents: 
69605diff
changeset | 908 | using inj_of_nat by (rule inj_compose) | 
| 64435 | 909 | then show "inj (of_nat :: nat \<Rightarrow> 'a star)" | 
| 910 | by (simp add: comp_def) | |
| 38621 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37765diff
changeset | 911 | qed | 
| 27468 | 912 | |
| 913 | instance star :: (ring_char_0) ring_char_0 .. | |
| 914 | ||
| 915 | ||
| 61975 | 916 | subsection \<open>Finite class\<close> | 
| 27468 | 917 | |
| 918 | lemma starset_finite: "finite A \<Longrightarrow> *s* A = star_of ` A" | |
| 64435 | 919 | by (erule finite_induct) simp_all | 
| 27468 | 920 | |
| 921 | instance star :: (finite) finite | |
| 70219 | 922 | proof intro_classes | 
| 923 | show "finite (UNIV::'a star set)" | |
| 924 | by (metis starset_UNIV finite finite_imageI starset_finite) | |
| 925 | qed | |
| 27468 | 926 | |
| 927 | end |