| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 29 Jun 2010 02:18:08 +0100 | |
| changeset 37594 | 32ad67684ee7 | 
| parent 36176 | 3fe7e97ccca8 | 
| child 37947 | 844977c7abeb | 
| permissions | -rw-r--r-- | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 24999 | 2  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
3  | 
header {* Type of target language numerals *}
 | 
| 24999 | 4  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
5  | 
theory Code_Numeral  | 
| 
33318
 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33296 
diff
changeset
 | 
6  | 
imports Nat_Numeral Nat_Transfer Divides  | 
| 24999 | 7  | 
begin  | 
8  | 
||
9  | 
text {*
 | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
10  | 
  Code numerals are isomorphic to HOL @{typ nat} but
 | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
11  | 
mapped to target-language builtin numerals.  | 
| 24999 | 12  | 
*}  | 
13  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
14  | 
subsection {* Datatype of target language numerals *}
 | 
| 24999 | 15  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
16  | 
typedef (open) code_numeral = "UNIV \<Colon> nat set"  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
17  | 
morphisms nat_of of_nat by rule  | 
| 24999 | 18  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
19  | 
lemma of_nat_nat_of [simp]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
20  | 
"of_nat (nat_of k) = k"  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
21  | 
by (rule nat_of_inverse)  | 
| 24999 | 22  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
23  | 
lemma nat_of_of_nat [simp]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
24  | 
"nat_of (of_nat n) = n"  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
25  | 
by (rule of_nat_inverse) (rule UNIV_I)  | 
| 24999 | 26  | 
|
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
27  | 
lemma [measure_function]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
28  | 
"is_measure nat_of" by (rule is_measure_trivial)  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
29  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
30  | 
lemma code_numeral:  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
31  | 
"(\<And>n\<Colon>code_numeral. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (of_nat n))"  | 
| 24999 | 32  | 
proof  | 
| 25767 | 33  | 
fix n :: nat  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
34  | 
assume "\<And>n\<Colon>code_numeral. PROP P n"  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
35  | 
then show "PROP P (of_nat n)" .  | 
| 24999 | 36  | 
next  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
37  | 
fix n :: code_numeral  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
38  | 
assume "\<And>n\<Colon>nat. PROP P (of_nat n)"  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
39  | 
then have "PROP P (of_nat (nat_of n))" .  | 
| 25767 | 40  | 
then show "PROP P n" by simp  | 
| 24999 | 41  | 
qed  | 
42  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
43  | 
lemma code_numeral_case:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
44  | 
assumes "\<And>n. k = of_nat n \<Longrightarrow> P"  | 
| 26140 | 45  | 
shows P  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
46  | 
by (rule assms [of "nat_of k"]) simp  | 
| 26140 | 47  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
48  | 
lemma code_numeral_induct_raw:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
49  | 
assumes "\<And>n. P (of_nat n)"  | 
| 26140 | 50  | 
shows "P k"  | 
51  | 
proof -  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
52  | 
from assms have "P (of_nat (nat_of k))" .  | 
| 26140 | 53  | 
then show ?thesis by simp  | 
54  | 
qed  | 
|
55  | 
||
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
56  | 
lemma nat_of_inject [simp]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
57  | 
"nat_of k = nat_of l \<longleftrightarrow> k = l"  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
58  | 
by (rule nat_of_inject)  | 
| 26140 | 59  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
60  | 
lemma of_nat_inject [simp]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
61  | 
"of_nat n = of_nat m \<longleftrightarrow> n = m"  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
62  | 
by (rule of_nat_inject) (rule UNIV_I)+  | 
| 26140 | 63  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
64  | 
instantiation code_numeral :: zero  | 
| 26140 | 65  | 
begin  | 
66  | 
||
| 28562 | 67  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
68  | 
"0 = of_nat 0"  | 
| 26140 | 69  | 
|
70  | 
instance ..  | 
|
71  | 
||
72  | 
end  | 
|
73  | 
||
74  | 
definition [simp]:  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
75  | 
"Suc_code_numeral k = of_nat (Suc (nat_of k))"  | 
| 26140 | 76  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
77  | 
rep_datatype "0 \<Colon> code_numeral" Suc_code_numeral  | 
| 26140 | 78  | 
proof -  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
79  | 
fix P :: "code_numeral \<Rightarrow> bool"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
80  | 
fix k :: code_numeral  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
81  | 
assume "P 0" then have init: "P (of_nat 0)" by simp  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
82  | 
assume "\<And>k. P k \<Longrightarrow> P (Suc_code_numeral k)"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
83  | 
then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc_code_numeral (of_nat n))" .  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
84  | 
then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Suc n))" by simp  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
85  | 
from init step have "P (of_nat (nat_of k))"  | 
| 34915 | 86  | 
    by (induct ("nat_of k")) simp_all
 | 
