| author | ballarin | 
| Tue, 16 Dec 2008 21:10:53 +0100 | |
| changeset 29237 | e90d9d51106b | 
| parent 28708 | a1a436f09ec6 | 
| child 29815 | 9e94b7078fa5 | 
| permissions | -rw-r--r-- | 
| 24999 | 1  | 
(* ID: $Id$  | 
2  | 
Author: Florian Haftmann, TU Muenchen  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Type of indices *}
 | 
|
6  | 
||
7  | 
theory Code_Index  | 
|
| 28228 | 8  | 
imports Plain "~~/src/HOL/Code_Eval" "~~/src/HOL/Presburger"  | 
| 24999 | 9  | 
begin  | 
10  | 
||
11  | 
text {*
 | 
|
| 25767 | 12  | 
  Indices are isomorphic to HOL @{typ nat} but
 | 
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
13  | 
mapped to target-language builtin integers.  | 
| 24999 | 14  | 
*}  | 
15  | 
||
16  | 
subsection {* Datatype of indices *}
 | 
|
17  | 
||
| 26140 | 18  | 
typedef index = "UNIV \<Colon> nat set"  | 
19  | 
morphisms nat_of_index index_of_nat by rule  | 
|
| 24999 | 20  | 
|
| 26140 | 21  | 
lemma index_of_nat_nat_of_index [simp]:  | 
22  | 
"index_of_nat (nat_of_index k) = k"  | 
|
23  | 
by (rule nat_of_index_inverse)  | 
|
| 24999 | 24  | 
|
| 26140 | 25  | 
lemma nat_of_index_index_of_nat [simp]:  | 
26  | 
"nat_of_index (index_of_nat n) = n"  | 
|
27  | 
by (rule index_of_nat_inverse)  | 
|
28  | 
(unfold index_def, rule UNIV_I)  | 
|
| 24999 | 29  | 
|
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
30  | 
lemma [measure_function]:  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
31  | 
"is_measure nat_of_index" by (rule is_measure_trivial)  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
32  | 
|
| 24999 | 33  | 
lemma index:  | 
| 25767 | 34  | 
"(\<And>n\<Colon>index. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (index_of_nat n))"  | 
| 24999 | 35  | 
proof  | 
| 25767 | 36  | 
fix n :: nat  | 
37  | 
assume "\<And>n\<Colon>index. PROP P n"  | 
|
38  | 
then show "PROP P (index_of_nat n)" .  | 
|
| 24999 | 39  | 
next  | 
| 25767 | 40  | 
fix n :: index  | 
41  | 
assume "\<And>n\<Colon>nat. PROP P (index_of_nat n)"  | 
|
42  | 
then have "PROP P (index_of_nat (nat_of_index n))" .  | 
|
43  | 
then show "PROP P n" by simp  | 
|
| 24999 | 44  | 
qed  | 
45  | 
||
| 26140 | 46  | 
lemma index_case:  | 
47  | 
assumes "\<And>n. k = index_of_nat n \<Longrightarrow> P"  | 
|
48  | 
shows P  | 
|
49  | 
by (rule assms [of "nat_of_index k"]) simp  | 
|
50  | 
||
| 26304 | 51  | 
lemma index_induct_raw:  | 
| 26140 | 52  | 
assumes "\<And>n. P (index_of_nat n)"  | 
53  | 
shows "P k"  | 
|
54  | 
proof -  | 
|
55  | 
from assms have "P (index_of_nat (nat_of_index k))" .  | 
|
56  | 
then show ?thesis by simp  | 
|
57  | 
qed  | 
|
58  | 
||
59  | 
lemma nat_of_index_inject [simp]:  | 
|
60  | 
"nat_of_index k = nat_of_index l \<longleftrightarrow> k = l"  | 
|
61  | 
by (rule nat_of_index_inject)  | 
|
62  | 
||
63  | 
lemma index_of_nat_inject [simp]:  | 
|
64  | 
"index_of_nat n = index_of_nat m \<longleftrightarrow> n = m"  | 
|
65  | 
by (auto intro!: index_of_nat_inject simp add: index_def)  | 
|
66  | 
||
67  | 
instantiation index :: zero  | 
|
68  | 
begin  | 
|
69  | 
||
| 28562 | 70  | 
definition [simp, code del]:  | 
| 26140 | 71  | 
"0 = index_of_nat 0"  | 
72  | 
||
73  | 
instance ..  | 
|
74  | 
||
75  | 
end  | 
|
76  | 
||
77  | 
definition [simp]:  | 
|
78  | 
"Suc_index k = index_of_nat (Suc (nat_of_index k))"  | 
|
79  | 
||
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
80  | 
rep_datatype "0 \<Colon> index" Suc_index  | 
| 26140 | 81  | 
proof -  | 
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
82  | 
fix P :: "index \<Rightarrow> bool"  | 
| 
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
83  | 
fix k :: index  | 
| 26140 | 84  | 
assume "P 0" then have init: "P (index_of_nat 0)" by simp  | 
85  | 
assume "\<And>k. P k \<Longrightarrow> P (Suc_index k)"  | 
|
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
86  | 
then have "\<And>n. P (index_of_nat n) \<Longrightarrow> P (Suc_index (index_of_nat n))" .  | 
| 26140 | 87  | 
then have step: "\<And>n. P (index_of_nat n) \<Longrightarrow> P (index_of_nat (Suc n))" by simp  | 
88  | 
from init step have "P (index_of_nat (nat_of_index k))"  | 
|
89  | 
by (induct "nat_of_index k") simp_all  | 
|
90  | 
then show "P k" by simp  | 
|
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
91  | 
qed simp_all  | 
| 26140 | 92  | 
|
| 28562 | 93  | 
lemmas [code del] = index.recs index.cases  | 
| 26140 | 94  | 
|
95  | 
declare index_case [case_names nat, cases type: index]  | 
|
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
96  | 
declare index.induct [case_names nat, induct type: index]  | 
| 26140 | 97  | 
|
| 28562 | 98  | 
lemma [code]:  | 
| 26140 | 99  | 
"index_size = nat_of_index"  | 
100  | 
proof (rule ext)  | 
|
101  | 
fix k  | 
|
102  | 
have "index_size k = nat_size (nat_of_index k)"  | 
|
103  | 
by (induct k rule: index.induct) (simp_all del: zero_index_def Suc_index_def, simp_all)  | 
|
104  | 
also have "nat_size (nat_of_index k) = nat_of_index k" by (induct "nat_of_index k") simp_all  | 
|
105  | 
finally show "index_size k = nat_of_index k" .  | 
|
106  | 
qed  | 
|
107  | 
||
| 28562 | 108  | 
lemma [code]:  | 
| 26140 | 109  | 
"size = nat_of_index"  | 
110  | 
proof (rule ext)  | 
|
111  | 
fix k  | 
|
112  | 
show "size k = nat_of_index k"  | 
|
113  | 
by (induct k) (simp_all del: zero_index_def Suc_index_def, simp_all)  | 
|
114  | 
qed  | 
|
115  | 
||
| 28562 | 116  | 
lemma [code]:  | 
| 
28346
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28228 
diff
changeset
 | 
