equal
deleted
inserted
replaced
11 begin |
11 begin |
12 |
12 |
13 subsection {* A Free Ultrafilter over the Naturals *} |
13 subsection {* A Free Ultrafilter over the Naturals *} |
14 |
14 |
15 definition |
15 definition |
16 FreeUltrafilterNat :: "nat set set" ("\<U>") |
16 FreeUltrafilterNat :: "nat set set" ("\<U>") where |
17 "\<U> = (SOME U. freeultrafilter U)" |
17 "\<U> = (SOME U. freeultrafilter U)" |
18 |
18 |
19 lemma freeultrafilter_FUFNat: "freeultrafilter \<U>" |
19 lemma freeultrafilter_FUFNat: "freeultrafilter \<U>" |
20 apply (unfold FreeUltrafilterNat_def) |
20 apply (unfold FreeUltrafilterNat_def) |
21 apply (rule someI_ex) |
21 apply (rule someI_ex) |
34 |
34 |
35 |
35 |
36 subsection {* Definition of @{text star} type constructor *} |
36 subsection {* Definition of @{text star} type constructor *} |
37 |
37 |
38 definition |
38 definition |
39 starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set" |
39 starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set" where |
40 "starrel = {(X,Y). {n. X n = Y n} \<in> \<U>}" |
40 "starrel = {(X,Y). {n. X n = Y n} \<in> \<U>}" |
41 |
41 |
42 typedef 'a star = "(UNIV :: (nat \<Rightarrow> 'a) set) // starrel" |
42 typedef 'a star = "(UNIV :: (nat \<Rightarrow> 'a) set) // starrel" |
43 by (auto intro: quotientI) |
43 by (auto intro: quotientI) |
44 |
44 |
45 definition |
45 definition |
46 star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star" |
46 star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star" where |
47 "star_n X = Abs_star (starrel `` {X})" |
47 "star_n X = Abs_star (starrel `` {X})" |
48 |
48 |
49 theorem star_cases [case_names star_n, cases type: star]: |
49 theorem star_cases [case_names star_n, cases type: star]: |
50 "(\<And>X. x = star_n X \<Longrightarrow> P) \<Longrightarrow> P" |
50 "(\<And>X. x = star_n X \<Longrightarrow> P) \<Longrightarrow> P" |
51 by (cases x, unfold star_n_def star_def, erule quotientE, fast) |
51 by (cases x, unfold star_n_def star_def, erule quotientE, fast) |
155 |
155 |
156 |
156 |
157 subsection {* Standard elements *} |
157 subsection {* Standard elements *} |
158 |
158 |
159 definition |
159 definition |
160 star_of :: "'a \<Rightarrow> 'a star" |
160 star_of :: "'a \<Rightarrow> 'a star" where |
161 "star_of x == star_n (\<lambda>n. x)" |
161 "star_of x == star_n (\<lambda>n. x)" |
162 |
162 |
163 Standard :: "'a star set" |
163 definition |
|
164 Standard :: "'a star set" where |
164 "Standard = range star_of" |
165 "Standard = range star_of" |
165 |
166 |
166 text {* Transfer tactic should remove occurrences of @{term star_of} *} |
167 text {* Transfer tactic should remove occurrences of @{term star_of} *} |
167 setup {* Transfer.add_const "StarDef.star_of" *} |
168 setup {* Transfer.add_const "StarDef.star_of" *} |
168 |
169 |
176 |
177 |
177 |
178 |
178 subsection {* Internal functions *} |
179 subsection {* Internal functions *} |
179 |
180 |
180 definition |
181 definition |
181 Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300) |
182 Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300) where |
182 "Ifun f \<equiv> \<lambda>x. Abs_star |
183 "Ifun f \<equiv> \<lambda>x. Abs_star |
183 (\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})" |
184 (\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})" |
184 |
185 |
185 lemma Ifun_congruent2: |
186 lemma Ifun_congruent2: |
186 "congruent2 starrel starrel (\<lambda>F X. starrel``{\<lambda>n. F n (X n)})" |
187 "congruent2 starrel starrel (\<lambda>F X. starrel``{\<lambda>n. F n (X n)})" |
205 by (auto simp add: Standard_def) |
206 by (auto simp add: Standard_def) |
206 |
207 |
207 text {* Nonstandard extensions of functions *} |
208 text {* Nonstandard extensions of functions *} |
208 |
209 |
209 definition |
210 definition |
210 starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a star \<Rightarrow> 'b star)" |
211 starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a star \<Rightarrow> 'b star)" ("*f* _" [80] 80) where |
211 ("*f* _" [80] 80) |
|
212 "starfun f == \<lambda>x. star_of f \<star> x" |
212 "starfun f == \<lambda>x. star_of f \<star> x" |
213 |
213 |
|
214 definition |
214 starfun2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a star \<Rightarrow> 'b star \<Rightarrow> 'c star)" |
215 starfun2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a star \<Rightarrow> 'b star \<Rightarrow> 'c star)" |
215 ("*f2* _" [80] 80) |
216 ("*f2* _" [80] 80) where |
216 "starfun2 f == \<lambda>x y. star_of f \<star> x \<star> y" |
217 "starfun2 f == \<lambda>x y. star_of f \<star> x \<star> y" |
217 |
218 |
218 declare starfun_def [transfer_unfold] |
219 declare starfun_def [transfer_unfold] |
219 declare starfun2_def [transfer_unfold] |
220 declare starfun2_def [transfer_unfold] |
220 |
221 |
240 |
241 |
241 |
242 |
242 subsection {* Internal predicates *} |
243 subsection {* Internal predicates *} |
243 |
244 |
244 definition |
245 definition |
245 unstar :: "bool star \<Rightarrow> bool" |
246 unstar :: "bool star \<Rightarrow> bool" where |
246 "unstar b = (b = star_of True)" |
247 "unstar b = (b = star_of True)" |
247 |
248 |
248 lemma unstar_star_n: "unstar (star_n P) = ({n. P n} \<in> \<U>)" |
249 lemma unstar_star_n: "unstar (star_n P) = ({n. P n} \<in> \<U>)" |
249 by (simp add: unstar_def star_of_def star_n_eq_iff) |
250 by (simp add: unstar_def star_of_def star_n_eq_iff) |
250 |
251 |
257 lemma transfer_unstar [transfer_intro]: |
258 lemma transfer_unstar [transfer_intro]: |
258 "p \<equiv> star_n P \<Longrightarrow> unstar p \<equiv> {n. P n} \<in> \<U>" |
259 "p \<equiv> star_n P \<Longrightarrow> unstar p \<equiv> {n. P n} \<in> \<U>" |
259 by (simp only: unstar_star_n) |
260 by (simp only: unstar_star_n) |
260 |
261 |
261 definition |
262 definition |
262 starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool" |
263 starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool" ("*p* _" [80] 80) where |
263 ("*p* _" [80] 80) |
|
264 "*p* P = (\<lambda>x. unstar (star_of P \<star> x))" |
264 "*p* P = (\<lambda>x. unstar (star_of P \<star> x))" |
265 |
265 |
266 starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool" |
266 definition |
267 ("*p2* _" [80] 80) |
267 starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool" ("*p2* _" [80] 80) where |
268 "*p2* P = (\<lambda>x y. unstar (star_of P \<star> x \<star> y))" |
268 "*p2* P = (\<lambda>x y. unstar (star_of P \<star> x \<star> y))" |
269 |
269 |
270 declare starP_def [transfer_unfold] |
270 declare starP_def [transfer_unfold] |
271 declare starP2_def [transfer_unfold] |
271 declare starP2_def [transfer_unfold] |
272 |
272 |
285 |
285 |
286 |
286 |
287 subsection {* Internal sets *} |
287 subsection {* Internal sets *} |
288 |
288 |
289 definition |
289 definition |
290 Iset :: "'a set star \<Rightarrow> 'a star set" |
290 Iset :: "'a set star \<Rightarrow> 'a star set" where |
291 "Iset A = {x. ( *p2* op \<in>) x A}" |
291 "Iset A = {x. ( *p2* op \<in>) x A}" |
292 |
292 |
293 lemma Iset_star_n: |
293 lemma Iset_star_n: |
294 "(star_n X \<in> Iset (star_n A)) = ({n. X n \<in> A n} \<in> \<U>)" |
294 "(star_n X \<in> Iset (star_n A)) = ({n. X n \<in> A n} \<in> \<U>)" |
295 by (simp add: Iset_def starP2_star_n) |
295 by (simp add: Iset_def starP2_star_n) |
327 by simp |
327 by simp |
328 |
328 |
329 text {* Nonstandard extensions of sets. *} |
329 text {* Nonstandard extensions of sets. *} |
330 |
330 |
331 definition |
331 definition |
332 starset :: "'a set \<Rightarrow> 'a star set" ("*s* _" [80] 80) |
332 starset :: "'a set \<Rightarrow> 'a star set" ("*s* _" [80] 80) where |
333 "starset A = Iset (star_of A)" |
333 "starset A = Iset (star_of A)" |
334 |
334 |
335 declare starset_def [transfer_unfold] |
335 declare starset_def [transfer_unfold] |
336 |
336 |
337 lemma starset_mem: "(star_of x \<in> *s* A) = (x \<in> A)" |
337 lemma starset_mem: "(star_of x \<in> *s* A) = (x \<in> A)" |