| 26140 | 87  | 
then show "P k" by simp  | 
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
26304 
diff
changeset
 | 
88  | 
qed simp_all  | 
| 26140 | 89  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
90  | 
declare code_numeral_case [case_names nat, cases type: code_numeral]  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
91  | 
declare code_numeral.induct [case_names nat, induct type: code_numeral]  | 
| 26140 | 92  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
93  | 
lemma code_numeral_decr [termination_simp]:  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
94  | 
"k \<noteq> of_nat 0 \<Longrightarrow> nat_of k - Suc 0 < nat_of k"  | 
| 
30245
 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 
haftmann 
parents: 
29823 
diff
changeset
 | 
95  | 
by (cases k) simp  | 
| 
 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 
haftmann 
parents: 
29823 
diff
changeset
 | 
96  | 
|
| 
 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 
haftmann 
parents: 
29823 
diff
changeset
 | 
97  | 
lemma [simp, code]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
98  | 
"code_numeral_size = nat_of"  | 
| 26140 | 99  | 
proof (rule ext)  | 
100  | 
fix k  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
101  | 
have "code_numeral_size k = nat_size (nat_of k)"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
102  | 
by (induct k rule: code_numeral.induct) (simp_all del: zero_code_numeral_def Suc_code_numeral_def, simp_all)  | 
| 34915 | 103  | 
  also have "nat_size (nat_of k) = nat_of k" by (induct ("nat_of k")) simp_all
 | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
104  | 
finally show "code_numeral_size k = nat_of k" .  | 
| 26140 | 105  | 
qed  | 
106  | 
||
| 
30245
 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 
haftmann 
parents: 
29823 
diff
changeset
 | 
107  | 
lemma [simp, code]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
108  | 
"size = nat_of"  | 
| 26140 | 109  | 
proof (rule ext)  | 
110  | 
fix k  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
111  | 
show "size k = nat_of k"  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
112  | 
by (induct k) (simp_all del: zero_code_numeral_def Suc_code_numeral_def, simp_all)  | 
| 26140 | 113  | 
qed  | 
114  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
115  | 
lemmas [code del] = code_numeral.recs code_numeral.cases  | 
| 
30245
 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 
haftmann 
parents: 
29823 
diff
changeset
 | 
116  | 
|
| 28562 | 117  | 
lemma [code]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
118  | 
"eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of k) (nat_of l)"  | 
| 
28346
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28228 
diff
changeset
 | 
119  | 
by (cases k, cases l) (simp add: eq)  | 
| 24999 | 120  | 
|
| 28351 | 121  | 
lemma [code nbe]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
122  | 
"eq_class.eq (k::code_numeral) k \<longleftrightarrow> True"  | 
| 28351 | 123  | 
by (rule HOL.eq_refl)  | 
124  | 
||
| 24999 | 125  | 
|
| 25767 | 126  | 
subsection {* Indices as datatype of ints *}
 | 