117  | 
"eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of_index k) (nat_of_index l)"  | 
| 
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28228 
diff
changeset
 | 
118  | 
by (cases k, cases l) (simp add: eq)  | 
| 24999 | 119  | 
|
| 28351 | 120  | 
lemma [code nbe]:  | 
121  | 
"eq_class.eq (k::index) k \<longleftrightarrow> True"  | 
|
122  | 
by (rule HOL.eq_refl)  | 
|
123  | 
||
| 24999 | 124  | 
|
| 25767 | 125  | 
subsection {* Indices as datatype of ints *}
 | 
126  | 
||
127  | 
instantiation index :: number  | 
|
128  | 
begin  | 
|
| 24999 | 129  | 
|
| 25767 | 130  | 
definition  | 
131  | 
"number_of = index_of_nat o nat"  | 
|
132  | 
||
133  | 
instance ..  | 
|
134  | 
||
135  | 
end  | 
|
| 24999 | 136  | 
|
| 26264 | 137  | 
lemma nat_of_index_number [simp]:  | 
138  | 
"nat_of_index (number_of k) = number_of k"  | 
|
139  | 
by (simp add: number_of_index_def nat_number_of_def number_of_is_id)  | 
|
140  | 
||
| 24999 | 141  | 
code_datatype "number_of \<Colon> int \<Rightarrow> index"  | 
142  | 
||
143  | 
||
144  | 
subsection {* Basic arithmetic *}
 | 
