| author | wenzelm | 
| Mon, 06 Oct 2014 19:55:49 +0200 | |
| changeset 58604 | 13dfea1621b2 | 
| parent 57512 | cc97b347b301 | 
| child 58825 | 2065f49da190 | 
| permissions | -rw-r--r-- | 
| 27468 | 1 | (* Title : HOL/Hyperreal/StarDef.thy | 
| 2 | Author : Jacques D. Fleuriot and Brian Huffman | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Construction of Star Types Using Ultrafilters *}
 | |
| 6 | ||
| 7 | theory StarDef | |
| 8 | imports Filter | |
| 9 | begin | |
| 10 | ||
| 11 | subsection {* A Free Ultrafilter over the Naturals *}
 | |
| 12 | ||
| 13 | definition | |
| 14 |   FreeUltrafilterNat :: "nat set set"  ("\<U>") where
 | |
| 15 | "\<U> = (SOME U. freeultrafilter U)" | |
| 16 | ||
| 17 | lemma freeultrafilter_FreeUltrafilterNat: "freeultrafilter \<U>" | |
| 18 | apply (unfold FreeUltrafilterNat_def) | |
| 46008 
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
 wenzelm parents: 
45694diff
changeset | 19 | apply (rule someI_ex) | 
| 27468 | 20 | apply (rule freeultrafilter_Ex) | 
| 54580 | 21 | apply (rule infinite_UNIV_nat) | 
| 27468 | 22 | done | 
| 23 | ||
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
30198diff
changeset | 24 | interpretation FreeUltrafilterNat: freeultrafilter FreeUltrafilterNat | 
| 27468 | 25 | by (rule freeultrafilter_FreeUltrafilterNat) | 
| 26 | ||
| 27 | text {* This rule takes the place of the old ultra tactic *}
 | |
| 28 | ||
| 29 | lemma ultra: | |
| 30 |   "\<lbrakk>{n. P n} \<in> \<U>; {n. P n \<longrightarrow> Q n} \<in> \<U>\<rbrakk> \<Longrightarrow> {n. Q n} \<in> \<U>"
 | |
| 31 | by (simp add: Collect_imp_eq | |
| 32 | FreeUltrafilterNat.Un_iff FreeUltrafilterNat.Compl_iff) | |
| 33 | ||
| 34 | ||
| 35 | subsection {* Definition of @{text star} type constructor *}
 | |
| 36 | ||
| 37 | definition | |
| 38 | starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set" where | |
| 39 |   "starrel = {(X,Y). {n. X n = Y n} \<in> \<U>}"
 | |
| 40 | ||
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 41 | definition "star = (UNIV :: (nat \<Rightarrow> 'a) set) // starrel" | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 42 | |
| 49834 | 43 | typedef 'a star = "star :: (nat \<Rightarrow> 'a) set set" | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 44 | unfolding star_def by (auto intro: quotientI) | 
| 27468 | 45 | |
| 46 | definition | |
| 47 | star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star" where | |
| 48 |   "star_n X = Abs_star (starrel `` {X})"
 | |
| 49 | ||
| 50 | theorem star_cases [case_names star_n, cases type: star]: | |
| 51 | "(\<And>X. x = star_n X \<Longrightarrow> P) \<Longrightarrow> P" | |
| 52 | by (cases x, unfold star_n_def star_def, erule quotientE, fast) | |
| 53 | ||
| 54 | lemma all_star_eq: "(\<forall>x. P x) = (\<forall>X. P (star_n X))" | |
| 55 | by (auto, rule_tac x=x in star_cases, simp) | |
| 56 | ||
| 57 | lemma ex_star_eq: "(\<exists>x. P x) = (\<exists>X. P (star_n X))" | |
| 58 | by (auto, rule_tac x=x in star_cases, auto) | |
| 59 | ||
| 60 | text {* Proving that @{term starrel} is an equivalence relation *}
 | |
| 61 | ||
| 62 | lemma starrel_iff [iff]: "((X,Y) \<in> starrel) = ({n. X n = Y n} \<in> \<U>)"
 | |
| 63 | by (simp add: starrel_def) | |
| 64 | ||
| 65 | lemma equiv_starrel: "equiv UNIV starrel" | |
| 40815 | 66 | proof (rule equivI) | 
| 30198 | 67 | show "refl starrel" by (simp add: refl_on_def) | 
| 27468 | 68 | show "sym starrel" by (simp add: sym_def eq_commute) | 
| 69 | show "trans starrel" by (auto intro: transI elim!: ultra) | |
| 70 | qed | |
| 71 | ||
| 72 | lemmas equiv_starrel_iff = | |
| 73 | eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I] | |
| 74 | ||
| 75 | lemma starrel_in_star: "starrel``{x} \<in> star"
 | |
| 76 | by (simp add: star_def quotientI) | |
| 77 | ||
| 78 | lemma star_n_eq_iff: "(star_n X = star_n Y) = ({n. X n = Y n} \<in> \<U>)"
 | |
| 79 | by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff) | |
| 80 | ||
| 81 | ||
| 82 | subsection {* Transfer principle *}
 | |
| 83 | ||
| 84 | text {* This introduction rule starts each transfer proof. *}
 | |
| 85 | lemma transfer_start: | |
| 86 |   "P \<equiv> {n. Q} \<in> \<U> \<Longrightarrow> Trueprop P \<equiv> Trueprop Q"
 | |
| 87 | by (subgoal_tac "P \<equiv> Q", simp, simp add: atomize_eq) | |
| 88 | ||
| 89 | text {*Initialize transfer tactic.*}
 | |
| 48891 | 90 | ML_file "transfer.ML" | 
| 47328 
9f11a3cd84b1
rename ML structure to avoid shadowing earlier name
 huffman parents: 