127  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
128  | 
instantiation code_numeral :: number  | 
| 25767 | 129  | 
begin  | 
| 24999 | 130  | 
|
| 25767 | 131  | 
definition  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
132  | 
"number_of = of_nat o nat"  | 
| 25767 | 133  | 
|
134  | 
instance ..  | 
|
135  | 
||
136  | 
end  | 
|
| 24999 | 137  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
138  | 
lemma nat_of_number [simp]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
139  | 
"nat_of (number_of k) = number_of k"  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
140  | 
by (simp add: number_of_code_numeral_def nat_number_of_def number_of_is_id)  | 
| 26264 | 141  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
142  | 
code_datatype "number_of \<Colon> int \<Rightarrow> code_numeral"  | 
| 24999 | 143  | 
|
144  | 
||
145  | 
subsection {* Basic arithmetic *}
 | 
|
146  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34944 
diff
changeset
 | 
147  | 
instantiation code_numeral :: "{minus, linordered_semidom, semiring_div, linorder}"
 | 
| 25767 | 148  | 
begin  | 
| 24999 | 149  | 
|
| 28562 | 150  | 
definition [simp, code del]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
151  | 
"(1\<Colon>code_numeral) = of_nat 1"  | 
| 24999 | 152  | 
|
| 28562 | 153  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
154  | 
"n + m = of_nat (nat_of n + nat_of m)"  | 
| 25767 | 155  | 
|
| 28562 | 156  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
157  | 
"n - m = of_nat (nat_of n - nat_of m)"  | 
| 25767 | 158  | 
|
| 28562 | 159  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
160  | 
"n * m = of_nat (nat_of n * nat_of m)"  | 
| 25767 | 161  | 
|
| 28562 | 162  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
163  | 
"n div m = of_nat (nat_of n div nat_of m)"  | 
| 24999 | 164  | 
|
| 28562 | 165  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
166  | 
"n mod m = of_nat (nat_of n mod nat_of m)"  | 
| 24999 | 167  | 
|
| 28562 | 168  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
169  | 
"n \<le> m \<longleftrightarrow> nat_of n \<le> nat_of m"  | 
| 24999 | 170  | 
|
| 28562 | 171  | 
definition [simp, code del]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
172  | 
"n < m \<longleftrightarrow> nat_of n < nat_of m"  | 
| 24999 | 173  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
174  | 
instance proof  | 
| 
33340
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
175  | 
qed (auto simp add: code_numeral left_distrib intro: mult_commute)  | 
| 
28708
 
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  | 
end  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
178  | 
|
| 
32069
 
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
 
haftmann 
parents: 
31998 
diff
changeset
 | 
179  | 
lemma zero_code_numeral_code [code, code_unfold]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
180  | 
"(0\<Colon>code_numeral) = Numeral0"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
181  | 
by (simp add: number_of_code_numeral_def Pls_def)  | 
| 
31998
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31377 
diff
changeset
 | 
182  | 
lemma [code_post]: "Numeral0 = (0\<Colon>code_numeral)"  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
183  | 
using zero_code_numeral_code ..  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
184  | 
|
| 
32069
 
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
 
haftmann 
parents: 
31998 
diff
changeset
 | 
185  | 
lemma one_code_numeral_code [code, code_unfold]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
186  | 
"(1\<Colon>code_numeral) = Numeral1"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
187  | 
by (simp add: number_of_code_numeral_def Pls_def Bit1_def)  | 
| 
31998
 
2c7a24f74db9
code attributes use common underscore convention
 
haftmann 
parents: 
31377 
diff
changeset
 | 
