author | mengj |
Wed, 19 Oct 2005 06:33:24 +0200 | |
changeset 17905 | 1574533861b1 |
parent 15140 | 322485b816ac |
child 19736 | d8d0f8f51d69 |
permissions | -rw-r--r-- |
11355 | 1 |
(* Title: HOL/Library/Nat_Infinity.thy |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb, TU Muenchen |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
4 |
*) |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
5 |
|
14706 | 6 |
header {* Natural numbers with infinity *} |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
7 |
|
15131 | 8 |
theory Nat_Infinity |
15140 | 9 |
imports Main |
15131 | 10 |
begin |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
11 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
12 |
subsection "Definitions" |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
13 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
14 |
text {* |
11355 | 15 |
We extend the standard natural numbers by a special value indicating |
16 |
infinity. This includes extending the ordering relations @{term "op |
|
17 |
<"} and @{term "op \<le>"}. |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
18 |
*} |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
19 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
20 |
datatype inat = Fin nat | Infty |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
21 |
|
14691 | 22 |
instance inat :: "{ord, zero}" .. |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
23 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
24 |
consts |
11355 | 25 |
iSuc :: "inat => inat" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
26 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
27 |
syntax (xsymbols) |
11355 | 28 |
Infty :: inat ("\<infinity>") |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
29 |
|
14565 | 30 |
syntax (HTML output) |
31 |
Infty :: inat ("\<infinity>") |
|
32 |
||
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
33 |
defs |
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11655
diff
changeset
|
34 |
Zero_inat_def: "0 == Fin 0" |
11355 | 35 |
iSuc_def: "iSuc i == case i of Fin n => Fin (Suc n) | \<infinity> => \<infinity>" |
36 |
iless_def: "m < n == |
|
37 |
case m of Fin m1 => (case n of Fin n1 => m1 < n1 | \<infinity> => True) |
|
38 |
| \<infinity> => False" |
|
39 |
ile_def: "(m::inat) \<le> n == \<not> (n < m)" |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
40 |
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11655
diff
changeset
|
41 |
lemmas inat_defs = Zero_inat_def iSuc_def iless_def ile_def |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
42 |
lemmas inat_splits = inat.split inat.split_asm |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
43 |
|
11355 | 44 |
text {* |
11357 | 45 |
Below is a not quite complete set of theorems. Use the method |
46 |
@{text "(simp add: inat_defs split:inat_splits, arith?)"} to prove |
|
47 |
new theorems or solve arithmetic subgoals involving @{typ inat} on |
|
48 |
the fly. |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
49 |
*} |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
50 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
51 |
subsection "Constructors" |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
52 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
53 |
lemma Fin_0: "Fin 0 = 0" |
11357 | 54 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
55 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
56 |
lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0" |
11357 | 57 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
58 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
59 |
lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>" |
11357 | 60 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
61 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
62 |
lemma iSuc_Fin [simp]: "iSuc (Fin n) = Fin (Suc n)" |
11357 | 63 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
64 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
65 |
lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>" |
11357 | 66 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
67 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
68 |
lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0" |
11357 | 69 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
70 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
71 |
lemma iSuc_inject [simp]: "(iSuc x = iSuc y) = (x = y)" |
11357 | 72 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
73 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
74 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
75 |
subsection "Ordering relations" |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
76 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
77 |
lemma Infty_ilessE [elim!]: "\<infinity> < Fin m ==> R" |
11357 | 78 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
79 |
|
11355 | 80 |
lemma iless_linear: "m < n \<or> m = n \<or> n < (m::inat)" |
11357 | 81 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
82 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
83 |
lemma iless_not_refl [simp]: "\<not> n < (n::inat)" |
11357 | 84 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
85 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
86 |
lemma iless_trans: "i < j ==> j < k ==> i < (k::inat)" |
11357 | 87 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
88 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
89 |
lemma iless_not_sym: "n < m ==> \<not> m < (n::inat)" |
11357 | 90 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
91 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
92 |
lemma Fin_iless_mono [simp]: "(Fin n < Fin m) = (n < m)" |
11357 | 93 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
94 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
95 |
lemma Fin_iless_Infty [simp]: "Fin n < \<infinity>" |
11357 | 96 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
97 |
|
11655 | 98 |
lemma Infty_eq [simp]: "(n < \<infinity>) = (n \<noteq> \<infinity>)" |
11357 | 99 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
100 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
101 |
lemma i0_eq [simp]: "((0::inat) < n) = (n \<noteq> 0)" |
11357 | 102 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
103 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
104 |
lemma i0_iless_iSuc [simp]: "0 < iSuc n" |
11357 | 105 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
106 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
107 |
lemma not_ilessi0 [simp]: "\<not> n < (0::inat)" |
11357 | 108 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
109 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
110 |
lemma Fin_iless: "n < Fin m ==> \<exists>k. n = Fin k" |
11357 | 111 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
112 |
|
11655 | 113 |
lemma iSuc_mono [simp]: "(iSuc n < iSuc m) = (n < m)" |
11357 | 114 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
115 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
116 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
117 |
(* ----------------------------------------------------------------------- *) |
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
118 |
|
11655 | 119 |
lemma ile_def2: "(m \<le> n) = (m < n \<or> m = (n::inat))" |
11357 | 120 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
121 |
|
11355 | 122 |
lemma ile_refl [simp]: "n \<le> (n::inat)" |
11357 | 123 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
124 |
|
11355 | 125 |
lemma ile_trans: "i \<le> j ==> j \<le> k ==> i \<le> (k::inat)" |
11357 | 126 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
127 |
|
11355 | 128 |
lemma ile_iless_trans: "i \<le> j ==> j < k ==> i < (k::inat)" |
11357 | 129 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
130 |
|
11355 | 131 |
lemma iless_ile_trans: "i < j ==> j \<le> k ==> i < (k::inat)" |
11357 | 132 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
133 |
|
11355 | 134 |
lemma Infty_ub [simp]: "n \<le> \<infinity>" |
11357 | 135 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
136 |
|
11355 | 137 |
lemma i0_lb [simp]: "(0::inat) \<le> n" |
11357 | 138 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
139 |
|
11355 | 140 |
lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m ==> R" |
11357 | 141 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
142 |
|
11355 | 143 |
lemma Fin_ile_mono [simp]: "(Fin n \<le> Fin m) = (n \<le> m)" |
11357 | 144 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
145 |
|
11355 | 146 |
lemma ilessI1: "n \<le> m ==> n \<noteq> m ==> n < (m::inat)" |
11357 | 147 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
148 |
|
11355 | 149 |
lemma ileI1: "m < n ==> iSuc m \<le> n" |
11357 | 150 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
151 |
|
11655 | 152 |
lemma Suc_ile_eq: "(Fin (Suc m) \<le> n) = (Fin m < n)" |
11357 | 153 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
154 |
|
11655 | 155 |
lemma iSuc_ile_mono [simp]: "(iSuc n \<le> iSuc m) = (n \<le> m)" |
11357 | 156 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
157 |
|
11655 | 158 |
lemma iless_Suc_eq [simp]: "(Fin m < iSuc n) = (Fin m \<le> n)" |
11357 | 159 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
160 |
|
11355 | 161 |
lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0" |
11357 | 162 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
163 |
|
11355 | 164 |
lemma ile_iSuc [simp]: "n \<le> iSuc n" |
11357 | 165 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
166 |
|
11355 | 167 |
lemma Fin_ile: "n \<le> Fin m ==> \<exists>k. n = Fin k" |
11357 | 168 |
by (simp add: inat_defs split:inat_splits, arith?) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
169 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
170 |
lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j" |
11355 | 171 |
apply (induct_tac k) |
172 |
apply (simp (no_asm) only: Fin_0) |
|
173 |
apply (fast intro: ile_iless_trans i0_lb) |
|
174 |
apply (erule exE) |
|
175 |
apply (drule spec) |
|
176 |
apply (erule exE) |
|
177 |
apply (drule ileI1) |
|
178 |
apply (rule iSuc_Fin [THEN subst]) |
|
179 |
apply (rule exI) |
|
180 |
apply (erule (1) ile_iless_trans) |
|
181 |
done |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
182 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
183 |
end |