| author | huffman | 
| Wed, 23 Nov 2011 07:00:01 +0100 | |
| changeset 45615 | c05e8209a3aa | 
| parent 45185 | 3a0c63c0ed48 | 
| child 45793 | 331ebffe0593 | 
| permissions | -rw-r--r-- | 
| 23854 | 1  | 
(* Title: HOL/Library/Efficient_Nat.thy  | 
| 25931 | 2  | 
Author: Stefan Berghofer, Florian Haftmann, TU Muenchen  | 
| 23854 | 3  | 
*)  | 
4  | 
||
| 25931 | 5  | 
header {* Implementation of natural numbers by target-language integers *}
 | 
| 23854 | 6  | 
|
7  | 
theory Efficient_Nat  | 
|
| 
31203
 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 
haftmann 
parents: 
31128 
diff
changeset
 | 
8  | 
imports Code_Integer Main  | 
| 23854 | 9  | 
begin  | 
10  | 
||
11  | 
text {*
 | 
|
| 25931 | 12  | 
When generating code for functions on natural numbers, the  | 
13  | 
  canonical representation using @{term "0::nat"} and
 | 
|
| 37388 | 14  | 
  @{term Suc} is unsuitable for computations involving large
 | 
| 25931 | 15  | 
numbers. The efficiency of the generated code can be improved  | 
16  | 
drastically by implementing natural numbers by target-language  | 
|
17  | 
integers. To do this, just include this theory.  | 
|
| 23854 | 18  | 
*}  | 
19  | 
||
| 25931 | 20  | 
subsection {* Basic arithmetic *}
 | 
| 23854 | 21  | 
|
22  | 
text {*
 | 
|
23  | 
Most standard arithmetic functions on natural numbers are implemented  | 
|
24  | 
using their counterparts on the integers:  | 
|
25  | 
*}  | 
|
26  | 
||
| 25931 | 27  | 
code_datatype number_nat_inst.number_of_nat  | 
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
28  | 
|
| 
32069
 
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
 
haftmann 
parents: 
31998 
diff
changeset
 | 
29  | 
lemma zero_nat_code [code, code_unfold_post]:  | 
| 25931 | 30  | 
"0 = (Numeral0 :: nat)"  | 
31  | 
by simp  | 
|
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
32  | 
|
| 
32069
 
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
 
haftmann 
parents: 
31998 
diff
changeset
 | 
33  | 
lemma one_nat_code [code, code_unfold_post]:  | 
| 25931 | 34  | 
"1 = (Numeral1 :: nat)"  | 
35  | 
by simp  | 
|
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
36  | 
|
| 25931 | 37  | 
lemma Suc_code [code]:  | 
38  | 
"Suc n = n + 1"  | 
|
39  | 
by simp  | 
|
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
40  | 
|
| 25931 | 41  | 
lemma plus_nat_code [code]:  | 
42  | 
"n + m = nat (of_nat n + of_nat m)"  | 
|
43  | 
by simp  | 
|
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
44  | 
|
| 25931 | 45  | 
lemma minus_nat_code [code]:  | 
46  | 
"n - m = nat (of_nat n - of_nat m)"  | 
|
47  | 
by simp  | 
|
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
48  | 
|
| 25931 | 49  | 
lemma times_nat_code [code]:  | 
50  | 
"n * m = nat (of_nat n * of_nat m)"  | 
|
51  | 
unfolding of_nat_mult [symmetric] by simp  | 
|
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
52  | 
|
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
53  | 
lemma divmod_nat_code [code]:  | 
| 40607 | 54  | 
"divmod_nat n m = map_pair nat nat (pdivmod (of_nat n) (of_nat m))"  | 
55  | 
by (simp add: map_pair_def split_def pdivmod_def nat_div_distrib nat_mod_distrib divmod_nat_div_mod)  | 
|
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
56  | 
|
| 25931 | 57  | 
lemma eq_nat_code [code]:  | 
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38774 
diff
changeset
 | 
58  | 
"HOL.equal n m \<longleftrightarrow> HOL.equal (of_nat n \<Colon> int) (of_nat m)"  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38774 
diff
changeset
 | 