188  | 
lemma [code_post]: "Numeral1 = (1\<Colon>code_numeral)"  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
189  | 
using one_code_numeral_code ..  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
190  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
191  | 
lemma plus_code_numeral_code [code nbe]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
192  | 
"of_nat n + of_nat m = of_nat (n + m)"  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
193  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
194  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
195  | 
definition subtract_code_numeral :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
196  | 
[simp, code del]: "subtract_code_numeral = op -"  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
197  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
198  | 
lemma subtract_code_numeral_code [code nbe]:  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
199  | 
"subtract_code_numeral (of_nat n) (of_nat m) = of_nat (n - m)"  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
200  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
201  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
202  | 
lemma minus_code_numeral_code [code]:  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
203  | 
"n - m = subtract_code_numeral n m"  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
204  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
205  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
206  | 
lemma times_code_numeral_code [code nbe]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
207  | 
"of_nat n * of_nat m = of_nat (n * m)"  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
208  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
209  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
210  | 
lemma less_eq_code_numeral_code [code nbe]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
211  | 
"of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m"  | 
| 25767 | 212  | 
by simp  | 
| 24999 | 213  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
214  | 
lemma less_code_numeral_code [code nbe]:  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
215  | 
"of_nat n < of_nat m \<longleftrightarrow> n < m"  | 
| 25767 | 216  | 
by simp  | 
| 24999 | 217  | 
|
| 31266 | 218  | 
lemma code_numeral_zero_minus_one:  | 
219  | 
"(0::code_numeral) - 1 = 0"  | 
|
220  | 
by simp  | 
|
221  | 
||
222  | 
lemma Suc_code_numeral_minus_one:  | 
|
223  | 
"Suc_code_numeral n - 1 = n"  | 
|
224  | 
by simp  | 
|
| 26140 | 225  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
226  | 
lemma of_nat_code [code]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
227  | 
"of_nat = Nat.of_nat"  | 
| 25918 | 228  | 
proof  | 
229  | 
fix n :: nat  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
230  | 
have "Nat.of_nat n = of_nat n"  | 
| 25918 | 231  | 
by (induct n) simp_all  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
232  | 
then show "of_nat n = Nat.of_nat n"  | 
| 25918 | 233  | 
by (rule sym)  | 
234  | 
qed  | 
|
235  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
236  | 
lemma code_numeral_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1"  | 
| 25928 | 237  | 
by (cases i) auto  | 
238  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
239  | 
definition nat_of_aux :: "code_numeral \<Rightarrow> nat \<Rightarrow> nat" where  | 
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
240  | 
"nat_of_aux i n = nat_of i + n"  | 
| 25928 | 241  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
242  | 
lemma nat_of_aux_code [code]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
243  | 
"nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Suc n))"  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
244  | 
by (auto simp add: nat_of_aux_def code_numeral_not_eq_zero)  | 
| 25928 | 245  | 
|
| 
29815
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
246  | 
lemma nat_of_code [code]:  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
247  | 
"nat_of i = nat_of_aux i 0"  | 
| 
 
9e94b7078fa5
mandatory prefix for index conversion operations
 
haftmann 
parents: 
28708 
diff
changeset
 | 
248  | 
by (simp add: nat_of_aux_def)  | 
| 25918 | 249  | 
|
| 35687 | 250  | 
definition div_mod_code_numeral :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
251  | 
[code del]: "div_mod_code_numeral n m = (n div m, n mod m)"  | 
| 26009 | 252  | 
|
| 28562 | 253  | 
lemma [code]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
254  | 
"div_mod_code_numeral n m = (if m = 0 then (0, n) else (n div m, n mod m))"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
255  | 
unfolding div_mod_code_numeral_def by auto  | 
| 26009 | 256  | 
|
| 28562 | 257  | 
lemma [code]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
258  | 
"n div m = fst (div_mod_code_numeral n m)"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
259  | 
unfolding div_mod_code_numeral_def by simp  | 
| 26009 | 260  | 
|
| 28562 | 261  | 
lemma [code]:  | 
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
262  | 
"n mod m = snd (div_mod_code_numeral n m)"  | 
| 
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
263  | 
unfolding div_mod_code_numeral_def by simp  | 
| 26009 | 264  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
265  | 
definition int_of :: "code_numeral \<Rightarrow> int" where  | 
| 31192 | 266  | 
"int_of = Nat.of_nat o nat_of"  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
267  | 
|
| 31192 | 268  | 
lemma int_of_code [code]:  | 
269  | 
"int_of k = (if k = 0 then 0  | 
|
270  | 
else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"  | 
|
| 
33340
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
271  | 
proof -  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
272  | 
have "(nat_of k div 2) * 2 + nat_of k mod 2 = nat_of k"  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
273  | 
by (rule mod_div_equality)  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
274  | 
then have "int ((nat_of k div 2) * 2 + nat_of k mod 2) = int (nat_of k)"  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
275  | 
by simp  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
276  | 
then have "int (nat_of k) = int (nat_of k div 2) * 2 + int (nat_of k mod 2)"  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
277  | 
unfolding int_mult zadd_int [symmetric] by simp  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
278  | 
then show ?thesis by (auto simp add: int_of_def mult_ac)  | 
| 
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
279  | 
qed  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
280  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36049 
diff
changeset
 | 
