| author | wenzelm | 
| Sun, 12 Feb 2006 21:34:25 +0100 | |
| changeset 19031 | 0059b5b195a2 | 
| parent 16417 | 9bc16273c2d4 | 
| child 24893 | b8ef7afe3a6b | 
| permissions | -rw-r--r-- | 
| 9654 
9754ba005b64
X-symbols for ordinal, cardinal, integer arithmetic
 paulson parents: 
9492diff
changeset | 1 | (* Title: ZF/Arith.thy | 
| 0 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1992 University of Cambridge | 
| 5 | ||
| 6 | *) | |
| 7 | ||
| 13163 | 8 | (*"Difference" is subtraction of natural numbers. | 
| 9 | There are no negative numbers; we have | |
| 10 | m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. | |
| 11 | Also, rec(m, 0, %z w.z) is pred(m). | |
| 12 | *) | |
| 13 | ||
| 13328 | 14 | header{*Arithmetic Operators and Their Definitions*} 
 | 
| 15 | ||
| 16417 | 16 | theory Arith imports Univ begin | 
| 6070 | 17 | |
| 13328 | 18 | text{*Proofs about elementary arithmetic: addition, multiplication, etc.*}
 | 
| 19 | ||
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 20 | constdefs | 
| 13163 | 21 | pred :: "i=>i" (*inverse of succ*) | 
| 13361 | 22 | "pred(y) == nat_case(0, %x. x, y)" | 
| 6070 | 23 | |
| 13163 | 24 | natify :: "i=>i" (*coerces non-nats to nats*) | 
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 25 | "natify == Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a)) | 
| 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 26 | else 0)" | 
| 6070 | 27 | |
| 0 | 28 | consts | 
| 13163 | 29 | raw_add :: "[i,i]=>i" | 
| 30 | raw_diff :: "[i,i]=>i" | |
| 31 | raw_mult :: "[i,i]=>i" | |
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 32 | |
| 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 33 | primrec | 
| 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 34 | "raw_add (0, n) = n" | 
| 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 35 | "raw_add (succ(m), n) = succ(raw_add(m, n))" | 
| 0 | 36 | |
| 6070 | 37 | primrec | 
| 13163 | 38 | raw_diff_0: "raw_diff(m, 0) = m" | 
| 39 | raw_diff_succ: "raw_diff(m, succ(n)) = | |
| 40 | nat_case(0, %x. x, raw_diff(m, n))" | |
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 41 | |
| 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 42 | primrec | 
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 43 | "raw_mult(0, n) = 0" | 
| 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 44 | "raw_mult(succ(m), n) = raw_add (n, raw_mult(m, n))" | 
| 13163 | 45 | |
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 46 | constdefs | 
| 13163 | 47 | add :: "[i,i]=>i" (infixl "#+" 65) | 
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 48 | "m #+ n == raw_add (natify(m), natify(n))" | 
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 49 | |
| 13163 | 50 | diff :: "[i,i]=>i" (infixl "#-" 65) | 
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 51 | "m #- n == raw_diff (natify(m), natify(n))" | 
| 0 | 52 | |
| 13163 | 53 | mult :: "[i,i]=>i" (infixl "#*" 70) | 
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 54 | "m #* n == raw_mult (natify(m), natify(n))" | 
| 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 55 | |
| 13163 | 56 | raw_div :: "[i,i]=>i" | 
| 57 | "raw_div (m, n) == | |
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 58 | transrec(m, %j f. if j<n | n=0 then 0 else succ(f`(j#-n)))" | 
| 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 59 | |
| 13163 | 60 | raw_mod :: "[i,i]=>i" | 
| 61 | "raw_mod (m, n) == | |
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 62 | transrec(m, %j f. if j<n | n=0 then j else f`(j#-n))" | 
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 63 | |
| 13163 | 64 | div :: "[i,i]=>i" (infixl "div" 70) | 
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 65 | "m div n == raw_div (natify(m), natify(n))" | 
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 66 | |
| 13163 | 67 | mod :: "[i,i]=>i" (infixl "mod" 70) | 
| 9492 
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
 paulson parents: 
9491diff
changeset | 68 | "m mod n == raw_mod (natify(m), natify(n))" | 
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6070diff
changeset | 69 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
9964diff
changeset | 70 | syntax (xsymbols) | 
| 13163 | 71 | "mult" :: "[i,i] => i" (infixr "#\<times>" 70) | 
| 9964 | 72 | |
| 73 | syntax (HTML output) | |
| 13163 | 74 | "mult" :: "[i, i] => i" (infixr "#\<times>" 70) | 
| 75 | ||
| 76 | ||
| 77 | declare rec_type [simp] | |
| 78 | nat_0_le [simp] | |
| 79 | ||
| 80 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 81 | lemma zero_lt_lemma: "[| 0<k; k \<in> nat |] ==> \<exists>j\<in>nat. k = succ(j)" | 
| 13163 | 82 | apply (erule rev_mp) | 
| 83 | apply (induct_tac "k", auto) | |
| 84 | done | |
| 85 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 86 | (* [| 0 < k; k \<in> nat; !!j. [| j \<in> nat; k = succ(j) |] ==> Q |] ==> Q *) | 
| 13163 | 87 | lemmas zero_lt_natE = zero_lt_lemma [THEN bexE, standard] | 
| 88 | ||
| 89 | ||
| 13356 | 90 | subsection{*@{text natify}, the Coercion to @{term nat}*}
 | 
| 13163 | 91 | |
| 92 | lemma pred_succ_eq [simp]: "pred(succ(y)) = y" | |
| 93 | by (unfold pred_def, auto) | |
| 94 | ||
| 95 | lemma natify_succ: "natify(succ(x)) = succ(natify(x))" | |
| 96 | by (rule natify_def [THEN def_Vrecursor, THEN trans], auto) | |
| 97 | ||
| 98 | lemma natify_0 [simp]: "natify(0) = 0" | |
| 99 | by (rule natify_def [THEN def_Vrecursor, THEN trans], auto) | |
| 100 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 101 | lemma natify_non_succ: "\<forall>z. x ~= succ(z) ==> natify(x) = 0" | 
| 13163 | 102 | by (rule natify_def [THEN def_Vrecursor, THEN trans], auto) | 
| 103 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 104 | lemma natify_in_nat [iff,TC]: "natify(x) \<in> nat" | 
| 13163 | 105 | apply (rule_tac a=x in eps_induct) | 
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 106 | apply (case_tac "\<exists>z. x = succ(z)") | 
| 13163 | 107 | apply (auto simp add: natify_succ natify_non_succ) | 
| 108 | done | |
| 109 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 110 | lemma natify_ident [simp]: "n \<in> nat ==> natify(n) = n" | 
| 13163 | 111 | apply (induct_tac "n") | 
| 112 | apply (auto simp add: natify_succ) | |
| 113 | done | |
| 114 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 115 | lemma natify_eqE: "[|natify(x) = y; x \<in> nat|] ==> x=y" | 
| 13163 | 116 | by auto | 
| 117 | ||
| 118 | ||
| 119 | (*** Collapsing rules: to remove natify from arithmetic expressions ***) | |
| 120 | ||
| 121 | lemma natify_idem [simp]: "natify(natify(x)) = natify(x)" | |
| 122 | by simp | |
| 123 | ||
| 124 | (** Addition **) | |
| 125 | ||
| 126 | lemma add_natify1 [simp]: "natify(m) #+ n = m #+ n" | |
| 127 | by (simp add: add_def) | |
| 128 | ||
| 129 | lemma add_natify2 [simp]: "m #+ natify(n) = m #+ n" | |
| 130 | by (simp add: add_def) | |
| 131 | ||
| 132 | (** Multiplication **) | |
| 133 | ||
| 134 | lemma mult_natify1 [simp]: "natify(m) #* n = m #* n" | |
| 135 | by (simp add: mult_def) | |
| 136 | ||
| 137 | lemma mult_natify2 [simp]: "m #* natify(n) = m #* n" | |
| 138 | by (simp add: mult_def) | |
| 139 | ||
| 140 | (** Difference **) | |
| 141 | ||
| 142 | lemma diff_natify1 [simp]: "natify(m) #- n = m #- n" | |
| 143 | by (simp add: diff_def) | |
| 144 | ||
| 145 | lemma diff_natify2 [simp]: "m #- natify(n) = m #- n" | |
| 146 | by (simp add: diff_def) | |
| 147 | ||
| 148 | (** Remainder **) | |
| 149 | ||
| 150 | lemma mod_natify1 [simp]: "natify(m) mod n = m mod n" | |
| 151 | by (simp add: mod_def) | |
| 152 | ||
| 153 | lemma mod_natify2 [simp]: "m mod natify(n) = m mod n" | |
| 154 | by (simp add: mod_def) | |
| 155 | ||
| 156 | ||
| 157 | (** Quotient **) | |
| 158 | ||
| 159 | lemma div_natify1 [simp]: "natify(m) div n = m div n" | |
| 160 | by (simp add: div_def) | |
| 161 | ||
| 162 | lemma div_natify2 [simp]: "m div natify(n) = m div n" | |
| 163 | by (simp add: div_def) | |
| 164 | ||
| 165 | ||
| 13356 | 166 | subsection{*Typing rules*}
 | 
| 13163 | 167 | |
| 168 | (** Addition **) | |
| 169 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 170 | lemma raw_add_type: "[| m\<in>nat; n\<in>nat |] ==> raw_add (m, n) \<in> nat" | 
| 13163 | 171 | by (induct_tac "m", auto) | 
| 172 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 173 | lemma add_type [iff,TC]: "m #+ n \<in> nat" | 
| 13163 | 174 | by (simp add: add_def raw_add_type) | 
| 175 | ||
| 176 | (** Multiplication **) | |
| 177 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 178 | lemma raw_mult_type: "[| m\<in>nat; n\<in>nat |] ==> raw_mult (m, n) \<in> nat" | 
| 13163 | 179 | apply (induct_tac "m") | 
| 180 | apply (simp_all add: raw_add_type) | |
| 181 | done | |
| 182 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 183 | lemma mult_type [iff,TC]: "m #* n \<in> nat" | 
| 13163 | 184 | by (simp add: mult_def raw_mult_type) | 
| 185 | ||
| 186 | ||
| 187 | (** Difference **) | |
| 188 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 189 | lemma raw_diff_type: "[| m\<in>nat; n\<in>nat |] ==> raw_diff (m, n) \<in> nat" | 
| 13173 | 190 | by (induct_tac "n", auto) | 
| 13163 | 191 | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 192 | lemma diff_type [iff,TC]: "m #- n \<in> nat" | 
| 13163 | 193 | by (simp add: diff_def raw_diff_type) | 
| 194 | ||
| 195 | lemma diff_0_eq_0 [simp]: "0 #- n = 0" | |
| 196 | apply (unfold diff_def) | |
| 197 | apply (rule natify_in_nat [THEN nat_induct], auto) | |
| 198 | done | |
| 199 | ||
| 200 | (*Must simplify BEFORE the induction: else we get a critical pair*) | |
| 201 | lemma diff_succ_succ [simp]: "succ(m) #- succ(n) = m #- n" | |
| 202 | apply (simp add: natify_succ diff_def) | |
| 13784 | 203 | apply (rule_tac x1 = n in natify_in_nat [THEN nat_induct], auto) | 
| 13163 | 204 | done | 
| 205 | ||
| 206 | (*This defining property is no longer wanted*) | |
| 207 | declare raw_diff_succ [simp del] | |
| 208 | ||
| 209 | (*Natify has weakened this law, compared with the older approach*) | |
| 210 | lemma diff_0 [simp]: "m #- 0 = natify(m)" | |
| 211 | by (simp add: diff_def) | |
| 212 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 213 | lemma diff_le_self: "m\<in>nat ==> (m #- n) le m" | 
| 13163 | 214 | apply (subgoal_tac " (m #- natify (n)) le m") | 
| 13784 | 215 | apply (rule_tac [2] m = m and n = "natify (n) " in diff_induct) | 
| 13163 | 216 | apply (erule_tac [6] leE) | 
| 217 | apply (simp_all add: le_iff) | |
| 218 | done | |
| 219 | ||
| 220 | ||
| 13356 | 221 | subsection{*Addition*}
 | 
| 13163 | 222 | |
| 223 | (*Natify has weakened this law, compared with the older approach*) | |
| 224 | lemma add_0_natify [simp]: "0 #+ m = natify(m)" | |
| 225 | by (simp add: add_def) | |
| 226 | ||
| 227 | lemma add_succ [simp]: "succ(m) #+ n = succ(m #+ n)" | |
| 228 | by (simp add: natify_succ add_def) | |
| 229 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 230 | lemma add_0: "m \<in> nat ==> 0 #+ m = m" | 
| 13163 | 231 | by simp | 
| 232 | ||
| 233 | (*Associative law for addition*) | |
| 234 | lemma add_assoc: "(m #+ n) #+ k = m #+ (n #+ k)" | |
| 235 | apply (subgoal_tac "(natify(m) #+ natify(n)) #+ natify(k) = | |
| 236 | natify(m) #+ (natify(n) #+ natify(k))") | |
| 237 | apply (rule_tac [2] n = "natify(m)" in nat_induct) | |
| 238 | apply auto | |
| 239 | done | |
| 240 | ||
| 241 | (*The following two lemmas are used for add_commute and sometimes | |
| 242 | elsewhere, since they are safe for rewriting.*) | |
| 243 | lemma add_0_right_natify [simp]: "m #+ 0 = natify(m)" | |
| 244 | apply (subgoal_tac "natify(m) #+ 0 = natify(m)") | |
| 245 | apply (rule_tac [2] n = "natify(m)" in nat_induct) | |
| 246 | apply auto | |
| 247 | done | |
| 248 | ||
| 249 | lemma add_succ_right [simp]: "m #+ succ(n) = succ(m #+ n)" | |
| 250 | apply (unfold add_def) | |
| 251 | apply (rule_tac n = "natify(m) " in nat_induct) | |
| 252 | apply (auto simp add: natify_succ) | |
| 253 | done | |
| 254 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 255 | lemma add_0_right: "m \<in> nat ==> m #+ 0 = m" | 
| 13163 | 256 | by auto | 
| 257 | ||
| 258 | (*Commutative law for addition*) | |
| 259 | lemma add_commute: "m #+ n = n #+ m" | |
| 260 | apply (subgoal_tac "natify(m) #+ natify(n) = natify(n) #+ natify(m) ") | |
| 261 | apply (rule_tac [2] n = "natify(m) " in nat_induct) | |
| 262 | apply auto | |
| 263 | done | |
| 264 | ||
| 265 | (*for a/c rewriting*) | |
| 266 | lemma add_left_commute: "m#+(n#+k)=n#+(m#+k)" | |
| 267 | apply (rule add_commute [THEN trans]) | |
| 268 | apply (rule add_assoc [THEN trans]) | |
| 269 | apply (rule add_commute [THEN subst_context]) | |
| 270 | done | |
| 271 | ||
| 272 | (*Addition is an AC-operator*) | |
| 273 | lemmas add_ac = add_assoc add_commute add_left_commute | |
| 274 | ||
| 275 | (*Cancellation law on the left*) | |
| 276 | lemma raw_add_left_cancel: | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 277 | "[| raw_add(k, m) = raw_add(k, n); k\<in>nat |] ==> m=n" | 
| 13163 | 278 | apply (erule rev_mp) | 
| 279 | apply (induct_tac "k", auto) | |
| 280 | done | |
| 281 | ||
| 282 | lemma add_left_cancel_natify: "k #+ m = k #+ n ==> natify(m) = natify(n)" | |
| 283 | apply (unfold add_def) | |
| 284 | apply (drule raw_add_left_cancel, auto) | |
| 285 | done | |
| 286 | ||
| 287 | lemma add_left_cancel: | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 288 | "[| i = j; i #+ m = j #+ n; m\<in>nat; n\<in>nat |] ==> m = n" | 
| 13163 | 289 | by (force dest!: add_left_cancel_natify) | 
| 290 | ||
| 291 | (*Thanks to Sten Agerholm*) | |
| 292 | lemma add_le_elim1_natify: "k#+m le k#+n ==> natify(m) le natify(n)" | |
| 293 | apply (rule_tac P = "natify(k) #+m le natify(k) #+n" in rev_mp) | |
| 294 | apply (rule_tac [2] n = "natify(k) " in nat_induct) | |
| 295 | apply auto | |
| 296 | done | |
| 297 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 298 | lemma add_le_elim1: "[| k#+m le k#+n; m \<in> nat; n \<in> nat |] ==> m le n" | 
| 13163 | 299 | by (drule add_le_elim1_natify, auto) | 
| 300 | ||
| 301 | lemma add_lt_elim1_natify: "k#+m < k#+n ==> natify(m) < natify(n)" | |
| 302 | apply (rule_tac P = "natify(k) #+m < natify(k) #+n" in rev_mp) | |
| 303 | apply (rule_tac [2] n = "natify(k) " in nat_induct) | |
| 304 | apply auto | |
| 305 | done | |
| 306 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 307 | lemma add_lt_elim1: "[| k#+m < k#+n; m \<in> nat; n \<in> nat |] ==> m < n" | 
| 13163 | 308 | by (drule add_lt_elim1_natify, auto) | 
| 309 | ||
| 15201 | 310 | lemma zero_less_add: "[| n \<in> nat; m \<in> nat |] ==> 0 < m #+ n <-> (0<m | 0<n)" | 
| 311 | by (induct_tac "n", auto) | |
| 312 | ||
| 13163 | 313 | |
| 13356 | 314 | subsection{*Monotonicity of Addition*}
 | 
| 13163 | 315 | |
| 316 | (*strict, in 1st argument; proof is by rule induction on 'less than'. | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 317 | Still need j\<in>nat, for consider j = omega. Then we can have i<omega, | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 318 | which is the same as i\<in>nat, but natify(j)=0, so the conclusion fails.*) | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 319 | lemma add_lt_mono1: "[| i<j; j\<in>nat |] ==> i#+k < j#+k" | 
| 13163 | 320 | apply (frule lt_nat_in_nat, assumption) | 
| 321 | apply (erule succ_lt_induct) | |
| 322 | apply (simp_all add: leI) | |
| 323 | done | |
| 324 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 325 | text{*strict, in second argument*}
 | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 326 | lemma add_lt_mono2: "[| i<j; j\<in>nat |] ==> k#+i < k#+j" | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 327 | by (simp add: add_commute [of k] add_lt_mono1) | 
| 13163 | 328 | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 329 | text{*A [clumsy] way of lifting < monotonicity to @{text "\<le>"} monotonicity*}
 | 
| 13163 | 330 | lemma Ord_lt_mono_imp_le_mono: | 
| 331 | assumes lt_mono: "!!i j. [| i<j; j:k |] ==> f(i) < f(j)" | |
| 332 | and ford: "!!i. i:k ==> Ord(f(i))" | |
| 333 | and leij: "i le j" | |
| 334 | and jink: "j:k" | |
| 335 | shows "f(i) le f(j)" | |
| 336 | apply (insert leij jink) | |
| 337 | apply (blast intro!: leCI lt_mono ford elim!: leE) | |
| 338 | done | |
| 339 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 340 | text{*@{text "\<le>"} monotonicity, 1st argument*}
 | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 341 | lemma add_le_mono1: "[| i le j; j\<in>nat |] ==> i#+k le j#+k" | 
| 13163 | 342 | apply (rule_tac f = "%j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck) | 
| 343 | apply (blast intro: add_lt_mono1 add_type [THEN nat_into_Ord])+ | |
| 344 | done | |
| 345 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 346 | text{*@{text "\<le>"} monotonicity, both arguments*}
 | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 347 | lemma add_le_mono: "[| i le j; k le l; j\<in>nat; l\<in>nat |] ==> i#+k le j#+l" | 
| 13163 | 348 | apply (rule add_le_mono1 [THEN le_trans], assumption+) | 
| 349 | apply (subst add_commute, subst add_commute, rule add_le_mono1, assumption+) | |
| 350 | done | |
| 351 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 352 | text{*Combinations of less-than and less-than-or-equals*}
 | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 353 | |
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 354 | lemma add_lt_le_mono: "[| i<j; k\<le>l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l" | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 355 | apply (rule add_lt_mono1 [THEN lt_trans2], assumption+) | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 356 | apply (subst add_commute, subst add_commute, rule add_le_mono1, assumption+) | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 357 | done | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 358 | |
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 359 | lemma add_le_lt_mono: "[| i\<le>j; k<l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l" | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 360 | by (subst add_commute, subst add_commute, erule add_lt_le_mono, assumption+) | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 361 | |
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 362 | text{*Less-than: in other words, strict in both arguments*}
 | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 363 | lemma add_lt_mono: "[| i<j; k<l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l" | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 364 | apply (rule add_lt_le_mono) | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 365 | apply (auto intro: leI) | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 366 | done | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 367 | |
| 13163 | 368 | (** Subtraction is the inverse of addition. **) | 
| 369 | ||
| 370 | lemma diff_add_inverse: "(n#+m) #- n = natify(m)" | |
| 371 | apply (subgoal_tac " (natify(n) #+ m) #- natify(n) = natify(m) ") | |
| 372 | apply (rule_tac [2] n = "natify(n) " in nat_induct) | |
| 373 | apply auto | |
| 374 | done | |
| 375 | ||
| 376 | lemma diff_add_inverse2: "(m#+n) #- n = natify(m)" | |
| 377 | by (simp add: add_commute [of m] diff_add_inverse) | |
| 378 | ||
| 379 | lemma diff_cancel: "(k#+m) #- (k#+n) = m #- n" | |
| 380 | apply (subgoal_tac "(natify(k) #+ natify(m)) #- (natify(k) #+ natify(n)) = | |
| 381 | natify(m) #- natify(n) ") | |
| 382 | apply (rule_tac [2] n = "natify(k) " in nat_induct) | |
| 383 | apply auto | |
| 384 | done | |
| 385 | ||
| 386 | lemma diff_cancel2: "(m#+k) #- (n#+k) = m #- n" | |
| 387 | by (simp add: add_commute [of _ k] diff_cancel) | |
| 388 | ||
| 389 | lemma diff_add_0: "n #- (n#+m) = 0" | |
| 390 | apply (subgoal_tac "natify(n) #- (natify(n) #+ natify(m)) = 0") | |
| 391 | apply (rule_tac [2] n = "natify(n) " in nat_induct) | |
| 392 | apply auto | |
| 393 | done | |
| 394 | ||
| 13361 | 395 | lemma pred_0 [simp]: "pred(0) = 0" | 
| 396 | by (simp add: pred_def) | |
| 397 | ||
| 398 | lemma eq_succ_imp_eq_m1: "[|i = succ(j); i\<in>nat|] ==> j = i #- 1 & j \<in>nat" | |
| 399 | by simp | |
| 400 | ||
| 401 | lemma pred_Un_distrib: | |
| 402 | "[|i\<in>nat; j\<in>nat|] ==> pred(i Un j) = pred(i) Un pred(j)" | |
| 403 | apply (erule_tac n=i in natE, simp) | |
| 404 | apply (erule_tac n=j in natE, simp) | |
| 405 | apply (simp add: succ_Un_distrib [symmetric]) | |
| 406 | done | |
| 407 | ||
| 408 | lemma pred_type [TC,simp]: | |
| 409 | "i \<in> nat ==> pred(i) \<in> nat" | |
| 410 | by (simp add: pred_def split: split_nat_case) | |
| 411 | ||
| 412 | lemma nat_diff_pred: "[|i\<in>nat; j\<in>nat|] ==> i #- succ(j) = pred(i #- j)"; | |
| 413 | apply (rule_tac m=i and n=j in diff_induct) | |
| 414 | apply (auto simp add: pred_def nat_imp_quasinat split: split_nat_case) | |
| 415 | done | |
| 416 | ||
| 417 | lemma diff_succ_eq_pred: "i #- succ(j) = pred(i #- j)"; | |
| 418 | apply (insert nat_diff_pred [of "natify(i)" "natify(j)"]) | |
| 419 | apply (simp add: natify_succ [symmetric]) | |
| 420 | done | |
| 421 | ||
| 422 | lemma nat_diff_Un_distrib: | |
| 423 | "[|i\<in>nat; j\<in>nat; k\<in>nat|] ==> (i Un j) #- k = (i#-k) Un (j#-k)" | |
| 424 | apply (rule_tac n=k in nat_induct) | |
| 425 | apply (simp_all add: diff_succ_eq_pred pred_Un_distrib) | |
| 426 | done | |
| 427 | ||
| 428 | lemma diff_Un_distrib: | |
| 429 | "[|i\<in>nat; j\<in>nat|] ==> (i Un j) #- k = (i#-k) Un (j#-k)" | |
| 430 | by (insert nat_diff_Un_distrib [of i j "natify(k)"], simp) | |
| 431 | ||
| 432 | text{*We actually prove @{term "i #- j #- k = i #- (j #+ k)"}*}
 | |
| 433 | lemma diff_diff_left [simplified]: | |
| 434 | "natify(i)#-natify(j)#-k = natify(i) #- (natify(j)#+k)"; | |
| 435 | by (rule_tac m="natify(i)" and n="natify(j)" in diff_induct, auto) | |
| 436 | ||
| 13163 | 437 | |
| 438 | (** Lemmas for the CancelNumerals simproc **) | |
| 439 | ||
| 440 | lemma eq_add_iff: "(u #+ m = u #+ n) <-> (0 #+ m = natify(n))" | |
| 441 | apply auto | |
| 442 | apply (blast dest: add_left_cancel_natify) | |
| 443 | apply (simp add: add_def) | |
| 444 | done | |
| 445 | ||
| 446 | lemma less_add_iff: "(u #+ m < u #+ n) <-> (0 #+ m < natify(n))" | |
| 447 | apply (auto simp add: add_lt_elim1_natify) | |
| 448 | apply (drule add_lt_mono1) | |
| 449 | apply (auto simp add: add_commute [of u]) | |
| 450 | done | |
| 451 | ||
| 452 | lemma diff_add_eq: "((u #+ m) #- (u #+ n)) = ((0 #+ m) #- n)" | |
| 453 | by (simp add: diff_cancel) | |
| 454 | ||
| 455 | (*To tidy up the result of a simproc. Only the RHS will be simplified.*) | |
| 456 | lemma eq_cong2: "u = u' ==> (t==u) == (t==u')" | |
| 457 | by auto | |
| 458 | ||
| 459 | lemma iff_cong2: "u <-> u' ==> (t==u) == (t==u')" | |
| 460 | by auto | |
| 461 | ||
| 462 | ||
| 13356 | 463 | subsection{*Multiplication*}
 | 
| 13163 | 464 | |
| 465 | lemma mult_0 [simp]: "0 #* m = 0" | |
| 466 | by (simp add: mult_def) | |
| 467 | ||
| 468 | lemma mult_succ [simp]: "succ(m) #* n = n #+ (m #* n)" | |
| 469 | by (simp add: add_def mult_def natify_succ raw_mult_type) | |
| 470 | ||
| 471 | (*right annihilation in product*) | |
| 472 | lemma mult_0_right [simp]: "m #* 0 = 0" | |
| 473 | apply (unfold mult_def) | |
| 474 | apply (rule_tac n = "natify(m) " in nat_induct) | |
| 475 | apply auto | |
| 476 | done | |
| 477 | ||
| 478 | (*right successor law for multiplication*) | |
| 479 | lemma mult_succ_right [simp]: "m #* succ(n) = m #+ (m #* n)" | |
| 480 | apply (subgoal_tac "natify(m) #* succ (natify(n)) = | |
| 481 | natify(m) #+ (natify(m) #* natify(n))") | |
| 482 | apply (simp (no_asm_use) add: natify_succ add_def mult_def) | |
| 483 | apply (rule_tac n = "natify(m) " in nat_induct) | |
| 484 | apply (simp_all add: add_ac) | |
| 485 | done | |
| 486 | ||
| 487 | lemma mult_1_natify [simp]: "1 #* n = natify(n)" | |
| 488 | by auto | |
| 489 | ||
| 490 | lemma mult_1_right_natify [simp]: "n #* 1 = natify(n)" | |
| 491 | by auto | |
| 492 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 493 | lemma mult_1: "n \<in> nat ==> 1 #* n = n" | 
| 13163 | 494 | by simp | 
| 495 | ||
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 496 | lemma mult_1_right: "n \<in> nat ==> n #* 1 = n" | 
| 13163 | 497 | by simp | 
| 498 | ||
| 499 | (*Commutative law for multiplication*) | |
| 500 | lemma mult_commute: "m #* n = n #* m" | |
| 501 | apply (subgoal_tac "natify(m) #* natify(n) = natify(n) #* natify(m) ") | |
| 502 | apply (rule_tac [2] n = "natify(m) " in nat_induct) | |
| 503 | apply auto | |
| 504 | done | |
| 505 | ||
| 506 | (*addition distributes over multiplication*) | |
| 507 | lemma add_mult_distrib: "(m #+ n) #* k = (m #* k) #+ (n #* k)" | |
| 508 | apply (subgoal_tac "(natify(m) #+ natify(n)) #* natify(k) = | |
| 509 | (natify(m) #* natify(k)) #+ (natify(n) #* natify(k))") | |
| 510 | apply (rule_tac [2] n = "natify(m) " in nat_induct) | |
| 511 | apply (simp_all add: add_assoc [symmetric]) | |
| 512 | done | |
| 513 | ||
| 514 | (*Distributive law on the left*) | |
| 515 | lemma add_mult_distrib_left: "k #* (m #+ n) = (k #* m) #+ (k #* n)" | |
| 516 | apply (subgoal_tac "natify(k) #* (natify(m) #+ natify(n)) = | |
| 517 | (natify(k) #* natify(m)) #+ (natify(k) #* natify(n))") | |
| 518 | apply (rule_tac [2] n = "natify(m) " in nat_induct) | |
| 519 | apply (simp_all add: add_ac) | |
| 520 | done | |
| 521 | ||
| 522 | (*Associative law for multiplication*) | |
| 523 | lemma mult_assoc: "(m #* n) #* k = m #* (n #* k)" | |
| 524 | apply (subgoal_tac "(natify(m) #* natify(n)) #* natify(k) = | |
| 525 | natify(m) #* (natify(n) #* natify(k))") | |
| 526 | apply (rule_tac [2] n = "natify(m) " in nat_induct) | |
| 527 | apply (simp_all add: add_mult_distrib) | |
| 528 | done | |
| 529 | ||
| 530 | (*for a/c rewriting*) | |
| 531 | lemma mult_left_commute: "m #* (n #* k) = n #* (m #* k)" | |
| 532 | apply (rule mult_commute [THEN trans]) | |
| 533 | apply (rule mult_assoc [THEN trans]) | |
| 534 | apply (rule mult_commute [THEN subst_context]) | |
| 535 | done | |
| 536 | ||
| 537 | lemmas mult_ac = mult_assoc mult_commute mult_left_commute | |
| 538 | ||
| 539 | ||
| 540 | lemma lt_succ_eq_0_disj: | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 541 | "[| m\<in>nat; n\<in>nat |] | 
| 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 542 | ==> (m < succ(n)) <-> (m = 0 | (\<exists>j\<in>nat. m = succ(j) & j < n))" | 
| 13163 | 543 | by (induct_tac "m", auto) | 
| 544 | ||
| 545 | lemma less_diff_conv [rule_format]: | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 546 | "[| j\<in>nat; k\<in>nat |] ==> \<forall>i\<in>nat. (i < j #- k) <-> (i #+ k < j)" | 
| 13784 | 547 | by (erule_tac m = k in diff_induct, auto) | 
| 13163 | 548 | |
| 549 | lemmas nat_typechecks = rec_type nat_0I nat_1I nat_succI Ord_nat | |
| 550 | ||
| 551 | ML | |
| 552 | {*
 | |
| 553 | val pred_def = thm "pred_def"; | |
| 554 | val raw_div_def = thm "raw_div_def"; | |
| 555 | val raw_mod_def = thm "raw_mod_def"; | |
| 556 | val div_def = thm "div_def"; | |
| 557 | val mod_def = thm "mod_def"; | |
| 558 | ||
| 559 | val zero_lt_natE = thm "zero_lt_natE"; | |
| 560 | val pred_succ_eq = thm "pred_succ_eq"; | |
| 561 | val natify_succ = thm "natify_succ"; | |
| 562 | val natify_0 = thm "natify_0"; | |
| 563 | val natify_non_succ = thm "natify_non_succ"; | |
| 564 | val natify_in_nat = thm "natify_in_nat"; | |
| 565 | val natify_ident = thm "natify_ident"; | |
| 566 | val natify_eqE = thm "natify_eqE"; | |
| 567 | val natify_idem = thm "natify_idem"; | |
| 568 | val add_natify1 = thm "add_natify1"; | |
| 569 | val add_natify2 = thm "add_natify2"; | |
| 570 | val mult_natify1 = thm "mult_natify1"; | |
| 571 | val mult_natify2 = thm "mult_natify2"; | |
| 572 | val diff_natify1 = thm "diff_natify1"; | |
| 573 | val diff_natify2 = thm "diff_natify2"; | |
| 574 | val mod_natify1 = thm "mod_natify1"; | |
| 575 | val mod_natify2 = thm "mod_natify2"; | |
| 576 | val div_natify1 = thm "div_natify1"; | |
| 577 | val div_natify2 = thm "div_natify2"; | |
| 578 | val raw_add_type = thm "raw_add_type"; | |
| 579 | val add_type = thm "add_type"; | |
| 580 | val raw_mult_type = thm "raw_mult_type"; | |
| 581 | val mult_type = thm "mult_type"; | |
| 582 | val raw_diff_type = thm "raw_diff_type"; | |
| 583 | val diff_type = thm "diff_type"; | |
| 584 | val diff_0_eq_0 = thm "diff_0_eq_0"; | |
| 585 | val diff_succ_succ = thm "diff_succ_succ"; | |
| 586 | val diff_0 = thm "diff_0"; | |
| 587 | val diff_le_self = thm "diff_le_self"; | |
| 588 | val add_0_natify = thm "add_0_natify"; | |
| 589 | val add_succ = thm "add_succ"; | |
| 590 | val add_0 = thm "add_0"; | |
| 591 | val add_assoc = thm "add_assoc"; | |
| 592 | val add_0_right_natify = thm "add_0_right_natify"; | |
| 593 | val add_succ_right = thm "add_succ_right"; | |
| 594 | val add_0_right = thm "add_0_right"; | |
| 595 | val add_commute = thm "add_commute"; | |
| 596 | val add_left_commute = thm "add_left_commute"; | |
| 597 | val raw_add_left_cancel = thm "raw_add_left_cancel"; | |
| 598 | val add_left_cancel_natify = thm "add_left_cancel_natify"; | |
| 599 | val add_left_cancel = thm "add_left_cancel"; | |
| 600 | val add_le_elim1_natify = thm "add_le_elim1_natify"; | |
| 601 | val add_le_elim1 = thm "add_le_elim1"; | |
| 602 | val add_lt_elim1_natify = thm "add_lt_elim1_natify"; | |
| 603 | val add_lt_elim1 = thm "add_lt_elim1"; | |
| 604 | val add_lt_mono1 = thm "add_lt_mono1"; | |
| 14060 
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
 paulson parents: 
13784diff
changeset | 605 | val add_lt_mono2 = thm "add_lt_mono2"; | 
| 13163 | 606 | val add_lt_mono = thm "add_lt_mono"; | 
| 607 | val Ord_lt_mono_imp_le_mono = thm "Ord_lt_mono_imp_le_mono"; | |
| 608 | val add_le_mono1 = thm "add_le_mono1"; | |
| 609 | val add_le_mono = thm "add_le_mono"; | |
| 610 | val diff_add_inverse = thm "diff_add_inverse"; | |
| 611 | val diff_add_inverse2 = thm "diff_add_inverse2"; | |
| 612 | val diff_cancel = thm "diff_cancel"; | |
| 613 | val diff_cancel2 = thm "diff_cancel2"; | |
| 614 | val diff_add_0 = thm "diff_add_0"; | |
| 615 | val eq_add_iff = thm "eq_add_iff"; | |
| 616 | val less_add_iff = thm "less_add_iff"; | |
| 617 | val diff_add_eq = thm "diff_add_eq"; | |
| 618 | val eq_cong2 = thm "eq_cong2"; | |
| 619 | val iff_cong2 = thm "iff_cong2"; | |
| 620 | val mult_0 = thm "mult_0"; | |
| 621 | val mult_succ = thm "mult_succ"; | |
| 622 | val mult_0_right = thm "mult_0_right"; | |
| 623 | val mult_succ_right = thm "mult_succ_right"; | |
| 624 | val mult_1_natify = thm "mult_1_natify"; | |
| 625 | val mult_1_right_natify = thm "mult_1_right_natify"; | |
| 626 | val mult_1 = thm "mult_1"; | |
| 627 | val mult_1_right = thm "mult_1_right"; | |
| 628 | val mult_commute = thm "mult_commute"; | |
| 629 | val add_mult_distrib = thm "add_mult_distrib"; | |
| 630 | val add_mult_distrib_left = thm "add_mult_distrib_left"; | |
| 631 | val mult_assoc = thm "mult_assoc"; | |
| 632 | val mult_left_commute = thm "mult_left_commute"; | |
| 633 | val lt_succ_eq_0_disj = thm "lt_succ_eq_0_disj"; | |
| 634 | val less_diff_conv = thm "less_diff_conv"; | |
| 635 | ||
| 636 | val add_ac = thms "add_ac"; | |
| 637 | val mult_ac = thms "mult_ac"; | |
| 638 | val nat_typechecks = thms "nat_typechecks"; | |
| 639 | *} | |
| 9654 
9754ba005b64
X-symbols for ordinal, cardinal, integer arithmetic
 paulson parents: 
9492diff
changeset | 640 | |
| 0 | 641 | end |