| author | wenzelm | 
| Sat, 01 Jun 2024 15:03:13 +0200 | |
| changeset 80229 | 5e32da8238e1 | 
| parent 80088 | 5afbf04418ec | 
| child 81639 | 7fa796a773a5 | 
| permissions | -rw-r--r-- | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1  | 
(* Title: HOL/Code_Numeral.thy  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
2  | 
Author: Florian Haftmann, TU Muenchen  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
3  | 
*)  | 
| 24999 | 4  | 
|
| 60758 | 5  | 
section \<open>Numeric types for code generation onto target language numerals only\<close>  | 
| 24999 | 6  | 
|
| 
31205
 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 
haftmann 
parents: 
31203 
diff
changeset
 | 
7  | 
theory Code_Numeral  | 
| 78669 | 8  | 
imports Lifting Bit_Operations  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
9  | 
begin  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
10  | 
|
| 60758 | 11  | 
subsection \<open>Type of target language integers\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
12  | 
|
| 61076 | 13  | 
typedef integer = "UNIV :: int set"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
14  | 
morphisms int_of_integer integer_of_int ..  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
15  | 
|
| 
59487
 
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
 
haftmann 
parents: 
58889 
diff
changeset
 | 
16  | 
setup_lifting type_definition_integer  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
17  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
18  | 
lemma integer_eq_iff:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
19  | 
"k = l \<longleftrightarrow> int_of_integer k = int_of_integer l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
20  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
21  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
22  | 
lemma integer_eqI:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
23  | 
"int_of_integer k = int_of_integer l \<Longrightarrow> k = l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
24  | 
using integer_eq_iff [of k l] by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
25  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
26  | 
lemma int_of_integer_integer_of_int [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
27  | 
"int_of_integer (integer_of_int k) = k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
28  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
29  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
30  | 
lemma integer_of_int_int_of_integer [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
31  | 
"integer_of_int (int_of_integer k) = k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
32  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
33  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
34  | 
instantiation integer :: ring_1  | 
| 24999 | 35  | 
begin  | 
36  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
37  | 
lift_definition zero_integer :: integer  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
38  | 
is "0 :: int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
39  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
40  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
41  | 
declare zero_integer.rep_eq [simp]  | 
| 24999 | 42  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
43  | 
lift_definition one_integer :: integer  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
44  | 
is "1 :: int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
45  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
46  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
47  | 
declare one_integer.rep_eq [simp]  | 
| 24999 | 48  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
49  | 
lift_definition plus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
50  | 
is "plus :: int \<Rightarrow> int \<Rightarrow> int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
51  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
52  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
53  | 
declare plus_integer.rep_eq [simp]  | 
| 24999 | 54  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
55  | 
lift_definition uminus_integer :: "integer \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
56  | 
is "uminus :: int \<Rightarrow> int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
57  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
58  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
59  | 
declare uminus_integer.rep_eq [simp]  | 
| 24999 | 60  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
61  | 
lift_definition minus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
62  | 
is "minus :: int \<Rightarrow> int \<Rightarrow> int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
63  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
64  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
65  | 
declare minus_integer.rep_eq [simp]  | 
| 24999 | 66  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
67  | 
lift_definition times_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
68  | 
is "times :: int \<Rightarrow> int \<Rightarrow> int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
69  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
70  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
71  | 
declare times_integer.rep_eq [simp]  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
72  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
73  | 
instance proof  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
74  | 
qed (transfer, simp add: algebra_simps)+  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
75  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
76  | 
end  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
77  | 
|
| 
64241
 
430d74089d4d
transfer rules for divides relation on integer and natural
 
haftmann 
parents: 
64178 
diff
changeset
 | 
78  | 
instance integer :: Rings.dvd ..  | 
| 
 
430d74089d4d
transfer rules for divides relation on integer and natural
 
haftmann 
parents: 
64178 
diff
changeset
 | 
79  | 
|
| 70927 | 80  | 
context  | 
81  | 
includes lifting_syntax  | 
|
82  | 
notes transfer_rule_numeral [transfer_rule]  | 
|
83  | 
begin  | 
|
| 
64241
 
430d74089d4d
transfer rules for divides relation on integer and natural
 
haftmann 
parents: 
64178 
diff
changeset
 | 
84  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
85  | 
lemma [transfer_rule]:  | 
| 70927 | 86  | 
"(pcr_integer ===> pcr_integer ===> (\<longleftrightarrow>)) (dvd) (dvd)"  | 
87  | 
by (unfold dvd_def) transfer_prover  | 
|
88  | 
||
89  | 
lemma [transfer_rule]:  | 
|
90  | 
"((\<longleftrightarrow>) ===> pcr_integer) of_bool of_bool"  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
91  | 
by (unfold of_bool_def) transfer_prover  | 
| 68010 | 92  | 
|
93  | 
lemma [transfer_rule]:  | 
|
| 70927 | 94  | 
"((=) ===> pcr_integer) int of_nat"  | 
| 64178 | 95  | 
by (rule transfer_rule_of_nat) transfer_prover+  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
96  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
97  | 
lemma [transfer_rule]:  | 
| 70927 | 98  | 
"((=) ===> pcr_integer) (\<lambda>k. k) of_int"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
99  | 
proof -  | 
| 70927 | 100  | 
have "((=) ===> pcr_integer) of_int of_int"  | 
| 64178 | 101  | 
by (rule transfer_rule_of_int) transfer_prover+  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
102  | 
then show ?thesis by (simp add: id_def)  | 
| 24999 | 103  | 
qed  | 
104  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
105  | 
lemma [transfer_rule]:  | 
| 70927 | 106  | 
"((=) ===> pcr_integer) numeral numeral"  | 
107  | 
by transfer_prover  | 
|
| 26140 | 108  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
109  | 
lemma [transfer_rule]:  | 
| 70927 | 110  | 
"((=) ===> (=) ===> pcr_integer) Num.sub Num.sub"  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
111  | 
by (unfold Num.sub_def) transfer_prover  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
112  | 
|
| 68010 | 113  | 
lemma [transfer_rule]:  | 
| 70927 | 114  | 
"(pcr_integer ===> (=) ===> pcr_integer) (^) (^)"  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
115  | 
by (unfold power_def) transfer_prover  | 
| 68010 | 116  | 
|
| 70927 | 117  | 
end  | 
118  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
119  | 
lemma int_of_integer_of_nat [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
120  | 
"int_of_integer (of_nat n) = of_nat n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
121  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
122  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
123  | 
lift_definition integer_of_nat :: "nat \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
124  | 
is "of_nat :: nat \<Rightarrow> int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
125  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
126  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
127  | 
lemma integer_of_nat_eq_of_nat [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
128  | 
"integer_of_nat = of_nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
129  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
130  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
131  | 
lemma int_of_integer_integer_of_nat [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
132  | 
"int_of_integer (integer_of_nat n) = of_nat n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
133  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
134  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
135  | 
lift_definition nat_of_integer :: "integer \<Rightarrow> nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
136  | 
is Int.nat  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
137  | 
.  | 
| 26140 | 138  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
139  | 
lemma nat_of_integer_of_nat [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
140  | 
"nat_of_integer (of_nat n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
141  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
142  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
143  | 
lemma int_of_integer_of_int [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
144  | 
"int_of_integer (of_int k) = k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
145  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
146  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
147  | 
lemma nat_of_integer_integer_of_nat [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
148  | 
"nat_of_integer (integer_of_nat n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
149  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
150  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
151  | 
lemma integer_of_int_eq_of_int [simp, code_abbrev]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
152  | 
"integer_of_int = of_int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
153  | 
by transfer (simp add: fun_eq_iff)  | 
| 26140 | 154  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
155  | 
lemma of_int_integer_of [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
156  | 
"of_int (int_of_integer k) = (k :: integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
157  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
158  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
159  | 
lemma int_of_integer_numeral [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
160  | 
"int_of_integer (numeral k) = numeral k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
161  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
162  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
163  | 
lemma int_of_integer_sub [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
164  | 
"int_of_integer (Num.sub k l) = Num.sub k l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
165  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
166  | 
|
| 66190 | 167  | 
definition integer_of_num :: "num \<Rightarrow> integer"  | 
168  | 
where [simp]: "integer_of_num = numeral"  | 
|
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
169  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
170  | 
lemma integer_of_num [code]:  | 
| 66190 | 171  | 
"integer_of_num Num.One = 1"  | 
172  | 
"integer_of_num (Num.Bit0 n) = (let k = integer_of_num n in k + k)"  | 
|
173  | 
"integer_of_num (Num.Bit1 n) = (let k = integer_of_num n in k + k + 1)"  | 
|
174  | 
by (simp_all only: integer_of_num_def numeral.simps Let_def)  | 
|
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
175  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
176  | 
lemma integer_of_num_triv:  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
177  | 
"integer_of_num Num.One = 1"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
178  | 
"integer_of_num (Num.Bit0 Num.One) = 2"  | 
| 66190 | 179  | 
by simp_all  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
180  | 
|
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
181  | 
instantiation integer :: equal  | 
| 26140 | 182  | 
begin  | 
183  | 
||
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
184  | 
lift_definition equal_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> bool\<close>  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
185  | 
is \<open>HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool\<close>  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
186  | 
.  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
187  | 
|
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
188  | 
instance  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
189  | 
by (standard; transfer) (fact equal_eq)  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
190  | 
|
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
191  | 
end  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
192  | 
|
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
193  | 
instantiation integer :: linordered_idom  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
194  | 
begin  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
195  | 
|
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
196  | 
lift_definition abs_integer :: \<open>integer \<Rightarrow> integer\<close>  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
197  | 
is \<open>abs :: int \<Rightarrow> int\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
198  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
199  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
200  | 
declare abs_integer.rep_eq [simp]  | 
| 26140 | 201  | 
|
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
202  | 
lift_definition sgn_integer :: \<open>integer \<Rightarrow> integer\<close>  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
203  | 
is \<open>sgn :: int \<Rightarrow> int\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
204  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
205  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
206  | 
declare sgn_integer.rep_eq [simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
207  | 
|
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
208  | 
lift_definition less_eq_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> bool\<close>  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
209  | 
is \<open>less_eq :: int \<Rightarrow> int \<Rightarrow> bool\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
210  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
211  | 
|
| 69946 | 212  | 
lemma integer_less_eq_iff:  | 
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
213  | 
\<open>k \<le> l \<longleftrightarrow> int_of_integer k \<le> int_of_integer l\<close>  | 
| 69946 | 214  | 
by (fact less_eq_integer.rep_eq)  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
215  | 
|
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
216  | 
lift_definition less_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> bool\<close>  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
217  | 
is \<open>less :: int \<Rightarrow> int \<Rightarrow> bool\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
218  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
219  | 
|
| 69946 | 220  | 
lemma integer_less_iff:  | 
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
221  | 
\<open>k < l \<longleftrightarrow> int_of_integer k < int_of_integer l\<close>  | 
| 69946 | 222  | 
by (fact less_integer.rep_eq)  | 
223  | 
||
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
224  | 
instance  | 
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
225  | 
by (standard; transfer)  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
226  | 
(simp_all add: algebra_simps less_le_not_le [symmetric] mult_strict_right_mono linear)  | 
| 26140 | 227  | 
|
228  | 
end  | 
|
229  | 
||
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
230  | 
instance integer :: discrete_linordered_semidom  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
231  | 
by (standard; transfer)  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
232  | 
(fact less_iff_succ_less_eq)  | 
| 
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
233  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
234  | 
context  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
235  | 
includes lifting_syntax  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
236  | 
begin  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
237  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
238  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
239  | 
\<open>(pcr_integer ===> pcr_integer ===> pcr_integer) min min\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
240  | 
by (unfold min_def) transfer_prover  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
241  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
242  | 
lemma [transfer_rule]:  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
243  | 
\<open>(pcr_integer ===> pcr_integer ===> pcr_integer) max max\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
244  | 
by (unfold max_def) transfer_prover  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
245  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
246  | 
end  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
247  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
248  | 
lemma int_of_integer_min [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
249  | 
"int_of_integer (min k l) = min (int_of_integer k) (int_of_integer l)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
250  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
251  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
252  | 
lemma int_of_integer_max [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
253  | 
"int_of_integer (max k l) = max (int_of_integer k) (int_of_integer l)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
254  | 
by transfer rule  | 
| 26140 | 255  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
256  | 
lemma nat_of_integer_non_positive [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
257  | 
"k \<le> 0 \<Longrightarrow> nat_of_integer k = 0"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
258  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
259  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
260  | 
lemma of_nat_of_integer [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
261  | 
"of_nat (nat_of_integer k) = max 0 k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
262  | 
by transfer auto  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
263  | 
|
| 66817 | 264  | 
instantiation integer :: unique_euclidean_ring  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
265  | 
begin  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
266  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
267  | 
lift_definition divide_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
268  | 
is "divide :: int \<Rightarrow> int \<Rightarrow> int"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
269  | 
.  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
270  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
271  | 
declare divide_integer.rep_eq [simp]  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
272  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
273  | 
lift_definition modulo_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
274  | 
is "modulo :: int \<Rightarrow> int \<Rightarrow> int"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
275  | 
.  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
276  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
277  | 
declare modulo_integer.rep_eq [simp]  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
278  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
279  | 
lift_definition euclidean_size_integer :: "integer \<Rightarrow> nat"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
280  | 
is "euclidean_size :: int \<Rightarrow> nat"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
281  | 
.  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
282  | 
|
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
283  | 
declare euclidean_size_integer.rep_eq [simp]  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
284  | 
|
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
285  | 
lift_definition division_segment_integer :: "integer \<Rightarrow> integer"  | 
| 
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
286  | 
is "division_segment :: int \<Rightarrow> int"  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
287  | 
.  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
288  | 
|
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
289  | 
declare division_segment_integer.rep_eq [simp]  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
290  | 
|
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
291  | 
instance  | 
| 76053 | 292  | 
apply (standard; transfer)  | 
293  | 
apply (use mult_le_mono2 [of 1] in \<open>auto simp add: sgn_mult_abs abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib  | 
|
294  | 
division_segment_mult division_segment_mod\<close>)  | 
|
295  | 
apply (simp add: division_segment_int_def split: if_splits)  | 
|
296  | 
done  | 
|
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
297  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
298  | 
end  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
299  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
300  | 
lemma [code]:  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
301  | 
"euclidean_size = nat_of_integer \<circ> abs"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
302  | 
by (simp add: fun_eq_iff nat_of_integer.rep_eq)  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
303  | 
|
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
304  | 
lemma [code]:  | 
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
305  | 
"division_segment (k :: integer) = (if k \<ge> 0 then 1 else - 1)"  | 
| 
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
306  | 
by transfer (simp add: division_segment_int_def)  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
307  | 
|
| 
78937
 
5e6b195eee83
slightly less technical formulation of very specific type class
 
haftmann 
parents: 
78935 
diff
changeset
 | 
308  | 
instance integer :: linordered_euclidean_semiring  | 
| 66839 | 309  | 
by (standard; transfer) (simp_all add: of_nat_div division_segment_int_def)  | 
| 66815 | 310  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
311  | 
instantiation integer :: ring_bit_operations  | 
| 71094 | 312  | 
begin  | 
313  | 
||
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
314  | 
lift_definition bit_integer :: \<open>integer \<Rightarrow> nat \<Rightarrow> bool\<close>  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
315  | 
is bit .  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
316  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
317  | 
lift_definition not_integer :: \<open>integer \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
318  | 
is not .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
319  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
320  | 
lift_definition and_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
321  | 
is \<open>and\<close> .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
322  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
323  | 
lift_definition or_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
324  | 
is or .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
325  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
326  | 
lift_definition xor_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
327  | 
is xor .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
328  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
329  | 
lift_definition mask_integer :: \<open>nat \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
330  | 
is mask .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
331  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
332  | 
lift_definition set_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
333  | 
is set_bit .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
334  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
335  | 
lift_definition unset_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
336  | 
is unset_bit .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
337  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
338  | 
lift_definition flip_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
339  | 
is flip_bit .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
340  | 
|
| 71094 | 341  | 
lift_definition push_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
342  | 
is push_bit .  | 
| 71094 | 343  | 
|
344  | 
lift_definition drop_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
|
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
345  | 
is drop_bit .  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
346  | 
|
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
347  | 
lift_definition take_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
348  | 
is take_bit .  | 
| 71094 | 349  | 
|
350  | 
instance by (standard; transfer)  | 
|
| 
79541
 
4f40225936d1
common type class for trivial properties on div/mod
 
haftmann 
parents: 
79531 
diff
changeset
 | 
351  | 
(fact bit_induct div_by_0 div_by_1 div_0 even_half_succ_eq  | 
| 
79673
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79588 
diff
changeset
 | 
352  | 
half_div_exp_eq even_double_div_exp_iff bits_mod_div_trivial  | 
| 
79531
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
353  | 
bit_iff_odd push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod  | 
| 
79072
 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 
haftmann 
parents: 
79031 
diff
changeset
 | 
354  | 
and_rec or_rec xor_rec mask_eq_exp_minus_1  | 
| 79489 | 355  | 
set_bit_eq_or unset_bit_eq_or_xor flip_bit_eq_xor not_eq_complement)+  | 
| 71094 | 356  | 
|
357  | 
end  | 
|
| 68010 | 358  | 
|
| 
79531
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
359  | 
|
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
360  | 
|
| 78955 | 361  | 
instance integer :: linordered_euclidean_semiring_bit_operations ..  | 
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
362  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
363  | 
context  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
364  | 
includes bit_operations_syntax  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
365  | 
begin  | 
| 71094 | 366  | 
|
367  | 
lemma [code]:  | 
|
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
368  | 
\<open>bit k n \<longleftrightarrow> odd (drop_bit n k)\<close>  | 
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
369  | 
\<open>NOT k = - k - 1\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
370  | 
\<open>mask n = 2 ^ n - (1 :: integer)\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
371  | 
\<open>set_bit n k = k OR push_bit n 1\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
372  | 
\<open>unset_bit n k = k AND NOT (push_bit n 1)\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
373  | 
\<open>flip_bit n k = k XOR push_bit n 1\<close>  | 
| 71094 | 374  | 
\<open>push_bit n k = k * 2 ^ n\<close>  | 
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
375  | 
\<open>drop_bit n k = k div 2 ^ n\<close>  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
376  | 
\<open>take_bit n k = k mod 2 ^ n\<close> for k :: integer  | 
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
377  | 
by (fact bit_iff_odd_drop_bit not_eq_complement mask_eq_exp_minus_1  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
378  | 
set_bit_eq_or unset_bit_eq_and_not flip_bit_eq_xor push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
379  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
380  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
381  | 
\<open>k AND l = (if k = 0 \<or> l = 0 then 0 else if k = - 1 then l else if l = - 1 then k  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
382  | 
else (k mod 2) * (l mod 2) + 2 * ((k div 2) AND (l div 2)))\<close> for k l :: integer  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
383  | 
by transfer (fact and_int_unfold)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
384  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
385  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
386  | 
\<open>k OR l = (if k = - 1 \<or> l = - 1 then - 1 else if k = 0 then l else if l = 0 then k  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
387  | 
else max (k mod 2) (l mod 2) + 2 * ((k div 2) OR (l div 2)))\<close> for k l :: integer  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
388  | 
by transfer (fact or_int_unfold)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
389  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
390  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
391  | 
\<open>k XOR l = (if k = - 1 then NOT l else if l = - 1 then NOT k else if k = 0 then l else if l = 0 then k  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
392  | 
else \<bar>k mod 2 - l mod 2\<bar> + 2 * ((k div 2) XOR (l div 2)))\<close> for k l :: integer  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
393  | 
by transfer (fact xor_int_unfold)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
394  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
395  | 
end  | 
| 68010 | 396  | 
|
| 
78937
 
5e6b195eee83
slightly less technical formulation of very specific type class
 
haftmann 
parents: 
78935 
diff
changeset
 | 
397  | 
instantiation integer :: linordered_euclidean_semiring_division  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
398  | 
begin  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
399  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
400  | 
definition divmod_integer :: "num \<Rightarrow> num \<Rightarrow> integer \<times> integer"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
401  | 
where  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
402  | 
divmod_integer'_def: "divmod_integer m n = (numeral m div numeral n, numeral m mod numeral n)"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
403  | 
|
| 75883 | 404  | 
definition divmod_step_integer :: "integer \<Rightarrow> integer \<times> integer \<Rightarrow> integer \<times> integer"  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
405  | 
where  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
406  | 
"divmod_step_integer l qr = (let (q, r) = qr  | 
| 75936 | 407  | 
in if \<bar>l\<bar> \<le> \<bar>r\<bar> then (2 * q + 1, r - l)  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
408  | 
else (2 * q, r))"  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
409  | 
|
| 75936 | 410  | 
instance by standard  | 
411  | 
(auto simp add: divmod_integer'_def divmod_step_integer_def integer_less_eq_iff)  | 
|
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
412  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
413  | 
end  | 
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
414  | 
|
| 
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
415  | 
declare divmod_algorithm_code [where ?'a = integer,  | 
| 69946 | 416  | 
folded integer_of_num_def, unfolded integer_of_num_triv,  | 
| 
61275
 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 
haftmann 
parents: 
61274 
diff
changeset
 | 
417  | 
code]  | 
| 
53069
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
52435 
diff
changeset
 | 
418  | 
|
| 
55427
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
419  | 
lemma integer_of_nat_0: "integer_of_nat 0 = 0"  | 
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
420  | 
by transfer simp  | 
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
421  | 
|
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
422  | 
lemma integer_of_nat_1: "integer_of_nat 1 = 1"  | 
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
423  | 
by transfer simp  | 
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
424  | 
|
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
425  | 
lemma integer_of_nat_numeral:  | 
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
426  | 
"integer_of_nat (numeral n) = numeral n"  | 
| 
 
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
 
Andreas Lochbihler 
parents: 
54489 
diff
changeset
 | 
427  | 
by transfer simp  | 
| 26140 | 428  | 
|
| 68010 | 429  | 
|
| 60758 | 430  | 
subsection \<open>Code theorems for target language integers\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
431  | 
|
| 60758 | 432  | 
text \<open>Constructors\<close>  | 
| 26140 | 433  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
434  | 
definition Pos :: "num \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
435  | 
where  | 
| 
61274
 
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
 
haftmann 
parents: 
61076 
diff
changeset
 | 
436  | 
[simp, code_post]: "Pos = numeral"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
437  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
438  | 
context  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
439  | 
includes lifting_syntax  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
440  | 
begin  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
441  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
442  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
443  | 
\<open>((=) ===> pcr_integer) numeral Pos\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
444  | 
by simp transfer_prover  | 
| 
30245
 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 
haftmann 
parents: 
29823 
diff
changeset
 | 
445  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
446  | 
end  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
447  | 
|
| 
61274
 
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
 
haftmann 
parents: 
61076 
diff
changeset
 | 
448  | 
lemma Pos_fold [code_unfold]:  | 
| 
 
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
 
haftmann 
parents: 
61076 
diff
changeset
 | 
449  | 
"numeral Num.One = Pos Num.One"  | 
| 
 
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
 
haftmann 
parents: 
61076 
diff
changeset
 | 
450  | 
"numeral (Num.Bit0 k) = Pos (Num.Bit0 k)"  | 
| 
 
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
 
haftmann 
parents: 
61076 
diff
changeset
 | 
451  | 
"numeral (Num.Bit1 k) = Pos (Num.Bit1 k)"  | 
| 
 
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
 
haftmann 
parents: 
61076 
diff
changeset
 | 
452  | 
by simp_all  | 
| 
 
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
 
haftmann 
parents: 
61076 
diff
changeset
 | 
453  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
454  | 
definition Neg :: "num \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
455  | 
where  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
53069 
diff
changeset
 | 
456  | 
[simp, code_abbrev]: "Neg n = - Pos n"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
457  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
458  | 
context  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
459  | 
includes lifting_syntax  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
460  | 
begin  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
461  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
462  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
463  | 
\<open>((=) ===> pcr_integer) (\<lambda>n. - numeral n) Neg\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
464  | 
by (unfold Neg_def) transfer_prover  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
465  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
466  | 
end  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
467  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
468  | 
code_datatype "0::integer" Pos Neg  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
469  | 
|
| 
64994
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
470  | 
|
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
471  | 
text \<open>A further pair of constructors for generated computations\<close>  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
472  | 
|
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
473  | 
context  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
474  | 
begin  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
475  | 
|
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
476  | 
qualified definition positive :: "num \<Rightarrow> integer"  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
477  | 
where [simp]: "positive = numeral"  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
478  | 
|
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
479  | 
qualified definition negative :: "num \<Rightarrow> integer"  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
480  | 
where [simp]: "negative = uminus \<circ> numeral"  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
481  | 
|
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
482  | 
lemma [code_computation_unfold]:  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
483  | 
"numeral = positive"  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
484  | 
"Pos = positive"  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
485  | 
"Neg = negative"  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
486  | 
by (simp_all add: fun_eq_iff)  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
487  | 
|
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
488  | 
end  | 
| 
 
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
 
haftmann 
parents: 
64848 
diff
changeset
 | 
489  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
490  | 
|
| 60758 | 491  | 
text \<open>Auxiliary operations\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
492  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
493  | 
lift_definition dup :: "integer \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
494  | 
is "\<lambda>k::int. k + k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
495  | 
.  | 
| 26140 | 496  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
497  | 
lemma dup_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
498  | 
"dup 0 = 0"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
499  | 
"dup (Pos n) = Pos (Num.Bit0 n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
500  | 
"dup (Neg n) = Neg (Num.Bit0 n)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
53069 
diff
changeset
 | 
501  | 
by (transfer, simp only: numeral_Bit0 minus_add_distrib)+  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
502  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
503  | 
lift_definition sub :: "num \<Rightarrow> num \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
504  | 
is "\<lambda>m n. numeral m - numeral n :: int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
505  | 
.  | 
| 26140 | 506  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
507  | 
lemma sub_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
508  | 
"sub Num.One Num.One = 0"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
509  | 
"sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
510  | 
"sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
511  | 
"sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
512  | 
"sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
513  | 
"sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
514  | 
"sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
515  | 
"sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
516  | 
"sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
517  | 
by (transfer, simp add: dbl_def dbl_inc_def dbl_dec_def)+  | 
| 28351 | 518  | 
|
| 24999 | 519  | 
|
| 60758 | 520  | 
text \<open>Implementations\<close>  | 
| 24999 | 521  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
522  | 
lemma one_integer_code [code, code_unfold]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
523  | 
"1 = Pos Num.One"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
524  | 
by simp  | 
| 24999 | 525  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
526  | 
lemma plus_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
527  | 
"k + 0 = (k::integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
528  | 
"0 + l = (l::integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
529  | 
"Pos m + Pos n = Pos (m + n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
530  | 
"Pos m + Neg n = sub m n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
531  | 
"Neg m + Pos n = sub n m"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
532  | 
"Neg m + Neg n = Neg (m + n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
533  | 
by (transfer, simp)+  | 
| 24999 | 534  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
535  | 
lemma uminus_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
536  | 
"uminus 0 = (0::integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
537  | 
"uminus (Pos m) = Neg m"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
538  | 
"uminus (Neg m) = Pos m"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
539  | 
by simp_all  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
540  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
541  | 
lemma minus_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
542  | 
"k - 0 = (k::integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
543  | 
"0 - l = uminus (l::integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
544  | 
"Pos m - Pos n = sub m n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
545  | 
"Pos m - Neg n = Pos (m + n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
546  | 
"Neg m - Pos n = Neg (m + n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
547  | 
"Neg m - Neg n = sub n m"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
548  | 
by (transfer, simp)+  | 
| 
46028
 
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
 
haftmann 
parents: 
45694 
diff
changeset
 | 
549  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
550  | 
lemma abs_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
551  | 
"\<bar>k\<bar> = (if (k::integer) < 0 then - k else k)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
552  | 
by simp  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46664 
diff
changeset
 | 
553  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
554  | 
lemma sgn_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
555  | 
"sgn k = (if k = 0 then 0 else if (k::integer) < 0 then - 1 else 1)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46664 
diff
changeset
 | 
556  | 
by simp  | 
| 
46028
 
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
 
haftmann 
parents: 
45694 
diff
changeset
 | 
557  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
558  | 
lemma times_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
559  | 
"k * 0 = (0::integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
560  | 
"0 * l = (0::integer)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
561  | 
"Pos m * Pos n = Pos (m * n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
562  | 
"Pos m * Neg n = Neg (m * n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
563  | 
"Neg m * Pos n = Neg (m * n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
564  | 
"Neg m * Neg n = Pos (m * n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
565  | 
by simp_all  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
566  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
567  | 
definition divmod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
568  | 
where  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
569  | 
"divmod_integer k l = (k div l, k mod l)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
570  | 
|
| 66801 | 571  | 
lemma fst_divmod_integer [simp]:  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
572  | 
"fst (divmod_integer k l) = k div l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
573  | 
by (simp add: divmod_integer_def)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
574  | 
|
| 66801 | 575  | 
lemma snd_divmod_integer [simp]:  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
576  | 
"snd (divmod_integer k l) = k mod l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
577  | 
by (simp add: divmod_integer_def)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
578  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
579  | 
definition divmod_abs :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
580  | 
where  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
581  | 
"divmod_abs k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
582  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
583  | 
lemma fst_divmod_abs [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
584  | 
"fst (divmod_abs k l) = \<bar>k\<bar> div \<bar>l\<bar>"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
585  | 
by (simp add: divmod_abs_def)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
586  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
587  | 
lemma snd_divmod_abs [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
588  | 
"snd (divmod_abs k l) = \<bar>k\<bar> mod \<bar>l\<bar>"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
589  | 
by (simp add: divmod_abs_def)  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
590  | 
|
| 
53069
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
52435 
diff
changeset
 | 
591  | 
lemma divmod_abs_code [code]:  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
52435 
diff
changeset
 | 
592  | 
"divmod_abs (Pos k) (Pos l) = divmod k l"  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
52435 
diff
changeset
 | 
593  | 
"divmod_abs (Neg k) (Neg l) = divmod k l"  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
52435 
diff
changeset
 | 
594  | 
"divmod_abs (Neg k) (Pos l) = divmod k l"  | 
| 
 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 
haftmann 
parents: 
52435 
diff
changeset
 | 
595  | 
"divmod_abs (Pos k) (Neg l) = divmod k l"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
596  | 
"divmod_abs j 0 = (0, \<bar>j\<bar>)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
597  | 
"divmod_abs 0 j = (0, 0)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
598  | 
by (simp_all add: prod_eq_iff)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
599  | 
|
| 69946 | 600  | 
lemma divmod_integer_eq_cases:  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
601  | 
"divmod_integer k l =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
602  | 
(if k = 0 then (0, 0) else if l = 0 then (0, k) else  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
603  | 
(apsnd \<circ> times \<circ> sgn) l (if sgn k = sgn l  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
604  | 
then divmod_abs k l  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
605  | 
else (let (r, s) = divmod_abs k l in  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
606  | 
if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
607  | 
proof -  | 
| 69946 | 608  | 
have *: "sgn k = sgn l \<longleftrightarrow> k = 0 \<and> l = 0 \<or> 0 < l \<and> 0 < k \<or> l < 0 \<and> k < 0" for k l :: int  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
609  | 
by (auto simp add: sgn_if)  | 
| 69946 | 610  | 
have **: "- k = l * q \<longleftrightarrow> k = - (l * q)" for k l q :: int  | 
611  | 
by auto  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
612  | 
show ?thesis  | 
| 69946 | 613  | 
by (simp add: divmod_integer_def divmod_abs_def)  | 
614  | 
(transfer, auto simp add: * ** not_less zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right)  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
615  | 
qed  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
616  | 
|
| 69946 | 617  | 
lemma divmod_integer_code [code]: \<^marker>\<open>contributor \<open>René Thiemann\<close>\<close> \<^marker>\<open>contributor \<open>Akihisa Yamada\<close>\<close>  | 
618  | 
"divmod_integer k l =  | 
|
619  | 
(if k = 0 then (0, 0)  | 
|
620  | 
else if l > 0 then  | 
|
621  | 
(if k > 0 then Code_Numeral.divmod_abs k l  | 
|
622  | 
else case Code_Numeral.divmod_abs k l of (r, s) \<Rightarrow>  | 
|
623  | 
if s = 0 then (- r, 0) else (- r - 1, l - s))  | 
|
624  | 
else if l = 0 then (0, k)  | 
|
625  | 
else apsnd uminus  | 
|
626  | 
(if k < 0 then Code_Numeral.divmod_abs k l  | 
|
627  | 
else case Code_Numeral.divmod_abs k l of (r, s) \<Rightarrow>  | 
|
628  | 
if s = 0 then (- r, 0) else (- r - 1, - l - s)))"  | 
|
629  | 
by (cases l "0 :: integer" rule: linorder_cases)  | 
|
630  | 
(auto split: prod.splits simp add: divmod_integer_eq_cases)  | 
|
631  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
632  | 
lemma div_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
633  | 
"k div l = fst (divmod_integer k l)"  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
634  | 
by simp  | 
| 
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
635  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
636  | 
lemma mod_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
637  | 
"k mod l = snd (divmod_integer k l)"  | 
| 25767 | 638  | 
by simp  | 
| 24999 | 639  | 
|
| 68028 | 640  | 
definition bit_cut_integer :: "integer \<Rightarrow> integer \<times> bool"  | 
641  | 
where "bit_cut_integer k = (k div 2, odd k)"  | 
|
642  | 
||
643  | 
lemma bit_cut_integer_code [code]:  | 
|
644  | 
"bit_cut_integer k = (if k = 0 then (0, False)  | 
|
645  | 
else let (r, s) = Code_Numeral.divmod_abs k 2  | 
|
646  | 
in (if k > 0 then r else - r - s, s = 1))"  | 
|
647  | 
proof -  | 
|
648  | 
have "bit_cut_integer k = (let (r, s) = divmod_integer k 2 in (r, s = 1))"  | 
|
649  | 
by (simp add: divmod_integer_def bit_cut_integer_def odd_iff_mod_2_eq_one)  | 
|
650  | 
then show ?thesis  | 
|
651  | 
by (simp add: divmod_integer_code) (auto simp add: split_def)  | 
|
652  | 
qed  | 
|
653  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
654  | 
lemma equal_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
655  | 
"HOL.equal 0 (0::integer) \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
656  | 
"HOL.equal 0 (Pos l) \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
657  | 
"HOL.equal 0 (Neg l) \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
658  | 
"HOL.equal (Pos k) 0 \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
659  | 
"HOL.equal (Pos k) (Pos l) \<longleftrightarrow> HOL.equal k l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
660  | 
"HOL.equal (Pos k) (Neg l) \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
661  | 
"HOL.equal (Neg k) 0 \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
662  | 
"HOL.equal (Neg k) (Pos l) \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
663  | 
"HOL.equal (Neg k) (Neg l) \<longleftrightarrow> HOL.equal k l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
664  | 
by (simp_all add: equal)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
665  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
666  | 
lemma equal_integer_refl [code nbe]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
667  | 
"HOL.equal (k::integer) k \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
668  | 
by (fact equal_refl)  | 
| 31266 | 669  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
670  | 
lemma less_eq_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
671  | 
"0 \<le> (0::integer) \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
672  | 
"0 \<le> Pos l \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
673  | 
"0 \<le> Neg l \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
674  | 
"Pos k \<le> 0 \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
675  | 
"Pos k \<le> Pos l \<longleftrightarrow> k \<le> l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
676  | 
"Pos k \<le> Neg l \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
677  | 
"Neg k \<le> 0 \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
678  | 
"Neg k \<le> Pos l \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
679  | 
"Neg k \<le> Neg l \<longleftrightarrow> l \<le> k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
680  | 
by simp_all  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
681  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
682  | 
lemma less_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
683  | 
"0 < (0::integer) \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
684  | 
"0 < Pos l \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
685  | 
"0 < Neg l \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
686  | 
"Pos k < 0 \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
687  | 
"Pos k < Pos l \<longleftrightarrow> k < l"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
688  | 
"Pos k < Neg l \<longleftrightarrow> False"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
689  | 
"Neg k < 0 \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
690  | 
"Neg k < Pos l \<longleftrightarrow> True"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
691  | 
"Neg k < Neg l \<longleftrightarrow> l < k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
692  | 
by simp_all  | 
| 26140 | 693  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
694  | 
lift_definition num_of_integer :: "integer \<Rightarrow> num"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
695  | 
is "num_of_nat \<circ> nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
696  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
697  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
698  | 
lemma num_of_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
699  | 
"num_of_integer k = (if k \<le> 1 then Num.One  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
700  | 
else let  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
701  | 
(l, j) = divmod_integer k 2;  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
702  | 
l' = num_of_integer l;  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
703  | 
l'' = l' + l'  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
704  | 
in if j = 0 then l'' else l'' + Num.One)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
705  | 
proof -  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
706  | 
  {
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
707  | 
assume "int_of_integer k mod 2 = 1"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
708  | 
then have "nat (int_of_integer k mod 2) = nat 1" by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
709  | 
moreover assume *: "1 < int_of_integer k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
710  | 
ultimately have **: "nat (int_of_integer k) mod 2 = 1" by (simp add: nat_mod_distrib)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
711  | 
have "num_of_nat (nat (int_of_integer k)) =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
712  | 
num_of_nat (2 * (nat (int_of_integer k) div 2) + nat (int_of_integer k) mod 2)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
713  | 
by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
714  | 
then have "num_of_nat (nat (int_of_integer k)) =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
715  | 
num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + nat (int_of_integer k) mod 2)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
716  | 
by (simp add: mult_2)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
717  | 
with ** have "num_of_nat (nat (int_of_integer k)) =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
718  | 
num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + 1)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
719  | 
by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
720  | 
}  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
721  | 
note aux = this  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
722  | 
show ?thesis  | 
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
54489 
diff
changeset
 | 
723  | 
by (auto simp add: num_of_integer_def nat_of_integer_def Let_def case_prod_beta  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
724  | 
not_le integer_eq_iff less_eq_integer_def  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
725  | 
nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
726  | 
mult_2 [where 'a=nat] aux add_One)  | 
| 25918 | 727  | 
qed  | 
728  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
729  | 
lemma nat_of_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
730  | 
"nat_of_integer k = (if k \<le> 0 then 0  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
731  | 
else let  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
732  | 
(l, j) = divmod_integer k 2;  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
733  | 
l' = nat_of_integer l;  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
734  | 
l'' = l' + l'  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
735  | 
in if j = 0 then l'' else l'' + 1)"  | 
| 
33340
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
736  | 
proof -  | 
| 66886 | 737  | 
obtain j where k: "k = integer_of_int j"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
738  | 
proof  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
739  | 
show "k = integer_of_int (int_of_integer k)" by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
740  | 
qed  | 
| 66886 | 741  | 
have *: "nat j mod 2 = nat_of_integer (of_int j mod 2)" if "j \<ge> 0"  | 
742  | 
using that by transfer (simp add: nat_mod_distrib)  | 
|
743  | 
from k show ?thesis  | 
|
744  | 
by (auto simp add: split_def Let_def nat_of_integer_def nat_div_distrib mult_2 [symmetric]  | 
|
745  | 
minus_mod_eq_mult_div [symmetric] *)  | 
|
| 
33340
 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 
haftmann 
parents: 
33318 
diff
changeset
 | 
746  | 
qed  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
747  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
748  | 
lemma int_of_integer_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
749  | 
"int_of_integer k = (if k < 0 then - (int_of_integer (- k))  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
750  | 
else if k = 0 then 0  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
751  | 
else let  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
752  | 
(l, j) = divmod_integer k 2;  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
753  | 
l' = 2 * int_of_integer l  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
754  | 
in if j = 0 then l' else l' + 1)"  | 
| 64246 | 755  | 
by (auto simp add: split_def Let_def integer_eq_iff minus_mod_eq_mult_div [symmetric])  | 
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
756  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
757  | 
lemma integer_of_int_code [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
758  | 
"integer_of_int k = (if k < 0 then - (integer_of_int (- k))  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
759  | 
else if k = 0 then 0  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
760  | 
else let  | 
| 
60868
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
761  | 
l = 2 * integer_of_int (k div 2);  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
762  | 
j = k mod 2  | 
| 
 
dd18c33c001e
direct bootstrap of integer division from natural division
 
haftmann 
parents: 
60758 
diff
changeset
 | 
763  | 
in if j = 0 then l else l + 1)"  | 
| 64246 | 764  | 
by (auto simp add: split_def Let_def integer_eq_iff minus_mod_eq_mult_div [symmetric])  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
765  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
766  | 
hide_const (open) Pos Neg sub dup divmod_abs  | 
| 
46547
 
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
 
huffman 
parents: 
46028 
diff
changeset
 | 
767  | 
|
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
768  | 
|
| 60758 | 769  | 
subsection \<open>Serializer setup for target language integers\<close>  | 
| 24999 | 770  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
771  | 
code_reserved Eval int Integer abs  | 
| 25767 | 772  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
773  | 
code_printing  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
774  | 
type_constructor integer \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
775  | 
(SML) "IntInf.int"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
776  | 
and (OCaml) "Z.t"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
777  | 
and (Haskell) "Integer"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
778  | 
and (Scala) "BigInt"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
779  | 
and (Eval) "int"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
780  | 
| class_instance integer :: equal \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
781  | 
(Haskell) -  | 
| 24999 | 782  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
783  | 
code_printing  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
784  | 
constant "0::integer" \<rightharpoonup>  | 
| 
58400
 
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
 
haftmann 
parents: 
58399 
diff
changeset
 | 
785  | 
(SML) "!(0/ :/ IntInf.int)"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
786  | 
and (OCaml) "Z.zero"  | 
| 
58400
 
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
 
haftmann 
parents: 
58399 
diff
changeset
 | 
787  | 
and (Haskell) "!(0/ ::/ Integer)"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
788  | 
and (Scala) "BigInt(0)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46664 
diff
changeset
 | 
789  | 
|
| 60758 | 790  | 
setup \<open>  | 
| 58399 | 791  | 
fold (fn target =>  | 
| 69593 | 792  | 
Numeral.add_code \<^const_name>\<open>Code_Numeral.Pos\<close> I Code_Printer.literal_numeral target  | 
793  | 
#> Numeral.add_code \<^const_name>\<open>Code_Numeral.Neg\<close> (~) Code_Printer.literal_numeral target)  | 
|
| 58399 | 794  | 
["SML", "OCaml", "Haskell", "Scala"]  | 
| 60758 | 795  | 
\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
796  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
797  | 
code_printing  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
798  | 
constant "plus :: integer \<Rightarrow> _ \<Rightarrow> _" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
799  | 
(SML) "IntInf.+ ((_), (_))"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
800  | 
and (OCaml) "Z.add"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
801  | 
and (Haskell) infixl 6 "+"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
802  | 
and (Scala) infixl 7 "+"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
803  | 
and (Eval) infixl 8 "+"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
804  | 
| constant "uminus :: integer \<Rightarrow> _" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
805  | 
(SML) "IntInf.~"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
806  | 
and (OCaml) "Z.neg"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
807  | 
and (Haskell) "negate"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
808  | 
and (Scala) "!(- _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
809  | 
and (Eval) "~/ _"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
810  | 
| constant "minus :: integer \<Rightarrow> _" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
811  | 
(SML) "IntInf.- ((_), (_))"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
812  | 
and (OCaml) "Z.sub"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
813  | 
and (Haskell) infixl 6 "-"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
814  | 
and (Scala) infixl 7 "-"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
815  | 
and (Eval) infixl 8 "-"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
816  | 
| constant Code_Numeral.dup \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
817  | 
(SML) "IntInf.*/ (2,/ (_))"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
818  | 
and (OCaml) "Z.shift'_left/ _/ 1"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
819  | 
and (Haskell) "!(2 * _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
820  | 
and (Scala) "!(2 * _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
821  | 
and (Eval) "!(2 * _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
822  | 
| constant Code_Numeral.sub \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
823  | 
(SML) "!(raise/ Fail/ \"sub\")"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
824  | 
and (OCaml) "failwith/ \"sub\""  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
825  | 
and (Haskell) "error/ \"sub\""  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
826  | 
and (Scala) "!sys.error(\"sub\")"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
827  | 
| constant "times :: integer \<Rightarrow> _ \<Rightarrow> _" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
828  | 
(SML) "IntInf.* ((_), (_))"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
829  | 
and (OCaml) "Z.mul"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
830  | 
and (Haskell) infixl 7 "*"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
831  | 
and (Scala) infixl 8 "*"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
832  | 
and (Eval) infixl 9 "*"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
833  | 
| constant Code_Numeral.divmod_abs \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
834  | 
(SML) "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
835  | 
and (OCaml) "!(fun k l ->/ if Z.equal Z.zero l then/ (Z.zero, l) else/ Z.div'_rem/ (Z.abs k)/ (Z.abs l))"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
836  | 
and (Haskell) "divMod/ (abs _)/ (abs _)"  | 
| 
80088
 
5afbf04418ec
avoid Scala if-expressions and thus make it work both for -new-syntax or -old-syntax;
 
wenzelm 
parents: 
79673 
diff
changeset
 | 
837  | 
    and (Scala) "!((k: BigInt) => (l: BigInt) =>/ l == 0 match { case true => (BigInt(0), k) case false => (k.abs '/% l.abs) })"
 | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
838  | 
and (Eval) "Integer.div'_mod/ (abs _)/ (abs _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
839  | 
| constant "HOL.equal :: integer \<Rightarrow> _ \<Rightarrow> bool" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
840  | 
(SML) "!((_ : IntInf.int) = _)"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
841  | 
and (OCaml) "Z.equal"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
842  | 
and (Haskell) infix 4 "=="  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
843  | 
and (Scala) infixl 5 "=="  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
844  | 
and (Eval) infixl 6 "="  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
845  | 
| constant "less_eq :: integer \<Rightarrow> _ \<Rightarrow> bool" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
846  | 
(SML) "IntInf.<= ((_), (_))"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
847  | 
and (OCaml) "Z.leq"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
848  | 
and (Haskell) infix 4 "<="  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
849  | 
and (Scala) infixl 4 "<="  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
850  | 
and (Eval) infixl 6 "<="  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
851  | 
| constant "less :: integer \<Rightarrow> _ \<Rightarrow> bool" \<rightharpoonup>  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
852  | 
(SML) "IntInf.< ((_), (_))"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
853  | 
and (OCaml) "Z.lt"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
854  | 
and (Haskell) infix 4 "<"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
855  | 
and (Scala) infixl 4 "<"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
856  | 
and (Eval) infixl 6 "<"  | 
| 
61857
 
542f2c6da692
add serialisation for abs on integer to target language operation
 
Andreas Lochbihler 
parents: 
61275 
diff
changeset
 | 
857  | 
| constant "abs :: integer \<Rightarrow> _" \<rightharpoonup>  | 
| 
 
542f2c6da692
add serialisation for abs on integer to target language operation
 
Andreas Lochbihler 
parents: 
61275 
diff
changeset
 | 
858  | 
(SML) "IntInf.abs"  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69593 
diff
changeset
 | 
859  | 
and (OCaml) "Z.abs"  | 
| 
61857
 
542f2c6da692
add serialisation for abs on integer to target language operation
 
Andreas Lochbihler 
parents: 
61275 
diff
changeset
 | 
860  | 
and (Haskell) "Prelude.abs"  | 
| 
 
542f2c6da692
add serialisation for abs on integer to target language operation
 
Andreas Lochbihler 
parents: 
61275 
diff
changeset
 | 
861  | 
and (Scala) "_.abs"  | 
| 
 
542f2c6da692
add serialisation for abs on integer to target language operation
 
Andreas Lochbihler 
parents: 
61275 
diff
changeset
 | 
862  | 
and (Eval) "abs"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
863  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
864  | 
code_identifier  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51375 
diff
changeset
 | 
865  | 
code_module Code_Numeral \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith  | 
| 
46547
 
d1dcb91a512e
use qualified constant names instead of suffixes (from Florian Haftmann)
 
huffman 
parents: 
46028 
diff
changeset
 | 
866  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
867  | 
|
| 60758 | 868  | 
subsection \<open>Type of target language naturals\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
869  | 
|
| 61076 | 870  | 
typedef natural = "UNIV :: nat set"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
871  | 
morphisms nat_of_natural natural_of_nat ..  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
872  | 
|
| 
59487
 
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
 
haftmann 
parents: 
58889 
diff
changeset
 | 
873  | 
setup_lifting type_definition_natural  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
874  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
875  | 
lemma natural_eq_iff [termination_simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
876  | 
"m = n \<longleftrightarrow> nat_of_natural m = nat_of_natural n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
877  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
878  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
879  | 
lemma natural_eqI:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
880  | 
"nat_of_natural m = nat_of_natural n \<Longrightarrow> m = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
881  | 
using natural_eq_iff [of m n] by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
882  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
883  | 
lemma nat_of_natural_of_nat_inverse [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
884  | 
"nat_of_natural (natural_of_nat n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
885  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
886  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
887  | 
lemma natural_of_nat_of_natural_inverse [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
888  | 
"natural_of_nat (nat_of_natural n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
889  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
890  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
891  | 
instantiation natural :: "{comm_monoid_diff, semiring_1}"
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
892  | 
begin  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
893  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
894  | 
lift_definition zero_natural :: natural  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
895  | 
is "0 :: nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
896  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
897  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
898  | 
declare zero_natural.rep_eq [simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
899  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
900  | 
lift_definition one_natural :: natural  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
901  | 
is "1 :: nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
902  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
903  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
904  | 
declare one_natural.rep_eq [simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
905  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
906  | 
lift_definition plus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
907  | 
is "plus :: nat \<Rightarrow> nat \<Rightarrow> nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
908  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
909  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
910  | 
declare plus_natural.rep_eq [simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
911  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
912  | 
lift_definition minus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
913  | 
is "minus :: nat \<Rightarrow> nat \<Rightarrow> nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
914  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
915  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
916  | 
declare minus_natural.rep_eq [simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
917  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
918  | 
lift_definition times_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
919  | 
is "times :: nat \<Rightarrow> nat \<Rightarrow> nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
920  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
921  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
922  | 
declare times_natural.rep_eq [simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
923  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
924  | 
instance proof  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
925  | 
qed (transfer, simp add: algebra_simps)+  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
926  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
927  | 
end  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
928  | 
|
| 
64241
 
430d74089d4d
transfer rules for divides relation on integer and natural
 
haftmann 
parents: 
64178 
diff
changeset
 | 
929  | 
instance natural :: Rings.dvd ..  | 
| 
 
430d74089d4d
transfer rules for divides relation on integer and natural
 
haftmann 
parents: 
64178 
diff
changeset
 | 
930  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
931  | 
context  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
932  | 
includes lifting_syntax  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
933  | 
begin  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
934  | 
|
| 
64241
 
430d74089d4d
transfer rules for divides relation on integer and natural
 
haftmann 
parents: 
64178 
diff
changeset
 | 
935  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
936  | 
\<open>(pcr_natural ===> pcr_natural ===> (\<longleftrightarrow>)) (dvd) (dvd)\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
937  | 
by (unfold dvd_def) transfer_prover  | 
| 
64241
 
430d74089d4d
transfer rules for divides relation on integer and natural
 
haftmann 
parents: 
64178 
diff
changeset
 | 
938  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
939  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
940  | 
\<open>((\<longleftrightarrow>) ===> pcr_natural) of_bool of_bool\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
941  | 
by (unfold of_bool_def) transfer_prover  | 
| 68010 | 942  | 
|
943  | 
lemma [transfer_rule]:  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
944  | 
\<open>((=) ===> pcr_natural) (\<lambda>n. n) of_nat\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
945  | 
proof -  | 
| 55945 | 946  | 
have "rel_fun HOL.eq pcr_natural (of_nat :: nat \<Rightarrow> nat) (of_nat :: nat \<Rightarrow> natural)"  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
947  | 
by (unfold of_nat_def) transfer_prover  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
948  | 
then show ?thesis by (simp add: id_def)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
949  | 
qed  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
950  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
951  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
952  | 
\<open>((=) ===> pcr_natural) numeral numeral\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
953  | 
proof -  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
954  | 
have \<open>((=) ===> pcr_natural) numeral (\<lambda>n. of_nat (numeral n))\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
955  | 
by transfer_prover  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
956  | 
then show ?thesis by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
957  | 
qed  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
958  | 
|
| 68010 | 959  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
960  | 
\<open>(pcr_natural ===> (=) ===> pcr_natural) (^) (^)\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
961  | 
by (unfold power_def) transfer_prover  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
962  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
963  | 
end  | 
| 68010 | 964  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
965  | 
lemma nat_of_natural_of_nat [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
966  | 
"nat_of_natural (of_nat n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
967  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
968  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
969  | 
lemma natural_of_nat_of_nat [simp, code_abbrev]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
970  | 
"natural_of_nat = of_nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
971  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
972  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
973  | 
lemma of_nat_of_natural [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
974  | 
"of_nat (nat_of_natural n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
975  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
976  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
977  | 
lemma nat_of_natural_numeral [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
978  | 
"nat_of_natural (numeral k) = numeral k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
979  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
980  | 
|
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
981  | 
instantiation natural :: "{linordered_semiring, equal}"
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
982  | 
begin  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
983  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
984  | 
lift_definition less_eq_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
985  | 
is "less_eq :: nat \<Rightarrow> nat \<Rightarrow> bool"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
986  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
987  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
988  | 
declare less_eq_natural.rep_eq [termination_simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
989  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
990  | 
lift_definition less_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
991  | 
is "less :: nat \<Rightarrow> nat \<Rightarrow> bool"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
992  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
993  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
994  | 
declare less_natural.rep_eq [termination_simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
995  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
996  | 
lift_definition equal_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
997  | 
is "HOL.equal :: nat \<Rightarrow> nat \<Rightarrow> bool"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
998  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
999  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1000  | 
instance proof  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1001  | 
qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] linear)+  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1002  | 
|
| 24999 | 1003  | 
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
 | 
1004  | 
|
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1005  | 
context  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1006  | 
includes lifting_syntax  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1007  | 
begin  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1008  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1009  | 
lemma [transfer_rule]:  | 
| 
71535
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1010  | 
\<open>(pcr_natural ===> pcr_natural ===> pcr_natural) min min\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1011  | 
by (unfold min_def) transfer_prover  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1012  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1013  | 
lemma [transfer_rule]:  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1014  | 
\<open>(pcr_natural ===> pcr_natural ===> pcr_natural) max max\<close>  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1015  | 
by (unfold max_def) transfer_prover  | 
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1016  | 
|
| 
 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 
haftmann 
parents: 
71424 
diff
changeset
 | 
1017  | 
end  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1018  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1019  | 
lemma nat_of_natural_min [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1020  | 
"nat_of_natural (min k l) = min (nat_of_natural k) (nat_of_natural l)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1021  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1022  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1023  | 
lemma nat_of_natural_max [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1024  | 
"nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1025  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1026  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1027  | 
instantiation natural :: unique_euclidean_semiring  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1028  | 
begin  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1029  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1030  | 
lift_definition divide_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1031  | 
is "divide :: nat \<Rightarrow> nat \<Rightarrow> nat"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1032  | 
.  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1033  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1034  | 
declare divide_natural.rep_eq [simp]  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1035  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1036  | 
lift_definition modulo_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1037  | 
is "modulo :: nat \<Rightarrow> nat \<Rightarrow> nat"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1038  | 
.  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1039  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1040  | 
declare modulo_natural.rep_eq [simp]  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1041  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1042  | 
lift_definition euclidean_size_natural :: "natural \<Rightarrow> nat"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1043  | 
is "euclidean_size :: nat \<Rightarrow> nat"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1044  | 
.  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1045  | 
|
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1046  | 
declare euclidean_size_natural.rep_eq [simp]  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1047  | 
|
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
1048  | 
lift_definition division_segment_natural :: "natural \<Rightarrow> natural"  | 
| 
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
1049  | 
is "division_segment :: nat \<Rightarrow> nat"  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1050  | 
.  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1051  | 
|
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
1052  | 
declare division_segment_natural.rep_eq [simp]  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1053  | 
|
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1054  | 
instance  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1055  | 
by (standard; transfer)  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1056  | 
(auto simp add: algebra_simps unit_factor_nat_def gr0_conv_Suc)  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1057  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1058  | 
end  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1059  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1060  | 
lemma [code]:  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1061  | 
"euclidean_size = nat_of_natural"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1062  | 
by (simp add: fun_eq_iff)  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1063  | 
|
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1064  | 
lemma [code]:  | 
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
1065  | 
"division_segment (n::natural) = 1"  | 
| 
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66836 
diff
changeset
 | 
1066  | 
by (simp add: natural_eq_iff)  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
66801 
diff
changeset
 | 
1067  | 
|
| 
78935
 
5e788ff7a489
explicit type class for discrete linordered semidoms
 
haftmann 
parents: 
78669 
diff
changeset
 | 
1068  | 
instance natural :: discrete_linordered_semidom  | 
| 
78937
 
5e6b195eee83
slightly less technical formulation of very specific type class
 
haftmann 
parents: 
78935 
diff
changeset
 | 
1069  | 
by (standard; transfer) (simp_all add: Suc_le_eq)  | 
| 67905 | 1070  | 
|
| 
78937
 
5e6b195eee83
slightly less technical formulation of very specific type class
 
haftmann 
parents: 
78935 
diff
changeset
 | 
1071  | 
instance natural :: linordered_euclidean_semiring  | 
| 66839 | 1072  | 
by (standard; transfer) simp_all  | 
| 66815 | 1073  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1074  | 
instantiation natural :: semiring_bit_operations  | 
| 71094 | 1075  | 
begin  | 
1076  | 
||
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1077  | 
lift_definition bit_natural :: \<open>natural \<Rightarrow> nat \<Rightarrow> bool\<close>  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1078  | 
is bit .  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1079  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1080  | 
lift_definition and_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1081  | 
is \<open>and\<close> .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1082  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1083  | 
lift_definition or_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1084  | 
is or .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1085  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1086  | 
lift_definition xor_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1087  | 
is xor .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1088  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1089  | 
lift_definition mask_natural :: \<open>nat \<Rightarrow> natural\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1090  | 
is mask .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1091  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1092  | 
lift_definition set_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1093  | 
is set_bit .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1094  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1095  | 
lift_definition unset_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1096  | 
is unset_bit .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1097  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1098  | 
lift_definition flip_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1099  | 
is flip_bit .  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1100  | 
|
| 71094 | 1101  | 
lift_definition push_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1102  | 
is push_bit .  | 
| 71094 | 1103  | 
|
1104  | 
lift_definition drop_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
|
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1105  | 
is drop_bit .  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1106  | 
|
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1107  | 
lift_definition take_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1108  | 
is take_bit .  | 
| 71094 | 1109  | 
|
1110  | 
instance by (standard; transfer)  | 
|
| 
79541
 
4f40225936d1
common type class for trivial properties on div/mod
 
haftmann 
parents: 
79531 
diff
changeset
 | 
1111  | 
(fact bit_induct div_by_0 div_by_1 div_0 even_half_succ_eq  | 
| 
79673
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79588 
diff
changeset
 | 
1112  | 
half_div_exp_eq even_double_div_exp_iff bits_mod_div_trivial  | 
| 
79531
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
1113  | 
bit_iff_odd push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
1114  | 
and_rec or_rec xor_rec mask_eq_exp_minus_1  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
1115  | 
set_bit_eq_or unset_bit_eq_or_xor flip_bit_eq_xor not_eq_complement)+  | 
| 71094 | 1116  | 
|
1117  | 
end  | 
|
| 68010 | 1118  | 
|
| 78955 | 1119  | 
instance natural :: linordered_euclidean_semiring_bit_operations ..  | 
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1120  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1121  | 
context  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1122  | 
includes bit_operations_syntax  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1123  | 
begin  | 
| 71094 | 1124  | 
|
1125  | 
lemma [code]:  | 
|
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1126  | 
\<open>bit m n \<longleftrightarrow> odd (drop_bit n m)\<close>  | 
| 79016 | 1127  | 
\<open>mask n = 2 ^ n - (1 :: natural)\<close>  | 
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1128  | 
\<open>set_bit n m = m OR push_bit n 1\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1129  | 
\<open>flip_bit n m = m XOR push_bit n 1\<close>  | 
| 71094 | 1130  | 
\<open>push_bit n m = m * 2 ^ n\<close>  | 
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1131  | 
\<open>drop_bit n m = m div 2 ^ n\<close>  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71535 
diff
changeset
 | 
1132  | 
\<open>take_bit n m = m mod 2 ^ n\<close> for m :: natural  | 
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1133  | 
by (fact bit_iff_odd_drop_bit mask_eq_exp_minus_1  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1134  | 
set_bit_eq_or flip_bit_eq_xor push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1135  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1136  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1137  | 
\<open>m AND n = (if m = 0 \<or> n = 0 then 0  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1138  | 
else (m mod 2) * (n mod 2) + 2 * ((m div 2) AND (n div 2)))\<close> for m n :: natural  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1139  | 
by transfer (fact and_nat_unfold)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1140  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1141  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1142  | 
\<open>m OR n = (if m = 0 then n else if n = 0 then m  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1143  | 
else max (m mod 2) (n mod 2) + 2 * ((m div 2) OR (n div 2)))\<close> for m n :: natural  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1144  | 
by transfer (fact or_nat_unfold)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1145  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1146  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1147  | 
\<open>m XOR n = (if m = 0 then n else if n = 0 then m  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1148  | 
else (m mod 2 + n mod 2) mod 2 + 2 * ((m div 2) XOR (n div 2)))\<close> for m n :: natural  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1149  | 
by transfer (fact xor_nat_unfold)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1150  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1151  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1152  | 
\<open>unset_bit 0 m = 2 * (m div 2)\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1153  | 
\<open>unset_bit (Suc n) m = m mod 2 + 2 * unset_bit n (m div 2)\<close> for m :: natural  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1154  | 
by (transfer; simp add: unset_bit_Suc)+  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1155  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1156  | 
end  | 
| 68010 | 1157  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1158  | 
lift_definition natural_of_integer :: "integer \<Rightarrow> natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1159  | 
is "nat :: int \<Rightarrow> nat"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1160  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1161  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1162  | 
lift_definition integer_of_natural :: "natural \<Rightarrow> integer"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1163  | 
is "of_nat :: nat \<Rightarrow> int"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1164  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1165  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1166  | 
lemma natural_of_integer_of_natural [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1167  | 
"natural_of_integer (integer_of_natural n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1168  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1169  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1170  | 
lemma integer_of_natural_of_integer [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1171  | 
"integer_of_natural (natural_of_integer k) = max 0 k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1172  | 
by transfer auto  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1173  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1174  | 
lemma int_of_integer_of_natural [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1175  | 
"int_of_integer (integer_of_natural n) = of_nat (nat_of_natural n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1176  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1177  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1178  | 
lemma integer_of_natural_of_nat [simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1179  | 
"integer_of_natural (of_nat n) = of_nat n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1180  | 
by transfer rule  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1181  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1182  | 
lemma [measure_function]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1183  | 
"is_measure nat_of_natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1184  | 
by (rule is_measure_trivial)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1185  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1186  | 
|
| 60758 | 1187  | 
subsection \<open>Inductive representation of target language naturals\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1188  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1189  | 
lift_definition Suc :: "natural \<Rightarrow> natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1190  | 
is Nat.Suc  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1191  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1192  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1193  | 
declare Suc.rep_eq [simp]  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1194  | 
|
| 
58306
 
117ba6cbe414
renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
 
blanchet 
parents: 
57512 
diff
changeset
 | 
1195  | 
old_rep_datatype "0::natural" Suc  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1196  | 
by (transfer, fact nat.induct nat.inject nat.distinct)+  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1197  | 
|
| 55416 | 1198  | 
lemma natural_cases [case_names nat, cases type: natural]:  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1199  | 
fixes m :: natural  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1200  | 
assumes "\<And>n. m = of_nat n \<Longrightarrow> P"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1201  | 
shows P  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1202  | 
using assms by transfer blast  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1203  | 
|
| 67332 | 1204  | 
instantiation natural :: size  | 
1205  | 
begin  | 
|
| 
58379
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1206  | 
|
| 67332 | 1207  | 
definition size_nat where [simp, code]: "size_nat = nat_of_natural"  | 
1208  | 
||
1209  | 
instance ..  | 
|
1210  | 
||
1211  | 
end  | 
|
| 
58379
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1212  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1213  | 
lemma natural_decr [termination_simp]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1214  | 
"n \<noteq> 0 \<Longrightarrow> nat_of_natural n - Nat.Suc 0 < nat_of_natural n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1215  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1216  | 
|
| 
58379
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1217  | 
lemma natural_zero_minus_one: "(0::natural) - 1 = 0"  | 
| 
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1218  | 
by (rule zero_diff)  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1219  | 
|
| 
58379
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1220  | 
lemma Suc_natural_minus_one: "Suc n - 1 = n"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1221  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1222  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1223  | 
hide_const (open) Suc  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1224  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1225  | 
|
| 60758 | 1226  | 
subsection \<open>Code refinement for target language naturals\<close>  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1227  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1228  | 
lift_definition Nat :: "integer \<Rightarrow> natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1229  | 
is nat  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1230  | 
.  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1231  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1232  | 
lemma [code_post]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1233  | 
"Nat 0 = 0"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1234  | 
"Nat 1 = 1"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1235  | 
"Nat (numeral k) = numeral k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1236  | 
by (transfer, simp)+  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1237  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1238  | 
lemma [code abstype]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1239  | 
"Nat (integer_of_natural n) = n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1240  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1241  | 
|
| 
63174
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1242  | 
lemma [code]:  | 
| 
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1243  | 
"natural_of_nat n = natural_of_integer (integer_of_nat n)"  | 
| 
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1244  | 
by transfer simp  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1245  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1246  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1247  | 
"integer_of_natural (natural_of_integer k) = max 0 k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1248  | 
by simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1249  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1250  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1251  | 
\<open>integer_of_natural (mask n) = mask n\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1252  | 
by transfer (simp add: mask_eq_exp_minus_1 of_nat_diff)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1253  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1254  | 
lemma [code_abbrev]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1255  | 
"natural_of_integer (Code_Numeral.Pos k) = numeral k"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1256  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1257  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1258  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1259  | 
"integer_of_natural 0 = 0"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1260  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1261  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1262  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1263  | 
"integer_of_natural 1 = 1"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1264  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1265  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1266  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1267  | 
"integer_of_natural (Code_Numeral.Suc n) = integer_of_natural n + 1"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1268  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1269  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1270  | 
lemma [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1271  | 
"nat_of_natural = nat_of_integer \<circ> integer_of_natural"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1272  | 
by transfer (simp add: fun_eq_iff)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1273  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1274  | 
lemma [code, code_unfold]:  | 
| 55416 | 1275  | 
"case_natural f g n = (if n = 0 then f else g (n - 1))"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1276  | 
by (cases n rule: natural.exhaust) (simp_all, simp add: Suc_def)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1277  | 
|
| 
55642
 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 
blanchet 
parents: 
55428 
diff
changeset
 | 
1278  | 
declare natural.rec [code del]  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1279  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1280  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1281  | 
"integer_of_natural (m + n) = integer_of_natural m + integer_of_natural n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1282  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1283  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1284  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1285  | 
"integer_of_natural (m - n) = max 0 (integer_of_natural m - integer_of_natural n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1286  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1287  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1288  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1289  | 
"integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n"  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1290  | 
by transfer simp  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64246 
diff
changeset
 | 
1291  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1292  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1293  | 
"integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1294  | 
by transfer (simp add: zdiv_int)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1295  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1296  | 
lemma [code abstract]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1297  | 
"integer_of_natural (m mod n) = integer_of_natural m mod integer_of_natural n"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1298  | 
by transfer (simp add: zmod_int)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1299  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1300  | 
lemma [code]:  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1301  | 
"HOL.equal m n \<longleftrightarrow> HOL.equal (integer_of_natural m) (integer_of_natural n)"  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1302  | 
by transfer (simp add: equal)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1303  | 
|
| 
58379
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1304  | 
lemma [code nbe]: "HOL.equal n (n::natural) \<longleftrightarrow> True"  | 
| 
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1305  | 
by (rule equal_class.equal_refl)  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1306  | 
|
| 
58379
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1307  | 
lemma [code]: "m \<le> n \<longleftrightarrow> integer_of_natural m \<le> integer_of_natural n"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1308  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1309  | 
|
| 
58379
 
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
 
blanchet 
parents: 
58377 
diff
changeset
 | 
1310  | 
lemma [code]: "m < n \<longleftrightarrow> integer_of_natural m < integer_of_natural n"  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1311  | 
by transfer simp  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1312  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1313  | 
hide_const (open) Nat  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1314  | 
|
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1315  | 
code_reflect Code_Numeral  | 
| 
63174
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1316  | 
datatypes natural  | 
| 
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1317  | 
functions "Code_Numeral.Suc" "0 :: natural" "1 :: natural"  | 
| 
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1318  | 
"plus :: natural \<Rightarrow> _" "minus :: natural \<Rightarrow> _"  | 
| 
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1319  | 
"times :: natural \<Rightarrow> _" "divide :: natural \<Rightarrow> _"  | 
| 
63950
 
cdc1e59aa513
syntactic type class for operation mod named after mod;
 
haftmann 
parents: 
63174 
diff
changeset
 | 
1320  | 
"modulo :: natural \<Rightarrow> _"  | 
| 
63174
 
57c0d60e491c
do not export abstract constructors in code_reflect
 
haftmann 
parents: 
61857 
diff
changeset
 | 
1321  | 
integer_of_natural natural_of_integer  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1322  | 
|
| 74101 | 1323  | 
lifting_update integer.lifting  | 
1324  | 
lifting_forget integer.lifting  | 
|
1325  | 
||
1326  | 
lifting_update natural.lifting  | 
|
1327  | 
lifting_forget natural.lifting  | 
|
1328  | 
||
1329  | 
end  |