281  | 
hide_const (open) of_nat nat_of int_of  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
282  | 
|
| 
36049
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
283  | 
subsubsection {* Lazy Evaluation of an indexed function *}
 | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
284  | 
|
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
285  | 
function iterate_upto :: "(code_numeral => 'a) => code_numeral => code_numeral => 'a Predicate.pred"  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
286  | 
where  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
287  | 
"iterate_upto f n m = Predicate.Seq (%u. if n > m then Predicate.Empty else Predicate.Insert (f n) (iterate_upto f (n + 1) m))"  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
288  | 
by pat_completeness auto  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
289  | 
|
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
290  | 
termination by (relation "measure (%(f, n, m). Code_Numeral.nat_of (m + 1 - n))") auto  | 
| 
 
0ce5b7a5c2fd
adding iterate_upto interface in compilations and iterate_upto functions in Isabelle theories for arithmetic setup of the predicate compiler
 
bulwahn 
parents: 
35687 
diff
changeset
 | 
291  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36049 
diff
changeset
 | 
292  | 
hide_const (open) iterate_upto  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
293  | 
|
| 28228 | 294  | 
subsection {* Code generator setup *}
 | 
| 24999 | 295  | 
|
| 25767 | 296  | 
text {* Implementation of indices by bounded integers *}
 | 
297  | 
||
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
298  | 
code_type code_numeral  | 
| 24999 | 299  | 
(SML "int")  | 
| 31377 | 300  | 
(OCaml "Big'_int.big'_int")  | 
| 25967 | 301  | 
(Haskell "Int")  | 
| 34886 | 302  | 
(Scala "Int")  | 
| 24999 | 303  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
304  | 
code_instance code_numeral :: eq  | 
| 24999 | 305  | 
(Haskell -)  | 
306  | 
||
307  | 
setup {*
 | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
308  | 
  fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
 | 
| 
34944
 
970e1466028d
code literals: distinguish numeral classes by different entries
 
haftmann 
parents: 
34917 
diff
changeset
 | 
309  | 
false Code_Printer.literal_naive_numeral) ["SML", "Haskell"]  | 
| 
34898
 
62d70417f8ce
allow individual printing of numerals during code serialization
 
haftmann 
parents: 
34886 
diff
changeset
 | 
310  | 
  #> Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
 | 
| 
34944
 
970e1466028d
code literals: distinguish numeral classes by different entries
 
haftmann 
parents: 
34917 
diff
changeset
 | 
311  | 
false Code_Printer.literal_numeral "OCaml"  | 
| 
34898
 
62d70417f8ce
allow individual printing of numerals during code serialization
 
haftmann 
parents: 
34886 
diff
changeset
 | 
312  | 
  #> Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
 | 
| 
34944
 
970e1466028d
code literals: distinguish numeral classes by different entries
 
haftmann 
parents: 
34917 
diff
changeset
 | 