59  | 
by (simp add: equal)  | 
| 28351 | 60  | 
|
61  | 
lemma eq_nat_refl [code nbe]:  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38774 
diff
changeset
 | 
62  | 
"HOL.equal (n::nat) n \<longleftrightarrow> True"  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38774 
diff
changeset
 | 
63  | 
by (rule equal_refl)  | 
| 
24715
 
f96d86cdbe5a
Efficient_Nat and Pretty_Int integrated with ML_Int
 
haftmann 
parents: 
24630 
diff
changeset
 | 
64  | 
|
| 25931 | 65  | 
lemma less_eq_nat_code [code]:  | 
66  | 
"n \<le> m \<longleftrightarrow> (of_nat n \<Colon> int) \<le> of_nat m"  | 
|
67  | 
by simp  | 
|
| 23854 | 68  | 
|
| 25931 | 69  | 
lemma less_nat_code [code]:  | 
70  | 
"n < m \<longleftrightarrow> (of_nat n \<Colon> int) < of_nat m"  | 
|
71  | 
by simp  | 
|
| 23854 | 72  | 
|
| 25931 | 73  | 
subsection {* Case analysis *}
 | 
| 23854 | 74  | 
|
75  | 
text {*
 | 
|
| 25931 | 76  | 
Case analysis on natural numbers is rephrased using a conditional  | 
77  | 
expression:  | 
|
| 23854 | 78  | 
*}  | 
79  | 
||
| 
31998
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31954 
diff
changeset
 | 
80  | 
lemma [code, code_unfold]:  | 
| 25931 | 81  | 
"nat_case = (\<lambda>f g n. if n = 0 then f else g (n - 1))"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39272 
diff
changeset
 | 
82  | 
by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc)  | 
| 25615 | 83  | 
|
| 23854 | 84  | 
|
85  | 
subsection {* Preprocessors *}
 | 
|
86  | 
||
87  | 
text {*
 | 
|
88  | 
  In contrast to @{term "Suc n"}, the term @{term "n + (1::nat)"} is no longer
 | 
|
89  | 
a constructor term. Therefore, all occurrences of this term in a position  | 
|
90  | 
where a pattern is expected (i.e.\ on the left-hand side of a recursion  | 
|
91  | 
equation or in the arguments of an inductive relation in an introduction  | 
|
92  | 
rule) must be eliminated.  | 
|
93  | 
This can be accomplished by applying the following transformation rules:  | 
|
94  | 
*}  | 
|
95  | 
||
| 29937 | 96  | 
lemma Suc_if_eq: "(\<And>n. f (Suc n) \<equiv> h n) \<Longrightarrow> f 0 \<equiv> g \<Longrightarrow>  | 
97  | 
f n \<equiv> if n = 0 then g else h (n - 1)"  | 
|
| 
31954
 
8db19c99b00a
Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
 
haftmann 
parents: 
31377 
diff
changeset
 | 
98  | 
by (rule eq_reflection) (cases n, simp_all)  | 
| 29937 | 99  | 
|
| 25931 | 100  | 
lemma Suc_clause: "(\<And>n. P n (Suc n)) \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> P (n - 1) n"  | 
| 
29932
 
a2594b5c945a
dropped clause_suc_preproc for generic code generator
 
haftmann 
parents: 
29815 
diff
changeset
 | 