|
145  | 
||
| 25767 | 146  | 
instantiation index :: "{minus, ordered_semidom, Divides.div, linorder}"
 | 
147  | 
begin  | 
|
| 24999 | 148  | 
|
| 28562 | 149  | 
definition [simp, code del]:  | 
| 25767 | 150  | 
"(1\<Colon>index) = index_of_nat 1"  | 
| 24999 | 151  | 
|
| 28562 | 152  | 
definition [simp, code del]:  | 
| 25767 | 153  | 
"n + m = index_of_nat (nat_of_index n + nat_of_index m)"  | 
154  | 
||
| 28562 | 155  | 
definition [simp, code del]:  | 
| 25767 | 156  | 
"n - m = index_of_nat (nat_of_index n - nat_of_index m)"  | 
157  | 
||
| 28562 | 158  | 
definition [simp, code del]:  | 
| 25767 | 159  | 
"n * m = index_of_nat (nat_of_index n * nat_of_index m)"  | 
160  | 
||
| 28562 | 161  | 
definition [simp, code del]:  | 
| 25767 | 162  | 
"n div m = index_of_nat (nat_of_index n div nat_of_index m)"  | 
| 24999 | 163  | 
|
| 28562 | 164  | 
definition [simp, code del]:  | 
| 25767 | 165  | 
"n mod m = index_of_nat (nat_of_index n mod nat_of_index m)"  | 
| 24999 | 166  | 
|
| 28562 | 167  | 
definition [simp, code del]:  | 
| 25767 | 168  | 
"n \<le> m \<longleftrightarrow> nat_of_index n \<le> nat_of_index m"  | 
| 24999 | 169  | 
|
| 28562 | 170  | 
definition [simp, code del]:  | 
| 25767 | 171  | 
"n < m \<longleftrightarrow> nat_of_index n < nat_of_index m"  | 
| 24999 | 172  | 
|
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
173  | 
instance by default (auto simp add: left_distrib index)  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
174  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
175  | 
end  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
176  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
177  | 
lemma zero_index_code [code inline, code]:  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
178  | 
"(0\<Colon>index) = Numeral0"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
179  | 
by (simp add: number_of_index_def Pls_def)  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
180  | 
lemma [code post]: "Numeral0 = (0\<Colon>index)"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
181  | 
using zero_index_code ..  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
182  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
183  | 
lemma one_index_code [code inline, code]:  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
184  | 
"(1\<Colon>index) = Numeral1"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
185  | 
by (simp add: number_of_index_def Pls_def Bit1_def)  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
186  | 
lemma [code post]: "Numeral1 = (1\<Colon>index)"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
187  | 
using one_index_code ..  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
188  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
189  | 
lemma plus_index_code [code nbe]:  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
190  | 
"index_of_nat n + index_of_nat m = index_of_nat (n + m)"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
191  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
192  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
193  | 
definition subtract_index :: "index \<Rightarrow> index \<Rightarrow> index" where  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
194  | 
[simp, code del]: "subtract_index = op -"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
195  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
196  | 
lemma subtract_index_code [code nbe]:  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
197  | 
"subtract_index (index_of_nat n) (index_of_nat m) = index_of_nat (n - m)"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
198  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
199  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
200  | 
lemma minus_index_code [code]:  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
201  | 
"n - m = subtract_index n m"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
202  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
203  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
204  | 
lemma times_index_code [code nbe]:  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
205  | 
"index_of_nat n * index_of_nat m = index_of_nat (n * m)"  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
206  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
207  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
208  | 
lemma less_eq_index_code [code nbe]:  | 
| 25767 | 209  | 
"index_of_nat n \<le> index_of_nat m \<longleftrightarrow> n \<le> m"  | 
210  | 
by simp  | 
|
| 24999 | 211  | 
|
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
212  | 
lemma less_index_code [code nbe]:  | 
| 25767 | 213  | 
"index_of_nat n < index_of_nat m \<longleftrightarrow> n < m"  | 
214  | 
by simp  | 
|
| 24999 | 215  | 
|
| 26140 | 216  | 
lemma Suc_index_minus_one: "Suc_index n - 1 = n" by simp  | 
217  | 
||
| 25928 | 218  | 
lemma index_of_nat_code [code]:  | 
| 25918 | 219  | 
"index_of_nat = of_nat"  | 
220  | 
proof  | 
|
221  | 
fix n :: nat  | 
|
222  | 
have "of_nat n = index_of_nat n"  | 
|
223  | 
by (induct n) simp_all  | 
|
224  | 
then show "index_of_nat n = of_nat n"  | 
|
225  | 
by (rule sym)  | 
|
226  | 
qed  | 
|
227  | 
||
| 25928 | 228  | 
lemma index_not_eq_zero: "i \<noteq> index_of_nat 0 \<longleftrightarrow> i \<ge> 1"  | 
229  | 
by (cases i) auto  | 
|
230  | 
||
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
231  | 
definition nat_of_index_aux :: "index \<Rightarrow> nat \<Rightarrow> nat" where  | 
| 25928 | 232  | 
"nat_of_index_aux i n = nat_of_index i + n"  | 
233  | 
||
234  | 
lemma nat_of_index_aux_code [code]:  | 
|
235  | 
"nat_of_index_aux i n = (if i = 0 then n else nat_of_index_aux (i - 1) (Suc n))"  | 
|
236  | 
by (auto simp add: nat_of_index_aux_def index_not_eq_zero)  | 
|
237  | 
||
238  | 
lemma nat_of_index_code [code]:  | 
|
239  | 
"nat_of_index i = nat_of_index_aux i 0"  | 
|
240  | 
by (simp add: nat_of_index_aux_def)  | 
|
| 25918 | 241  | 
|
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
242  | 
definition div_mod_index :: "index \<Rightarrow> index \<Rightarrow> index \<times> index" where  | 
| 28562 | 243  | 
[code del]: "div_mod_index n m = (n div m, n mod m)"  | 
| 26009 | 244  | 
|
| 28562 | 245  | 
lemma [code]:  | 
| 26009 | 246  | 
"div_mod_index n m = (if m = 0 then (0, n) else (n div m, n mod m))"  | 
247  | 
unfolding div_mod_index_def by auto  | 
|
248  | 
||
| 28562 | 249  | 
lemma [code]:  | 
| 26009 | 250  | 
"n div m = fst (div_mod_index n m)"  | 
251  | 
unfolding div_mod_index_def by simp  | 
|
252  | 
||
| 28562 | 253  | 
lemma [code]:  | 
| 26009 | 254  | 
"n mod m = snd (div_mod_index n m)"  | 
255  | 
unfolding div_mod_index_def by simp  | 
|
256  | 
||
257  | 
||
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
258  | 
subsection {* ML interface *}
 | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