47108diff
changeset | 91 | setup Transfer_Principle.setup | 
| 27468 | 92 | |
| 47432 | 93 | method_setup transfer = {*
 | 
| 94 | Attrib.thms >> (fn ths => fn ctxt => | |
| 95 | SIMPLE_METHOD' (Transfer_Principle.transfer_tac ctxt ths)) | |
| 96 | *} "transfer principle" | |
| 97 | ||
| 98 | ||
| 27468 | 99 | text {* Transfer introduction rules. *}
 | 
| 100 | ||
| 101 | lemma transfer_ex [transfer_intro]: | |
| 102 |   "\<lbrakk>\<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
 | |
| 103 |     \<Longrightarrow> \<exists>x::'a star. p x \<equiv> {n. \<exists>x. P n x} \<in> \<U>"
 | |
| 104 | by (simp only: ex_star_eq FreeUltrafilterNat.Collect_ex) | |
| 105 | ||
| 106 | lemma transfer_all [transfer_intro]: | |
| 107 |   "\<lbrakk>\<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
 | |
| 108 |     \<Longrightarrow> \<forall>x::'a star. p x \<equiv> {n. \<forall>x. P n x} \<in> \<U>"
 | |
| 109 | by (simp only: all_star_eq FreeUltrafilterNat.Collect_all) | |
| 110 | ||
| 111 | lemma transfer_not [transfer_intro]: | |
| 112 |   "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>\<rbrakk> \<Longrightarrow> \<not> p \<equiv> {n. \<not> P n} \<in> \<U>"
 | |
| 113 | by (simp only: FreeUltrafilterNat.Collect_not) | |
| 114 | ||
| 115 | lemma transfer_conj [transfer_intro]: | |
| 116 |   "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
 | |
| 117 |     \<Longrightarrow> p \<and> q \<equiv> {n. P n \<and> Q n} \<in> \<U>"
 | |
| 118 | by (simp only: FreeUltrafilterNat.Collect_conj) | |
| 119 | ||
| 120 | lemma transfer_disj [transfer_intro]: | |
| 121 |   "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
 | |
| 122 |     \<Longrightarrow> p \<or> q \<equiv> {n. P n \<or> Q n} \<in> \<U>"
 | |
| 123 | by (simp only: FreeUltrafilterNat.Collect_disj) | |
| 124 | ||
| 125 | lemma transfer_imp [transfer_intro]: | |
| 126 |   "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
 | |
| 127 |     \<Longrightarrow> p \<longrightarrow> q \<equiv> {n. P n \<longrightarrow> Q n} \<in> \<U>"
 | |
| 128 | by (simp only: imp_conv_disj transfer_disj transfer_not) | |
| 129 | ||
| 130 | lemma transfer_iff [transfer_intro]: | |
| 131 |   "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
 | |
| 132 |     \<Longrightarrow> p = q \<equiv> {n. P n = Q n} \<in> \<U>"
 | |
| 133 | by (simp only: iff_conv_conj_imp transfer_conj transfer_imp) | |
| 134 | ||
| 135 | lemma transfer_if_bool [transfer_intro]: | |
| 136 |   "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; x \<equiv> {n. X n} \<in> \<U>; y \<equiv> {n. Y n} \<in> \<U>\<rbrakk>
 | |
| 137 |     \<Longrightarrow> (if p then x else y) \<equiv> {n. if P n then X n else Y n} \<in> \<U>"
 | |
| 138 | by (simp only: if_bool_eq_conj transfer_conj transfer_imp transfer_not) | |
| 139 | ||
| 140 | lemma transfer_eq [transfer_intro]: | |
| 141 |   "\<lbrakk>x \<equiv> star_n X; y \<equiv> star_n Y\<rbrakk> \<Longrightarrow> x = y \<equiv> {n. X n = Y n} \<in> \<U>"
 | |
| 142 | by (simp only: star_n_eq_iff) | |
| 143 | ||
| 144 | lemma transfer_if [transfer_intro]: | |
| 145 |   "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; x \<equiv> star_n X; y \<equiv> star_n Y\<rbrakk>
 | |
| 146 | \<Longrightarrow> (if p then x else y) \<equiv> star_n (\<lambda>n. if P n then X n else Y n)" | |
| 147 | apply (rule eq_reflection) | |
| 148 | apply (auto simp add: star_n_eq_iff transfer_not elim!: ultra) | |
| 149 | done | |
| 150 | ||
| 151 | lemma transfer_fun_eq [transfer_intro]: | |
| 152 | "\<lbrakk>\<And>X. f (star_n X) = g (star_n X) | |
| 153 |     \<equiv> {n. F n (X n) = G n (X n)} \<in> \<U>\<rbrakk>
 | |
| 154 |       \<Longrightarrow> f = g \<equiv> {n. F n = G n} \<in> \<U>"
 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 155 | by (simp only: fun_eq_iff transfer_all) | 
| 27468 | 156 | |
| 157 | lemma transfer_star_n [transfer_intro]: "star_n X \<equiv> star_n (\<lambda>n. X n)" | |
| 158 | by (rule reflexive) | |
| 159 | ||
| 160 | lemma transfer_bool [transfer_intro]: "p \<equiv> {n. p} \<in> \<U>"
 | |
| 161 | by (simp add: atomize_eq) | |
| 162 | ||
| 163 | ||
| 164 | subsection {* Standard elements *}
 | |
| 165 | ||
| 166 | definition | |
| 167 | star_of :: "'a \<Rightarrow> 'a star" where | |
| 168 | "star_of x == star_n (\<lambda>n. x)" | |
| 169 | ||
| 170 | definition | |
| 171 | Standard :: "'a star set" where | |
| 172 | "Standard = range star_of" | |
| 173 | ||
| 174 | text {* Transfer tactic should remove occurrences of @{term star_of} *}
 | |
| 56256 | 175 | setup {* Transfer_Principle.add_const @{const_name star_of} *}
 | 
| 27468 | 176 | |
| 177 | declare star_of_def [transfer_intro] | |
| 178 | ||
| 179 | lemma star_of_inject: "(star_of x = star_of y) = (x = y)" | |
| 180 | by (transfer, rule refl) | |
| 181 | ||
| 182 | lemma Standard_star_of [simp]: "star_of x \<in> Standard" | |
| 183 | by (simp add: Standard_def) | |
| 184 | ||
| 185 | ||
| 186 | subsection {* Internal functions *}
 | |
| 187 | ||
| 188 | definition | |
| 189 |   Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300) where
 | |
| 190 | "Ifun f \<equiv> \<lambda>x. Abs_star | |
| 191 |        (\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})"
 | |
| 192 | ||
| 193 | lemma Ifun_congruent2: | |
| 194 |   "congruent2 starrel starrel (\<lambda>F X. starrel``{\<lambda>n. F n (X n)})"
 | |
| 195 | by (auto simp add: congruent2_def equiv_starrel_iff elim!: ultra) | |
| 196 | ||
| 197 | lemma Ifun_star_n: "star_n F \<star> star_n X = star_n (\<lambda>n. F n (X n))" | |
| 198 | by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star | |
| 199 | UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2]) | |
| 200 | ||
| 201 | text {* Transfer tactic should remove occurrences of @{term Ifun} *}
 | |
| 56256 | 202 | setup {* Transfer_Principle.add_const @{const_name Ifun} *}
 | 
| 27468 | 203 | |
| 204 | lemma transfer_Ifun [transfer_intro]: | |
| 205 | "\<lbrakk>f \<equiv> star_n F; x \<equiv> star_n X\<rbrakk> \<Longrightarrow> f \<star> x \<equiv> star_n (\<lambda>n. F n (X n))" | |
| 206 | by (simp only: Ifun_star_n) | |
| 207 | ||
| 208 | lemma Ifun_star_of [simp]: "star_of f \<star> star_of x = star_of (f x)" | |
| 209 | by (transfer, rule refl) | |
| 210 | ||
| 211 | lemma Standard_Ifun [simp]: | |
| 212 | "\<lbrakk>f \<in> Standard; x \<in> Standard\<rbrakk> \<Longrightarrow> f \<star> x \<in> Standard" | |
| 213 | by (auto simp add: Standard_def) | |
| 214 | ||
| 215 | text {* Nonstandard extensions of functions *}
 | |
| 216 | ||
| 217 | definition | |
| 218 |   starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a star \<Rightarrow> 'b star)"  ("*f* _" [80] 80) where
 | |
| 219 | "starfun f == \<lambda>x. star_of f \<star> x" | |
| 220 | ||
| 221 | definition | |
| 222 |   starfun2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a star \<Rightarrow> 'b star \<Rightarrow> 'c star)"
 | |
