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