313  | 
false Code_Printer.literal_naive_numeral "Scala"  | 
| 24999 | 314  | 
*}  | 
315  | 
||
| 25918 | 316  | 
code_reserved SML Int int  | 
| 34886 | 317  | 
code_reserved Scala Int  | 
| 24999 | 318  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
319  | 
code_const "op + \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"  | 
| 25928 | 320  | 
(SML "Int.+/ ((_),/ (_))")  | 
| 31377 | 321  | 
(OCaml "Big'_int.add'_big'_int")  | 
| 24999 | 322  | 
(Haskell infixl 6 "+")  | 
| 34886 | 323  | 
(Scala infixl 7 "+")  | 
| 24999 | 324  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
325  | 
code_const "subtract_code_numeral \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"  | 
| 25918 | 326  | 
(SML "Int.max/ (_/ -/ _,/ 0 : int)")  | 
| 31377 | 327  | 
(OCaml "Big'_int.max'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)/ Big'_int.zero'_big'_int")  | 
| 25918 | 328  | 
(Haskell "max/ (_/ -/ _)/ (0 :: Int)")  | 
| 34886 | 329  | 
(Scala "!(_/ -/ _).max(0)")  | 
| 24999 | 330  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
331  | 
code_const "op * \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"  | 
| 25928 | 332  | 
(SML "Int.*/ ((_),/ (_))")  | 
| 31377 | 333  | 
(OCaml "Big'_int.mult'_big'_int")  | 
| 24999 | 334  | 
(Haskell infixl 7 "*")  | 
| 34886 | 335  | 
(Scala infixl 8 "*")  | 
| 24999 | 336  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
337  | 
code_const div_mod_code_numeral  | 
| 
34898
 
62d70417f8ce
allow individual printing of numerals during code serialization
 
haftmann 
parents: 
34886 
diff
changeset
 | 
338  | 
(SML "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (n div m, n mod m))")  | 
| 
 
62d70417f8ce
allow individual printing of numerals during code serialization
 
haftmann 
parents: 
34886 
diff
changeset
 | 
339  | 
(OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")  | 
| 26009 | 340  | 
(Haskell "divMod")  | 
| 
34898
 
62d70417f8ce
allow individual printing of numerals during code serialization
 
haftmann 
parents: 
34886 
diff
changeset
 | 
341  | 
(Scala "!((n: Int) => (m: Int) =>/ if (m == 0)/ (0, n) else/ (n '/ m, n % m))")  | 
| 25928 | 342  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
343  | 
code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"  | 
| 24999 | 344  | 
(SML "!((_ : Int.int) = _)")  | 
| 31377 | 345  | 
(OCaml "Big'_int.eq'_big'_int")  | 
| 24999 | 346  | 
(Haskell infixl 4 "==")  | 
| 34886 | 347  | 
(Scala infixl 5 "==")  | 
| 24999 | 348  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
349  | 
code_const "op \<le> \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"  | 
| 25928 | 350  | 
(SML "Int.<=/ ((_),/ (_))")  | 
| 31377 | 351  | 
(OCaml "Big'_int.le'_big'_int")  | 
| 24999 | 352  | 
(Haskell infix 4 "<=")  | 
| 
34898
 
62d70417f8ce
allow individual printing of numerals during code serialization
 
haftmann 
parents: 
34886 
diff
changeset
 | 
353  | 
(Scala infixl 4 "<=")  | 
| 24999 | 354  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
355  | 
code_const "op < \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"  | 
| 25928 | 356  | 
(SML "Int.</ ((_),/ (_))")  | 
| 31377 | 357  | 
(OCaml "Big'_int.lt'_big'_int")  | 
| 24999 | 358  | 
(Haskell infix 4 "<")  | 
| 
34898
 
62d70417f8ce
allow individual printing of numerals during code serialization
 
haftmann 
parents: 
34886 
diff
changeset
 | 
359  | 
(Scala infixl 4 "<")  | 
| 24999 | 360  | 
|
361  | 
end  |