| 223 |     ("*f2* _" [80] 80) where
 | |
| 224 | "starfun2 f == \<lambda>x y. star_of f \<star> x \<star> y" | |
| 225 | ||
| 226 | declare starfun_def [transfer_unfold] | |
| 227 | declare starfun2_def [transfer_unfold] | |
| 228 | ||
| 229 | lemma starfun_star_n: "( *f* f) (star_n X) = star_n (\<lambda>n. f (X n))" | |
| 230 | by (simp only: starfun_def star_of_def Ifun_star_n) | |
| 231 | ||
| 232 | lemma starfun2_star_n: | |
| 233 | "( *f2* f) (star_n X) (star_n Y) = star_n (\<lambda>n. f (X n) (Y n))" | |
| 234 | by (simp only: starfun2_def star_of_def Ifun_star_n) | |
| 235 | ||
| 236 | lemma starfun_star_of [simp]: "( *f* f) (star_of x) = star_of (f x)" | |
| 237 | by (transfer, rule refl) | |
| 238 | ||
| 239 | lemma starfun2_star_of [simp]: "( *f2* f) (star_of x) = *f* f x" | |
| 240 | by (transfer, rule refl) | |
| 241 | ||
| 242 | lemma Standard_starfun [simp]: "x \<in> Standard \<Longrightarrow> starfun f x \<in> Standard" | |
| 243 | by (simp add: starfun_def) | |
| 244 | ||
| 245 | lemma Standard_starfun2 [simp]: | |
| 246 | "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> starfun2 f x y \<in> Standard" | |
| 247 | by (simp add: starfun2_def) | |
| 248 | ||
| 249 | lemma Standard_starfun_iff: | |
| 250 | assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y" | |
| 251 | shows "(starfun f x \<in> Standard) = (x \<in> Standard)" | |
| 252 | proof | |
| 253 | assume "x \<in> Standard" | |
| 254 | thus "starfun f x \<in> Standard" by simp | |
| 255 | next | |
| 256 | have inj': "\<And>x y. starfun f x = starfun f y \<Longrightarrow> x = y" | |
| 257 | using inj by transfer | |
| 258 | assume "starfun f x \<in> Standard" | |
| 259 | then obtain b where b: "starfun f x = star_of b" | |
| 260 | unfolding Standard_def .. | |
| 261 | hence "\<exists>x. starfun f x = star_of b" .. | |
| 262 | hence "\<exists>a. f a = b" by transfer | |
| 263 | then obtain a where "f a = b" .. | |
| 264 | hence "starfun f (star_of a) = star_of b" by transfer | |
| 265 | with b have "starfun f x = starfun f (star_of a)" by simp | |
| 266 | hence "x = star_of a" by (rule inj') | |
| 267 | thus "x \<in> Standard" | |
| 268 | unfolding Standard_def by auto | |
| 269 | qed | |
| 270 | ||
| 271 | lemma Standard_starfun2_iff: | |
| 272 | assumes inj: "\<And>a b a' b'. f a b = f a' b' \<Longrightarrow> a = a' \<and> b = b'" | |
| 273 | shows "(starfun2 f x y \<in> Standard) = (x \<in> Standard \<and> y \<in> Standard)" | |
| 274 | proof | |
| 275 | assume "x \<in> Standard \<and> y \<in> Standard" | |
| 276 | thus "starfun2 f x y \<in> Standard" by simp | |
| 277 | next | |
| 278 | have inj': "\<And>x y z w. starfun2 f x y = starfun2 f z w \<Longrightarrow> x = z \<and> y = w" | |
| 279 | using inj by transfer | |
| 280 | assume "starfun2 f x y \<in> Standard" | |
| 281 | then obtain c where c: "starfun2 f x y = star_of c" | |
| 282 | unfolding Standard_def .. | |
| 283 | hence "\<exists>x y. starfun2 f x y = star_of c" by auto | |
| 284 | hence "\<exists>a b. f a b = c" by transfer | |
| 285 | then obtain a b where "f a b = c" by auto | |
| 286 | hence "starfun2 f (star_of a) (star_of b) = star_of c" | |
| 287 | by transfer | |
| 288 | with c have "starfun2 f x y = starfun2 f (star_of a) (star_of b)" | |
| 289 | by simp | |
| 290 | hence "x = star_of a \<and> y = star_of b" | |
| 291 | by (rule inj') | |
| 292 | thus "x \<in> Standard \<and> y \<in> Standard" | |
| 293 | unfolding Standard_def by auto | |
| 294 | qed | |
| 295 | ||
| 296 | ||
| 297 | subsection {* Internal predicates *}
 | |
| 298 | ||
| 299 | definition unstar :: "bool star \<Rightarrow> bool" where | |
| 37765 | 300 | "unstar b \<longleftrightarrow> b = star_of True" | 
| 27468 | 301 | |
| 302 | lemma unstar_star_n: "unstar (star_n P) = ({n. P n} \<in> \<U>)"
 | |
| 303 | by (simp add: unstar_def star_of_def star_n_eq_iff) | |
| 304 | ||
| 305 | lemma unstar_star_of [simp]: "unstar (star_of p) = p" | |
| 306 | by (simp add: unstar_def star_of_inject) | |
| 307 | ||
| 308 | text {* Transfer tactic should remove occurrences of @{term unstar} *}
 | |
| 56256 | 309 | setup {* Transfer_Principle.add_const @{const_name unstar} *}
 | 
| 27468 | 310 | |
| 311 | lemma transfer_unstar [transfer_intro]: | |
| 312 |   "p \<equiv> star_n P \<Longrightarrow> unstar p \<equiv> {n. P n} \<in> \<U>"
 | |
| 313 | by (simp only: unstar_star_n) | |
| 314 | ||
| 315 | definition | |
| 316 |   starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool"  ("*p* _" [80] 80) where
 | |
| 317 | "*p* P = (\<lambda>x. unstar (star_of P \<star> x))" | |
| 318 | ||
| 319 | definition | |
| 320 |   starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool"  ("*p2* _" [80] 80) where
 | |
| 321 | "*p2* P = (\<lambda>x y. unstar (star_of P \<star> x \<star> y))" | |
| 322 | ||
| 323 | declare starP_def [transfer_unfold] | |
| 324 | declare starP2_def [transfer_unfold] | |
| 325 | ||
| 326 | lemma starP_star_n: "( *p* P) (star_n X) = ({n. P (X n)} \<in> \<U>)"
 | |
| 327 | by (simp only: starP_def star_of_def Ifun_star_n unstar_star_n) | |
| 328 | ||
| 329 | lemma starP2_star_n: | |
| 330 |   "( *p2* P) (star_n X) (star_n Y) = ({n. P (X n) (Y n)} \<in> \<U>)"
 | |
| 331 | by (simp only: starP2_def star_of_def Ifun_star_n unstar_star_n) | |
| 332 | ||
| 333 | lemma starP_star_of [simp]: "( *p* P) (star_of x) = P x" | |
| 334 | by (transfer, rule refl) | |
| 335 | ||
| 336 | lemma starP2_star_of [simp]: "( *p2* P) (star_of x) = *p* P x" | |
| 337 | by (transfer, rule refl) | |
| 338 | ||
| 339 | ||
| 340 | subsection {* Internal sets *}
 | |
| 341 | ||
| 342 | definition | |
| 343 | Iset :: "'a set star \<Rightarrow> 'a star set" where | |
| 344 |   "Iset A = {x. ( *p2* op \<in>) x A}"
 | |
| 345 | ||
| 346 | lemma Iset_star_n: | |
| 347 |   "(star_n X \<in> Iset (star_n A)) = ({n. X n \<in> A n} \<in> \<U>)"
 | |
| 348 | by (simp add: Iset_def starP2_star_n) | |
| 349 | ||
| 350 | text {* Transfer tactic should remove occurrences of @{term Iset} *}
 | |
| 56256 | 351 | setup {* Transfer_Principle.add_const @{const_name Iset} *}
 | 
| 27468 | 352 | |
| 353 | lemma transfer_mem [transfer_intro]: | |
| 354 | "\<lbrakk>x \<equiv> star_n X; a \<equiv> Iset (star_n A)\<rbrakk> | |
| 355 |     \<Longrightarrow> x \<in> a \<equiv> {n. X n \<in> A n} \<in> \<U>"
 | |
| 356 | by (simp only: Iset_star_n) | |
| 357 | ||
| 358 | lemma transfer_Collect [transfer_intro]: | |
| 359 |   "\<lbrakk>\<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
 | |
| 360 | \<Longrightarrow> Collect p \<equiv> Iset (star_n (\<lambda>n. Collect (P n)))" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 361 | by (simp add: atomize_eq set_eq_iff all_star_eq Iset_star_n) | 
| 27468 | 362 | |
| 363 | lemma transfer_set_eq [transfer_intro]: | |
| 364 | "\<lbrakk>a \<equiv> Iset (star_n A); b \<equiv> Iset (star_n B)\<rbrakk> | |
| 365 |     \<Longrightarrow> a = b \<equiv> {n. A n = B n} \<in> \<U>"
 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 366 | by (simp only: set_eq_iff transfer_all transfer_iff transfer_mem) | 
| 27468 | 367 | |
| 368 | lemma transfer_ball [transfer_intro]: | |
| 369 |   "\<lbrakk>a \<equiv> Iset (star_n A); \<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
 | |
| 370 |     \<Longrightarrow> \<forall>x\<in>a. p x \<equiv> {n. \<forall>x\<in>A n. P n x} \<in> \<U>"
 | |
| 371 | by (simp only: Ball_def transfer_all transfer_imp transfer_mem) | |
| 372 | ||
| 373 | lemma transfer_bex [transfer_intro]: | |
| 374 |   "\<lbrakk>a \<equiv> Iset (star_n A); \<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
 | |
| 375 |     \<Longrightarrow> \<exists>x\<in>a. p x \<equiv> {n. \<exists>x\<in>A n. P n x} \<in> \<U>"
 | |
| 376 | by (simp only: Bex_def transfer_ex transfer_conj transfer_mem) | |
| 377 | ||
| 378 | lemma transfer_Iset [transfer_intro]: | |
| 379 | "\<lbrakk>a \<equiv> star_n A\<rbrakk> \<Longrightarrow> Iset a \<equiv> Iset (star_n (\<lambda>n. A n))" | |
| 380 | by simp | |
| 381 | ||
| 382 | text {* Nonstandard extensions of sets. *}
 | |
| 383 | ||
| 384 | definition | |
| 385 |   starset :: "'a set \<Rightarrow> 'a star set" ("*s* _" [80] 80) where
 | |
| 386 | "starset A = Iset (star_of A)" | |
| 387 | ||
| 388 | declare starset_def [transfer_unfold] | |
| 389 | ||
| 390 | lemma starset_mem: "(star_of x \<in> *s* A) = (x \<in> A)" | |
| 391 | by (transfer, rule refl) | |
| 392 | ||
| 393 | lemma starset_UNIV: "*s* (UNIV::'a set) = (UNIV::'a star set)" | |
| 394 | by (transfer UNIV_def, rule refl) | |
| 395 | ||
| 396 | lemma starset_empty: "*s* {} = {}"
 | |
| 397 | by (transfer empty_def, rule refl) | |
| 398 | ||
| 399 | lemma starset_insert: "*s* (insert x A) = insert (star_of x) ( *s* A)" | |
| 400 | by (transfer insert_def Un_def, rule refl) | |
| 401 | ||
| 402 | lemma starset_Un: "*s* (A \<union> B) = *s* A \<union> *s* B" | |
| 403 | by (transfer Un_def, rule refl) | |
| 404 | ||
| 405 | lemma starset_Int: "*s* (A \<inter> B) = *s* A \<inter> *s* B" | |
| 406 | by (transfer Int_def, rule refl) | |
| 407 | ||
| 408 | lemma starset_Compl: "*s* -A = -( *s* A)" | |
| 409 | by (transfer Compl_eq, rule refl) | |
| 410 | ||
| 411 | lemma starset_diff: "*s* (A - B) = *s* A - *s* B" | |
| 412 | by (transfer set_diff_eq, rule refl) | |
| 413 | ||
| 414 | lemma starset_image: "*s* (f ` A) = ( *f* f) ` ( *s* A)" | |
| 415 | by (transfer image_def, rule refl) | |
| 416 | ||
| 417 | lemma starset_vimage: "*s* (f -` A) = ( *f* f) -` ( *s* A)" | |
| 418 | by (transfer vimage_def, rule refl) | |
| 419 | ||
| 420 | lemma starset_subset: "( *s* A \<subseteq> *s* B) = (A \<subseteq> B)" | |
| 421 | by (transfer subset_eq, rule refl) | |
| 422 | ||
| 423 | lemma starset_eq: "( *s* A = *s* B) = (A = B)" | |
| 424 | by (transfer, rule refl) | |
| 425 | ||
| 426 | lemmas starset_simps [simp] = | |
| 427 | starset_mem starset_UNIV | |
| 428 | starset_empty starset_insert | |
| 429 | starset_Un starset_Int | |
| 430 | starset_Compl starset_diff | |
| 431 | starset_image starset_vimage | |
| 432 | starset_subset starset_eq | |
| 433 | ||
| 434 | ||
| 435 | subsection {* Syntactic classes *}
 | |
| 436 | ||
| 437 | instantiation star :: (zero) zero | |
| 438 | begin | |
| 439 | ||
| 440 | definition | |
| 37765 | 441 | star_zero_def: "0 \<equiv> star_of 0" | 
| 27468 | 442 | |
| 443 | instance .. | |
| 444 | ||
| 445 | end | |
| 446 | ||
| 447 | instantiation star :: (one) one | |
| 448 | begin | |
| 449 | ||
| 450 | definition | |
| 37765 | 451 | star_one_def: "1 \<equiv> star_of 1" | 
| 27468 | 452 | |
| 453 | instance .. | |
| 454 | ||
| 455 | end | |
| 456 | ||
| 457 | instantiation star :: (plus) plus | |
| 458 | begin | |
| 459 | ||
| 460 | definition | |
| 37765 | 461 | star_add_def: "(op +) \<equiv> *f2* (op +)" | 
| 27468 | 462 | |
| 463 | instance .. | |
| 464 | ||
| 465 | end | |
| 466 | ||
| 467 | instantiation star :: (times) times | |
| 468 | begin | |
| 469 | ||
| 470 | definition | |
| 37765 | 471 | star_mult_def: "(op *) \<equiv> *f2* (op *)" | 
| 27468 | 472 | |
| 473 | instance .. | |
| 474 | ||
| 475 | end | |
| 476 | ||
| 477 | instantiation star :: (uminus) uminus | |
| 478 | begin | |
| 479 | ||
| 480 | definition | |
| 37765 | 481 | star_minus_def: "uminus \<equiv> *f* uminus" | 
| 27468 | 482 | |
| 483 | instance .. | |
| 484 | ||
| 485 | end | |
| 486 | ||
| 487 | instantiation star :: (minus) minus | |
| 488 | begin | |
| 489 | ||
| 490 | definition | |
| 37765 | 491 | star_diff_def: "(op -) \<equiv> *f2* (op -)" | 
| 27468 | 492 | |
| 493 | instance .. | |
| 494 | ||
| 495 | end | |
| 496 | ||
| 497 | instantiation star :: (abs) abs | |
| 498 | begin | |
| 499 | ||
| 500 | definition | |
| 501 | star_abs_def: "abs \<equiv> *f* abs" | |
| 502 | ||
| 503 | instance .. | |
| 504 | ||
| 505 | end | |
| 506 | ||
| 507 | instantiation star :: (sgn) sgn | |
| 508 | begin | |
| 509 | ||
| 510 | definition | |
| 511 | star_sgn_def: "sgn \<equiv> *f* sgn" | |
| 512 | ||
| 513 | instance .. | |
| 514 | ||
| 515 | end | |
| 516 | ||
| 517 | instantiation star :: (inverse) inverse | |
| 518 | begin | |
| 519 | ||
| 520 | definition | |
| 521 | star_divide_def: "(op /) \<equiv> *f2* (op /)" | |
| 522 | ||
| 523 | definition | |
| 524 | star_inverse_def: "inverse \<equiv> *f* inverse" | |
| 525 | ||
| 526 | instance .. | |
| 527 | ||
| 528 | end | |
| 529 | ||
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
35043diff
changeset | 530 | 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 | 531 | |
| 27468 | 532 | instantiation star :: (Divides.div) Divides.div | 
| 533 | begin | |
| 534 | ||
| 535 | definition | |
| 536 | star_div_def: "(op div) \<equiv> *f2* (op div)" | |
| 537 | ||
| 538 | definition | |
| 539 | star_mod_def: "(op mod) \<equiv> *f2* (op mod)" | |
| 540 | ||
| 541 | instance .. | |
| 542 | ||
| 543 | end | |
| 544 | ||
| 545 | instantiation star :: (ord) ord | |
| 546 | begin | |
| 547 | ||
| 548 | definition | |
| 549 | star_le_def: "(op \<le>) \<equiv> *p2* (op \<le>)" | |
| 550 | ||
| 551 | definition | |
| 552 | star_less_def: "(op <) \<equiv> *p2* (op <)" | |
| 553 | ||
| 554 | instance .. | |
| 555 | ||
| 556 | end | |
| 557 | ||
| 558 | lemmas star_class_defs [transfer_unfold] = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 559 | star_zero_def star_one_def | 
| 27468 | 560 | star_add_def star_diff_def star_minus_def | 
| 561 | star_mult_def star_divide_def star_inverse_def | |
| 562 | star_le_def star_less_def star_abs_def star_sgn_def | |
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 563 | star_div_def star_mod_def | 
| 27468 | 564 | |
| 565 | text {* Class operations preserve standard elements *}
 | |
| 566 | ||
| 567 | lemma Standard_zero: "0 \<in> Standard" | |
| 568 | by (simp add: star_zero_def) | |
| 569 | ||
| 570 | lemma Standard_one: "1 \<in> Standard" | |
| 571 | by (simp add: star_one_def) | |
| 572 | ||
| 573 | lemma Standard_add: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x + y \<in> Standard" | |
| 574 | by (simp add: star_add_def) | |
| 575 | ||
| 576 | lemma Standard_diff: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x - y \<in> Standard" | |
| 577 | by (simp add: star_diff_def) | |
| 578 | ||
| 579 | lemma Standard_minus: "x \<in> Standard \<Longrightarrow> - x \<in> Standard" | |
| 580 | by (simp add: star_minus_def) | |
| 581 | ||
| 582 | lemma Standard_mult: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x * y \<in> Standard" | |
| 583 | by (simp add: star_mult_def) | |
| 584 | ||
| 585 | lemma Standard_divide: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x / y \<in> Standard" | |
| 586 | by (simp add: star_divide_def) | |
| 587 | ||
| 588 | lemma Standard_inverse: "x \<in> Standard \<Longrightarrow> inverse x \<in> Standard" | |
| 589 | by (simp add: star_inverse_def) | |
| 590 | ||
| 591 | lemma Standard_abs: "x \<in> Standard \<Longrightarrow> abs x \<in> Standard" | |
| 592 | by (simp add: star_abs_def) | |
| 593 | ||
| 594 | lemma Standard_div: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x div y \<in> Standard" | |
| 595 | by (simp add: star_div_def) | |
| 596 | ||
| 597 | lemma Standard_mod: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x mod y \<in> Standard" | |
| 598 | by (simp add: star_mod_def) | |
| 599 | ||
| 600 | lemmas Standard_simps [simp] = | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 601 | Standard_zero Standard_one | 
| 27468 | 602 | Standard_add Standard_diff Standard_minus | 
| 603 | Standard_mult Standard_divide Standard_inverse | |
| 604 | Standard_abs Standard_div Standard_mod | |
| 605 | ||
| 606 | text {* @{term star_of} preserves class operations *}
 | |
| 607 | ||
| 608 | lemma star_of_add: "star_of (x + y) = star_of x + star_of y" | |
| 609 | by transfer (rule refl) | |
| 610 | ||
| 611 | lemma star_of_diff: "star_of (x - y) = star_of x - star_of y" | |
| 612 | by transfer (rule refl) | |
| 613 | ||
| 614 | lemma star_of_minus: "star_of (-x) = - star_of x" | |
| 615 | by transfer (rule refl) | |
| 616 | ||
| 617 | lemma star_of_mult: "star_of (x * y) = star_of x * star_of y" | |
| 618 | by transfer (rule refl) | |
| 619 | ||
| 620 | lemma star_of_divide: "star_of (x / y) = star_of x / star_of y" | |
| 621 | by transfer (rule refl) | |
| 622 | ||
| 623 | lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)" | |
| 624 | by transfer (rule refl) | |
| 625 | ||
| 626 | lemma star_of_div: "star_of (x div y) = star_of x div star_of y" | |
| 627 | by transfer (rule refl) | |
| 628 | ||
| 629 | lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y" | |
| 630 | by transfer (rule refl) | |
| 631 | ||
| 632 | lemma star_of_abs: "star_of (abs x) = abs (star_of x)" | |
| 633 | by transfer (rule refl) | |
| 634 | ||
| 635 | text {* @{term star_of} preserves numerals *}
 | |
| 636 | ||
| 637 | lemma star_of_zero: "star_of 0 = 0" | |
| 638 | by transfer (rule refl) | |
| 639 | ||
| 640 | lemma star_of_one: "star_of 1 = 1" | |
| 641 | by transfer (rule refl) | |
| 642 | ||
| 643 | text {* @{term star_of} preserves orderings *}
 | |
| 644 | ||
| 645 | lemma star_of_less: "(star_of x < star_of y) = (x < y)" | |
| 646 | by transfer (rule refl) | |
| 647 | ||
| 648 | lemma star_of_le: "(star_of x \<le> star_of y) = (x \<le> y)" | |
| 649 | by transfer (rule refl) | |
| 650 | ||
| 651 | lemma star_of_eq: "(star_of x = star_of y) = (x = y)" | |
| 652 | by transfer (rule refl) | |
| 653 | ||
| 654 | text{*As above, for 0*}
 | |
| 655 | ||
| 656 | lemmas star_of_0_less = star_of_less [of 0, simplified star_of_zero] | |
| 657 | lemmas star_of_0_le = star_of_le [of 0, simplified star_of_zero] | |
| 658 | lemmas star_of_0_eq = star_of_eq [of 0, simplified star_of_zero] | |
| 659 | ||
| 660 | lemmas star_of_less_0 = star_of_less [of _ 0, simplified star_of_zero] | |
| 661 | lemmas star_of_le_0 = star_of_le [of _ 0, simplified star_of_zero] | |
| 662 | lemmas star_of_eq_0 = star_of_eq [of _ 0, simplified star_of_zero] | |
| 663 | ||
| 664 | text{*As above, for 1*}
 | |
| 665 | ||
| 666 | lemmas star_of_1_less = star_of_less [of 1, simplified star_of_one] | |
| 667 | lemmas star_of_1_le = star_of_le [of 1, simplified star_of_one] | |
| 668 | lemmas star_of_1_eq = star_of_eq [of 1, simplified star_of_one] | |
| 669 | ||
| 670 | lemmas star_of_less_1 = star_of_less [of _ 1, simplified star_of_one] | |
| 671 | lemmas star_of_le_1 = star_of_le [of _ 1, simplified star_of_one] | |
| 672 | lemmas star_of_eq_1 = star_of_eq [of _ 1, simplified star_of_one] | |
| 673 | ||
| 674 | lemmas star_of_simps [simp] = | |
| 675 | star_of_add star_of_diff star_of_minus | |
| 676 | star_of_mult star_of_divide star_of_inverse | |
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 677 | star_of_div star_of_mod star_of_abs | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 678 | star_of_zero star_of_one | 
| 27468 | 679 | star_of_less star_of_le star_of_eq | 
| 680 | star_of_0_less star_of_0_le star_of_0_eq | |
| 681 | star_of_less_0 star_of_le_0 star_of_eq_0 | |
| 682 | star_of_1_less star_of_1_le star_of_1_eq | |
| 683 | star_of_less_1 star_of_le_1 star_of_eq_1 | |
| 684 | ||
| 685 | subsection {* Ordering and lattice classes *}
 | |
| 686 | ||
| 687 | instance star :: (order) order | |
| 688 | apply (intro_classes) | |
| 27682 | 689 | apply (transfer, rule less_le_not_le) | 
| 27468 | 690 | apply (transfer, rule order_refl) | 
| 691 | apply (transfer, erule (1) order_trans) | |
| 692 | apply (transfer, erule (1) order_antisym) | |
| 693 | done | |
| 694 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 695 | instantiation star :: (semilattice_inf) semilattice_inf | 
| 27468 | 696 | begin | 
| 697 | ||
| 698 | definition | |
| 699 | star_inf_def [transfer_unfold]: "inf \<equiv> *f2* inf" | |
| 700 | ||
| 701 | instance | |
| 702 | by default (transfer star_inf_def, auto)+ | |
| 703 | ||
| 704 | end | |
| 705 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 706 | instantiation star :: (semilattice_sup) semilattice_sup | 
| 27468 | 707 | begin | 
| 708 | ||
| 709 | definition | |
| 710 | star_sup_def [transfer_unfold]: "sup \<equiv> *f2* sup" | |
| 711 | ||
| 712 | instance | |
| 713 | by default (transfer star_sup_def, auto)+ | |
| 714 | ||
| 715 | end | |
| 716 | ||
| 717 | instance star :: (lattice) lattice .. | |
| 718 | ||
| 719 | instance star :: (distrib_lattice) distrib_lattice | |
| 720 | by default (transfer, auto simp add: sup_inf_distrib1) | |
| 721 | ||
| 722 | lemma Standard_inf [simp]: | |
| 723 | "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> inf x y \<in> Standard" | |
| 724 | by (simp add: star_inf_def) | |
| 725 | ||
| 726 | lemma Standard_sup [simp]: | |
| 727 | "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> sup x y \<in> Standard" | |
| 728 | by (simp add: star_sup_def) | |
| 729 | ||
| 730 | lemma star_of_inf [simp]: "star_of (inf x y) = inf (star_of x) (star_of y)" | |
| 731 | by transfer (rule refl) | |
| 732 | ||
| 733 | lemma star_of_sup [simp]: "star_of (sup x y) = sup (star_of x) (star_of y)" | |
| 734 | by transfer (rule refl) | |
| 735 | ||
| 736 | instance star :: (linorder) linorder | |
| 737 | by (intro_classes, transfer, rule linorder_linear) | |
| 738 | ||
| 739 | lemma star_max_def [transfer_unfold]: "max = *f2* max" | |
| 740 | apply (rule ext, rule ext) | |
| 741 | apply (unfold max_def, transfer, fold max_def) | |
| 742 | apply (rule refl) | |
| 743 | done | |
| 744 | ||
| 745 | lemma star_min_def [transfer_unfold]: "min = *f2* min" | |
| 746 | apply (rule ext, rule ext) | |
| 747 | apply (unfold min_def, transfer, fold min_def) | |
| 748 | apply (rule refl) | |
| 749 | done | |
| 750 | ||
| 751 | lemma Standard_max [simp]: | |
| 752 | "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> max x y \<in> Standard" | |
| 753 | by (simp add: star_max_def) | |
| 754 | ||
| 755 | lemma Standard_min [simp]: | |
| 756 | "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> min x y \<in> Standard" | |
| 757 | by (simp add: star_min_def) | |
| 758 | ||
| 759 | lemma star_of_max [simp]: "star_of (max x y) = max (star_of x) (star_of y)" | |
| 760 | by transfer (rule refl) | |
| 761 | ||
| 762 | lemma star_of_min [simp]: "star_of (min x y) = min (star_of x) (star_of y)" | |
| 763 | by transfer (rule refl) | |
| 764 | ||
| 765 | ||
| 766 | subsection {* Ordered group classes *}
 | |
| 767 | ||
| 768 | instance star :: (semigroup_add) semigroup_add | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56256diff
changeset | 769 | by (intro_classes, transfer, rule add.assoc) | 
| 27468 | 770 | |
| 771 | instance star :: (ab_semigroup_add) ab_semigroup_add | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56256diff
changeset | 772 | by (intro_classes, transfer, rule add.commute) | 
| 27468 | 773 | |
| 774 | instance star :: (semigroup_mult) semigroup_mult | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56256diff
changeset | 775 | by (intro_classes, transfer, rule mult.assoc) | 
| 27468 | 776 | |
| 777 | instance star :: (ab_semigroup_mult) ab_semigroup_mult | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56256diff
changeset | 778 | by (intro_classes, transfer, rule mult.commute) | 
| 27468 | 779 | |
| 780 | instance star :: (comm_monoid_add) comm_monoid_add | |
| 28059 | 781 | by (intro_classes, transfer, rule comm_monoid_add_class.add_0) | 
| 27468 | 782 | |
| 783 | instance star :: (monoid_mult) monoid_mult | |
| 784 | apply (intro_classes) | |
| 785 | apply (transfer, rule mult_1_left) | |
| 786 | apply (transfer, rule mult_1_right) | |
| 787 | done | |
| 788 | ||
| 789 | instance star :: (comm_monoid_mult) comm_monoid_mult | |
| 790 | by (intro_classes, transfer, rule mult_1) | |
| 791 | ||
| 792 | instance star :: (cancel_semigroup_add) cancel_semigroup_add | |
| 793 | apply (intro_classes) | |
| 794 | apply (transfer, erule add_left_imp_eq) | |
| 795 | apply (transfer, erule add_right_imp_eq) | |
| 796 | done | |
| 797 | ||
| 798 | instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add | |
| 799 | by (intro_classes, transfer, rule add_imp_eq) | |
| 800 | ||
| 29904 | 801 | instance star :: (cancel_comm_monoid_add) cancel_comm_monoid_add .. | 
| 802 | ||
| 27468 | 803 | instance star :: (ab_group_add) ab_group_add | 
| 804 | apply (intro_classes) | |
| 805 | apply (transfer, rule left_minus) | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
49962diff
changeset | 806 | apply (transfer, rule diff_conv_add_uminus) | 
| 27468 | 807 | done | 
| 808 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 809 | instance star :: (ordered_ab_semigroup_add) ordered_ab_semigroup_add | 
| 27468 | 810 | by (intro_classes, transfer, rule add_left_mono) | 
| 811 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 812 | instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add .. | 
| 27468 | 813 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 814 | instance star :: (ordered_ab_semigroup_add_imp_le) ordered_ab_semigroup_add_imp_le | 
| 27468 | 815 | by (intro_classes, transfer, rule add_le_imp_le_left) | 
| 816 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 817 | instance star :: (ordered_comm_monoid_add) ordered_comm_monoid_add .. | 
| 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 818 | instance star :: (ordered_ab_group_add) ordered_ab_group_add .. | 
| 27468 | 819 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 820 | instance star :: (ordered_ab_group_add_abs) ordered_ab_group_add_abs | 
| 27468 | 821 | by intro_classes (transfer, | 
| 822 | simp add: abs_ge_self abs_leI abs_triangle_ineq)+ | |
| 823 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 824 | instance star :: (linordered_cancel_ab_semigroup_add) linordered_cancel_ab_semigroup_add .. | 
| 27468 | 825 | |
| 826 | ||
| 827 | subsection {* Ring and field classes *}
 | |
| 828 | ||
| 829 | instance star :: (semiring) semiring | |
| 830 | apply (intro_classes) | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49834diff
changeset | 831 | apply (transfer, rule distrib_right) | 
| 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49834diff
changeset | 832 | apply (transfer, rule distrib_left) | 
| 27468 | 833 | done | 
| 834 | ||
| 835 | instance star :: (semiring_0) semiring_0 | |
| 836 | by intro_classes (transfer, simp)+ | |
| 837 | ||
| 838 | instance star :: (semiring_0_cancel) semiring_0_cancel .. | |
| 839 | ||
| 840 | instance star :: (comm_semiring) comm_semiring | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49834diff
changeset | 841 | by (intro_classes, transfer, rule distrib_right) | 
| 27468 | 842 | |
| 843 | instance star :: (comm_semiring_0) comm_semiring_0 .. | |
| 844 | instance star :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. | |
| 845 | ||
| 846 | instance star :: (zero_neq_one) zero_neq_one | |
| 847 | by (intro_classes, transfer, rule zero_neq_one) | |
| 848 | ||
| 849 | instance star :: (semiring_1) semiring_1 .. | |
| 850 | instance star :: (comm_semiring_1) comm_semiring_1 .. | |
| 851 | ||
| 852 | instance star :: (no_zero_divisors) no_zero_divisors | |
| 853 | by (intro_classes, transfer, rule no_zero_divisors) | |
| 854 | ||
| 855 | instance star :: (semiring_1_cancel) semiring_1_cancel .. | |
| 856 | instance star :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. | |
| 857 | instance star :: (ring) ring .. | |
| 858 | instance star :: (comm_ring) comm_ring .. | |
| 859 | instance star :: (ring_1) ring_1 .. | |
| 860 | instance star :: (comm_ring_1) comm_ring_1 .. | |
| 861 | instance star :: (ring_no_zero_divisors) ring_no_zero_divisors .. | |
| 862 | instance star :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors .. | |
| 863 | instance star :: (idom) idom .. | |
| 864 | ||
| 865 | instance star :: (division_ring) division_ring | |
| 866 | apply (intro_classes) | |
| 867 | apply (transfer, erule left_inverse) | |
| 868 | apply (transfer, erule right_inverse) | |
| 35083 | 869 | apply (transfer, fact divide_inverse) | 
| 27468 | 870 | done | 
| 871 | ||
| 36412 | 872 | instance star :: (division_ring_inverse_zero) division_ring_inverse_zero | 
| 873 | by (intro_classes, transfer, rule inverse_zero) | |
| 874 | ||
| 27468 | 875 | instance star :: (field) field | 
| 876 | apply (intro_classes) | |
| 877 | apply (transfer, erule left_inverse) | |
| 878 | apply (transfer, rule divide_inverse) | |
| 879 | done | |
| 880 | ||
| 36412 | 881 | instance star :: (field_inverse_zero) field_inverse_zero | 
| 882 | apply intro_classes | |
| 883 | apply (rule inverse_zero) | |
| 884 | done | |
| 27468 | 885 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 886 | instance star :: (ordered_semiring) ordered_semiring | 
| 27468 | 887 | apply (intro_classes) | 
| 888 | apply (transfer, erule (1) mult_left_mono) | |
| 889 | apply (transfer, erule (1) mult_right_mono) | |
| 890 | done | |
| 891 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 892 | instance star :: (ordered_cancel_semiring) ordered_cancel_semiring .. | 
| 27468 | 893 | |
| 35043 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 haftmann parents: 
35035diff
changeset | 894 | instance star :: (linordered_semiring_strict) linordered_semiring_strict | 
| 27468 | 895 | apply (intro_classes) | 
| 896 | apply (transfer, erule (1) mult_strict_left_mono) | |
| 897 | apply (transfer, erule (1) mult_strict_right_mono) | |
| 898 | done | |
| 899 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 900 | instance star :: (ordered_comm_semiring) ordered_comm_semiring | 
| 38642 
8fa437809c67
dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
 haftmann parents: 
38621diff
changeset | 901 | by (intro_classes, transfer, rule mult_left_mono) | 
| 27468 | 902 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 903 | instance star :: (ordered_cancel_comm_semiring) ordered_cancel_comm_semiring .. | 
| 27468 | 904 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 905 | instance star :: (linordered_comm_semiring_strict) linordered_comm_semiring_strict | 
| 38642 
8fa437809c67
dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
 haftmann parents: 
38621diff
changeset | 906 | by (intro_classes, transfer, rule mult_strict_left_mono) | 
| 27468 | 907 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 908 | instance star :: (ordered_ring) ordered_ring .. | 
| 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 909 | instance star :: (ordered_ring_abs) ordered_ring_abs | 
| 27468 | 910 | by intro_classes (transfer, rule abs_eq_mult) | 
| 911 | ||
| 912 | instance star :: (abs_if) abs_if | |
| 913 | by (intro_classes, transfer, rule abs_if) | |
| 914 | ||
| 915 | instance star :: (sgn_if) sgn_if | |
| 916 | by (intro_classes, transfer, rule sgn_if) | |
| 917 | ||
| 35043 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 haftmann parents: 
35035diff
changeset | 918 | 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 | 919 | instance star :: (ordered_comm_ring) ordered_comm_ring .. | 
| 27468 | 920 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 921 | instance star :: (linordered_semidom) linordered_semidom | 
| 27468 | 922 | by (intro_classes, transfer, rule zero_less_one) | 
| 923 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 924 | instance star :: (linordered_idom) linordered_idom .. | 
| 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
31021diff
changeset | 925 | instance star :: (linordered_field) linordered_field .. | 
| 36414 | 926 | instance star :: (linordered_field_inverse_zero) linordered_field_inverse_zero .. | 
| 27468 | 927 | |
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 928 | |
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 929 | subsection {* Power *}
 | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 930 | |
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 931 | lemma star_power_def [transfer_unfold]: | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 932 | "(op ^) \<equiv> \<lambda>x n. ( *f* (\<lambda>x. x ^ n)) x" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 933 | proof (rule eq_reflection, rule ext, rule ext) | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 934 | fix n :: nat | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 935 | show "\<And>x::'a star. x ^ n = ( *f* (\<lambda>x. x ^ n)) x" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 936 | proof (induct n) | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 937 | case 0 | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 938 | have "\<And>x::'a star. ( *f* (\<lambda>x. 1)) x = 1" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 939 | by transfer simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 940 | then show ?case by simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 941 | next | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 942 | case (Suc n) | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 943 | have "\<And>x::'a star. x * ( *f* (\<lambda>x\<Colon>'a. x ^ n)) x = ( *f* (\<lambda>x\<Colon>'a. x * x ^ n)) x" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 944 | by transfer simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 945 | with Suc show ?case by simp | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 946 | qed | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 947 | qed | 
| 27468 | 948 | |
| 30968 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 949 | lemma Standard_power [simp]: "x \<in> Standard \<Longrightarrow> x ^ n \<in> Standard" | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 950 | by (simp add: star_power_def) | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 951 | |
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 952 | 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 | 953 | by transfer (rule refl) | 
| 
10fef94f40fc
adaptions due to rearrangment of power operation
 haftmann parents: 
30729diff
changeset | 954 | |
| 27468 | 955 | |
| 956 | subsection {* Number classes *}
 | |
| 957 | ||
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 958 | instance star :: (numeral) numeral .. | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 959 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 960 | lemma star_numeral_def [transfer_unfold]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 961 | "numeral k = star_of (numeral k)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 962 | by (induct k, simp_all only: numeral.simps star_of_one star_of_add) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 963 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 964 | lemma Standard_numeral [simp]: "numeral k \<in> Standard" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 965 | by (simp add: star_numeral_def) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 966 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 967 | lemma star_of_numeral [simp]: "star_of (numeral k) = numeral k" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 968 | by transfer (rule refl) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 969 | |
| 27468 | 970 | lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)" | 
| 971 | by (induct n, simp_all) | |
| 972 | ||
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 973 | lemmas star_of_compare_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 974 | star_of_less [of "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 975 | star_of_le [of "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 976 | star_of_eq [of "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 977 | star_of_less [of _ "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 978 | star_of_le [of _ "numeral k", simplified star_of_numeral] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46008diff
changeset | 979 | star_of_eq [of _ "numeral k", simplified star_of_numeral] | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 980 | star_of_less [of "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 981 | star_of_le [of "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 982 | star_of_eq [of "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 983 | star_of_less [of _ "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 984 | star_of_le [of _ "- numeral k", simplified star_of_numeral] | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 985 | 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 | 986 | |
| 27468 | 987 | lemma Standard_of_nat [simp]: "of_nat n \<in> Standard" | 
| 988 | by (simp add: star_of_nat_def) | |
| 989 | ||
| 990 | lemma star_of_of_nat [simp]: "star_of (of_nat n) = of_nat n" | |
| 991 | by transfer (rule refl) | |
| 992 | ||
| 993 | lemma star_of_int_def [transfer_unfold]: "of_int z = star_of (of_int z)" | |
| 994 | by (rule_tac z=z in int_diff_cases, simp) | |
| 995 | ||
| 996 | lemma Standard_of_int [simp]: "of_int z \<in> Standard" | |
| 997 | by (simp add: star_of_int_def) | |
| 998 | ||
| 999 | lemma star_of_of_int [simp]: "star_of (of_int z) = of_int z" | |
| 1000 | by transfer (rule refl) | |
| 1001 | ||
| 38621 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37765diff
changeset | 1002 | instance star :: (semiring_char_0) semiring_char_0 proof | 
| 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37765diff
changeset | 1003 | have "inj (star_of :: 'a \<Rightarrow> 'a star)" by (rule injI) simp | 
| 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37765diff
changeset | 1004 | then have "inj (star_of \<circ> of_nat :: nat \<Rightarrow> 'a star)" using inj_of_nat by (rule inj_comp) | 
| 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37765diff
changeset | 1005 | then show "inj (of_nat :: nat \<Rightarrow> 'a star)" by (simp add: comp_def) | 
| 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37765diff
changeset | 1006 | qed | 
| 27468 | 1007 | |
| 1008 | instance star :: (ring_char_0) ring_char_0 .. | |
| 1009 | ||
| 1010 | ||
| 1011 | subsection {* Finite class *}
 | |
| 1012 | ||
| 1013 | lemma starset_finite: "finite A \<Longrightarrow> *s* A = star_of ` A" | |
| 1014 | by (erule finite_induct, simp_all) | |
| 1015 | ||
| 1016 | instance star :: (finite) finite | |
| 1017 | apply (intro_classes) | |
| 1018 | apply (subst starset_UNIV [symmetric]) | |
| 1019 | apply (subst starset_finite [OF finite]) | |
| 1020 | apply (rule finite_imageI [OF finite]) | |
| 1021 | done | |
| 1022 | ||
| 1023 | end |