| author | haftmann | 
| Fri, 05 Feb 2010 14:33:50 +0100 | |
| changeset 35028 | 108662d50512 | 
| parent 29920 | b95f5b8b93dd | 
| child 37765 | 26bdfb7b680b | 
| permissions | -rw-r--r-- | 
| 27468 | 1  | 
(* Title : HyperNat.thy  | 
2  | 
Author : Jacques D. Fleuriot  | 
|
3  | 
Copyright : 1998 University of Cambridge  | 
|
4  | 
||
5  | 
Converted to Isar and polished by lcp  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
header{*Hypernatural numbers*}
 | 
|
9  | 
||
10  | 
theory HyperNat  | 
|
11  | 
imports StarDef  | 
|
12  | 
begin  | 
|
13  | 
||
14  | 
types hypnat = "nat star"  | 
|
15  | 
||
16  | 
abbreviation  | 
|
17  | 
hypnat_of_nat :: "nat => nat star" where  | 
|
18  | 
"hypnat_of_nat == star_of"  | 
|
19  | 
||
20  | 
definition  | 
|
21  | 
hSuc :: "hypnat => hypnat" where  | 
|
| 28562 | 22  | 
hSuc_def [transfer_unfold, code del]: "hSuc = *f* Suc"  | 
| 27468 | 23  | 
|
24  | 
subsection{*Properties Transferred from Naturals*}
 | 
|
25  | 
||
26  | 
lemma hSuc_not_zero [iff]: "\<And>m. hSuc m \<noteq> 0"  | 
|
27  | 
by transfer (rule Suc_not_Zero)  | 
|
28  | 
||
29  | 
lemma zero_not_hSuc [iff]: "\<And>m. 0 \<noteq> hSuc m"  | 
|
30  | 
by transfer (rule Zero_not_Suc)  | 
|
31  | 
||
32  | 
lemma hSuc_hSuc_eq [iff]: "\<And>m n. (hSuc m = hSuc n) = (m = n)"  | 
|
33  | 
by transfer (rule nat.inject)  | 
|
34  | 
||
35  | 
lemma zero_less_hSuc [iff]: "\<And>n. 0 < hSuc n"  | 
|
36  | 
by transfer (rule zero_less_Suc)  | 
|
37  | 
||
38  | 
lemma hypnat_minus_zero [simp]: "!!z. z - z = (0::hypnat)"  | 
|
39  | 
by transfer (rule diff_self_eq_0)  | 
|
40  | 
||
41  | 
lemma hypnat_diff_0_eq_0 [simp]: "!!n. (0::hypnat) - n = 0"  | 
|
42  | 
by transfer (rule diff_0_eq_0)  | 
|
43  | 
||
44  | 
lemma hypnat_add_is_0 [iff]: "!!m n. (m+n = (0::hypnat)) = (m=0 & n=0)"  | 
|
45  | 
by transfer (rule add_is_0)  | 
|
46  | 
||
47  | 
lemma hypnat_diff_diff_left: "!!i j k. (i::hypnat) - j - k = i - (j+k)"  | 
|
48  | 
by transfer (rule diff_diff_left)  | 
|
49  | 
||
50  | 
lemma hypnat_diff_commute: "!!i j k. (i::hypnat) - j - k = i-k-j"  | 
|
51  | 
by transfer (rule diff_commute)  | 
|
52  | 
||
53  | 
lemma hypnat_diff_add_inverse [simp]: "!!m n. ((n::hypnat) + m) - n = m"  | 
|
54  | 
by transfer (rule diff_add_inverse)  | 
|
55  | 
||
56  | 
lemma hypnat_diff_add_inverse2 [simp]: "!!m n. ((m::hypnat) + n) - n = m"  | 
|
57  | 
by transfer (rule diff_add_inverse2)  | 
|
58  | 
||
59  | 
lemma hypnat_diff_cancel [simp]: "!!k m n. ((k::hypnat) + m) - (k+n) = m - n"  | 
|
60  | 
by transfer (rule diff_cancel)  | 
|
61  | 
||
62  | 
lemma hypnat_diff_cancel2 [simp]: "!!k m n. ((m::hypnat) + k) - (n+k) = m - n"  | 
|
63  | 
by transfer (rule diff_cancel2)  | 
|
64  | 
||
65  | 
lemma hypnat_diff_add_0 [simp]: "!!m n. (n::hypnat) - (n+m) = (0::hypnat)"  | 
|
66  | 
by transfer (rule diff_add_0)  | 
|
67  | 
||
68  | 
lemma hypnat_diff_mult_distrib: "!!k m n. ((m::hypnat) - n) * k = (m * k) - (n * k)"  | 
|
69  | 
by transfer (rule diff_mult_distrib)  | 
|
70  | 
||
71  | 
lemma hypnat_diff_mult_distrib2: "!!k m n. (k::hypnat) * (m - n) = (k * m) - (k * n)"  | 
|
72  | 
by transfer (rule diff_mult_distrib2)  | 
|
73  | 
||
74  | 
lemma hypnat_le_zero_cancel [iff]: "!!n. (n \<le> (0::hypnat)) = (n = 0)"  | 
|
75  | 
by transfer (rule le_0_eq)  | 
|
76  | 
||
77  | 
lemma hypnat_mult_is_0 [simp]: "!!m n. (m*n = (0::hypnat)) = (m=0 | n=0)"  | 
|
78  | 
by transfer (rule mult_is_0)  | 
|
79  | 
||
80  | 
lemma hypnat_diff_is_0_eq [simp]: "!!m n. ((m::hypnat) - n = 0) = (m \<le> n)"  | 
|
81  | 
by transfer (rule diff_is_0_eq)  | 
|
82  | 
||
83  | 
lemma hypnat_not_less0 [iff]: "!!n. ~ n < (0::hypnat)"  | 
|
84  | 
by transfer (rule not_less0)  | 
|
85  | 
||
86  | 
lemma hypnat_less_one [iff]:  | 
|
87  | 
"!!n. (n < (1::hypnat)) = (n=0)"  | 
|
88  | 
by transfer (rule less_one)  | 
|
89  | 
||
90  | 
lemma hypnat_add_diff_inverse: "!!m n. ~ m<n ==> n+(m-n) = (m::hypnat)"  | 
|
91  | 
by transfer (rule add_diff_inverse)  | 
|
92  | 
||
93  | 
lemma hypnat_le_add_diff_inverse [simp]: "!!m n. n \<le> m ==> n+(m-n) = (m::hypnat)"  | 
|
94  | 
by transfer (rule le_add_diff_inverse)  | 
|
95  | 
||
96  | 
lemma hypnat_le_add_diff_inverse2 [simp]: "!!m n. n\<le>m ==> (m-n)+n = (m::hypnat)"  | 
|
97  | 
by transfer (rule le_add_diff_inverse2)  | 
|
98  | 
||
99  | 
declare hypnat_le_add_diff_inverse2 [OF order_less_imp_le]  | 
|
100  | 
||
101  | 
lemma hypnat_le0 [iff]: "!!n. (0::hypnat) \<le> n"  | 
|
102  | 
by transfer (rule le0)  | 
|
103  | 
||
104  | 
lemma hypnat_le_add1 [simp]: "!!x n. (x::hypnat) \<le> x + n"  | 
|
105  | 
by transfer (rule le_add1)  | 
|
106  | 
||
107  | 
lemma hypnat_add_self_le [simp]: "!!x n. (x::hypnat) \<le> n + x"  | 
|
108  | 
by transfer (rule le_add2)  | 
|
109  | 
||
110  | 
lemma hypnat_add_one_self_less [simp]: "(x::hypnat) < x + (1::hypnat)"  | 
|
111  | 
by (insert add_strict_left_mono [OF zero_less_one], auto)  | 
|
112  | 
||
113  | 
lemma hypnat_neq0_conv [iff]: "!!n. (n \<noteq> 0) = (0 < (n::hypnat))"  | 
|
114  | 
by transfer (rule neq0_conv)  | 
|
115  | 
||
116  | 
lemma hypnat_gt_zero_iff: "((0::hypnat) < n) = ((1::hypnat) \<le> n)"  | 
|
117  | 
by (auto simp add: linorder_not_less [symmetric])  | 
|
118  | 
||
119  | 
lemma hypnat_gt_zero_iff2: "(0 < n) = (\<exists>m. n = m + (1::hypnat))"  | 
|
120  | 
apply safe  | 
|
121  | 
apply (rule_tac x = "n - (1::hypnat) " in exI)  | 
|
122  | 
apply (simp add: hypnat_gt_zero_iff)  | 
|
123  | 
apply (insert add_le_less_mono [OF _ zero_less_one, of 0], auto)  | 
|
124  | 
done  | 
|
125  | 
||
126  | 
lemma hypnat_add_self_not_less: "~ (x + y < (x::hypnat))"  | 
|
127  | 
by (simp add: linorder_not_le [symmetric] add_commute [of x])  | 
|
128  | 
||
129  | 
lemma hypnat_diff_split:  | 
|
130  | 
"P(a - b::hypnat) = ((a<b --> P 0) & (ALL d. a = b + d --> P d))"  | 
|
131  | 
    -- {* elimination of @{text -} on @{text hypnat} *}
 | 
|
132  | 
proof (cases "a<b" rule: case_split)  | 
|
133  | 
case True  | 
|
134  | 
thus ?thesis  | 
|
135  | 
by (auto simp add: hypnat_add_self_not_less order_less_imp_le  | 
|
136  | 
hypnat_diff_is_0_eq [THEN iffD2])  | 
|
137  | 
next  | 
|
138  | 
case False  | 
|
139  | 
thus ?thesis  | 
|
140  | 
by (auto simp add: linorder_not_less dest: order_le_less_trans)  | 
|
141  | 
qed  | 
|
142  | 
||
143  | 
subsection{*Properties of the set of embedded natural numbers*}
 | 
|
144  | 
||
145  | 
lemma of_nat_eq_star_of [simp]: "of_nat = star_of"  | 
|
146  | 
proof  | 
|
147  | 
fix n :: nat  | 
|
148  | 
show "of_nat n = star_of n" by transfer simp  | 
|
149  | 
qed  | 
|
150  | 
||
151  | 
lemma Nats_eq_Standard: "(Nats :: nat star set) = Standard"  | 
|
152  | 
by (auto simp add: Nats_def Standard_def)  | 
|
153  | 
||
154  | 
lemma hypnat_of_nat_mem_Nats [simp]: "hypnat_of_nat n \<in> Nats"  | 
|
155  | 
by (simp add: Nats_eq_Standard)  | 
|
156  | 
||
157  | 
lemma hypnat_of_nat_one [simp]: "hypnat_of_nat (Suc 0) = (1::hypnat)"  | 
|
158  | 
by transfer simp  | 
|
159  | 
||
160  | 
lemma hypnat_of_nat_Suc [simp]:  | 
|
161  | 
"hypnat_of_nat (Suc n) = hypnat_of_nat n + (1::hypnat)"  | 
|
162  | 
by transfer simp  | 
|
163  | 
||
164  | 
lemma of_nat_eq_add [rule_format]:  | 
|
165  | 
"\<forall>d::hypnat. of_nat m = of_nat n + d --> d \<in> range of_nat"  | 
|
166  | 
apply (induct n)  | 
|
167  | 
apply (auto simp add: add_assoc)  | 
|
168  | 
apply (case_tac x)  | 
|
169  | 
apply (auto simp add: add_commute [of 1])  | 
|
170  | 
done  | 
|
171  | 
||
172  | 
lemma Nats_diff [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> (a-b :: hypnat) \<in> Nats"  | 
|
173  | 
by (simp add: Nats_eq_Standard)  | 
|
174  | 
||
175  | 
||
176  | 
subsection{*Infinite Hypernatural Numbers -- @{term HNatInfinite}*}
 | 
|
177  | 
||
178  | 
definition  | 
|
179  | 
(* the set of infinite hypernatural numbers *)  | 
|
180  | 
HNatInfinite :: "hypnat set" where  | 
|
181  | 
  "HNatInfinite = {n. n \<notin> Nats}"
 | 
|
182  | 
||
183  | 
lemma Nats_not_HNatInfinite_iff: "(x \<in> Nats) = (x \<notin> HNatInfinite)"  | 
|
184  | 
by (simp add: HNatInfinite_def)  | 
|
185  | 
||
186  | 
lemma HNatInfinite_not_Nats_iff: "(x \<in> HNatInfinite) = (x \<notin> Nats)"  | 
|
187  | 
by (simp add: HNatInfinite_def)  | 
|
188  | 
||
189  | 
lemma star_of_neq_HNatInfinite: "N \<in> HNatInfinite \<Longrightarrow> star_of n \<noteq> N"  | 
|
190  | 
by (auto simp add: HNatInfinite_def Nats_eq_Standard)  | 
|
191  | 
||
192  | 
lemma star_of_Suc_lessI:  | 
|
193  | 
"\<And>N. \<lbrakk>star_of n < N; star_of (Suc n) \<noteq> N\<rbrakk> \<Longrightarrow> star_of (Suc n) < N"  | 
|
194  | 
by transfer (rule Suc_lessI)  | 
|
195  | 
||
196  | 
lemma star_of_less_HNatInfinite:  | 
|
197  | 
assumes N: "N \<in> HNatInfinite"  | 
|
198  | 
shows "star_of n < N"  | 
|
199  | 
proof (induct n)  | 
|
200  | 
case 0  | 
|
201  | 
from N have "star_of 0 \<noteq> N" by (rule star_of_neq_HNatInfinite)  | 
|
202  | 
thus "star_of 0 < N" by simp  | 
|
203  | 
next  | 
|
204  | 
case (Suc n)  | 
|
205  | 
from N have "star_of (Suc n) \<noteq> N" by (rule star_of_neq_HNatInfinite)  | 
|
206  | 
with Suc show "star_of (Suc n) < N" by (rule star_of_Suc_lessI)  | 
|
207  | 
qed  | 
|
208  | 
||
209  | 
lemma star_of_le_HNatInfinite: "N \<in> HNatInfinite \<Longrightarrow> star_of n \<le> N"  | 
|
210  | 
by (rule star_of_less_HNatInfinite [THEN order_less_imp_le])  | 
|
211  | 
||
212  | 
subsubsection {* Closure Rules *}
 | 
|
213  | 
||
214  | 
lemma Nats_less_HNatInfinite: "\<lbrakk>x \<in> Nats; y \<in> HNatInfinite\<rbrakk> \<Longrightarrow> x < y"  | 
|
215  | 
by (auto simp add: Nats_def star_of_less_HNatInfinite)  | 
|
216  | 
||
217  | 
lemma Nats_le_HNatInfinite: "\<lbrakk>x \<in> Nats; y \<in> HNatInfinite\<rbrakk> \<Longrightarrow> x \<le> y"  | 
|
218  | 
by (rule Nats_less_HNatInfinite [THEN order_less_imp_le])  | 
|
219  | 
||
220  | 
lemma zero_less_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 0 < x"  | 
|
221  | 
by (simp add: Nats_less_HNatInfinite)  | 
|
222  | 
||
223  | 
lemma one_less_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 1 < x"  | 
|
224  | 
by (simp add: Nats_less_HNatInfinite)  | 
|
225  | 
||
226  | 
lemma one_le_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 1 \<le> x"  | 
|
227  | 
by (simp add: Nats_le_HNatInfinite)  | 
|
228  | 
||
229  | 
lemma zero_not_mem_HNatInfinite [simp]: "0 \<notin> HNatInfinite"  | 
|
230  | 
by (simp add: HNatInfinite_def)  | 
|
231  | 
||
232  | 
lemma Nats_downward_closed:  | 
|
233  | 
"\<lbrakk>x \<in> Nats; (y::hypnat) \<le> x\<rbrakk> \<Longrightarrow> y \<in> Nats"  | 
|
234  | 
apply (simp only: linorder_not_less [symmetric])  | 
|
235  | 
apply (erule contrapos_np)  | 
|
236  | 
apply (drule HNatInfinite_not_Nats_iff [THEN iffD2])  | 
|
237  | 
apply (erule (1) Nats_less_HNatInfinite)  | 
|
238  | 
done  | 
|
239  | 
||
240  | 
lemma HNatInfinite_upward_closed:  | 
|
241  | 
"\<lbrakk>x \<in> HNatInfinite; x \<le> y\<rbrakk> \<Longrightarrow> y \<in> HNatInfinite"  | 
|
242  | 
apply (simp only: HNatInfinite_not_Nats_iff)  | 
|
243  | 
apply (erule contrapos_nn)  | 
|
244  | 
apply (erule (1) Nats_downward_closed)  | 
|
245  | 
done  | 
|
246  | 
||
247  | 
lemma HNatInfinite_add: "x \<in> HNatInfinite \<Longrightarrow> x + y \<in> HNatInfinite"  | 
|
248  | 
apply (erule HNatInfinite_upward_closed)  | 
|
249  | 
apply (rule hypnat_le_add1)  | 
|
250  | 
done  | 
|
251  | 
||
252  | 
lemma HNatInfinite_add_one: "x \<in> HNatInfinite \<Longrightarrow> x + 1 \<in> HNatInfinite"  | 
|
253  | 
by (rule HNatInfinite_add)  | 
|
254  | 
||
255  | 
lemma HNatInfinite_diff:  | 
|
256  | 
"\<lbrakk>x \<in> HNatInfinite; y \<in> Nats\<rbrakk> \<Longrightarrow> x - y \<in> HNatInfinite"  | 
|
257  | 
apply (frule (1) Nats_le_HNatInfinite)  | 
|
258  | 
apply (simp only: HNatInfinite_not_Nats_iff)  | 
|
259  | 
apply (erule contrapos_nn)  | 
|
260  | 
apply (drule (1) Nats_add, simp)  | 
|
261  | 
done  | 
|
262  | 
||
263  | 
lemma HNatInfinite_is_Suc: "x \<in> HNatInfinite ==> \<exists>y. x = y + (1::hypnat)"  | 
|
264  | 
apply (rule_tac x = "x - (1::hypnat) " in exI)  | 
|
265  | 
apply (simp add: Nats_le_HNatInfinite)  | 
|
266  | 
done  | 
|
267  | 
||
268  | 
||
269  | 
subsection{*Existence of an infinite hypernatural number*}
 | 
|
270  | 
||
271  | 
definition  | 
|
272  | 
(* omega is in fact an infinite hypernatural number = [<1,2,3,...>] *)  | 
|
273  | 
whn :: hypnat where  | 
|
274  | 
hypnat_omega_def: "whn = star_n (%n::nat. n)"  | 
|
275  | 
||
276  | 
lemma hypnat_of_nat_neq_whn: "hypnat_of_nat n \<noteq> whn"  | 
|
277  | 
by (simp add: hypnat_omega_def star_of_def star_n_eq_iff)  | 
|
278  | 
||
279  | 
lemma whn_neq_hypnat_of_nat: "whn \<noteq> hypnat_of_nat n"  | 
|
280  | 
by (simp add: hypnat_omega_def star_of_def star_n_eq_iff)  | 
|
281  | 
||
282  | 
lemma whn_not_Nats [simp]: "whn \<notin> Nats"  | 
|
283  | 
by (simp add: Nats_def image_def whn_neq_hypnat_of_nat)  | 
|
284  | 
||
285  | 
lemma HNatInfinite_whn [simp]: "whn \<in> HNatInfinite"  | 
|
286  | 
by (simp add: HNatInfinite_def)  | 
|
287  | 
||
288  | 
lemma lemma_unbounded_set [simp]: "{n::nat. m < n} \<in> FreeUltrafilterNat"
 | 
|
| 29920 | 289  | 
apply (insert finite_atMost [of m])  | 
| 27468 | 290  | 
apply (drule FreeUltrafilterNat.finite)  | 
291  | 
apply (drule FreeUltrafilterNat.not_memD)  | 
|
| 29920 | 292  | 
apply (simp add: Collect_neg_eq [symmetric] linorder_not_le atMost_def)  | 
| 27468 | 293  | 
done  | 
294  | 
||
295  | 
lemma Compl_Collect_le: "- {n::nat. N \<le> n} = {n. n < N}"
 | 
|
296  | 
by (simp add: Collect_neg_eq [symmetric] linorder_not_le)  | 
|
297  | 
||
298  | 
lemma hypnat_of_nat_eq:  | 
|
299  | 
"hypnat_of_nat m = star_n (%n::nat. m)"  | 
|
300  | 
by (simp add: star_of_def)  | 
|
301  | 
||
302  | 
lemma SHNat_eq: "Nats = {n. \<exists>N. n = hypnat_of_nat N}"
 | 
|
303  | 
by (simp add: Nats_def image_def)  | 
|
304  | 
||
305  | 
lemma Nats_less_whn: "n \<in> Nats \<Longrightarrow> n < whn"  | 
|
306  | 
by (simp add: Nats_less_HNatInfinite)  | 
|
307  | 
||
308  | 
lemma Nats_le_whn: "n \<in> Nats \<Longrightarrow> n \<le> whn"  | 
|
309  | 
by (simp add: Nats_le_HNatInfinite)  | 
|
310  | 
||
311  | 
lemma hypnat_of_nat_less_whn [simp]: "hypnat_of_nat n < whn"  | 
|
312  | 
by (simp add: Nats_less_whn)  | 
|
313  | 
||
314  | 
lemma hypnat_of_nat_le_whn [simp]: "hypnat_of_nat n \<le> whn"  | 
|
315  | 
by (simp add: Nats_le_whn)  | 
|
316  | 
||
317  | 
lemma hypnat_zero_less_hypnat_omega [simp]: "0 < whn"  | 
|
318  | 
by (simp add: Nats_less_whn)  | 
|
319  | 
||
320  | 
lemma hypnat_one_less_hypnat_omega [simp]: "1 < whn"  | 
|
321  | 
by (simp add: Nats_less_whn)  | 
|
322  | 
||
323  | 
||
324  | 
subsubsection{*Alternative characterization of the set of infinite hypernaturals*}
 | 
|
325  | 
||
326  | 
text{* @{term "HNatInfinite = {N. \<forall>n \<in> Nats. n < N}"}*}
 | 
|
327  | 
||
328  | 
(*??delete? similar reasoning in hypnat_omega_gt_SHNat above*)  | 
|
329  | 
lemma HNatInfinite_FreeUltrafilterNat_lemma:  | 
|
330  | 
  assumes "\<forall>N::nat. {n. f n \<noteq> N} \<in> FreeUltrafilterNat"
 | 
|
331  | 
  shows "{n. N < f n} \<in> FreeUltrafilterNat"
 | 
|
332  | 
apply (induct N)  | 
|
333  | 
using assms  | 
|
334  | 
apply (drule_tac x = 0 in spec, simp)  | 
|
335  | 
using assms  | 
|
336  | 
apply (drule_tac x = "Suc N" in spec)  | 
|
337  | 
apply (elim ultra, auto)  | 
|
338  | 
done  | 
|
339  | 
||
340  | 
lemma HNatInfinite_iff: "HNatInfinite = {N. \<forall>n \<in> Nats. n < N}"
 | 
|
341  | 
apply (safe intro!: Nats_less_HNatInfinite)  | 
|
342  | 
apply (auto simp add: HNatInfinite_def)  | 
|
343  | 
done  | 
|
344  | 
||
345  | 
||
346  | 
subsubsection{*Alternative Characterization of @{term HNatInfinite} using 
 | 
|
347  | 
Free Ultrafilter*}  | 
|
348  | 
||
349  | 
lemma HNatInfinite_FreeUltrafilterNat:  | 
|
350  | 
     "star_n X \<in> HNatInfinite ==> \<forall>u. {n. u < X n}:  FreeUltrafilterNat"
 | 
|
351  | 
apply (auto simp add: HNatInfinite_iff SHNat_eq)  | 
|
352  | 
apply (drule_tac x="star_of u" in spec, simp)  | 
|
353  | 
apply (simp add: star_of_def star_less_def starP2_star_n)  | 
|
354  | 
done  | 
|
355  | 
||
356  | 
lemma FreeUltrafilterNat_HNatInfinite:  | 
|
357  | 
     "\<forall>u. {n. u < X n}:  FreeUltrafilterNat ==> star_n X \<in> HNatInfinite"
 | 
|
358  | 
by (auto simp add: star_less_def starP2_star_n HNatInfinite_iff SHNat_eq hypnat_of_nat_eq)  | 
|
359  | 
||
360  | 
lemma HNatInfinite_FreeUltrafilterNat_iff:  | 
|
361  | 
     "(star_n X \<in> HNatInfinite) = (\<forall>u. {n. u < X n}:  FreeUltrafilterNat)"
 | 
|
362  | 
by (rule iffI [OF HNatInfinite_FreeUltrafilterNat  | 
|
363  | 
FreeUltrafilterNat_HNatInfinite])  | 
|
364  | 
||
365  | 
subsection {* Embedding of the Hypernaturals into other types *}
 | 
|
366  | 
||
367  | 
definition  | 
|
368  | 
of_hypnat :: "hypnat \<Rightarrow> 'a::semiring_1_cancel star" where  | 
|
| 28562 | 369  | 
of_hypnat_def [transfer_unfold, code del]: "of_hypnat = *f* of_nat"  | 
| 27468 | 370  | 
|
371  | 
lemma of_hypnat_0 [simp]: "of_hypnat 0 = 0"  | 
|
372  | 
by transfer (rule of_nat_0)  | 
|
373  | 
||
374  | 
lemma of_hypnat_1 [simp]: "of_hypnat 1 = 1"  | 
|
375  | 
by transfer (rule of_nat_1)  | 
|
376  | 
||
377  | 
lemma of_hypnat_hSuc: "\<And>m. of_hypnat (hSuc m) = 1 + of_hypnat m"  | 
|
378  | 
by transfer (rule of_nat_Suc)  | 
|
379  | 
||
380  | 
lemma of_hypnat_add [simp]:  | 
|
381  | 
"\<And>m n. of_hypnat (m + n) = of_hypnat m + of_hypnat n"  | 
|
382  | 
by transfer (rule of_nat_add)  | 
|
383  | 
||
384  | 
lemma of_hypnat_mult [simp]:  | 
|
385  | 
"\<And>m n. of_hypnat (m * n) = of_hypnat m * of_hypnat n"  | 
|
386  | 
by transfer (rule of_nat_mult)  | 
|
387  | 
||
388  | 
lemma of_hypnat_less_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
389  | 
"\<And>m n. (of_hypnat m < (of_hypnat n::'a::linordered_semidom star)) = (m < n)"  | 
| 27468 | 390  | 
by transfer (rule of_nat_less_iff)  | 
391  | 
||
392  | 
lemma of_hypnat_0_less_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
393  | 
"\<And>n. (0 < (of_hypnat n::'a::linordered_semidom star)) = (0 < n)"  | 
| 27468 | 394  | 
by transfer (rule of_nat_0_less_iff)  | 
395  | 
||
396  | 
lemma of_hypnat_less_0_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
397  | 
"\<And>m. \<not> (of_hypnat m::'a::linordered_semidom star) < 0"  | 
| 27468 | 398  | 
by transfer (rule of_nat_less_0_iff)  | 
399  | 
||
400  | 
lemma of_hypnat_le_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
401  | 
"\<And>m n. (of_hypnat m \<le> (of_hypnat n::'a::linordered_semidom star)) = (m \<le> n)"  | 
| 27468 | 402  | 
by transfer (rule of_nat_le_iff)  | 
403  | 
||
404  | 
lemma of_hypnat_0_le_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
405  | 
"\<And>n. 0 \<le> (of_hypnat n::'a::linordered_semidom star)"  | 
| 27468 | 406  | 
by transfer (rule of_nat_0_le_iff)  | 
407  | 
||
408  | 
lemma of_hypnat_le_0_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
409  | 
"\<And>m. ((of_hypnat m::'a::linordered_semidom star) \<le> 0) = (m = 0)"  | 
| 27468 | 410  | 
by transfer (rule of_nat_le_0_iff)  | 
411  | 
||
412  | 
lemma of_hypnat_eq_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
413  | 
"\<And>m n. (of_hypnat m = (of_hypnat n::'a::linordered_semidom star)) = (m = n)"  | 
| 27468 | 414  | 
by transfer (rule of_nat_eq_iff)  | 
415  | 
||
416  | 
lemma of_hypnat_eq_0_iff [simp]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
417  | 
"\<And>m. ((of_hypnat m::'a::linordered_semidom star) = 0) = (m = 0)"  | 
| 27468 | 418  | 
by transfer (rule of_nat_eq_0_iff)  | 
419  | 
||
420  | 
lemma HNatInfinite_of_hypnat_gt_zero:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
29920 
diff
changeset
 | 
421  | 
"N \<in> HNatInfinite \<Longrightarrow> (0::'a::linordered_semidom star) < of_hypnat N"  | 
| 27468 | 422  | 
by (rule ccontr, simp add: linorder_not_less)  | 
423  | 
||
424  | 
end  |