101  | 
by (cases n) simp_all  | 
| 23854 | 102  | 
|
103  | 
text {*
 | 
|
104  | 
The rules above are built into a preprocessor that is plugged into  | 
|
105  | 
the code generator. Since the preprocessor for introduction rules  | 
|
106  | 
does not know anything about modes, some of the modes that worked  | 
|
107  | 
for the canonical representation of natural numbers may no longer work.  | 
|
108  | 
*}  | 
|
109  | 
||
110  | 
(*<*)  | 
|
| 27609 | 111  | 
setup {*
 | 
112  | 
let  | 
|
| 23854 | 113  | 
|
| 
31954
 
8db19c99b00a
Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
 
haftmann 
parents: 
31377 
diff
changeset
 | 
114  | 
fun remove_suc thy thms =  | 
| 23854 | 115  | 
let  | 
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
40607 
diff
changeset
 | 
116  | 
val vname = singleton (Name.variant_list (map fst  | 
| 
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
40607 
diff
changeset
 | 
117  | 
(fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n";  | 
| 23854 | 118  | 
val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT));  | 
119  | 
fun lhs_of th = snd (Thm.dest_comb  | 
|
| 
31954
 
8db19c99b00a
Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
 
haftmann 
parents: 
31377 
diff
changeset
 | 
120  | 
(fst (Thm.dest_comb (cprop_of th))));  | 
| 
 
8db19c99b00a
Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
 
haftmann 
parents: 
31377 
diff
changeset
 | 
121  | 
fun rhs_of th = snd (Thm.dest_comb (cprop_of th));  | 
| 23854 | 122  | 
fun find_vars ct = (case term_of ct of  | 
| 
29932
 
a2594b5c945a
dropped clause_suc_preproc for generic code generator
 
haftmann 
parents: 
29815 
diff
changeset
 | 
123  | 
        (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
 | 
| 23854 | 124  | 
| _ $ _ =>  | 
125  | 
let val (ct1, ct2) = Thm.dest_comb ct  | 
|
126  | 
in  | 
|
127  | 
map (apfst (fn ct => Thm.capply ct ct2)) (find_vars ct1) @  | 
|
128  | 
map (apfst (Thm.capply ct1)) (find_vars ct2)  | 
|
129  | 
end  | 
|
130  | 
| _ => []);  | 
|
131  | 
val eqs = maps  | 
|
132  | 
(fn th => map (pair th) (find_vars (lhs_of th))) thms;  | 
|
133  | 
fun mk_thms (th, (ct, cv')) =  | 
|
134  | 
let  | 
|
135  | 
val th' =  | 
|
136  | 
Thm.implies_elim  | 
|
137  | 
(Conv.fconv_rule (Thm.beta_conversion true)  | 
|
138  | 
(Drule.instantiate'  | 
|
139  | 
[SOME (ctyp_of_term ct)] [SOME (Thm.cabs cv ct),  | 
|
140  | 
SOME (Thm.cabs cv' (rhs_of th)), NONE, SOME cv']  | 
|
| 
31954
 
8db19c99b00a
Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
 
haftmann 
parents: 
31377 
diff
changeset
 | 
141  | 
               @{thm Suc_if_eq})) (Thm.forall_intr cv' th)
 | 
| 23854 | 142  | 
in  | 
143  | 
case map_filter (fn th'' =>  | 
|
144  | 
SOME (th'', singleton  | 
|
| 
36603
 
d5d6111761a6
renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36176 
diff
changeset
 | 
145  | 
(Variable.trade (K (fn [th'''] => [th''' RS th']))  | 
| 
 
d5d6111761a6
renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36176 
diff
changeset
 | 
146  | 
(Variable.global_thm_context th'')) th'')  | 
| 23854 | 147  | 
handle THM _ => NONE) thms of  | 
148  | 
[] => NONE  | 
|
149  | 
| thps =>  | 
|
150  | 
let val (ths1, ths2) = split_list thps  | 
|
151  | 
in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end  | 
|
152  | 
end  | 
|
| 29937 | 153  | 
in get_first mk_thms eqs end;  | 
154  | 
||
| 
34893
 
ecdc526af73a
function transformer preprocessor applies to both code generators
 
haftmann 
parents: 
33364 
diff
changeset
 | 
155  | 
fun eqn_suc_base_preproc thy thms =  | 
| 29937 | 156  | 
let  | 
| 
31954
 
8db19c99b00a
Stefan Berghofer's code generator uses Pure equality instead of HOL equality now
 
haftmann 
parents: 
31377 
diff
changeset
 | 
157  | 
val dest = fst o Logic.dest_equals o prop_of;  | 
| 29937 | 158  | 
    val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc});
 | 
159  | 
in  | 
|
160  | 
if forall (can dest) thms andalso exists (contains_suc o dest) thms  | 
|
| 32348 | 161  | 
then thms |> perhaps_loop (remove_suc thy) |> (Option.map o map) Drule.zero_var_indexes  | 
| 29937 | 162  | 
else NONE  | 
| 23854 | 163  | 
end;  | 
164  | 
||
| 
34893
 
ecdc526af73a
function transformer preprocessor applies to both code generators
 
haftmann 
parents: 
33364 
diff
changeset
 | 
165  | 
val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc;  | 
| 23854 | 166  | 
|
| 27609 | 167  | 
in  | 
168  | 
||
| 
34893
 
ecdc526af73a
function transformer preprocessor applies to both code generators
 
haftmann 
parents: 
33364 
diff
changeset
 | 
169  | 
  Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc)
 | 
| 27609 | 170  | 
|
171  | 
end;  | 
|
| 23854 | 172  | 
*}  | 
173  | 
(*>*)  | 
|
174  | 
||
| 27609 | 175  | 
|
| 25931 | 176  | 
subsection {* Target language setup *}
 | 
177  | 
||
178  | 
text {*
 | 
|
| 25967 | 179  | 
  For ML, we map @{typ nat} to target language integers, where we
 | 
| 34899 | 180  | 
ensure that values are always non-negative.  | 
| 25931 | 181  | 
*}  | 
182  | 
||
183  | 
code_type nat  | 
|
| 27496 | 184  | 
(SML "IntInf.int")  | 
| 25931 | 185  | 
(OCaml "Big'_int.big'_int")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
186  | 
(Eval "int")  | 
| 25931 | 187  | 
|
188  | 
text {*
 | 
|
| 
45185
 
3a0c63c0ed48
removing old code generator setup for efficient natural numbers; cleaned typo
 
bulwahn 
parents: 
43653 
diff
changeset
 | 
189  | 
  For Haskell and Scala we define our own @{typ nat} type.  The reason
 | 
| 34899 | 190  | 
  is that we have to distinguish type class instances for @{typ nat}
 | 
191  | 
  and @{typ int}.
 | 
|
| 25967 | 192  | 
*}  | 
193  | 
||
| 38774 | 194  | 
code_include Haskell "Nat"  | 
195  | 
{*newtype Nat = Nat Integer deriving (Eq, Show, Read);
 | 
|
| 25967 | 196  | 
|
197  | 
instance Num Nat where {
 | 
|
198  | 
fromInteger k = Nat (if k >= 0 then k else 0);  | 
|
199  | 
Nat n + Nat m = Nat (n + m);  | 
|
200  | 
Nat n - Nat m = fromInteger (n - m);  | 
|
201  | 
Nat n * Nat m = Nat (n * m);  | 
|
202  | 
abs n = n;  | 
|
203  | 
signum _ = 1;  | 
|
204  | 
negate n = error "negate Nat";  | 
|
205  | 
};  | 
|
206  | 
||
207  | 
instance Ord Nat where {
 | 
|
208  | 
Nat n <= Nat m = n <= m;  | 
|
209  | 
Nat n < Nat m = n < m;  | 
|
210  | 
};  | 
|
211  | 
||
212  | 
instance Real Nat where {
 | 
|
213  | 
toRational (Nat n) = toRational n;  | 
|
214  | 
};  | 
|
215  | 
||
216  | 
instance Enum Nat where {
 | 
|
217  | 
toEnum k = fromInteger (toEnum k);  | 
|
218  | 
fromEnum (Nat n) = fromEnum n;  | 
|
219  | 
};  | 
|
220  | 
||
221  | 
instance Integral Nat where {
 | 
|
222  | 
toInteger (Nat n) = n;  | 
|
223  | 
divMod n m = quotRem n m;  | 
|
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
224  | 
quotRem (Nat n) (Nat m)  | 
| 
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
225  | 
| (m == 0) = (0, Nat n)  | 
| 
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
226  | 
| otherwise = (Nat k, Nat l) where (k, l) = quotRem n m;  | 
| 25967 | 227  | 
};  | 
228  | 
*}  | 
|
229  | 
||
230  | 
code_reserved Haskell Nat  | 
|
231  | 
||
| 38774 | 232  | 
code_include Scala "Nat"  | 
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38857 
diff
changeset
 | 
233  | 
{*object Nat {
 | 
| 34899 | 234  | 
|
235  | 
def apply(numeral: BigInt): Nat = new Nat(numeral max 0)  | 
|
236  | 
def apply(numeral: Int): Nat = Nat(BigInt(numeral))  | 
|
237  | 
def apply(numeral: String): Nat = Nat(BigInt(numeral))  | 
|
238  | 
||
239  | 
}  | 
|
240  | 
||
241  | 
class Nat private(private val value: BigInt) {
 | 
|
242  | 
||
243  | 
override def hashCode(): Int = this.value.hashCode()  | 
|
244  | 
||
245  | 
  override def equals(that: Any): Boolean = that match {
 | 
|
246  | 
case that: Nat => this equals that  | 
|
247  | 
case _ => false  | 
|
248  | 
}  | 
|
249  | 
||
250  | 
override def toString(): String = this.value.toString  | 
|
251  | 
||
252  | 
def equals(that: Nat): Boolean = this.value == that.value  | 
|
253  | 
||
254  | 
def as_BigInt: BigInt = this.value  | 
|
| 39781 | 255  | 
def as_Int: Int = if (this.value >= scala.Int.MinValue && this.value <= scala.Int.MaxValue)  | 
| 
34944
 
970e1466028d
code literals: distinguish numeral classes by different entries
 
haftmann 
parents: 
34902 
diff
changeset
 | 
256  | 
this.value.intValue  | 
| 37969 | 257  | 
    else error("Int value out of range: " + this.value.toString)
 | 
| 34899 | 258  | 
|
259  | 
def +(that: Nat): Nat = new Nat(this.value + that.value)  | 
|
| 37223 | 260  | 
def -(that: Nat): Nat = Nat(this.value - that.value)  | 
| 34899 | 261  | 
def *(that: Nat): Nat = new Nat(this.value * that.value)  | 
262  | 
||
263  | 
def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this)  | 
|
264  | 
    else {
 | 
|
265  | 
val (k, l) = this.value /% that.value  | 
|
266  | 
(new Nat(k), new Nat(l))  | 
|
267  | 
}  | 
|
268  | 
||
269  | 
def <=(that: Nat): Boolean = this.value <= that.value  | 
|
270  | 
||
271  | 
def <(that: Nat): Boolean = this.value < that.value  | 
|
272  | 
||
273  | 
}  | 
|
274  | 
*}  | 
|
275  | 
||
276  | 
code_reserved Scala Nat  | 
|
277  | 
||
| 25967 | 278  | 
code_type nat  | 
| 29793 | 279  | 
(Haskell "Nat.Nat")  | 
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38857 
diff
changeset
 | 
280  | 
(Scala "Nat")  | 
| 25967 | 281  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38774 
diff
changeset
 | 
282  | 
code_instance nat :: equal  | 
| 25967 | 283  | 
(Haskell -)  | 
284  | 
||
285  | 
text {*
 | 
|
| 25931 | 286  | 
Natural numerals.  | 
287  | 
*}  | 
|
288  | 
||
| 
32069
 
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
 
haftmann 
parents: 
31998 
diff
changeset
 | 
289  | 
lemma [code_unfold_post]:  | 
| 25931 | 290  | 
"nat (number_of i) = number_nat_inst.number_of_nat i"  | 
291  | 
  -- {* this interacts as desired with @{thm nat_number_of_def} *}
 | 
|
292  | 
by (simp add: number_nat_inst.number_of_nat)  | 
|
293  | 
||
294  | 
setup {*
 | 
|
295  | 
  fold (Numeral.add_code @{const_name number_nat_inst.number_of_nat}
 | 
|
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
296  | 
false Code_Printer.literal_positive_numeral) ["SML", "OCaml", "Haskell", "Scala"]  | 
| 25931 | 297  | 
*}  | 
298  | 
||
299  | 
text {*
 | 
|
300  | 
Since natural numbers are implemented  | 
|
| 25967 | 301  | 
  using integers in ML, the coercion function @{const "of_nat"} of type
 | 
| 25931 | 302  | 
  @{typ "nat \<Rightarrow> int"} is simply implemented by the identity function.
 | 
| 37388 | 303  | 
  For the @{const nat} function for converting an integer to a natural
 | 
| 25931 | 304  | 
number, we give a specific implementation using an ML function that  | 
305  | 
returns its input value, provided that it is non-negative, and otherwise  | 
|
306  | 
  returns @{text "0"}.
 | 
|
307  | 
*}  | 
|
308  | 
||
| 32073 | 309  | 
definition int :: "nat \<Rightarrow> int" where  | 
| 28562 | 310  | 
[code del]: "int = of_nat"  | 
| 25931 | 311  | 
|
| 28562 | 312  | 
lemma int_code' [code]:  | 
| 25931 | 313  | 
"int (number_of l) = (if neg (number_of l \<Colon> int) then 0 else number_of l)"  | 
314  | 
unfolding int_nat_number_of [folded int_def] ..  | 
|
315  | 
||
| 28562 | 316  | 
lemma nat_code' [code]:  | 
| 25931 | 317  | 
"nat (number_of l) = (if neg (number_of l \<Colon> int) then 0 else number_of l)"  | 
| 28969 | 318  | 
unfolding nat_number_of_def number_of_is_id neg_def by simp  | 
| 25931 | 319  | 
|
| 
32069
 
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
 
haftmann 
parents: 
31998 
diff
changeset
 | 
320  | 
lemma of_nat_int [code_unfold_post]:  | 
| 25931 | 321  | 
"of_nat = int" by (simp add: int_def)  | 
322  | 
||
| 32073 | 323  | 
lemma of_nat_aux_int [code_unfold]:  | 
324  | 
"of_nat_aux (\<lambda>i. i + 1) k 0 = int k"  | 
|
325  | 
by (simp add: int_def Nat.of_nat_code)  | 
|
326  | 
||
| 25931 | 327  | 
code_const int  | 
328  | 
(SML "_")  | 
|
329  | 
(OCaml "_")  | 
|
330  | 
||
| 25967 | 331  | 
code_const nat  | 
| 43653 | 332  | 
(SML "IntInf.max/ (0,/ _)")  | 
| 25967 | 333  | 
(OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
334  | 
(Eval "Integer.max/ _/ 0")  | 
| 25967 | 335  | 
|
| 35689 | 336  | 
text {* For Haskell and Scala, things are slightly different again. *}
 | 
| 25967 | 337  | 
|
338  | 
code_const int and nat  | 
|
339  | 
(Haskell "toInteger" and "fromInteger")  | 
|
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38857 
diff
changeset
 | 
340  | 
(Scala "!_.as'_BigInt" and "Nat")  | 
| 25931 | 341  | 
|
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
342  | 
text {* Conversion from and to code numerals. *}
 | 
| 25931 | 343  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
344  | 
code_const Code_Numeral.of_nat  | 
| 25967 | 345  | 
(SML "IntInf.toInt")  | 
| 31377 | 346  | 
(OCaml "_")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
347  | 
(Haskell "!(fromInteger/ ./ toInteger)")  | 
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38857 
diff
changeset
 | 
348  | 
(Scala "!Natural(_.as'_BigInt)")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
349  | 
(Eval "_")  | 
| 25967 | 350  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
351  | 
code_const Code_Numeral.nat_of  | 
| 25931 | 352  | 
(SML "IntInf.fromInt")  | 
| 31377 | 353  | 
(OCaml "_")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
354  | 
(Haskell "!(fromInteger/ ./ toInteger)")  | 
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38857 
diff
changeset
 | 
355  | 
(Scala "!Nat(_.as'_BigInt)")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
356  | 
(Eval "_")  | 
| 25931 | 357  | 
|
358  | 
text {* Using target language arithmetic operations whenever appropriate *}
 | 
|
359  | 
||
360  | 
code_const "op + \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"  | 
|
361  | 
(SML "IntInf.+ ((_), (_))")  | 
|
362  | 
(OCaml "Big'_int.add'_big'_int")  | 
|
363  | 
(Haskell infixl 6 "+")  | 
|
| 34899 | 364  | 
(Scala infixl 7 "+")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
365  | 
(Eval infixl 8 "+")  | 
| 34899 | 366  | 
|
367  | 
code_const "op - \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"  | 
|
368  | 
(Haskell infixl 6 "-")  | 
|
369  | 
(Scala infixl 7 "-")  | 
|
| 25931 | 370  | 
|
371  | 
code_const "op * \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"  | 
|
372  | 
(SML "IntInf.* ((_), (_))")  | 
|
373  | 
(OCaml "Big'_int.mult'_big'_int")  | 
|
374  | 
(Haskell infixl 7 "*")  | 
|
| 34899 | 375  | 
(Scala infixl 8 "*")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
376  | 
(Eval infixl 9 "*")  | 
| 25931 | 377  | 
|
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
378  | 
code_const divmod_nat  | 
| 26009 | 379  | 
(SML "IntInf.divMod/ ((_),/ (_))")  | 
380  | 
(OCaml "Big'_int.quomod'_big'_int")  | 
|
381  | 
(Haskell "divMod")  | 
|
| 34899 | 382  | 
(Scala infixl 8 "/%")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
383  | 
(Eval "Integer.div'_mod")  | 
| 25931 | 384  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38774 
diff
changeset
 | 
385  | 
code_const "HOL.equal \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"  | 
| 25931 | 386  | 
(SML "!((_ : IntInf.int) = _)")  | 
387  | 
(OCaml "Big'_int.eq'_big'_int")  | 
|
| 39272 | 388  | 
(Haskell infix 4 "==")  | 
| 34899 | 389  | 
(Scala infixl 5 "==")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
390  | 
(Eval infixl 6 "=")  | 
| 25931 | 391  | 
|
392  | 
code_const "op \<le> \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"  | 
|
393  | 
(SML "IntInf.<= ((_), (_))")  | 
|
394  | 
(OCaml "Big'_int.le'_big'_int")  | 
|
395  | 
(Haskell infix 4 "<=")  | 
|
| 34899 | 396  | 
(Scala infixl 4 "<=")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
397  | 
(Eval infixl 6 "<=")  | 
| 25931 | 398  | 
|
399  | 
code_const "op < \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"  | 
|
400  | 
(SML "IntInf.< ((_), (_))")  | 
|
401  | 
(OCaml "Big'_int.lt'_big'_int")  | 
|
402  | 
(Haskell infix 4 "<")  | 
|
| 34899 | 403  | 
(Scala infixl 4 "<")  | 
| 
37958
 
9728342bcd56
another refinement chapter in the neverending numeral story
 
haftmann 
parents: 
37947 
diff
changeset
 | 
404  | 
(Eval infixl 6 "<")  | 
| 25931 | 405  | 
|
406  | 
||
| 28228 | 407  | 
text {* Evaluation *}
 | 
408  | 
||
| 28562 | 409  | 
lemma [code, code del]:  | 
| 32657 | 410  | 
"(Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term) = Code_Evaluation.term_of" ..  | 
| 28228 | 411  | 
|
| 32657 | 412  | 
code_const "Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term"  | 
| 28228 | 413  | 
(SML "HOLogic.mk'_number/ HOLogic.natT")  | 
414  | 
||
415  | 
||
| 25931 | 416  | 
text {* Module names *}
 | 
| 23854 | 417  | 
|
418  | 
code_modulename SML  | 
|
| 33364 | 419  | 
Efficient_Nat Arith  | 
| 23854 | 420  | 
|
421  | 
code_modulename OCaml  | 
|
| 33364 | 422  | 
Efficient_Nat Arith  | 
| 23854 | 423  | 
|
424  | 
code_modulename Haskell  | 
|
| 33364 | 425  | 
Efficient_Nat Arith  | 
| 23854 | 426  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
35689 
diff
changeset
 | 
427  | 
hide_const int  | 
| 23854 | 428  | 
|
429  | 
end  |