| author | wenzelm | 
| Tue, 18 Dec 2007 22:21:41 +0100 | |
| changeset 25703 | 832073e402ae | 
| parent 25615 | b337edd55a07 | 
| child 25767 | 852bce03412a | 
| permissions | -rw-r--r-- | 
| 23854 | 1 | (* Title: HOL/Library/Efficient_Nat.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Stefan Berghofer, TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Implementation of natural numbers by integers *}
 | |
| 7 | ||
| 8 | theory Efficient_Nat | |
| 24994 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24715diff
changeset | 9 | imports Main Code_Integer | 
| 23854 | 10 | begin | 
| 11 | ||
| 12 | text {*
 | |
| 13 | When generating code for functions on natural numbers, the canonical | |
| 14 | representation using @{term "0::nat"} and @{term "Suc"} is unsuitable for
 | |
| 15 | computations involving large numbers. The efficiency of the generated | |
| 16 | code can be improved drastically by implementing natural numbers by | |
| 17 | integers. To do this, just include this theory. | |
| 18 | *} | |
| 19 | ||
| 20 | subsection {* Logical rewrites *}
 | |
| 21 | ||
| 22 | text {*
 | |
| 23 | An int-to-nat conversion | |
| 24 |   restricted to non-negative ints (in contrast to @{const nat}).
 | |
| 25 | Note that this restriction has no logical relevance and | |
| 26 | is just a kind of proof hint -- nothing prevents you from | |
| 27 |   writing nonsense like @{term "nat_of_int (-4)"}
 | |
| 28 | *} | |
| 29 | ||
| 30 | definition | |
| 31 | nat_of_int :: "int \<Rightarrow> nat" where | |
| 32 | "k \<ge> 0 \<Longrightarrow> nat_of_int k = nat k" | |
| 33 | ||
| 34 | definition | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 35 | int_of_nat :: "nat \<Rightarrow> int" where | 
| 25488 | 36 | [code func del]: "int_of_nat n = of_nat n" | 
| 23854 | 37 | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 38 | lemma int_of_nat_Suc [simp]: | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 39 | "int_of_nat (Suc n) = 1 + int_of_nat n" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 40 | unfolding int_of_nat_def by simp | 
| 23854 | 41 | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 42 | lemma int_of_nat_add: | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 43 | "int_of_nat (m + n) = int_of_nat m + int_of_nat n" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 44 | unfolding int_of_nat_def by (rule of_nat_add) | 
| 23854 | 45 | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 46 | lemma int_of_nat_mult: | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 47 | "int_of_nat (m * n) = int_of_nat m * int_of_nat n" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 48 | unfolding int_of_nat_def by (rule of_nat_mult) | 
| 23854 | 49 | |
| 50 | lemma nat_of_int_of_number_of: | |
| 51 | fixes k | |
| 52 | assumes "k \<ge> 0" | |
| 53 | shows "number_of k = nat_of_int (number_of k)" | |
| 54 | unfolding nat_of_int_def [OF assms] nat_number_of_def number_of_is_id .. | |
| 55 | ||
| 56 | lemma nat_of_int_of_number_of_aux: | |
| 57 | fixes k | |
| 58 | assumes "Numeral.Pls \<le> k \<equiv> True" | |
| 59 | shows "k \<ge> 0" | |
| 60 | using assms unfolding Pls_def by simp | |
| 61 | ||
| 62 | lemma nat_of_int_int: | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 63 | "nat_of_int (int_of_nat n) = n" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 64 | using nat_of_int_def int_of_nat_def by simp | 
| 23854 | 65 | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 66 | lemma eq_nat_of_int: "int_of_nat n = x \<Longrightarrow> n = nat_of_int x" | 
| 23854 | 67 | by (erule subst, simp only: nat_of_int_int) | 
| 68 | ||
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24222diff
changeset | 69 | code_datatype nat_of_int | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24222diff
changeset | 70 | |
| 23854 | 71 | text {*
 | 
| 72 | Case analysis on natural numbers is rephrased using a conditional | |
| 73 | expression: | |
| 74 | *} | |
| 75 | ||
| 76 | lemma [code unfold, code inline del]: | |
| 25615 | 77 | "nat_case = (\<lambda>f g n. if n = 0 then f else g (n - 1))" | 
| 23854 | 78 | proof - | 
| 79 | have rewrite: "\<And>f g n. nat_case f g n = (if n = 0 then f else g (n - 1))" | |
| 80 | proof - | |
| 81 | fix f g n | |
| 82 | show "nat_case f g n = (if n = 0 then f else g (n - 1))" | |
| 83 | by (cases n) simp_all | |
| 84 | qed | |
| 25615 | 85 | show "nat_case = (\<lambda>f g n. if n = 0 then f else g (n - 1))" | 
| 23854 | 86 | by (rule eq_reflection ext rewrite)+ | 
| 87 | qed | |
| 88 | ||
| 89 | lemma [code inline]: | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 90 | "nat_case = (\<lambda>f g n. if n = 0 then f else g (nat_of_int (int_of_nat n - 1)))" | 
| 23854 | 91 | proof (rule ext)+ | 
| 92 | fix f g n | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 93 | show "nat_case f g n = (if n = 0 then f else g (nat_of_int (int_of_nat n - 1)))" | 
| 23854 | 94 | by (cases n) (simp_all add: nat_of_int_int) | 
| 95 | qed | |
| 96 | ||
| 97 | text {*
 | |
| 98 | Most standard arithmetic functions on natural numbers are implemented | |
| 99 | using their counterparts on the integers: | |
| 100 | *} | |
| 101 | ||
| 102 | lemma [code func]: "0 = nat_of_int 0" | |
| 103 | by (simp add: nat_of_int_def) | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 104 | |
| 23854 | 105 | lemma [code func, code inline]: "1 = nat_of_int 1" | 
| 106 | by (simp add: nat_of_int_def) | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 107 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 108 | lemma [code func]: "Suc n = nat_of_int (int_of_nat n + 1)" | 
| 23854 | 109 | by (simp add: eq_nat_of_int) | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 110 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 111 | lemma [code]: "m + n = nat (int_of_nat m + int_of_nat n)" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 112 | by (simp add: int_of_nat_def nat_eq_iff2) | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 113 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 114 | lemma [code func, code inline]: "m + n = nat_of_int (int_of_nat m + int_of_nat n)" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 115 | by (simp add: eq_nat_of_int int_of_nat_add) | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 116 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 117 | lemma [code, code inline]: "m - n = nat (int_of_nat m - int_of_nat n)" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 118 | by (simp add: int_of_nat_def nat_eq_iff2 of_nat_diff) | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 119 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 120 | lemma [code]: "m * n = nat (int_of_nat m * int_of_nat n)" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 121 | unfolding int_of_nat_def | 
| 23854 | 122 | by (simp add: of_nat_mult [symmetric] del: of_nat_mult) | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 123 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 124 | lemma [code func, code inline]: "m * n = nat_of_int (int_of_nat m * int_of_nat n)" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 125 | by (simp add: eq_nat_of_int int_of_nat_mult) | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 126 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 127 | lemma [code]: "m div n = nat (int_of_nat m div int_of_nat n)" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 128 | unfolding int_of_nat_def zdiv_int [symmetric] by simp | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 129 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 130 | lemma div_nat_code [code func]: | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 131 | "m div k = nat_of_int (fst (divAlg (int_of_nat m, int_of_nat k)))" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 132 | unfolding div_def [symmetric] int_of_nat_def zdiv_int [symmetric] | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 133 | unfolding int_of_nat_def [symmetric] nat_of_int_int .. | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 134 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 135 | lemma [code]: "m mod n = nat (int_of_nat m mod int_of_nat n)" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 136 | unfolding int_of_nat_def zmod_int [symmetric] by simp | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 137 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 138 | lemma mod_nat_code [code func]: | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 139 | "m mod k = nat_of_int (snd (divAlg (int_of_nat m, int_of_nat k)))" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 140 | unfolding mod_def [symmetric] int_of_nat_def zmod_int [symmetric] | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 141 | unfolding int_of_nat_def [symmetric] nat_of_int_int .. | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 142 | |
| 25615 | 143 | lemma [code, code inline]: "m < n \<longleftrightarrow> int_of_nat m < int_of_nat n" | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 144 | unfolding int_of_nat_def by simp | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 145 | |
| 25615 | 146 | lemma [code func, code inline]: "m \<le> n \<longleftrightarrow> int_of_nat m \<le> int_of_nat n" | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 147 | unfolding int_of_nat_def by simp | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 148 | |
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 149 | lemma [code func, code inline]: "m = n \<longleftrightarrow> int_of_nat m = int_of_nat n" | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 150 | unfolding int_of_nat_def by simp | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 151 | |
| 23854 | 152 | lemma [code func]: "nat k = (if k < 0 then 0 else nat_of_int k)" | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 153 | by (cases "k < 0") (simp, simp add: nat_of_int_def) | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 154 | |
| 23854 | 155 | lemma [code func]: | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 156 | "int_aux n i = (if int_of_nat n = 0 then i else int_aux (nat_of_int (int_of_nat n - 1)) (i + 1))" | 
| 23854 | 157 | proof - | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 158 | have "0 < n \<Longrightarrow> int_of_nat n = 1 + int_of_nat (nat_of_int (int_of_nat n - 1))" | 
| 23854 | 159 | proof - | 
| 160 | assume prem: "n > 0" | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 161 | then have "int_of_nat n - 1 \<ge> 0" unfolding int_of_nat_def by auto | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 162 | then have "nat_of_int (int_of_nat n - 1) = nat (int_of_nat n - 1)" by (simp add: nat_of_int_def) | 
| 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 163 | with prem show "int_of_nat n = 1 + int_of_nat (nat_of_int (int_of_nat n - 1))" unfolding int_of_nat_def by simp | 
| 23854 | 164 | qed | 
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 165 | then show ?thesis unfolding int_aux_def int_of_nat_def by auto | 
| 23854 | 166 | qed | 
| 167 | ||
| 24994 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24715diff
changeset | 168 | lemma index_of_nat_code [code func, code inline]: | 
| 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24715diff
changeset | 169 | "index_of_nat n = index_of_int (int_of_nat n)" | 
| 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24715diff
changeset | 170 | unfolding index_of_nat_def int_of_nat_def by simp | 
| 23854 | 171 | |
| 24994 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24715diff
changeset | 172 | lemma nat_of_index_code [code func, code inline]: | 
| 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24715diff
changeset | 173 | "nat_of_index k = nat (int_of_index k)" | 
| 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24715diff
changeset | 174 | unfolding nat_of_index_def by simp | 
| 23854 | 175 | |
| 176 | ||
| 177 | subsection {* Code generator setup for basic functions *}
 | |