259  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
260  | 
ML {*
 | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
261  | 
structure Index =  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
262  | 
struct  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
263  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
264  | 
fun mk k = HOLogic.mk_number @{typ index} k;
 | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
265  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
266  | 
end;  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
267  | 
*}  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
268  | 
|
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
269  | 
|
| 28228 | 270  | 
subsection {* Code generator setup *}
 | 
| 24999 | 271  | 
|
| 25767 | 272  | 
text {* Implementation of indices by bounded integers *}
 | 
273  | 
||
| 24999 | 274  | 
code_type index  | 
275  | 
(SML "int")  | 
|
276  | 
(OCaml "int")  | 
|
| 25967 | 277  | 
(Haskell "Int")  | 
| 24999 | 278  | 
|
279  | 
code_instance index :: eq  | 
|
280  | 
(Haskell -)  | 
|
281  | 
||
282  | 
setup {*
 | 
|
| 25928 | 283  | 
  fold (Numeral.add_code @{const_name number_index_inst.number_of_index}
 | 
284  | 
false false) ["SML", "OCaml", "Haskell"]  | 
|
| 24999 | 285  | 
*}  | 
286  | 
||
| 25918 | 287  | 
code_reserved SML Int int  | 
288  | 
code_reserved OCaml Pervasives int  | 
|
| 24999 | 289  | 
|
290  | 
code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index"  | 
|
| 25928 | 291  | 
(SML "Int.+/ ((_),/ (_))")  | 
| 25967 | 292  | 
(OCaml "Pervasives.( + )")  | 
| 24999 | 293  | 
(Haskell infixl 6 "+")  | 
294  | 
||
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
295  | 
code_const "subtract_index \<Colon> index \<Rightarrow> index \<Rightarrow> index"  | 
| 25918 | 296  | 
(SML "Int.max/ (_/ -/ _,/ 0 : int)")  | 
297  | 
(OCaml "Pervasives.max/ (_/ -/ _)/ (0 : int) ")  | 
|
298  | 
(Haskell "max/ (_/ -/ _)/ (0 :: Int)")  | 
|
| 24999 | 299  | 
|
300  | 
code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index"  | 
|
| 25928 | 301  | 
(SML "Int.*/ ((_),/ (_))")  | 
| 25967 | 302  | 
(OCaml "Pervasives.( * )")  | 
| 24999 | 303  | 
(Haskell infixl 7 "*")  | 
304  | 
||
| 26009 | 305  | 
code_const div_mod_index  | 
306  | 
(SML "(fn n => fn m =>/ (n div m, n mod m))")  | 
|
307  | 
(OCaml "(fun n -> fun m ->/ (n '/ m, n mod m))")  | 
|
308  | 
(Haskell "divMod")  | 
|
| 25928 | 309  | 
|
| 
28346
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28228 
diff
changeset
 | 
