author | webertj |
Fri, 19 Oct 2012 15:12:52 +0200 | |
changeset 49962 | a8cc904a6820 |
parent 49834 | b27bbb021df1 |
child 51000 | c9adb50f74ad |
permissions | -rw-r--r-- |
43919 | 1 |
(* Title: HOL/Library/Extended_Nat.thy |
27110 | 2 |
Author: David von Oheimb, TU Muenchen; Florian Haftmann, TU Muenchen |
41853 | 3 |
Contributions: David Trachtenherz, 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 |
|
43919 | 6 |
header {* Extended natural numbers (i.e. with infinity) *} |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
7 |
|
43919 | 8 |
theory Extended_Nat |
30663
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
haftmann
parents:
29668
diff
changeset
|
9 |
imports Main |
15131 | 10 |
begin |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
11 |
|
43921 | 12 |
class infinity = |
13 |
fixes infinity :: "'a" |
|
14 |
||
15 |
notation (xsymbols) |
|
16 |
infinity ("\<infinity>") |
|
17 |
||
18 |
notation (HTML output) |
|
19 |
infinity ("\<infinity>") |
|
20 |
||
27110 | 21 |
subsection {* Type definition *} |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
22 |
|
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
23 |
text {* |
11355 | 24 |
We extend the standard natural numbers by a special value indicating |
27110 | 25 |
infinity. |
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 |
|
49834 | 28 |
typedef enat = "UNIV :: nat option set" .. |
43921 | 29 |
|
43924 | 30 |
definition enat :: "nat \<Rightarrow> enat" where |
31 |
"enat n = Abs_enat (Some n)" |
|
43921 | 32 |
|
33 |
instantiation enat :: infinity |
|
34 |
begin |
|
35 |
definition "\<infinity> = Abs_enat None" |
|
36 |
instance proof qed |
|
37 |
end |
|
38 |
||
43924 | 39 |
rep_datatype enat "\<infinity> :: enat" |
43921 | 40 |
proof - |
43924 | 41 |
fix P i assume "\<And>j. P (enat j)" "P \<infinity>" |
43921 | 42 |
then show "P i" |
43 |
proof induct |
|
44 |
case (Abs_enat y) then show ?case |
|
45 |
by (cases y rule: option.exhaust) |
|
43924 | 46 |
(auto simp: enat_def infinity_enat_def) |
43921 | 47 |
qed |
43924 | 48 |
qed (auto simp add: enat_def infinity_enat_def Abs_enat_inject) |
19736 | 49 |
|
43924 | 50 |
declare [[coercion "enat::nat\<Rightarrow>enat"]] |
19736 | 51 |
|
45934 | 52 |
lemmas enat2_cases = enat.exhaust[case_product enat.exhaust] |
53 |
lemmas enat3_cases = enat.exhaust[case_product enat.exhaust enat.exhaust] |
|
54 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
55 |
lemma not_infinity_eq [iff]: "(x \<noteq> \<infinity>) = (EX i. x = enat i)" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
56 |
by (cases x) auto |
31084 | 57 |
|
43924 | 58 |
lemma not_enat_eq [iff]: "(ALL y. x ~= enat y) = (x = \<infinity>)" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
59 |
by (cases x) auto |
31077 | 60 |
|
43924 | 61 |
primrec the_enat :: "enat \<Rightarrow> nat" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
62 |
where "the_enat (enat n) = n" |
41855 | 63 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
64 |
|
27110 | 65 |
subsection {* Constructors and numbers *} |
66 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
67 |
instantiation enat :: "{zero, one}" |
25594 | 68 |
begin |
69 |
||
70 |
definition |
|
43924 | 71 |
"0 = enat 0" |
25594 | 72 |
|
73 |
definition |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
74 |
"1 = enat 1" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
75 |
|
25594 | 76 |
instance .. |
77 |
||
78 |
end |
|
79 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
80 |
definition eSuc :: "enat \<Rightarrow> enat" where |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
81 |
"eSuc i = (case i of enat n \<Rightarrow> enat (Suc n) | \<infinity> \<Rightarrow> \<infinity>)" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
82 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
83 |
lemma enat_0 [code_post]: "enat 0 = 0" |
43919 | 84 |
by (simp add: zero_enat_def) |
27110 | 85 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
86 |
lemma enat_1 [code_post]: "enat 1 = 1" |
43919 | 87 |
by (simp add: one_enat_def) |
27110 | 88 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
89 |
lemma one_eSuc: "1 = eSuc 0" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
90 |
by (simp add: zero_enat_def one_enat_def eSuc_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
91 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
92 |
lemma infinity_ne_i0 [simp]: "(\<infinity>::enat) \<noteq> 0" |
43919 | 93 |
by (simp add: zero_enat_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
94 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
95 |
lemma i0_ne_infinity [simp]: "0 \<noteq> (\<infinity>::enat)" |
43919 | 96 |
by (simp add: zero_enat_def) |
27110 | 97 |
|
43919 | 98 |
lemma zero_one_enat_neq [simp]: |
99 |
"\<not> 0 = (1\<Colon>enat)" |
|
100 |
"\<not> 1 = (0\<Colon>enat)" |
|
101 |
unfolding zero_enat_def one_enat_def by simp_all |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
102 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
103 |
lemma infinity_ne_i1 [simp]: "(\<infinity>::enat) \<noteq> 1" |
43919 | 104 |
by (simp add: one_enat_def) |
27110 | 105 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
106 |
lemma i1_ne_infinity [simp]: "1 \<noteq> (\<infinity>::enat)" |
43919 | 107 |
by (simp add: one_enat_def) |
27110 | 108 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
109 |
lemma eSuc_enat: "eSuc (enat n) = enat (Suc n)" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
110 |
by (simp add: eSuc_def) |
27110 | 111 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
112 |
lemma eSuc_infinity [simp]: "eSuc \<infinity> = \<infinity>" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
113 |
by (simp add: eSuc_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
114 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
115 |
lemma eSuc_ne_0 [simp]: "eSuc n \<noteq> 0" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
116 |
by (simp add: eSuc_def zero_enat_def split: enat.splits) |
27110 | 117 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
118 |
lemma zero_ne_eSuc [simp]: "0 \<noteq> eSuc n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
119 |
by (rule eSuc_ne_0 [symmetric]) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
120 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
121 |
lemma eSuc_inject [simp]: "eSuc m = eSuc n \<longleftrightarrow> m = n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
122 |
by (simp add: eSuc_def split: enat.splits) |
27110 | 123 |
|
124 |
subsection {* Addition *} |
|
125 |
||
43919 | 126 |
instantiation enat :: comm_monoid_add |
27110 | 127 |
begin |
128 |
||
38167 | 129 |
definition [nitpick_simp]: |
43924 | 130 |
"m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | enat m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | enat n \<Rightarrow> enat (m + n)))" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
131 |
|
43919 | 132 |
lemma plus_enat_simps [simp, code]: |
43921 | 133 |
fixes q :: enat |
43924 | 134 |
shows "enat m + enat n = enat (m + n)" |
43921 | 135 |
and "\<infinity> + q = \<infinity>" |
136 |
and "q + \<infinity> = \<infinity>" |
|
43919 | 137 |
by (simp_all add: plus_enat_def split: enat.splits) |
27110 | 138 |
|
139 |
instance proof |
|
43919 | 140 |
fix n m q :: enat |
27110 | 141 |
show "n + m + q = n + (m + q)" |
45934 | 142 |
by (cases n m q rule: enat3_cases) auto |
27110 | 143 |
show "n + m = m + n" |
45934 | 144 |
by (cases n m rule: enat2_cases) auto |
27110 | 145 |
show "0 + n = n" |
43919 | 146 |
by (cases n) (simp_all add: zero_enat_def) |
26089 | 147 |
qed |
148 |
||
27110 | 149 |
end |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
150 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
151 |
lemma eSuc_plus_1: |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
152 |
"eSuc n = n + 1" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
153 |
by (cases n) (simp_all add: eSuc_enat one_enat_def) |
27110 | 154 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
155 |
lemma plus_1_eSuc: |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
156 |
"1 + q = eSuc q" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
157 |
"q + 1 = eSuc q" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
158 |
by (simp_all add: eSuc_plus_1 add_ac) |
41853 | 159 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
160 |
lemma iadd_Suc: "eSuc m + n = eSuc (m + n)" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
161 |
by (simp_all add: eSuc_plus_1 add_ac) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
162 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
163 |
lemma iadd_Suc_right: "m + eSuc n = eSuc (m + n)" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
164 |
by (simp only: add_commute[of m] iadd_Suc) |
41853 | 165 |
|
43919 | 166 |
lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \<and> n = 0)" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
167 |
by (cases m, cases n, simp_all add: zero_enat_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
168 |
|
29014 | 169 |
subsection {* Multiplication *} |
170 |
||
43919 | 171 |
instantiation enat :: comm_semiring_1 |
29014 | 172 |
begin |
173 |
||
43919 | 174 |
definition times_enat_def [nitpick_simp]: |
43924 | 175 |
"m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | enat m \<Rightarrow> |
176 |
(case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | enat n \<Rightarrow> enat (m * n)))" |
|
29014 | 177 |
|
43919 | 178 |
lemma times_enat_simps [simp, code]: |
43924 | 179 |
"enat m * enat n = enat (m * n)" |
43921 | 180 |
"\<infinity> * \<infinity> = (\<infinity>::enat)" |
43924 | 181 |
"\<infinity> * enat n = (if n = 0 then 0 else \<infinity>)" |
182 |
"enat m * \<infinity> = (if m = 0 then 0 else \<infinity>)" |
|
43919 | 183 |
unfolding times_enat_def zero_enat_def |
184 |
by (simp_all split: enat.split) |
|
29014 | 185 |
|
186 |
instance proof |
|
43919 | 187 |
fix a b c :: enat |
29014 | 188 |
show "(a * b) * c = a * (b * c)" |
43919 | 189 |
unfolding times_enat_def zero_enat_def |
190 |
by (simp split: enat.split) |
|
29014 | 191 |
show "a * b = b * a" |
43919 | 192 |
unfolding times_enat_def zero_enat_def |
193 |
by (simp split: enat.split) |
|
29014 | 194 |
show "1 * a = a" |
43919 | 195 |
unfolding times_enat_def zero_enat_def one_enat_def |
196 |
by (simp split: enat.split) |
|
29014 | 197 |
show "(a + b) * c = a * c + b * c" |
43919 | 198 |
unfolding times_enat_def zero_enat_def |
49962
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents:
49834
diff
changeset
|
199 |
by (simp split: enat.split add: distrib_right) |
29014 | 200 |
show "0 * a = 0" |
43919 | 201 |
unfolding times_enat_def zero_enat_def |
202 |
by (simp split: enat.split) |
|
29014 | 203 |
show "a * 0 = 0" |
43919 | 204 |
unfolding times_enat_def zero_enat_def |
205 |
by (simp split: enat.split) |
|
206 |
show "(0::enat) \<noteq> 1" |
|
207 |
unfolding zero_enat_def one_enat_def |
|
29014 | 208 |
by simp |
209 |
qed |
|
210 |
||
211 |
end |
|
212 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
213 |
lemma mult_eSuc: "eSuc m * n = n + m * n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
214 |
unfolding eSuc_plus_1 by (simp add: algebra_simps) |
29014 | 215 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
216 |
lemma mult_eSuc_right: "m * eSuc n = m + m * n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
217 |
unfolding eSuc_plus_1 by (simp add: algebra_simps) |
29014 | 218 |
|
43924 | 219 |
lemma of_nat_eq_enat: "of_nat n = enat n" |
29023 | 220 |
apply (induct n) |
43924 | 221 |
apply (simp add: enat_0) |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
222 |
apply (simp add: plus_1_eSuc eSuc_enat) |
29023 | 223 |
done |
224 |
||
43919 | 225 |
instance enat :: semiring_char_0 proof |
43924 | 226 |
have "inj enat" by (rule injI) simp |
227 |
then show "inj (\<lambda>n. of_nat n :: enat)" by (simp add: of_nat_eq_enat) |
|
38621
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
haftmann
parents:
38167
diff
changeset
|
228 |
qed |
29023 | 229 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
230 |
lemma imult_is_0 [simp]: "((m::enat) * n = 0) = (m = 0 \<or> n = 0)" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
231 |
by (auto simp add: times_enat_def zero_enat_def split: enat.split) |
41853 | 232 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
233 |
lemma imult_is_infinity: "((a::enat) * b = \<infinity>) = (a = \<infinity> \<and> b \<noteq> 0 \<or> b = \<infinity> \<and> a \<noteq> 0)" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
234 |
by (auto simp add: times_enat_def zero_enat_def split: enat.split) |
41853 | 235 |
|
236 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
237 |
subsection {* Numerals *} |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
238 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
239 |
lemma numeral_eq_enat: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
240 |
"numeral k = enat (numeral k)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
241 |
using of_nat_eq_enat [of "numeral k"] by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
242 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
243 |
lemma enat_numeral [code_abbrev]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
244 |
"enat (numeral k) = numeral k" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
245 |
using numeral_eq_enat .. |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
246 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
247 |
lemma infinity_ne_numeral [simp]: "(\<infinity>::enat) \<noteq> numeral k" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
248 |
by (simp add: numeral_eq_enat) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
249 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
250 |
lemma numeral_ne_infinity [simp]: "numeral k \<noteq> (\<infinity>::enat)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
251 |
by (simp add: numeral_eq_enat) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
252 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
253 |
lemma eSuc_numeral [simp]: "eSuc (numeral k) = numeral (k + Num.One)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
254 |
by (simp only: eSuc_plus_1 numeral_plus_one) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
255 |
|
41853 | 256 |
subsection {* Subtraction *} |
257 |
||
43919 | 258 |
instantiation enat :: minus |
41853 | 259 |
begin |
260 |
||
43919 | 261 |
definition diff_enat_def: |
43924 | 262 |
"a - b = (case a of (enat x) \<Rightarrow> (case b of (enat y) \<Rightarrow> enat (x - y) | \<infinity> \<Rightarrow> 0) |
41853 | 263 |
| \<infinity> \<Rightarrow> \<infinity>)" |
264 |
||
265 |
instance .. |
|
266 |
||
267 |
end |
|
268 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
269 |
lemma idiff_enat_enat [simp, code]: "enat a - enat b = enat (a - b)" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
270 |
by (simp add: diff_enat_def) |
41853 | 271 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
272 |
lemma idiff_infinity [simp, code]: "\<infinity> - n = (\<infinity>::enat)" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
273 |
by (simp add: diff_enat_def) |
41853 | 274 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
275 |
lemma idiff_infinity_right [simp, code]: "enat a - \<infinity> = 0" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
276 |
by (simp add: diff_enat_def) |
41853 | 277 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
278 |
lemma idiff_0 [simp]: "(0::enat) - n = 0" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
279 |
by (cases n, simp_all add: zero_enat_def) |
41853 | 280 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
281 |
lemmas idiff_enat_0 [simp] = idiff_0 [unfolded zero_enat_def] |
41853 | 282 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
283 |
lemma idiff_0_right [simp]: "(n::enat) - 0 = n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
284 |
by (cases n) (simp_all add: zero_enat_def) |
41853 | 285 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
286 |
lemmas idiff_enat_0_right [simp] = idiff_0_right [unfolded zero_enat_def] |
41853 | 287 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
288 |
lemma idiff_self [simp]: "n \<noteq> \<infinity> \<Longrightarrow> (n::enat) - n = 0" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
289 |
by (auto simp: zero_enat_def) |
41853 | 290 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
291 |
lemma eSuc_minus_eSuc [simp]: "eSuc n - eSuc m = n - m" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
292 |
by (simp add: eSuc_def split: enat.split) |
41855 | 293 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
294 |
lemma eSuc_minus_1 [simp]: "eSuc n - 1 = n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
295 |
by (simp add: one_enat_def eSuc_enat[symmetric] zero_enat_def[symmetric]) |
41855 | 296 |
|
43924 | 297 |
(*lemmas idiff_self_eq_0_enat = idiff_self_eq_0[unfolded zero_enat_def]*) |
41853 | 298 |
|
27110 | 299 |
subsection {* Ordering *} |
300 |
||
43919 | 301 |
instantiation enat :: linordered_ab_semigroup_add |
27110 | 302 |
begin |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
303 |
|
38167 | 304 |
definition [nitpick_simp]: |
43924 | 305 |
"m \<le> n = (case n of enat n1 \<Rightarrow> (case m of enat m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False) |
27110 | 306 |
| \<infinity> \<Rightarrow> True)" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
307 |
|
38167 | 308 |
definition [nitpick_simp]: |
43924 | 309 |
"m < n = (case m of enat m1 \<Rightarrow> (case n of enat n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True) |
27110 | 310 |
| \<infinity> \<Rightarrow> False)" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
311 |
|
43919 | 312 |
lemma enat_ord_simps [simp]: |
43924 | 313 |
"enat m \<le> enat n \<longleftrightarrow> m \<le> n" |
314 |
"enat m < enat n \<longleftrightarrow> m < n" |
|
43921 | 315 |
"q \<le> (\<infinity>::enat)" |
316 |
"q < (\<infinity>::enat) \<longleftrightarrow> q \<noteq> \<infinity>" |
|
317 |
"(\<infinity>::enat) \<le> q \<longleftrightarrow> q = \<infinity>" |
|
318 |
"(\<infinity>::enat) < q \<longleftrightarrow> False" |
|
43919 | 319 |
by (simp_all add: less_eq_enat_def less_enat_def split: enat.splits) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
320 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
321 |
lemma numeral_le_enat_iff[simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
322 |
shows "numeral m \<le> enat n \<longleftrightarrow> numeral m \<le> n" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
323 |
by (auto simp: numeral_eq_enat) |
45934 | 324 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
325 |
lemma numeral_less_enat_iff[simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
326 |
shows "numeral m < enat n \<longleftrightarrow> numeral m < n" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
327 |
by (auto simp: numeral_eq_enat) |
45934 | 328 |
|
43919 | 329 |
lemma enat_ord_code [code]: |
43924 | 330 |
"enat m \<le> enat n \<longleftrightarrow> m \<le> n" |
331 |
"enat m < enat n \<longleftrightarrow> m < n" |
|
43921 | 332 |
"q \<le> (\<infinity>::enat) \<longleftrightarrow> True" |
43924 | 333 |
"enat m < \<infinity> \<longleftrightarrow> True" |
334 |
"\<infinity> \<le> enat n \<longleftrightarrow> False" |
|
43921 | 335 |
"(\<infinity>::enat) < q \<longleftrightarrow> False" |
27110 | 336 |
by simp_all |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
337 |
|
27110 | 338 |
instance by default |
43919 | 339 |
(auto simp add: less_eq_enat_def less_enat_def plus_enat_def split: enat.splits) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
340 |
|
27110 | 341 |
end |
342 |
||
43919 | 343 |
instance enat :: ordered_comm_semiring |
29014 | 344 |
proof |
43919 | 345 |
fix a b c :: enat |
29014 | 346 |
assume "a \<le> b" and "0 \<le> c" |
347 |
thus "c * a \<le> c * b" |
|
43919 | 348 |
unfolding times_enat_def less_eq_enat_def zero_enat_def |
349 |
by (simp split: enat.splits) |
|
29014 | 350 |
qed |
351 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
352 |
(* BH: These equations are already proven generally for any type in |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
353 |
class linordered_semidom. However, enat is not in that class because |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
354 |
it does not have the cancellation property. Would it be worthwhile to |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
355 |
a generalize linordered_semidom to a new class that includes enat? *) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
356 |
|
43919 | 357 |
lemma enat_ord_number [simp]: |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
358 |
"(numeral m \<Colon> enat) \<le> numeral n \<longleftrightarrow> (numeral m \<Colon> nat) \<le> numeral n" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
359 |
"(numeral m \<Colon> enat) < numeral n \<longleftrightarrow> (numeral m \<Colon> nat) < numeral n" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
360 |
by (simp_all add: numeral_eq_enat) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
361 |
|
43919 | 362 |
lemma i0_lb [simp]: "(0\<Colon>enat) \<le> n" |
363 |
by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
364 |
|
43919 | 365 |
lemma ile0_eq [simp]: "n \<le> (0\<Colon>enat) \<longleftrightarrow> n = 0" |
366 |
by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
367 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
368 |
lemma infinity_ileE [elim!]: "\<infinity> \<le> enat m \<Longrightarrow> R" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
369 |
by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
370 |
|
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
371 |
lemma infinity_ilessE [elim!]: "\<infinity> < enat m \<Longrightarrow> R" |
27110 | 372 |
by simp |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
373 |
|
43919 | 374 |
lemma not_iless0 [simp]: "\<not> n < (0\<Colon>enat)" |
375 |
by (simp add: zero_enat_def less_enat_def split: enat.splits) |
|
27110 | 376 |
|
43919 | 377 |
lemma i0_less [simp]: "(0\<Colon>enat) < n \<longleftrightarrow> n \<noteq> 0" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
378 |
by (simp add: zero_enat_def less_enat_def split: enat.splits) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
379 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
380 |
lemma eSuc_ile_mono [simp]: "eSuc n \<le> eSuc m \<longleftrightarrow> n \<le> m" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
381 |
by (simp add: eSuc_def less_eq_enat_def split: enat.splits) |
27110 | 382 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
383 |
lemma eSuc_mono [simp]: "eSuc n < eSuc m \<longleftrightarrow> n < m" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
384 |
by (simp add: eSuc_def less_enat_def split: enat.splits) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
385 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
386 |
lemma ile_eSuc [simp]: "n \<le> eSuc n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
387 |
by (simp add: eSuc_def less_eq_enat_def split: enat.splits) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
388 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
389 |
lemma not_eSuc_ilei0 [simp]: "\<not> eSuc n \<le> 0" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
390 |
by (simp add: zero_enat_def eSuc_def less_eq_enat_def split: enat.splits) |
27110 | 391 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
392 |
lemma i0_iless_eSuc [simp]: "0 < eSuc n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
393 |
by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.splits) |
27110 | 394 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
395 |
lemma iless_eSuc0[simp]: "(n < eSuc 0) = (n = 0)" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
396 |
by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.split) |
41853 | 397 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
398 |
lemma ileI1: "m < n \<Longrightarrow> eSuc m \<le> n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
399 |
by (simp add: eSuc_def less_eq_enat_def less_enat_def split: enat.splits) |
27110 | 400 |
|
43924 | 401 |
lemma Suc_ile_eq: "enat (Suc m) \<le> n \<longleftrightarrow> enat m < n" |
27110 | 402 |
by (cases n) auto |
403 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
404 |
lemma iless_Suc_eq [simp]: "enat m < eSuc n \<longleftrightarrow> enat m \<le> n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
405 |
by (auto simp add: eSuc_def less_enat_def split: enat.splits) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
406 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
407 |
lemma imult_infinity: "(0::enat) < n \<Longrightarrow> \<infinity> * n = \<infinity>" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
408 |
by (simp add: zero_enat_def less_enat_def split: enat.splits) |
41853 | 409 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
410 |
lemma imult_infinity_right: "(0::enat) < n \<Longrightarrow> n * \<infinity> = \<infinity>" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
411 |
by (simp add: zero_enat_def less_enat_def split: enat.splits) |
41853 | 412 |
|
43919 | 413 |
lemma enat_0_less_mult_iff: "(0 < (m::enat) * n) = (0 < m \<and> 0 < n)" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
414 |
by (simp only: i0_less imult_is_0, simp) |
41853 | 415 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
416 |
lemma mono_eSuc: "mono eSuc" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
417 |
by (simp add: mono_def) |
41853 | 418 |
|
419 |
||
43919 | 420 |
lemma min_enat_simps [simp]: |
43924 | 421 |
"min (enat m) (enat n) = enat (min m n)" |
27110 | 422 |
"min q 0 = 0" |
423 |
"min 0 q = 0" |
|
43921 | 424 |
"min q (\<infinity>::enat) = q" |
425 |
"min (\<infinity>::enat) q = q" |
|
27110 | 426 |
by (auto simp add: min_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
427 |
|
43919 | 428 |
lemma max_enat_simps [simp]: |
43924 | 429 |
"max (enat m) (enat n) = enat (max m n)" |
27110 | 430 |
"max q 0 = q" |
431 |
"max 0 q = q" |
|
43921 | 432 |
"max q \<infinity> = (\<infinity>::enat)" |
433 |
"max \<infinity> q = (\<infinity>::enat)" |
|
27110 | 434 |
by (simp_all add: max_def) |
435 |
||
43924 | 436 |
lemma enat_ile: "n \<le> enat m \<Longrightarrow> \<exists>k. n = enat k" |
27110 | 437 |
by (cases n) simp_all |
438 |
||
43924 | 439 |
lemma enat_iless: "n < enat m \<Longrightarrow> \<exists>k. n = enat k" |
27110 | 440 |
by (cases n) simp_all |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
441 |
|
43924 | 442 |
lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. enat k < Y j" |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
443 |
apply (induct_tac k) |
43924 | 444 |
apply (simp (no_asm) only: enat_0) |
27110 | 445 |
apply (fast intro: le_less_trans [OF i0_lb]) |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
446 |
apply (erule exE) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
447 |
apply (drule spec) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
448 |
apply (erule exE) |
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
449 |
apply (drule ileI1) |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
450 |
apply (rule eSuc_enat [THEN subst]) |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
451 |
apply (rule exI) |
27110 | 452 |
apply (erule (1) le_less_trans) |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
453 |
done |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
454 |
|
43919 | 455 |
instantiation enat :: "{bot, top}" |
29337 | 456 |
begin |
457 |
||
43919 | 458 |
definition bot_enat :: enat where |
459 |
"bot_enat = 0" |
|
29337 | 460 |
|
43919 | 461 |
definition top_enat :: enat where |
462 |
"top_enat = \<infinity>" |
|
29337 | 463 |
|
464 |
instance proof |
|
43919 | 465 |
qed (simp_all add: bot_enat_def top_enat_def) |
29337 | 466 |
|
467 |
end |
|
468 |
||
43924 | 469 |
lemma finite_enat_bounded: |
470 |
assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> enat n" |
|
42993 | 471 |
shows "finite A" |
472 |
proof (rule finite_subset) |
|
43924 | 473 |
show "finite (enat ` {..n})" by blast |
42993 | 474 |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44019
diff
changeset
|
475 |
have "A \<subseteq> {..enat n}" using le_fin by fastforce |
43924 | 476 |
also have "\<dots> \<subseteq> enat ` {..n}" |
42993 | 477 |
by (rule subsetI) (case_tac x, auto) |
43924 | 478 |
finally show "A \<subseteq> enat ` {..n}" . |
42993 | 479 |
qed |
480 |
||
26089 | 481 |
|
45775 | 482 |
subsection {* Cancellation simprocs *} |
483 |
||
484 |
lemma enat_add_left_cancel: "a + b = a + c \<longleftrightarrow> a = (\<infinity>::enat) \<or> b = c" |
|
485 |
unfolding plus_enat_def by (simp split: enat.split) |
|
486 |
||
487 |
lemma enat_add_left_cancel_le: "a + b \<le> a + c \<longleftrightarrow> a = (\<infinity>::enat) \<or> b \<le> c" |
|
488 |
unfolding plus_enat_def by (simp split: enat.split) |
|
489 |
||
490 |
lemma enat_add_left_cancel_less: "a + b < a + c \<longleftrightarrow> a \<noteq> (\<infinity>::enat) \<and> b < c" |
|
491 |
unfolding plus_enat_def by (simp split: enat.split) |
|
492 |
||
493 |
ML {* |
|
494 |
structure Cancel_Enat_Common = |
|
495 |
struct |
|
496 |
(* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) |
|
497 |
fun find_first_t _ _ [] = raise TERM("find_first_t", []) |
|
498 |
| find_first_t past u (t::terms) = |
|
499 |
if u aconv t then (rev past @ terms) |
|
500 |
else find_first_t (t::past) u terms |
|
501 |
||
502 |
val mk_sum = Arith_Data.long_mk_sum |
|
503 |
val dest_sum = Arith_Data.dest_sum |
|
504 |
val find_first = find_first_t [] |
|
505 |
val trans_tac = Numeral_Simprocs.trans_tac |
|
506 |
val norm_ss = HOL_basic_ss addsimps |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
507 |
@{thms add_ac add_0_left add_0_right} |
45775 | 508 |
fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) |
509 |
fun simplify_meta_eq ss cancel_th th = |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
510 |
Arith_Data.simplify_meta_eq [] ss |
45775 | 511 |
([th, cancel_th] MRS trans) |
512 |
fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) |
|
513 |
end |
|
514 |
||
515 |
structure Eq_Enat_Cancel = ExtractCommonTermFun |
|
516 |
(open Cancel_Enat_Common |
|
517 |
val mk_bal = HOLogic.mk_eq |
|
518 |
val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} @{typ enat} |
|
519 |
fun simp_conv _ _ = SOME @{thm enat_add_left_cancel} |
|
520 |
) |
|
521 |
||
522 |
structure Le_Enat_Cancel = ExtractCommonTermFun |
|
523 |
(open Cancel_Enat_Common |
|
524 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less_eq} |
|
525 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} @{typ enat} |
|
526 |
fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_le} |
|
527 |
) |
|
528 |
||
529 |
structure Less_Enat_Cancel = ExtractCommonTermFun |
|
530 |
(open Cancel_Enat_Common |
|
531 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less} |
|
532 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} @{typ enat} |
|
533 |
fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_less} |
|
534 |
) |
|
535 |
*} |
|
536 |
||
537 |
simproc_setup enat_eq_cancel |
|
538 |
("(l::enat) + m = n" | "(l::enat) = m + n") = |
|
539 |
{* fn phi => fn ss => fn ct => Eq_Enat_Cancel.proc ss (term_of ct) *} |
|
540 |
||
541 |
simproc_setup enat_le_cancel |
|
542 |
("(l::enat) + m \<le> n" | "(l::enat) \<le> m + n") = |
|
543 |
{* fn phi => fn ss => fn ct => Le_Enat_Cancel.proc ss (term_of ct) *} |
|
544 |
||
545 |
simproc_setup enat_less_cancel |
|
546 |
("(l::enat) + m < n" | "(l::enat) < m + n") = |
|
547 |
{* fn phi => fn ss => fn ct => Less_Enat_Cancel.proc ss (term_of ct) *} |
|
548 |
||
549 |
text {* TODO: add regression tests for these simprocs *} |
|
550 |
||
551 |
text {* TODO: add simprocs for combining and cancelling numerals *} |
|
552 |
||
553 |
||
27110 | 554 |
subsection {* Well-ordering *} |
26089 | 555 |
|
43924 | 556 |
lemma less_enatE: |
557 |
"[| n < enat m; !!k. n = enat k ==> k < m ==> P |] ==> P" |
|
26089 | 558 |
by (induct n) auto |
559 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
560 |
lemma less_infinityE: |
43924 | 561 |
"[| n < \<infinity>; !!k. n = enat k ==> P |] ==> P" |
26089 | 562 |
by (induct n) auto |
563 |
||
43919 | 564 |
lemma enat_less_induct: |
565 |
assumes prem: "!!n. \<forall>m::enat. m < n --> P m ==> P n" shows "P n" |
|
26089 | 566 |
proof - |
43924 | 567 |
have P_enat: "!!k. P (enat k)" |
26089 | 568 |
apply (rule nat_less_induct) |
569 |
apply (rule prem, clarify) |
|
43924 | 570 |
apply (erule less_enatE, simp) |
26089 | 571 |
done |
572 |
show ?thesis |
|
573 |
proof (induct n) |
|
574 |
fix nat |
|
43924 | 575 |
show "P (enat nat)" by (rule P_enat) |
26089 | 576 |
next |
43921 | 577 |
show "P \<infinity>" |
26089 | 578 |
apply (rule prem, clarify) |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
579 |
apply (erule less_infinityE) |
43924 | 580 |
apply (simp add: P_enat) |
26089 | 581 |
done |
582 |
qed |
|
583 |
qed |
|
584 |
||
43919 | 585 |
instance enat :: wellorder |
26089 | 586 |
proof |
27823 | 587 |
fix P and n |
43919 | 588 |
assume hyp: "(\<And>n\<Colon>enat. (\<And>m\<Colon>enat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)" |
589 |
show "P n" by (blast intro: enat_less_induct hyp) |
|
26089 | 590 |
qed |
591 |
||
42993 | 592 |
subsection {* Complete Lattice *} |
593 |
||
43919 | 594 |
instantiation enat :: complete_lattice |
42993 | 595 |
begin |
596 |
||
43919 | 597 |
definition inf_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where |
598 |
"inf_enat \<equiv> min" |
|
42993 | 599 |
|
43919 | 600 |
definition sup_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where |
601 |
"sup_enat \<equiv> max" |
|
42993 | 602 |
|
43919 | 603 |
definition Inf_enat :: "enat set \<Rightarrow> enat" where |
604 |
"Inf_enat A \<equiv> if A = {} then \<infinity> else (LEAST x. x \<in> A)" |
|
42993 | 605 |
|
43919 | 606 |
definition Sup_enat :: "enat set \<Rightarrow> enat" where |
607 |
"Sup_enat A \<equiv> if A = {} then 0 |
|
42993 | 608 |
else if finite A then Max A |
609 |
else \<infinity>" |
|
610 |
instance proof |
|
43919 | 611 |
fix x :: "enat" and A :: "enat set" |
42993 | 612 |
{ assume "x \<in> A" then show "Inf A \<le> x" |
43919 | 613 |
unfolding Inf_enat_def by (auto intro: Least_le) } |
42993 | 614 |
{ assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A" |
43919 | 615 |
unfolding Inf_enat_def |
42993 | 616 |
by (cases "A = {}") (auto intro: LeastI2_ex) } |
617 |
{ assume "x \<in> A" then show "x \<le> Sup A" |
|
43919 | 618 |
unfolding Sup_enat_def by (cases "finite A") auto } |
42993 | 619 |
{ assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x" |
43924 | 620 |
unfolding Sup_enat_def using finite_enat_bounded by auto } |
43919 | 621 |
qed (simp_all add: inf_enat_def sup_enat_def) |
42993 | 622 |
end |
623 |
||
43978 | 624 |
instance enat :: complete_linorder .. |
27110 | 625 |
|
626 |
subsection {* Traditional theorem names *} |
|
627 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
628 |
lemmas enat_defs = zero_enat_def one_enat_def eSuc_def |
43919 | 629 |
plus_enat_def less_eq_enat_def less_enat_def |
27110 | 630 |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
631 |
end |