| 178 | ||
| 179 | text {*
 | |
| 180 |   @{typ nat} is no longer a datatype but embedded into the integers.
 | |
| 181 | *} | |
| 182 | ||
| 183 | code_type nat | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24423diff
changeset | 184 | (SML "int") | 
| 23854 | 185 | (OCaml "Big'_int.big'_int") | 
| 186 | (Haskell "Integer") | |
| 187 | ||
| 188 | types_code | |
| 189 |   nat ("int")
 | |
| 190 | attach (term_of) {*
 | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24423diff
changeset | 191 | val term_of_nat = HOLogic.mk_number HOLogic.natT; | 
| 23854 | 192 | *} | 
| 193 | attach (test) {*
 | |
| 194 | fun gen_nat i = random_range 0 i; | |
| 195 | *} | |
| 196 | ||
| 197 | consts_code | |
| 198 |   "0 \<Colon> nat" ("0")
 | |
| 199 |   Suc ("(_ + 1)")
 | |
| 200 | ||
| 201 | text {*
 | |
| 202 | Since natural numbers are implemented | |
| 203 |   using integers, the coercion function @{const "int"} of type
 | |
| 204 |   @{typ "nat \<Rightarrow> int"} is simply implemented by the identity function,
 | |
| 205 |   likewise @{const nat_of_int} of type @{typ "int \<Rightarrow> nat"}.
 | |
| 206 |   For the @{const "nat"} function for converting an integer to a natural
 | |
| 207 | number, we give a specific implementation using an ML function that | |
| 208 | returns its input value, provided that it is non-negative, and otherwise | |
| 209 |   returns @{text "0"}.
 | |
| 210 | *} | |
| 211 | ||
| 212 | consts_code | |
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 213 |   int_of_nat ("(_)")
 | 
