author | noschinl |
Mon, 12 Mar 2012 15:12:22 +0100 | |
changeset 46883 | eec472dae593 |
parent 46664 | 1f6c140f9c72 |
child 47108 | 2a1953f0d20d |
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" |
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
44821
diff
changeset
|
17 |
morphisms nat_of of_nat .. |
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 |
||
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
74 |
definition Suc where [simp]: |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
75 |
"Suc k = of_nat (Nat.Suc (nat_of k))" |
26140 | 76 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
77 |
rep_datatype "0 \<Colon> code_numeral" Suc |
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 |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
82 |
assume "\<And>k. P k \<Longrightarrow> P (Suc k)" |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
83 |
then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc (of_nat n))" . |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
84 |
then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Nat.Suc n))" by simp |
29815
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]: |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
94 |
"k \<noteq> of_nat 0 \<Longrightarrow> nat_of k - Nat.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)" |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
102 |
by (induct k rule: code_numeral.induct) (simp_all del: zero_code_numeral_def Suc_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" |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
112 |
by (induct k) (simp_all del: zero_code_numeral_def Suc_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]: |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37958
diff
changeset
|
118 |
"HOL.equal k l \<longleftrightarrow> HOL.equal (nat_of k) (nat_of l)" |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37958
diff
changeset
|
119 |
by (cases k, cases l) (simp add: equal) |
24999 | 120 |
|
28351 | 121 |
lemma [code nbe]: |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37958
diff
changeset
|
122 |
"HOL.equal (k::code_numeral) k \<longleftrightarrow> True" |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37958
diff
changeset
|
123 |
by (rule equal_refl) |
28351 | 124 |
|
24999 | 125 |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
126 |
subsection {* Code numerals as datatype of ints *} |
25767 | 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 |
|
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
179 |
lemma zero_code_numeral_code [code]: |
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) |
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
182 |
|
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
183 |
lemma [code_abbrev]: "Numeral0 = (0\<Colon>code_numeral)" |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
184 |
using zero_code_numeral_code .. |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
185 |
|
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
186 |
lemma one_code_numeral_code [code]: |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
187 |
"(1\<Colon>code_numeral) = Numeral1" |
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
188 |
by (simp add: number_of_code_numeral_def Pls_def Bit1_def) |
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
189 |
|
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
190 |
lemma [code_abbrev]: "Numeral1 = (1\<Colon>code_numeral)" |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
191 |
using one_code_numeral_code .. |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
192 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
193 |
lemma plus_code_numeral_code [code nbe]: |
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
194 |
"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
|
195 |
by simp |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
196 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
197 |
definition subtract :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
198 |
[simp]: "subtract = minus" |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
199 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
200 |
lemma subtract_code [code nbe]: |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
201 |
"subtract (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
|
202 |
by simp |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
203 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
204 |
lemma minus_code_numeral_code [code]: |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
205 |
"minus = subtract" |
28708
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 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
208 |
lemma times_code_numeral_code [code nbe]: |
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
209 |
"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
|
210 |
by simp |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
211 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
212 |
lemma less_eq_code_numeral_code [code nbe]: |
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
213 |
"of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m" |
25767 | 214 |
by simp |
24999 | 215 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
216 |
lemma less_code_numeral_code [code nbe]: |
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
217 |
"of_nat n < of_nat m \<longleftrightarrow> n < m" |
25767 | 218 |
by simp |
24999 | 219 |
|
31266 | 220 |
lemma code_numeral_zero_minus_one: |
221 |
"(0::code_numeral) - 1 = 0" |
|
222 |
by simp |
|
223 |
||
224 |
lemma Suc_code_numeral_minus_one: |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
225 |
"Suc n - 1 = n" |
31266 | 226 |
by simp |
26140 | 227 |
|
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
228 |
lemma of_nat_code [code]: |
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
229 |
"of_nat = Nat.of_nat" |
25918 | 230 |
proof |
231 |
fix n :: nat |
|
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
232 |
have "Nat.of_nat n = of_nat n" |
25918 | 233 |
by (induct n) simp_all |
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
234 |
then show "of_nat n = Nat.of_nat n" |
25918 | 235 |
by (rule sym) |
236 |
qed |
|
237 |
||
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
238 |
lemma code_numeral_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1" |
25928 | 239 |
by (cases i) auto |
240 |
||
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
241 |
definition nat_of_aux :: "code_numeral \<Rightarrow> nat \<Rightarrow> nat" where |
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
242 |
"nat_of_aux i n = nat_of i + n" |
25928 | 243 |
|
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
244 |
lemma nat_of_aux_code [code]: |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
245 |
"nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Nat.Suc n))" |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
246 |
by (auto simp add: nat_of_aux_def code_numeral_not_eq_zero) |
25928 | 247 |
|
29815
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
248 |
lemma nat_of_code [code]: |
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
249 |
"nat_of i = nat_of_aux i 0" |
9e94b7078fa5
mandatory prefix for index conversion operations
haftmann
parents:
28708
diff
changeset
|
250 |
by (simp add: nat_of_aux_def) |
25918 | 251 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
252 |
definition div_mod :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
253 |
[code del]: "div_mod n m = (n div m, n mod m)" |
26009 | 254 |
|
28562 | 255 |
lemma [code]: |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
256 |
"div_mod n m = (if m = 0 then (0, n) else (n div m, n mod m))" |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
257 |
unfolding div_mod_def by auto |
26009 | 258 |
|
28562 | 259 |
lemma [code]: |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
260 |
"n div m = fst (div_mod n m)" |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
261 |
unfolding div_mod_def by simp |
26009 | 262 |
|
28562 | 263 |
lemma [code]: |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
264 |
"n mod m = snd (div_mod n m)" |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
265 |
unfolding div_mod_def by simp |
26009 | 266 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
267 |
definition int_of :: "code_numeral \<Rightarrow> int" where |
31192 | 268 |
"int_of = Nat.of_nat o nat_of" |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
269 |
|
31192 | 270 |
lemma int_of_code [code]: |
271 |
"int_of k = (if k = 0 then 0 |
|
272 |
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
|
273 |
proof - |
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents:
33318
diff
changeset
|
274 |
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
|
275 |
by (rule mod_div_equality) |
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents:
33318
diff
changeset
|
276 |
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
|
277 |
by simp |
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents:
33318
diff
changeset
|
278 |
then have "int (nat_of k) = int (nat_of k div 2) * 2 + int (nat_of k mod 2)" |
44821 | 279 |
unfolding of_nat_mult of_nat_add by simp |
33340
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents:
33318
diff
changeset
|
280 |
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
|
281 |
qed |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
282 |
|
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
283 |
|
46664
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
haftmann
parents:
46638
diff
changeset
|
284 |
hide_const (open) of_nat nat_of Suc subtract int_of |
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
285 |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
286 |
|
28228 | 287 |
subsection {* Code generator setup *} |
24999 | 288 |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
289 |
text {* Implementation of code numerals by bounded integers *} |
25767 | 290 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
291 |
code_type code_numeral |
24999 | 292 |
(SML "int") |
31377 | 293 |
(OCaml "Big'_int.big'_int") |
37947 | 294 |
(Haskell "Integer") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
295 |
(Scala "BigInt") |
24999 | 296 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37958
diff
changeset
|
297 |
code_instance code_numeral :: equal |
24999 | 298 |
(Haskell -) |
299 |
||
300 |
setup {* |
|
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
301 |
Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} |
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
302 |
false Code_Printer.literal_naive_numeral "SML" |
37947 | 303 |
#> fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
304 |
false Code_Printer.literal_numeral) ["OCaml", "Haskell", "Scala"] |
24999 | 305 |
*} |
306 |
||
25918 | 307 |
code_reserved SML Int int |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
308 |
code_reserved Eval Integer |
24999 | 309 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
310 |
code_const "plus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" |
25928 | 311 |
(SML "Int.+/ ((_),/ (_))") |
31377 | 312 |
(OCaml "Big'_int.add'_big'_int") |
24999 | 313 |
(Haskell infixl 6 "+") |
34886 | 314 |
(Scala infixl 7 "+") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
315 |
(Eval infixl 8 "+") |
24999 | 316 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
317 |
code_const "Code_Numeral.subtract \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" |
25918 | 318 |
(SML "Int.max/ (_/ -/ _,/ 0 : int)") |
31377 | 319 |
(OCaml "Big'_int.max'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)/ Big'_int.zero'_big'_int") |
37947 | 320 |
(Haskell "max/ (_/ -/ _)/ (0 :: Integer)") |
34886 | 321 |
(Scala "!(_/ -/ _).max(0)") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
322 |
(Eval "Integer.max/ (_/ -/ _)/ 0") |
24999 | 323 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
324 |
code_const "times \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" |
25928 | 325 |
(SML "Int.*/ ((_),/ (_))") |
31377 | 326 |
(OCaml "Big'_int.mult'_big'_int") |
24999 | 327 |
(Haskell infixl 7 "*") |
34886 | 328 |
(Scala infixl 8 "*") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
329 |
(Eval infixl 8 "*") |
24999 | 330 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
331 |
code_const Code_Numeral.div_mod |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
332 |
(SML "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Int.div (n, m), Int.mod (n, m)))") |
34898
62d70417f8ce
allow individual printing of numerals during code serialization
haftmann
parents:
34886
diff
changeset
|
333 |
(OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") |
26009 | 334 |
(Haskell "divMod") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
335 |
(Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") |
39818 | 336 |
(Eval "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Integer.div'_mod n m))") |
25928 | 337 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
37958
diff
changeset
|
338 |
code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" |
24999 | 339 |
(SML "!((_ : Int.int) = _)") |
31377 | 340 |
(OCaml "Big'_int.eq'_big'_int") |
39272 | 341 |
(Haskell infix 4 "==") |
34886 | 342 |
(Scala infixl 5 "==") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
343 |
(Eval "!((_ : int) = _)") |
24999 | 344 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
345 |
code_const "less_eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" |
25928 | 346 |
(SML "Int.<=/ ((_),/ (_))") |
31377 | 347 |
(OCaml "Big'_int.le'_big'_int") |
24999 | 348 |
(Haskell infix 4 "<=") |
34898
62d70417f8ce
allow individual printing of numerals during code serialization
haftmann
parents:
34886
diff
changeset
|
349 |
(Scala infixl 4 "<=") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
350 |
(Eval infixl 6 "<=") |
24999 | 351 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
352 |
code_const "less \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool" |
25928 | 353 |
(SML "Int.</ ((_),/ (_))") |
31377 | 354 |
(OCaml "Big'_int.lt'_big'_int") |
24999 | 355 |
(Haskell infix 4 "<") |
34898
62d70417f8ce
allow individual printing of numerals during code serialization
haftmann
parents:
34886
diff
changeset
|
356 |
(Scala infixl 4 "<") |
37958
9728342bcd56
another refinement chapter in the neverending numeral story
haftmann
parents:
37947
diff
changeset
|
357 |
(Eval infixl 6 "<") |
24999 | 358 |
|
46547
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
359 |
code_modulename SML |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
360 |
Code_Numeral Arith |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
361 |
|
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
362 |
code_modulename OCaml |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
363 |
Code_Numeral Arith |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
364 |
|
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
365 |
code_modulename Haskell |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
366 |
Code_Numeral Arith |
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
huffman
parents:
46028
diff
changeset
|
367 |
|
24999 | 368 |
end |
46664
1f6c140f9c72
moved predicate relations and conversion rules between set and predicate relations from Predicate.thy to Relation.thy; moved Predicate.thy upwards in theory hierarchy
haftmann
parents:
46638
diff
changeset
|
369 |