310  | 
code_const "eq_class.eq \<Colon> index \<Rightarrow> index \<Rightarrow> bool"  | 
| 24999 | 311  | 
(SML "!((_ : Int.int) = _)")  | 
| 25967 | 312  | 
(OCaml "!((_ : int) = _)")  | 
| 24999 | 313  | 
(Haskell infixl 4 "==")  | 
314  | 
||
315  | 
code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool"  | 
|
| 25928 | 316  | 
(SML "Int.<=/ ((_),/ (_))")  | 
| 25967 | 317  | 
(OCaml "!((_ : int) <= _)")  | 
| 24999 | 318  | 
(Haskell infix 4 "<=")  | 
319  | 
||
320  | 
code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool"  | 
|
| 25928 | 321  | 
(SML "Int.</ ((_),/ (_))")  | 
| 25967 | 322  | 
(OCaml "!((_ : int) < _)")  | 
| 24999 | 323  | 
(Haskell infix 4 "<")  | 
324  | 
||
| 28228 | 325  | 
text {* Evaluation *}
 | 
326  | 
||
| 28562 | 327  | 
lemma [code, code del]:  | 
| 28228 | 328  | 
"(Code_Eval.term_of \<Colon> index \<Rightarrow> term) = Code_Eval.term_of" ..  | 
329  | 
||
330  | 
code_const "Code_Eval.term_of \<Colon> index \<Rightarrow> term"  | 
|
331  | 
(SML "HOLogic.mk'_number/ HOLogic.indexT/ (IntInf.fromInt/ _)")  | 
|
332  | 
||
| 24999 | 333  | 
end  |