| 23854 | 214 |   nat ("\<module>nat")
 | 
| 215 | attach {*
 | |
| 216 | fun nat i = if i < 0 then 0 else i; | |
| 217 | *} | |
| 218 | ||
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 219 | code_const int_of_nat | 
| 23854 | 220 | (SML "_") | 
| 221 | (OCaml "_") | |
| 222 | (Haskell "_") | |
| 223 | ||
| 224 | code_const nat_of_int | |
| 225 | (SML "_") | |
| 226 | (OCaml "_") | |
| 227 | (Haskell "_") | |
| 228 | ||
| 25615 | 229 | text {* Using target language div and mod *}
 | 
| 230 | ||
| 231 | code_const "op div \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat" | |
| 232 | (SML "IntInf.div ((_), (_))") | |
| 233 | (OCaml "Big'_int.div'_big'_int") | |
| 234 | (Haskell "div") | |
| 235 | ||
| 236 | code_const "op mod \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat" | |
| 237 | (SML "IntInf.mod ((_), (_))") | |
| 238 | (OCaml "Big'_int.mod'_big'_int") | |
| 239 | (Haskell "mod") | |
| 240 | ||
| 23854 | 241 | |
| 242 | subsection {* Preprocessors *}
 | |
| 243 | ||
| 244 | text {*
 | |
| 245 |   Natural numerals should be expressed using @{const nat_of_int}.
 | |
| 246 | *} | |
| 247 | ||
| 248 | lemmas [code inline del] = nat_number_of_def | |
| 249 | ||
| 250 | ML {*
 | |
| 251 | fun nat_of_int_of_number_of thy cts = | |
| 252 | let | |
| 253 | val simplify_less = Simplifier.rewrite | |
| 254 |       (HOL_basic_ss addsimps (@{thms less_numeral_code} @ @{thms less_eq_numeral_code}));
 | |
| 255 | fun mk_rew (t, ty) = | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24423diff
changeset | 256 | if ty = HOLogic.natT andalso 0 <= HOLogic.dest_numeral t then | 
| 23854 | 257 |         Thm.capply @{cterm "(op \<le>) Numeral.Pls"} (Thm.cterm_of thy t)
 | 
| 258 | |> simplify_less | |
| 259 |         |> (fn thm => @{thm nat_of_int_of_number_of_aux} OF [thm])
 | |
| 260 |         |> (fn thm => @{thm nat_of_int_of_number_of} OF [thm])
 | |
| 261 |         |> (fn thm => @{thm eq_reflection} OF [thm])
 | |
| 262 | |> SOME | |
| 263 | else NONE | |
| 264 | in | |
| 265 | fold (HOLogic.add_numerals o Thm.term_of) cts [] | |
| 266 | |> map_filter mk_rew | |
| 267 | end; | |
| 268 | *} | |
| 269 | ||
| 270 | setup {*
 | |
| 24222 | 271 |   Code.add_inline_proc ("nat_of_int_of_number_of", nat_of_int_of_number_of)
 | 
| 23854 | 272 | *} | 
| 273 | ||
| 274 | text {*
 | |
| 275 |   In contrast to @{term "Suc n"}, the term @{term "n + (1::nat)"} is no longer
 | |
| 276 | a constructor term. Therefore, all occurrences of this term in a position | |
| 277 | where a pattern is expected (i.e.\ on the left-hand side of a recursion | |
| 278 | equation or in the arguments of an inductive relation in an introduction | |
| 279 | rule) must be eliminated. | |
| 280 | This can be accomplished by applying the following transformation rules: | |
| 281 | *} | |
| 282 | ||
| 283 | theorem Suc_if_eq: "(\<And>n. f (Suc n) = h n) \<Longrightarrow> f 0 = g \<Longrightarrow> | |
| 284 | f n = (if n = 0 then g else h (n - 1))" | |
| 285 | by (case_tac n) simp_all | |
| 286 | ||
| 287 | theorem Suc_clause: "(\<And>n. P n (Suc n)) \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> P (n - 1) n" | |
| 288 | by (case_tac n) simp_all | |
| 289 | ||
| 290 | text {*
 | |
| 291 | The rules above are built into a preprocessor that is plugged into | |
| 292 | the code generator. Since the preprocessor for introduction rules | |
| 293 | does not know anything about modes, some of the modes that worked | |
| 294 | for the canonical representation of natural numbers may no longer work. | |
| 295 | *} | |
| 296 | ||
| 297 | (*<*) | |
| 298 | ||
| 299 | ML {*
 | |
| 300 | fun remove_suc thy thms = | |
| 301 | let | |
| 302 | val vname = Name.variant (map fst | |
| 303 | (fold (Term.add_varnames o Thm.full_prop_of) thms [])) "x"; | |
| 304 | val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT)); | |
| 305 | fun lhs_of th = snd (Thm.dest_comb | |
| 306 | (fst (Thm.dest_comb (snd (Thm.dest_comb (cprop_of th)))))); | |
| 307 | fun rhs_of th = snd (Thm.dest_comb (snd (Thm.dest_comb (cprop_of th)))); | |
| 308 | fun find_vars ct = (case term_of ct of | |
| 309 |         (Const ("Suc", _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
 | |
| 310 | | _ $ _ => | |
| 311 | let val (ct1, ct2) = Thm.dest_comb ct | |
| 312 | in | |
| 313 | map (apfst (fn ct => Thm.capply ct ct2)) (find_vars ct1) @ | |
| 314 | map (apfst (Thm.capply ct1)) (find_vars ct2) | |
| 315 | end | |
| 316 | | _ => []); | |
| 317 | val eqs = maps | |
| 318 | (fn th => map (pair th) (find_vars (lhs_of th))) thms; | |
| 319 | fun mk_thms (th, (ct, cv')) = | |
| 320 | let | |
| 321 | val th' = | |
| 322 | Thm.implies_elim | |
| 323 | (Conv.fconv_rule (Thm.beta_conversion true) | |
| 324 | (Drule.instantiate' | |
| 325 | [SOME (ctyp_of_term ct)] [SOME (Thm.cabs cv ct), | |
| 326 | SOME (Thm.cabs cv' (rhs_of th)), NONE, SOME cv'] | |
| 24222 | 327 |                @{thm Suc_if_eq})) (Thm.forall_intr cv' th)
 | 
| 23854 | 328 | in | 
| 329 | case map_filter (fn th'' => | |
| 330 | SOME (th'', singleton | |
| 331 | (Variable.trade (K (fn [th'''] => [th''' RS th'])) (Variable.thm_context th'')) th'') | |
| 332 | handle THM _ => NONE) thms of | |
| 333 | [] => NONE | |
| 334 | | thps => | |
| 335 | let val (ths1, ths2) = split_list thps | |
| 336 | in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end | |
| 337 | end | |
| 338 | in | |
| 339 | case get_first mk_thms eqs of | |
| 340 | NONE => thms | |
| 341 | | SOME x => remove_suc thy x | |
| 342 | end; | |
| 343 | ||
| 344 | fun eqn_suc_preproc thy ths = | |
| 345 | let | |
| 24222 | 346 | val dest = fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of; | 
| 347 |     fun contains_suc t = member (op =) (term_consts t) @{const_name Suc};
 | |
| 23854 | 348 | in | 
| 349 | if forall (can dest) ths andalso | |
| 350 | exists (contains_suc o dest) ths | |
| 351 | then remove_suc thy ths else ths | |
| 352 | end; | |
| 353 | ||
| 354 | fun remove_suc_clause thy thms = | |
| 355 | let | |
| 356 | val vname = Name.variant (map fst | |
| 357 | (fold (Term.add_varnames o Thm.full_prop_of) thms [])) "x"; | |
| 24222 | 358 |     fun find_var (t as Const (@{const_name Suc}, _) $ (v as Var _)) = SOME (t, v)
 | 
| 23854 | 359 | | find_var (t $ u) = (case find_var t of NONE => find_var u | x => x) | 
| 360 | | find_var _ = NONE; | |
| 361 | fun find_thm th = | |
| 362 | let val th' = Conv.fconv_rule ObjectLogic.atomize th | |
| 363 | in Option.map (pair (th, th')) (find_var (prop_of th')) end | |
| 364 | in | |
| 365 | case get_first find_thm thms of | |
| 366 | NONE => thms | |
| 367 | | SOME ((th, th'), (Sucv, v)) => | |
| 368 | let | |
| 369 | val cert = cterm_of (Thm.theory_of_thm th); | |
| 370 | val th'' = ObjectLogic.rulify (Thm.implies_elim | |
| 371 | (Conv.fconv_rule (Thm.beta_conversion true) | |
| 372 | (Drule.instantiate' [] | |
| 373 |                 [SOME (cert (lambda v (Abs ("x", HOLogic.natT,
 | |
| 374 | abstract_over (Sucv, | |
| 375 | HOLogic.dest_Trueprop (prop_of th')))))), | |
| 24222 | 376 |                  SOME (cert v)] @{thm Suc_clause}))
 | 
| 23854 | 377 | (Thm.forall_intr (cert v) th')) | 
| 378 | in | |
| 379 | remove_suc_clause thy (map (fn th''' => | |
| 380 | if (op = o pairself prop_of) (th''', th) then th'' else th''') thms) | |
| 381 | end | |
| 382 | end; | |
| 383 | ||
| 384 | fun clause_suc_preproc thy ths = | |
| 385 | let | |
| 386 | val dest = fst o HOLogic.dest_mem o HOLogic.dest_Trueprop | |
| 387 | in | |
| 388 | if forall (can (dest o concl_of)) ths andalso | |
| 389 | exists (fn th => member (op =) (foldr add_term_consts | |
| 390 | [] (map_filter (try dest) (concl_of th :: prems_of th))) "Suc") ths | |
| 391 | then remove_suc_clause thy ths else ths | |
| 392 | end; | |
| 393 | ||
| 394 | fun lift_obj_eq f thy = | |
| 395 |   map (fn thm => thm RS @{thm meta_eq_to_obj_eq})
 | |
| 396 | #> f thy | |
| 397 |   #> map (fn thm => thm RS @{thm eq_reflection})
 | |
| 398 | #> map (Conv.fconv_rule Drule.beta_eta_conversion) | |
| 399 | *} | |
| 400 | ||
| 401 | setup {*
 | |
| 402 | Codegen.add_preprocessor eqn_suc_preproc | |
| 403 | #> Codegen.add_preprocessor clause_suc_preproc | |
| 24222 | 404 |   #> Code.add_preproc ("eqn_Suc", lift_obj_eq eqn_suc_preproc)
 | 
| 405 |   #> Code.add_preproc ("clause_Suc", lift_obj_eq clause_suc_preproc)
 | |
| 23854 | 406 | *} | 
| 407 | (*>*) | |
| 408 | ||
| 409 | ||
| 410 | subsection {* Module names *}
 | |
| 411 | ||
| 412 | code_modulename SML | |
| 413 | Nat Integer | |
| 414 | Divides Integer | |
| 415 | Efficient_Nat Integer | |
| 416 | ||
| 417 | code_modulename OCaml | |
| 418 | Nat Integer | |
| 419 | Divides Integer | |
| 420 | Efficient_Nat Integer | |
| 421 | ||
| 422 | code_modulename Haskell | |
| 423 | Nat Integer | |
| 24195 | 424 | Divides Integer | 
| 23854 | 425 | Efficient_Nat Integer | 
| 426 | ||
| 24715 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 haftmann parents: 
24630diff
changeset | 427 | hide const nat_of_int int_of_nat | 
| 23854 | 428 | |
| 429 | end |