author | wenzelm |
Fri, 21 Mar 2025 22:26:18 +0100 | |
changeset 82317 | 231b6d8231c6 |
parent 81816 | bee084ecd18c |
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 |
|
60500 | 6 |
section \<open>Extended natural numbers (i.e. with infinity)\<close> |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
7 |
|
43919 | 8 |
theory Extended_Nat |
60636
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
9 |
imports Main Countable Order_Continuity |
15131 | 10 |
begin |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
11 |
|
43921 | 12 |
class infinity = |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
78099
diff
changeset
|
13 |
fixes infinity :: "'a" (\<open>\<infinity>\<close>) |
43921 | 14 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
15 |
context |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
16 |
fixes f :: "nat \<Rightarrow> 'a::{canonically_ordered_monoid_add, linorder_topology, complete_linorder}" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
17 |
begin |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
18 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
19 |
lemma sums_SUP[simp, intro]: "f sums (SUP n. \<Sum>i<n. f i)" |
64267 | 20 |
unfolding sums_def by (intro LIMSEQ_SUP monoI sum_mono2 zero_le) auto |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
21 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
22 |
lemma suminf_eq_SUP: "suminf f = (SUP n. \<Sum>i<n. f i)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
23 |
using sums_SUP by (rule sums_unique[symmetric]) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
24 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
25 |
end |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
26 |
|
60500 | 27 |
subsection \<open>Type definition\<close> |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
28 |
|
60500 | 29 |
text \<open> |
11355 | 30 |
We extend the standard natural numbers by a special value indicating |
27110 | 31 |
infinity. |
60500 | 32 |
\<close> |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
33 |
|
49834 | 34 |
typedef enat = "UNIV :: nat option set" .. |
54415 | 35 |
|
69593 | 36 |
text \<open>TODO: introduce enat as coinductive datatype, enat is just \<^const>\<open>of_nat\<close>\<close> |
54415 | 37 |
|
43924 | 38 |
definition enat :: "nat \<Rightarrow> enat" where |
39 |
"enat n = Abs_enat (Some n)" |
|
62374
cb27a55d868a
remove lattice syntax from countable complete lattice
hoelzl
parents:
61631
diff
changeset
|
40 |
|
43921 | 41 |
instantiation enat :: infinity |
42 |
begin |
|
60679 | 43 |
|
44 |
definition "\<infinity> = Abs_enat None" |
|
45 |
instance .. |
|
46 |
||
43921 | 47 |
end |
54415 | 48 |
|
49 |
instance enat :: countable |
|
50 |
proof |
|
51 |
show "\<exists>to_nat::enat \<Rightarrow> nat. inj to_nat" |
|
52 |
by (rule exI[of _ "to_nat \<circ> Rep_enat"]) (simp add: inj_on_def Rep_enat_inject) |
|
53 |
qed |
|
62374
cb27a55d868a
remove lattice syntax from countable complete lattice
hoelzl
parents:
61631
diff
changeset
|
54 |
|
58306
117ba6cbe414
renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents:
57514
diff
changeset
|
55 |
old_rep_datatype enat "\<infinity> :: enat" |
43921 | 56 |
proof - |
43924 | 57 |
fix P i assume "\<And>j. P (enat j)" "P \<infinity>" |
43921 | 58 |
then show "P i" |
59 |
proof induct |
|
60 |
case (Abs_enat y) then show ?case |
|
61 |
by (cases y rule: option.exhaust) |
|
43924 | 62 |
(auto simp: enat_def infinity_enat_def) |
43921 | 63 |
qed |
43924 | 64 |
qed (auto simp add: enat_def infinity_enat_def Abs_enat_inject) |
19736 | 65 |
|
43924 | 66 |
declare [[coercion "enat::nat\<Rightarrow>enat"]] |
19736 | 67 |
|
45934 | 68 |
lemmas enat2_cases = enat.exhaust[case_product enat.exhaust] |
69 |
lemmas enat3_cases = enat.exhaust[case_product enat.exhaust enat.exhaust] |
|
70 |
||
54416 | 71 |
lemma not_infinity_eq [iff]: "(x \<noteq> \<infinity>) = (\<exists>i. x = enat i)" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
72 |
by (cases x) auto |
31084 | 73 |
|
54416 | 74 |
lemma not_enat_eq [iff]: "(\<forall>y. x \<noteq> enat y) = (x = \<infinity>)" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
75 |
by (cases x) auto |
31077 | 76 |
|
62376
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
77 |
lemma enat_ex_split: "(\<exists>c::enat. P c) \<longleftrightarrow> P \<infinity> \<or> (\<exists>c::nat. P c)" |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
78 |
by (metis enat.exhaust) |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
79 |
|
43924 | 80 |
primrec the_enat :: "enat \<Rightarrow> nat" |
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
81 |
where "the_enat (enat n) = n" |
41855 | 82 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
83 |
|
60500 | 84 |
subsection \<open>Constructors and numbers\<close> |
27110 | 85 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
86 |
instantiation enat :: zero_neq_one |
25594 | 87 |
begin |
88 |
||
89 |
definition |
|
43924 | 90 |
"0 = enat 0" |
25594 | 91 |
|
92 |
definition |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
93 |
"1 = enat 1" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
94 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
95 |
instance |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
96 |
proof qed (simp add: zero_enat_def one_enat_def) |
25594 | 97 |
|
98 |
end |
|
99 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
100 |
definition eSuc :: "enat \<Rightarrow> enat" where |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
101 |
"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
|
102 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
103 |
lemma enat_0 [code_post]: "enat 0 = 0" |
43919 | 104 |
by (simp add: zero_enat_def) |
27110 | 105 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
106 |
lemma enat_1 [code_post]: "enat 1 = 1" |
43919 | 107 |
by (simp add: one_enat_def) |
27110 | 108 |
|
54416 | 109 |
lemma enat_0_iff: "enat x = 0 \<longleftrightarrow> x = 0" "0 = enat x \<longleftrightarrow> x = 0" |
110 |
by (auto simp add: zero_enat_def) |
|
111 |
||
112 |
lemma enat_1_iff: "enat x = 1 \<longleftrightarrow> x = 1" "1 = enat x \<longleftrightarrow> x = 1" |
|
113 |
by (auto simp add: one_enat_def) |
|
114 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
115 |
lemma one_eSuc: "1 = eSuc 0" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
116 |
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
|
117 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
118 |
lemma infinity_ne_i0 [simp]: "(\<infinity>::enat) \<noteq> 0" |
43919 | 119 |
by (simp add: zero_enat_def) |
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 i0_ne_infinity [simp]: "0 \<noteq> (\<infinity>::enat)" |
43919 | 122 |
by (simp add: zero_enat_def) |
27110 | 123 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
124 |
lemma zero_one_enat_neq: |
61076 | 125 |
"\<not> 0 = (1::enat)" |
126 |
"\<not> 1 = (0::enat)" |
|
43919 | 127 |
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
|
128 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
129 |
lemma infinity_ne_i1 [simp]: "(\<infinity>::enat) \<noteq> 1" |
43919 | 130 |
by (simp add: one_enat_def) |
27110 | 131 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
132 |
lemma i1_ne_infinity [simp]: "1 \<noteq> (\<infinity>::enat)" |
43919 | 133 |
by (simp add: one_enat_def) |
27110 | 134 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
135 |
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
|
136 |
by (simp add: eSuc_def) |
27110 | 137 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
138 |
lemma eSuc_infinity [simp]: "eSuc \<infinity> = \<infinity>" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
139 |
by (simp add: eSuc_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
140 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
141 |
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
|
142 |
by (simp add: eSuc_def zero_enat_def split: enat.splits) |
27110 | 143 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
144 |
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
|
145 |
by (rule eSuc_ne_0 [symmetric]) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
146 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
147 |
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
|
148 |
by (simp add: eSuc_def split: enat.splits) |
27110 | 149 |
|
59000 | 150 |
lemma eSuc_enat_iff: "eSuc x = enat y \<longleftrightarrow> (\<exists>n. y = Suc n \<and> x = enat n)" |
151 |
by (cases y) (auto simp: enat_0 eSuc_enat[symmetric]) |
|
152 |
||
153 |
lemma enat_eSuc_iff: "enat y = eSuc x \<longleftrightarrow> (\<exists>n. y = Suc n \<and> enat n = x)" |
|
154 |
by (cases y) (auto simp: enat_0 eSuc_enat[symmetric]) |
|
155 |
||
60500 | 156 |
subsection \<open>Addition\<close> |
27110 | 157 |
|
43919 | 158 |
instantiation enat :: comm_monoid_add |
27110 | 159 |
begin |
160 |
||
38167 | 161 |
definition [nitpick_simp]: |
43924 | 162 |
"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
|
163 |
|
43919 | 164 |
lemma plus_enat_simps [simp, code]: |
43921 | 165 |
fixes q :: enat |
43924 | 166 |
shows "enat m + enat n = enat (m + n)" |
43921 | 167 |
and "\<infinity> + q = \<infinity>" |
168 |
and "q + \<infinity> = \<infinity>" |
|
43919 | 169 |
by (simp_all add: plus_enat_def split: enat.splits) |
27110 | 170 |
|
60679 | 171 |
instance |
172 |
proof |
|
43919 | 173 |
fix n m q :: enat |
27110 | 174 |
show "n + m + q = n + (m + q)" |
45934 | 175 |
by (cases n m q rule: enat3_cases) auto |
27110 | 176 |
show "n + m = m + n" |
45934 | 177 |
by (cases n m rule: enat2_cases) auto |
27110 | 178 |
show "0 + n = n" |
43919 | 179 |
by (cases n) (simp_all add: zero_enat_def) |
26089 | 180 |
qed |
181 |
||
27110 | 182 |
end |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
183 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
184 |
lemma eSuc_plus_1: |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
185 |
"eSuc n = n + 1" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
186 |
by (cases n) (simp_all add: eSuc_enat one_enat_def) |
62374
cb27a55d868a
remove lattice syntax from countable complete lattice
hoelzl
parents:
61631
diff
changeset
|
187 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
188 |
lemma plus_1_eSuc: |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
189 |
"1 + q = eSuc q" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
190 |
"q + 1 = eSuc q" |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
191 |
by (simp_all add: eSuc_plus_1 ac_simps) |
41853 | 192 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
193 |
lemma iadd_Suc: "eSuc m + n = eSuc (m + n)" |
81816 | 194 |
by (simp add: eSuc_plus_1 ac_simps) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
195 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
196 |
lemma iadd_Suc_right: "m + eSuc n = eSuc (m + n)" |
81816 | 197 |
by (metis add.commute iadd_Suc) |
41853 | 198 |
|
60500 | 199 |
subsection \<open>Multiplication\<close> |
29014 | 200 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
201 |
instantiation enat :: "{comm_semiring_1, semiring_no_zero_divisors}" |
29014 | 202 |
begin |
203 |
||
43919 | 204 |
definition times_enat_def [nitpick_simp]: |
43924 | 205 |
"m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | enat m \<Rightarrow> |
206 |
(case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | enat n \<Rightarrow> enat (m * n)))" |
|
29014 | 207 |
|
43919 | 208 |
lemma times_enat_simps [simp, code]: |
43924 | 209 |
"enat m * enat n = enat (m * n)" |
43921 | 210 |
"\<infinity> * \<infinity> = (\<infinity>::enat)" |
43924 | 211 |
"\<infinity> * enat n = (if n = 0 then 0 else \<infinity>)" |
212 |
"enat m * \<infinity> = (if m = 0 then 0 else \<infinity>)" |
|
43919 | 213 |
unfolding times_enat_def zero_enat_def |
214 |
by (simp_all split: enat.split) |
|
29014 | 215 |
|
60679 | 216 |
instance |
217 |
proof |
|
43919 | 218 |
fix a b c :: enat |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
219 |
show distr: "(a + b) * c = a * c + b * c" |
43919 | 220 |
unfolding times_enat_def zero_enat_def |
49962
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents:
49834
diff
changeset
|
221 |
by (simp split: enat.split add: distrib_right) |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
222 |
show "a * (b + c) = a * b + a * c" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
223 |
by (cases a b c rule: enat3_cases) (auto simp: times_enat_def zero_enat_def distrib_left) |
81816 | 224 |
qed (auto simp: times_enat_def zero_enat_def one_enat_def split: enat.split) |
29014 | 225 |
|
226 |
end |
|
227 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
228 |
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
|
229 |
unfolding eSuc_plus_1 by (simp add: algebra_simps) |
29014 | 230 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
231 |
lemma mult_eSuc_right: "m * eSuc n = m + m * n" |
81816 | 232 |
by (metis mult.commute mult_eSuc) |
29014 | 233 |
|
43924 | 234 |
lemma of_nat_eq_enat: "of_nat n = enat n" |
81816 | 235 |
by (induct n) (auto simp: enat_0 plus_1_eSuc eSuc_enat) |
29023 | 236 |
|
60679 | 237 |
instance enat :: semiring_char_0 |
238 |
proof |
|
43924 | 239 |
have "inj enat" by (rule injI) simp |
240 |
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
|
241 |
qed |
29023 | 242 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
243 |
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
|
244 |
by (auto simp add: times_enat_def zero_enat_def split: enat.split) |
41853 | 245 |
|
60500 | 246 |
subsection \<open>Numerals\<close> |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
247 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
248 |
lemma numeral_eq_enat: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
249 |
"numeral k = enat (numeral k)" |
81816 | 250 |
by (metis of_nat_eq_enat of_nat_numeral) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
251 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
252 |
lemma enat_numeral [code_abbrev]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
253 |
"enat (numeral k) = numeral k" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
254 |
using numeral_eq_enat .. |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
255 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
256 |
lemma infinity_ne_numeral [simp]: "(\<infinity>::enat) \<noteq> numeral k" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
257 |
by (simp add: numeral_eq_enat) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
258 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
259 |
lemma numeral_ne_infinity [simp]: "numeral k \<noteq> (\<infinity>::enat)" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
260 |
by (simp add: numeral_eq_enat) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
261 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
262 |
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
|
263 |
by (simp only: eSuc_plus_1 numeral_plus_one) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
264 |
|
60500 | 265 |
subsection \<open>Subtraction\<close> |
41853 | 266 |
|
43919 | 267 |
instantiation enat :: minus |
41853 | 268 |
begin |
269 |
||
43919 | 270 |
definition diff_enat_def: |
43924 | 271 |
"a - b = (case a of (enat x) \<Rightarrow> (case b of (enat y) \<Rightarrow> enat (x - y) | \<infinity> \<Rightarrow> 0) |
41853 | 272 |
| \<infinity> \<Rightarrow> \<infinity>)" |
273 |
||
274 |
instance .. |
|
275 |
||
276 |
end |
|
277 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
278 |
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
|
279 |
by (simp add: diff_enat_def) |
41853 | 280 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
281 |
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
|
282 |
by (simp add: diff_enat_def) |
41853 | 283 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
284 |
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
|
285 |
by (simp add: diff_enat_def) |
41853 | 286 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
287 |
lemma idiff_0 [simp]: "(0::enat) - n = 0" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
288 |
by (cases n, simp_all add: zero_enat_def) |
41853 | 289 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
290 |
lemmas idiff_enat_0 [simp] = idiff_0 [unfolded zero_enat_def] |
41853 | 291 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
292 |
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
|
293 |
by (cases n) (simp_all add: zero_enat_def) |
41853 | 294 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
295 |
lemmas idiff_enat_0_right [simp] = idiff_0_right [unfolded zero_enat_def] |
41853 | 296 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
297 |
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
|
298 |
by (auto simp: zero_enat_def) |
41853 | 299 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
300 |
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
|
301 |
by (simp add: eSuc_def split: enat.split) |
41855 | 302 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
303 |
lemma eSuc_minus_1 [simp]: "eSuc n - 1 = n" |
68406 | 304 |
by (simp add: one_enat_def flip: eSuc_enat zero_enat_def) |
41855 | 305 |
|
43924 | 306 |
(*lemmas idiff_self_eq_0_enat = idiff_self_eq_0[unfolded zero_enat_def]*) |
41853 | 307 |
|
60500 | 308 |
subsection \<open>Ordering\<close> |
27110 | 309 |
|
43919 | 310 |
instantiation enat :: linordered_ab_semigroup_add |
27110 | 311 |
begin |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
312 |
|
38167 | 313 |
definition [nitpick_simp]: |
43924 | 314 |
"m \<le> n = (case n of enat n1 \<Rightarrow> (case m of enat m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False) |
27110 | 315 |
| \<infinity> \<Rightarrow> True)" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
316 |
|
38167 | 317 |
definition [nitpick_simp]: |
43924 | 318 |
"m < n = (case m of enat m1 \<Rightarrow> (case n of enat n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True) |
27110 | 319 |
| \<infinity> \<Rightarrow> False)" |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
320 |
|
43919 | 321 |
lemma enat_ord_simps [simp]: |
43924 | 322 |
"enat m \<le> enat n \<longleftrightarrow> m \<le> n" |
323 |
"enat m < enat n \<longleftrightarrow> m < n" |
|
43921 | 324 |
"q \<le> (\<infinity>::enat)" |
325 |
"q < (\<infinity>::enat) \<longleftrightarrow> q \<noteq> \<infinity>" |
|
326 |
"(\<infinity>::enat) \<le> q \<longleftrightarrow> q = \<infinity>" |
|
327 |
"(\<infinity>::enat) < q \<longleftrightarrow> False" |
|
43919 | 328 |
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
|
329 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
330 |
lemma numeral_le_enat_iff[simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
331 |
shows "numeral m \<le> enat n \<longleftrightarrow> numeral m \<le> n" |
81816 | 332 |
by (auto simp: numeral_eq_enat) |
45934 | 333 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
334 |
lemma numeral_less_enat_iff[simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
335 |
shows "numeral m < enat n \<longleftrightarrow> numeral m < n" |
81816 | 336 |
by (auto simp: numeral_eq_enat) |
45934 | 337 |
|
43919 | 338 |
lemma enat_ord_code [code]: |
43924 | 339 |
"enat m \<le> enat n \<longleftrightarrow> m \<le> n" |
340 |
"enat m < enat n \<longleftrightarrow> m < n" |
|
43921 | 341 |
"q \<le> (\<infinity>::enat) \<longleftrightarrow> True" |
43924 | 342 |
"enat m < \<infinity> \<longleftrightarrow> True" |
343 |
"\<infinity> \<le> enat n \<longleftrightarrow> False" |
|
43921 | 344 |
"(\<infinity>::enat) < q \<longleftrightarrow> False" |
27110 | 345 |
by simp_all |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
346 |
|
60679 | 347 |
instance |
348 |
by standard (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
|
349 |
|
27110 | 350 |
end |
351 |
||
62376
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
352 |
instance enat :: dioid |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
353 |
proof |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
354 |
fix a b :: enat show "(a \<le> b) = (\<exists>c. b = a + c)" |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
355 |
by (cases a b rule: enat2_cases) (auto simp: le_iff_add enat_ex_split) |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
356 |
qed |
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents:
62374
diff
changeset
|
357 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
358 |
instance enat :: "{linordered_nonzero_semiring, strict_ordered_comm_monoid_add}" |
29014 | 359 |
proof |
43919 | 360 |
fix a b c :: enat |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
361 |
show "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow>c * a \<le> c * b" |
43919 | 362 |
unfolding times_enat_def less_eq_enat_def zero_enat_def |
363 |
by (simp split: enat.splits) |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
364 |
show "a < b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d" for a b c d :: enat |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
365 |
by (cases a b c d rule: enat2_cases[case_product enat2_cases]) auto |
67689
2c38ffd6ec71
type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
366 |
show "a < b \<Longrightarrow> a + 1 < b + 1" |
2c38ffd6ec71
type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents:
67091
diff
changeset
|
367 |
by (metis add_right_mono eSuc_minus_1 eSuc_plus_1 less_le) |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
368 |
qed (simp add: zero_enat_def one_enat_def) |
29014 | 369 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
370 |
(* 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
|
371 |
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
|
372 |
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
|
373 |
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
|
374 |
|
69800 | 375 |
lemma add_diff_assoc_enat: "z \<le> y \<Longrightarrow> x + (y - z) = x + y - (z::enat)" |
376 |
by(cases x)(auto simp add: diff_enat_def split: enat.split) |
|
377 |
||
43919 | 378 |
lemma enat_ord_number [simp]: |
61076 | 379 |
"(numeral m :: enat) \<le> numeral n \<longleftrightarrow> (numeral m :: nat) \<le> numeral n" |
380 |
"(numeral m :: enat) < numeral n \<longleftrightarrow> (numeral m :: nat) < numeral n" |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
381 |
by (simp_all add: numeral_eq_enat) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
382 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
383 |
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
|
384 |
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
|
385 |
|
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
386 |
lemma infinity_ilessE [elim!]: "\<infinity> < enat m \<Longrightarrow> R" |
27110 | 387 |
by simp |
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 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
|
390 |
by (simp add: eSuc_def less_eq_enat_def split: enat.splits) |
62374
cb27a55d868a
remove lattice syntax from countable complete lattice
hoelzl
parents:
61631
diff
changeset
|
391 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
392 |
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
|
393 |
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
|
394 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
395 |
lemma ile_eSuc [simp]: "n \<le> eSuc n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
396 |
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
|
397 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
398 |
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
|
399 |
by (simp add: zero_enat_def eSuc_def less_eq_enat_def split: enat.splits) |
27110 | 400 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
401 |
lemma i0_iless_eSuc [simp]: "0 < eSuc n" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
402 |
by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.splits) |
27110 | 403 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
404 |
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
|
405 |
by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.split) |
41853 | 406 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
407 |
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
|
408 |
by (simp add: eSuc_def less_eq_enat_def less_enat_def split: enat.splits) |
27110 | 409 |
|
43924 | 410 |
lemma Suc_ile_eq: "enat (Suc m) \<le> n \<longleftrightarrow> enat m < n" |
27110 | 411 |
by (cases n) auto |
412 |
||
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
413 |
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
|
414 |
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
|
415 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
416 |
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
|
417 |
by (simp add: zero_enat_def less_enat_def split: enat.splits) |
41853 | 418 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
419 |
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
|
420 |
by (simp add: zero_enat_def less_enat_def split: enat.splits) |
41853 | 421 |
|
43919 | 422 |
lemma enat_0_less_mult_iff: "(0 < (m::enat) * n) = (0 < m \<and> 0 < n)" |
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
423 |
by (simp only: zero_less_iff_neq_zero mult_eq_0_iff, simp) |
41853 | 424 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
425 |
lemma mono_eSuc: "mono eSuc" |
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
426 |
by (simp add: mono_def) |
41853 | 427 |
|
43919 | 428 |
lemma min_enat_simps [simp]: |
43924 | 429 |
"min (enat m) (enat n) = enat (min m n)" |
27110 | 430 |
"min q 0 = 0" |
431 |
"min 0 q = 0" |
|
43921 | 432 |
"min q (\<infinity>::enat) = q" |
433 |
"min (\<infinity>::enat) q = q" |
|
27110 | 434 |
by (auto simp add: min_def) |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
435 |
|
43919 | 436 |
lemma max_enat_simps [simp]: |
43924 | 437 |
"max (enat m) (enat n) = enat (max m n)" |
27110 | 438 |
"max q 0 = q" |
439 |
"max 0 q = q" |
|
43921 | 440 |
"max q \<infinity> = (\<infinity>::enat)" |
441 |
"max \<infinity> q = (\<infinity>::enat)" |
|
27110 | 442 |
by (simp_all add: max_def) |
443 |
||
43924 | 444 |
lemma enat_ile: "n \<le> enat m \<Longrightarrow> \<exists>k. n = enat k" |
27110 | 445 |
by (cases n) simp_all |
446 |
||
43924 | 447 |
lemma enat_iless: "n < enat m \<Longrightarrow> \<exists>k. n = enat k" |
27110 | 448 |
by (cases n) simp_all |
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
449 |
|
61631
4f7ef088c4ed
add lemmas for extended nats and reals
Andreas Lochbihler
parents:
61384
diff
changeset
|
450 |
lemma iadd_le_enat_iff: |
4f7ef088c4ed
add lemmas for extended nats and reals
Andreas Lochbihler
parents:
61384
diff
changeset
|
451 |
"x + y \<le> enat n \<longleftrightarrow> (\<exists>y' x'. x = enat x' \<and> y = enat y' \<and> x' + y' \<le> n)" |
4f7ef088c4ed
add lemmas for extended nats and reals
Andreas Lochbihler
parents:
61384
diff
changeset
|
452 |
by(cases x y rule: enat.exhaust[case_product enat.exhaust]) simp_all |
4f7ef088c4ed
add lemmas for extended nats and reals
Andreas Lochbihler
parents:
61384
diff
changeset
|
453 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
454 |
lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j \<Longrightarrow> \<exists>j. enat k < Y j" |
81816 | 455 |
proof (induction k) |
456 |
case 0 |
|
457 |
then show ?case |
|
458 |
using enat_0 zero_less_iff_neq_zero by fastforce |
|
459 |
next |
|
460 |
case (Suc k) |
|
461 |
then show ?case |
|
462 |
by (meson Suc_ile_eq order_le_less_trans) |
|
463 |
qed |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
464 |
|
60636
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
465 |
lemma eSuc_max: "eSuc (max x y) = max (eSuc x) (eSuc y)" |
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
466 |
by (simp add: eSuc_def split: enat.split) |
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
467 |
|
62374
cb27a55d868a
remove lattice syntax from countable complete lattice
hoelzl
parents:
61631
diff
changeset
|
468 |
lemma eSuc_Max: |
60636
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
469 |
assumes "finite A" "A \<noteq> {}" |
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
470 |
shows "eSuc (Max A) = Max (eSuc ` A)" |
81816 | 471 |
by (simp add: assms mono_Max_commute mono_eSuc) |
60636
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
472 |
|
52729
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents:
51717
diff
changeset
|
473 |
instantiation enat :: "{order_bot, order_top}" |
29337 | 474 |
begin |
475 |
||
60679 | 476 |
definition bot_enat :: enat where "bot_enat = 0" |
477 |
definition top_enat :: enat where "top_enat = \<infinity>" |
|
29337 | 478 |
|
60679 | 479 |
instance |
480 |
by standard (simp_all add: bot_enat_def top_enat_def) |
|
29337 | 481 |
|
482 |
end |
|
483 |
||
43924 | 484 |
lemma finite_enat_bounded: |
485 |
assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> enat n" |
|
42993 | 486 |
shows "finite A" |
487 |
proof (rule finite_subset) |
|
43924 | 488 |
show "finite (enat ` {..n})" by blast |
81816 | 489 |
have "A \<subseteq> enat ` {..n}" |
490 |
using enat_ile le_fin by fastforce |
|
491 |
then show "A \<subseteq> enat ` {..n}" . |
|
42993 | 492 |
qed |
493 |
||
26089 | 494 |
|
60500 | 495 |
subsection \<open>Cancellation simprocs\<close> |
45775 | 496 |
|
69803 | 497 |
lemma add_diff_cancel_enat[simp]: "x \<noteq> \<infinity> \<Longrightarrow> x + y - x = (y::enat)" |
81816 | 498 |
by (metis add.commute add.right_neutral add_diff_assoc_enat idiff_self order_refl) |
69803 | 499 |
|
45775 | 500 |
lemma enat_add_left_cancel: "a + b = a + c \<longleftrightarrow> a = (\<infinity>::enat) \<or> b = c" |
501 |
unfolding plus_enat_def by (simp split: enat.split) |
|
502 |
||
503 |
lemma enat_add_left_cancel_le: "a + b \<le> a + c \<longleftrightarrow> a = (\<infinity>::enat) \<or> b \<le> c" |
|
504 |
unfolding plus_enat_def by (simp split: enat.split) |
|
505 |
||
506 |
lemma enat_add_left_cancel_less: "a + b < a + c \<longleftrightarrow> a \<noteq> (\<infinity>::enat) \<and> b < c" |
|
507 |
unfolding plus_enat_def by (simp split: enat.split) |
|
508 |
||
69801 | 509 |
lemma plus_eq_infty_iff_enat: "(m::enat) + n = \<infinity> \<longleftrightarrow> m=\<infinity> \<or> n=\<infinity>" |
81816 | 510 |
using enat_add_left_cancel by fastforce |
69800 | 511 |
|
60500 | 512 |
ML \<open> |
45775 | 513 |
structure Cancel_Enat_Common = |
514 |
struct |
|
515 |
(* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) |
|
516 |
fun find_first_t _ _ [] = raise TERM("find_first_t", []) |
|
517 |
| find_first_t past u (t::terms) = |
|
518 |
if u aconv t then (rev past @ terms) |
|
519 |
else find_first_t (t::past) u terms |
|
520 |
||
69593 | 521 |
fun dest_summing (Const (\<^const_name>\<open>Groups.plus\<close>, _) $ t $ u, ts) = |
51366
abdcf1a7cabf
avoid using Arith_Data.dest_sum in extended-nat simprocs (it treats 'x - y' as 'x + - y', which is not valid for enat)
huffman
parents:
51301
diff
changeset
|
522 |
dest_summing (t, dest_summing (u, ts)) |
abdcf1a7cabf
avoid using Arith_Data.dest_sum in extended-nat simprocs (it treats 'x - y' as 'x + - y', which is not valid for enat)
huffman
parents:
51301
diff
changeset
|
523 |
| dest_summing (t, ts) = t :: ts |
abdcf1a7cabf
avoid using Arith_Data.dest_sum in extended-nat simprocs (it treats 'x - y' as 'x + - y', which is not valid for enat)
huffman
parents:
51301
diff
changeset
|
524 |
|
45775 | 525 |
val mk_sum = Arith_Data.long_mk_sum |
51366
abdcf1a7cabf
avoid using Arith_Data.dest_sum in extended-nat simprocs (it treats 'x - y' as 'x + - y', which is not valid for enat)
huffman
parents:
51301
diff
changeset
|
526 |
fun dest_sum t = dest_summing (t, []) |
45775 | 527 |
val find_first = find_first_t [] |
528 |
val trans_tac = Numeral_Simprocs.trans_tac |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51366
diff
changeset
|
529 |
val norm_ss = |
69593 | 530 |
simpset_of (put_simpset HOL_basic_ss \<^context> |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
531 |
addsimps @{thms ac_simps add_0_left add_0_right}) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51366
diff
changeset
|
532 |
fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51366
diff
changeset
|
533 |
fun simplify_meta_eq ctxt cancel_th th = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51366
diff
changeset
|
534 |
Arith_Data.simplify_meta_eq [] ctxt |
45775 | 535 |
([th, cancel_th] MRS trans) |
536 |
fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) |
|
537 |
end |
|
538 |
||
539 |
structure Eq_Enat_Cancel = ExtractCommonTermFun |
|
540 |
(open Cancel_Enat_Common |
|
541 |
val mk_bal = HOLogic.mk_eq |
|
69593 | 542 |
val dest_bal = HOLogic.dest_bin \<^const_name>\<open>HOL.eq\<close> \<^typ>\<open>enat\<close> |
45775 | 543 |
fun simp_conv _ _ = SOME @{thm enat_add_left_cancel} |
544 |
) |
|
545 |
||
546 |
structure Le_Enat_Cancel = ExtractCommonTermFun |
|
547 |
(open Cancel_Enat_Common |
|
69593 | 548 |
val mk_bal = HOLogic.mk_binrel \<^const_name>\<open>Orderings.less_eq\<close> |
549 |
val dest_bal = HOLogic.dest_bin \<^const_name>\<open>Orderings.less_eq\<close> \<^typ>\<open>enat\<close> |
|
45775 | 550 |
fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_le} |
551 |
) |
|
552 |
||
553 |
structure Less_Enat_Cancel = ExtractCommonTermFun |
|
554 |
(open Cancel_Enat_Common |
|
69593 | 555 |
val mk_bal = HOLogic.mk_binrel \<^const_name>\<open>Orderings.less\<close> |
556 |
val dest_bal = HOLogic.dest_bin \<^const_name>\<open>Orderings.less\<close> \<^typ>\<open>enat\<close> |
|
45775 | 557 |
fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_less} |
558 |
) |
|
60500 | 559 |
\<close> |
45775 | 560 |
|
561 |
simproc_setup enat_eq_cancel |
|
562 |
("(l::enat) + m = n" | "(l::enat) = m + n") = |
|
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
69861
diff
changeset
|
563 |
\<open>K (fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct))\<close> |
45775 | 564 |
|
565 |
simproc_setup enat_le_cancel |
|
566 |
("(l::enat) + m \<le> n" | "(l::enat) \<le> m + n") = |
|
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
69861
diff
changeset
|
567 |
\<open>K (fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct))\<close> |
45775 | 568 |
|
569 |
simproc_setup enat_less_cancel |
|
570 |
("(l::enat) + m < n" | "(l::enat) < m + n") = |
|
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
69861
diff
changeset
|
571 |
\<open>K (fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct))\<close> |
45775 | 572 |
|
60500 | 573 |
text \<open>TODO: add regression tests for these simprocs\<close> |
45775 | 574 |
|
60500 | 575 |
text \<open>TODO: add simprocs for combining and cancelling numerals\<close> |
45775 | 576 |
|
60500 | 577 |
subsection \<open>Well-ordering\<close> |
26089 | 578 |
|
43924 | 579 |
lemma less_enatE: |
81816 | 580 |
"\<lbrakk>n < enat m; \<And>k. \<lbrakk>n = enat k; k < m\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
581 |
using enat_iless enat_ord_simps(2) by blast |
|
26089 | 582 |
|
44019
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
huffman
parents:
43978
diff
changeset
|
583 |
lemma less_infinityE: |
81816 | 584 |
"\<lbrakk>n < \<infinity>; \<And>k. n = enat k \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
585 |
by auto |
|
26089 | 586 |
|
43919 | 587 |
lemma enat_less_induct: |
81816 | 588 |
assumes "\<And>n. \<forall>m::enat. m < n \<longrightarrow> P m \<Longrightarrow> P n" |
589 |
shows "P n" |
|
26089 | 590 |
proof - |
81816 | 591 |
have "P (enat k)" for k |
592 |
by (induction k rule: less_induct) (metis less_enatE assms) |
|
593 |
then show ?thesis |
|
594 |
by (metis enat.exhaust less_infinityE assms) |
|
26089 | 595 |
qed |
596 |
||
43919 | 597 |
instance enat :: wellorder |
26089 | 598 |
proof |
27823 | 599 |
fix P and n |
61076 | 600 |
assume hyp: "(\<And>n::enat. (\<And>m::enat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)" |
43919 | 601 |
show "P n" by (blast intro: enat_less_induct hyp) |
26089 | 602 |
qed |
603 |
||
60500 | 604 |
subsection \<open>Complete Lattice\<close> |
42993 | 605 |
|
43919 | 606 |
instantiation enat :: complete_lattice |
42993 | 607 |
begin |
608 |
||
43919 | 609 |
definition inf_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where |
56777 | 610 |
"inf_enat = min" |
42993 | 611 |
|
43919 | 612 |
definition sup_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where |
56777 | 613 |
"sup_enat = max" |
42993 | 614 |
|
43919 | 615 |
definition Inf_enat :: "enat set \<Rightarrow> enat" where |
56777 | 616 |
"Inf_enat A = (if A = {} then \<infinity> else (LEAST x. x \<in> A))" |
42993 | 617 |
|
43919 | 618 |
definition Sup_enat :: "enat set \<Rightarrow> enat" where |
56777 | 619 |
"Sup_enat A = (if A = {} then 0 else if finite A then Max A else \<infinity>)" |
81332 | 620 |
|
56777 | 621 |
instance |
622 |
proof |
|
43919 | 623 |
fix x :: "enat" and A :: "enat set" |
81332 | 624 |
show "x \<in> A \<Longrightarrow> Inf A \<le> x" |
625 |
unfolding Inf_enat_def by (auto intro: Least_le) |
|
626 |
show "(\<And>y. y \<in> A \<Longrightarrow> x \<le> y) \<Longrightarrow> x \<le> Inf A" |
|
627 |
unfolding Inf_enat_def |
|
628 |
by (cases "A = {}") (auto intro: LeastI2_ex) |
|
629 |
show "x \<in> A \<Longrightarrow> x \<le> Sup A" |
|
630 |
unfolding Sup_enat_def by (cases "finite A") auto |
|
631 |
show "(\<And>y. y \<in> A \<Longrightarrow> y \<le> x) \<Longrightarrow> Sup A \<le> x" |
|
632 |
unfolding Sup_enat_def using finite_enat_bounded by auto |
|
633 |
qed (simp_all add: inf_enat_def sup_enat_def bot_enat_def top_enat_def Inf_enat_def Sup_enat_def) |
|
634 |
||
42993 | 635 |
end |
636 |
||
43978 | 637 |
instance enat :: complete_linorder .. |
27110 | 638 |
|
60636
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
639 |
lemma eSuc_Sup: "A \<noteq> {} \<Longrightarrow> eSuc (Sup A) = Sup (eSuc ` A)" |
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
640 |
by(auto simp add: Sup_enat_def eSuc_Max inj_on_def dest: finite_imageD) |
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
641 |
|
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
642 |
lemma sup_continuous_eSuc: "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. eSuc (f x))" |
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69803
diff
changeset
|
643 |
using eSuc_Sup [of "_ ` UNIV"] by (auto simp: sup_continuous_def image_comp) |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69803
diff
changeset
|
644 |
|
60636
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents:
60500
diff
changeset
|
645 |
|
60500 | 646 |
subsection \<open>Traditional theorem names\<close> |
27110 | 647 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
45934
diff
changeset
|
648 |
lemmas enat_defs = zero_enat_def one_enat_def eSuc_def |
43919 | 649 |
plus_enat_def less_eq_enat_def less_enat_def |
27110 | 650 |
|
62378
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
651 |
lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \<and> n = 0)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
652 |
by (rule add_eq_0_iff_both_eq_0) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
653 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
654 |
lemma i0_lb : "(0::enat) \<le> n" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
655 |
by (rule zero_le) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
656 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
657 |
lemma ile0_eq: "n \<le> (0::enat) \<longleftrightarrow> n = 0" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
658 |
by (rule le_zero_eq) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
659 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
660 |
lemma not_iless0: "\<not> n < (0::enat)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
661 |
by (rule not_less_zero) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
662 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
663 |
lemma i0_less[simp]: "(0::enat) < n \<longleftrightarrow> n \<noteq> 0" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
664 |
by (rule zero_less_iff_neq_zero) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
665 |
|
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
666 |
lemma imult_is_0: "((m::enat) * n = 0) = (m = 0 \<or> n = 0)" |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
667 |
by (rule mult_eq_0_iff) |
85ed00c1fe7c
generalize more theorems to support enat and ennreal
hoelzl
parents:
62376
diff
changeset
|
668 |
|
11351
c5c403d30c77
added Library/Nat_Infinity.thy and Library/Continuity.thy
oheimb
parents:
diff
changeset
|
669 |
end |