author | wenzelm |
Sun, 11 Aug 2019 22:36:34 +0200 | |
changeset 70503 | f0b2635ee17f |
parent 70340 | 7383930fc946 |
child 70927 | cc204e10385c |
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 |
66836 | 8 |
imports Divides Lifting |
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 |
|
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
80 |
lemma [transfer_rule]: |
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
81 |
"rel_fun pcr_integer (rel_fun pcr_integer HOL.iff) Rings.dvd Rings.dvd" |
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
82 |
unfolding dvd_def by transfer_prover |
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
83 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
84 |
lemma [transfer_rule]: |
68010 | 85 |
"rel_fun (=) pcr_integer (of_bool :: bool \<Rightarrow> int) (of_bool :: bool \<Rightarrow> integer)" |
86 |
by (unfold of_bool_def [abs_def]) transfer_prover |
|
87 |
||
88 |
lemma [transfer_rule]: |
|
89 |
"rel_fun (=) pcr_integer (of_nat :: nat \<Rightarrow> int) (of_nat :: nat \<Rightarrow> integer)" |
|
64178 | 90 |
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
|
91 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
92 |
lemma [transfer_rule]: |
68010 | 93 |
"rel_fun (=) pcr_integer (\<lambda>k :: int. k :: int) (of_int :: int \<Rightarrow> integer)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
94 |
proof - |
55945 | 95 |
have "rel_fun HOL.eq pcr_integer (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> integer)" |
64178 | 96 |
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
|
97 |
then show ?thesis by (simp add: id_def) |
24999 | 98 |
qed |
99 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
100 |
lemma [transfer_rule]: |
55945 | 101 |
"rel_fun HOL.eq pcr_integer (numeral :: num \<Rightarrow> int) (numeral :: num \<Rightarrow> integer)" |
64178 | 102 |
by (rule transfer_rule_numeral) transfer_prover+ |
26140 | 103 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
104 |
lemma [transfer_rule]: |
55945 | 105 |
"rel_fun HOL.eq (rel_fun HOL.eq pcr_integer) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> int) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> integer)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
106 |
by (unfold Num.sub_def [abs_def]) transfer_prover |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
107 |
|
68010 | 108 |
lemma [transfer_rule]: |
109 |
"rel_fun pcr_integer (rel_fun (=) pcr_integer) (power :: _ \<Rightarrow> _ \<Rightarrow> int) (power :: _ \<Rightarrow> _ \<Rightarrow> integer)" |
|
110 |
by (unfold power_def [abs_def]) transfer_prover |
|
111 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
112 |
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
|
113 |
"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
|
114 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
115 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
116 |
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
|
117 |
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
|
118 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
119 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
120 |
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
|
121 |
"integer_of_nat = of_nat" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
122 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
123 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
124 |
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
|
125 |
"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
|
126 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
127 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
128 |
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
|
129 |
is Int.nat |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
130 |
. |
26140 | 131 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
132 |
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
|
133 |
"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
|
134 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
135 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
136 |
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
|
137 |
"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
|
138 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
139 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
140 |
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
|
141 |
"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
|
142 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
143 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
144 |
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
|
145 |
"integer_of_int = of_int" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
146 |
by transfer (simp add: fun_eq_iff) |
26140 | 147 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
148 |
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
|
149 |
"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
|
150 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
151 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
152 |
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
|
153 |
"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
|
154 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
155 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
156 |
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
|
157 |
"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
|
158 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
159 |
|
66190 | 160 |
definition integer_of_num :: "num \<Rightarrow> integer" |
161 |
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
|
162 |
|
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
163 |
lemma integer_of_num [code]: |
66190 | 164 |
"integer_of_num Num.One = 1" |
165 |
"integer_of_num (Num.Bit0 n) = (let k = integer_of_num n in k + k)" |
|
166 |
"integer_of_num (Num.Bit1 n) = (let k = integer_of_num n in k + k + 1)" |
|
167 |
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
|
168 |
|
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
169 |
lemma integer_of_num_triv: |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
170 |
"integer_of_num Num.One = 1" |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
171 |
"integer_of_num (Num.Bit0 Num.One) = 2" |
66190 | 172 |
by simp_all |
61275
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
173 |
|
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
174 |
instantiation integer :: "{linordered_idom, equal}" |
26140 | 175 |
begin |
176 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
177 |
lift_definition abs_integer :: "integer \<Rightarrow> integer" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
178 |
is "abs :: int \<Rightarrow> int" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
179 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
180 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
181 |
declare abs_integer.rep_eq [simp] |
26140 | 182 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
183 |
lift_definition sgn_integer :: "integer \<Rightarrow> integer" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
184 |
is "sgn :: int \<Rightarrow> int" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
185 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
186 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
187 |
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
|
188 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
189 |
lift_definition less_eq_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
190 |
is "less_eq :: int \<Rightarrow> int \<Rightarrow> bool" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
191 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
192 |
|
69946 | 193 |
lemma integer_less_eq_iff: |
194 |
"k \<le> l \<longleftrightarrow> int_of_integer k \<le> int_of_integer l" |
|
195 |
by (fact less_eq_integer.rep_eq) |
|
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
196 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
197 |
lift_definition less_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
198 |
is "less :: int \<Rightarrow> int \<Rightarrow> bool" |
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 |
|
69946 | 201 |
lemma integer_less_iff: |
202 |
"k < l \<longleftrightarrow> int_of_integer k < int_of_integer l" |
|
203 |
by (fact less_integer.rep_eq) |
|
204 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
205 |
lift_definition equal_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
206 |
is "HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
207 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
208 |
|
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
209 |
instance |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
210 |
by standard (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+ |
26140 | 211 |
|
212 |
end |
|
213 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
214 |
lemma [transfer_rule]: |
55945 | 215 |
"rel_fun pcr_integer (rel_fun pcr_integer pcr_integer) (min :: _ \<Rightarrow> _ \<Rightarrow> int) (min :: _ \<Rightarrow> _ \<Rightarrow> integer)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
216 |
by (unfold min_def [abs_def]) transfer_prover |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
217 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
218 |
lemma [transfer_rule]: |
55945 | 219 |
"rel_fun pcr_integer (rel_fun pcr_integer pcr_integer) (max :: _ \<Rightarrow> _ \<Rightarrow> int) (max :: _ \<Rightarrow> _ \<Rightarrow> integer)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
220 |
by (unfold max_def [abs_def]) transfer_prover |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
221 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
222 |
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
|
223 |
"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
|
224 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
225 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
226 |
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
|
227 |
"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
|
228 |
by transfer rule |
26140 | 229 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
230 |
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
|
231 |
"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
|
232 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
233 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
234 |
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
|
235 |
"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
|
236 |
by transfer auto |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
237 |
|
66817 | 238 |
instantiation integer :: unique_euclidean_ring |
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
239 |
begin |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
240 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
241 |
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
|
242 |
is "divide :: int \<Rightarrow> int \<Rightarrow> int" |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
243 |
. |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
244 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
245 |
declare divide_integer.rep_eq [simp] |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
246 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
247 |
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
|
248 |
is "modulo :: int \<Rightarrow> int \<Rightarrow> int" |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
249 |
. |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
250 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
251 |
declare modulo_integer.rep_eq [simp] |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
252 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
253 |
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
|
254 |
is "euclidean_size :: int \<Rightarrow> nat" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
255 |
. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
256 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
257 |
declare euclidean_size_integer.rep_eq [simp] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
258 |
|
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
259 |
lift_definition division_segment_integer :: "integer \<Rightarrow> integer" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
260 |
is "division_segment :: int \<Rightarrow> int" |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
261 |
. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
262 |
|
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
263 |
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
|
264 |
|
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
265 |
instance |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
266 |
by (standard; transfer) |
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
267 |
(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 |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
268 |
division_segment_mult division_segment_mod intro: div_eqI\<close>) |
64592
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 |
end |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
271 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
272 |
lemma [code]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
273 |
"euclidean_size = nat_of_integer \<circ> abs" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
274 |
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
|
275 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
276 |
lemma [code]: |
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
277 |
"division_segment (k :: integer) = (if k \<ge> 0 then 1 else - 1)" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
278 |
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
|
279 |
|
70340 | 280 |
instance integer :: unique_euclidean_ring_with_nat |
66839 | 281 |
by (standard; transfer) (simp_all add: of_nat_div division_segment_int_def) |
66815 | 282 |
|
68010 | 283 |
lemma [transfer_rule]: |
284 |
"rel_fun (=) (rel_fun pcr_integer pcr_integer) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> int) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> integer)" |
|
285 |
by (unfold push_bit_eq_mult [abs_def]) transfer_prover |
|
286 |
||
287 |
lemma [transfer_rule]: |
|
288 |
"rel_fun (=) (rel_fun pcr_integer pcr_integer) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> int) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> integer)" |
|
289 |
by (unfold take_bit_eq_mod [abs_def]) transfer_prover |
|
290 |
||
291 |
lemma [transfer_rule]: |
|
292 |
"rel_fun (=) (rel_fun pcr_integer pcr_integer) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> int) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> integer)" |
|
293 |
by (unfold drop_bit_eq_div [abs_def]) transfer_prover |
|
294 |
||
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
295 |
instantiation integer :: unique_euclidean_semiring_numeral |
61275
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
296 |
begin |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
297 |
|
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
298 |
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
|
299 |
where |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
300 |
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
|
301 |
|
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
302 |
definition divmod_step_integer :: "num \<Rightarrow> integer \<times> integer \<Rightarrow> integer \<times> integer" |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
303 |
where |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
304 |
"divmod_step_integer l qr = (let (q, r) = qr |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
305 |
in if r \<ge> numeral l then (2 * q + 1, r - numeral l) |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
306 |
else (2 * q, r))" |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
307 |
|
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
308 |
instance proof |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
309 |
show "divmod m n = (numeral m div numeral n :: integer, numeral m mod numeral n)" |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
310 |
for m n by (fact divmod_integer'_def) |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
311 |
show "divmod_step l qr = (let (q, r) = qr |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
312 |
in if r \<ge> numeral l then (2 * q + 1, r - numeral l) |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
313 |
else (2 * q, r))" for l and qr :: "integer \<times> integer" |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
314 |
by (fact divmod_step_integer_def) |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
315 |
qed (transfer, |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
316 |
fact le_add_diff_inverse2 |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
317 |
unique_euclidean_semiring_numeral_class.div_less |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
318 |
unique_euclidean_semiring_numeral_class.mod_less |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
319 |
unique_euclidean_semiring_numeral_class.div_positive |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
320 |
unique_euclidean_semiring_numeral_class.mod_less_eq_dividend |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
321 |
unique_euclidean_semiring_numeral_class.pos_mod_bound |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
322 |
unique_euclidean_semiring_numeral_class.pos_mod_sign |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
323 |
unique_euclidean_semiring_numeral_class.mod_mult2_eq |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
324 |
unique_euclidean_semiring_numeral_class.div_mult2_eq |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
325 |
unique_euclidean_semiring_numeral_class.discrete)+ |
61275
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
326 |
|
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
327 |
end |
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
328 |
|
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents:
61274
diff
changeset
|
329 |
declare divmod_algorithm_code [where ?'a = integer, |
69946 | 330 |
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
|
331 |
code] |
53069
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents:
52435
diff
changeset
|
332 |
|
55427
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
Andreas Lochbihler
parents:
54489
diff
changeset
|
333 |
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
|
334 |
by transfer simp |
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
Andreas Lochbihler
parents:
54489
diff
changeset
|
335 |
|
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
Andreas Lochbihler
parents:
54489
diff
changeset
|
336 |
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
|
337 |
by transfer simp |
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
Andreas Lochbihler
parents:
54489
diff
changeset
|
338 |
|
ff54d22fe357
make integer_of_char and char_of_integer work with NBE and code_simp
Andreas Lochbihler
parents:
54489
diff
changeset
|
339 |
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
|
340 |
"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
|
341 |
by transfer simp |
26140 | 342 |
|
68010 | 343 |
|
60758 | 344 |
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
|
345 |
|
60758 | 346 |
text \<open>Constructors\<close> |
26140 | 347 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
348 |
definition Pos :: "num \<Rightarrow> integer" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
349 |
where |
61274
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
haftmann
parents:
61076
diff
changeset
|
350 |
[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
|
351 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
352 |
lemma [transfer_rule]: |
55945 | 353 |
"rel_fun HOL.eq pcr_integer numeral Pos" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
354 |
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
|
355 |
|
61274
0261eec37233
more selective preprocessing allows bare "numeral" occurences to be retained as real function in generated code
haftmann
parents:
61076
diff
changeset
|
356 |
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
|
357 |
"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
|
358 |
"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
|
359 |
"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
|
360 |
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
|
361 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
362 |
definition Neg :: "num \<Rightarrow> integer" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
363 |
where |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
53069
diff
changeset
|
364 |
[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
|
365 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
366 |
lemma [transfer_rule]: |
55945 | 367 |
"rel_fun HOL.eq pcr_integer (\<lambda>n. - numeral n) Neg" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
53069
diff
changeset
|
368 |
by (simp add: Neg_def [abs_def]) transfer_prover |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
369 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
370 |
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
|
371 |
|
64994
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
372 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
373 |
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
|
374 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
375 |
context |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
376 |
begin |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
377 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
378 |
qualified definition positive :: "num \<Rightarrow> integer" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
379 |
where [simp]: "positive = numeral" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
380 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
381 |
qualified definition negative :: "num \<Rightarrow> integer" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
382 |
where [simp]: "negative = uminus \<circ> numeral" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
383 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
384 |
lemma [code_computation_unfold]: |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
385 |
"numeral = positive" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
386 |
"Pos = positive" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
387 |
"Neg = negative" |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
388 |
by (simp_all add: fun_eq_iff) |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
389 |
|
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
390 |
end |
6e4c05e8edbb
computation preprocessing rules to allow literals as input for computations
haftmann
parents:
64848
diff
changeset
|
391 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
392 |
|
60758 | 393 |
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
|
394 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
395 |
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
|
396 |
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
|
397 |
. |
26140 | 398 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
399 |
lemma dup_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
400 |
"dup 0 = 0" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
401 |
"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
|
402 |
"dup (Neg n) = Neg (Num.Bit0 n)" |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
53069
diff
changeset
|
403 |
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
|
404 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
405 |
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
|
406 |
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
|
407 |
. |
26140 | 408 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
409 |
lemma sub_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
410 |
"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
|
411 |
"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
|
412 |
"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
|
413 |
"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
|
414 |
"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
|
415 |
"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
|
416 |
"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
|
417 |
"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
|
418 |
"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
|
419 |
by (transfer, simp add: dbl_def dbl_inc_def dbl_dec_def)+ |
28351 | 420 |
|
24999 | 421 |
|
60758 | 422 |
text \<open>Implementations\<close> |
24999 | 423 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
424 |
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
|
425 |
"1 = Pos Num.One" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
426 |
by simp |
24999 | 427 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
428 |
lemma plus_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
429 |
"k + 0 = (k::integer)" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
430 |
"0 + l = (l::integer)" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
431 |
"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
|
432 |
"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
|
433 |
"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
|
434 |
"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
|
435 |
by (transfer, simp)+ |
24999 | 436 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
437 |
lemma uminus_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
438 |
"uminus 0 = (0::integer)" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
439 |
"uminus (Pos m) = Neg m" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
440 |
"uminus (Neg m) = Pos m" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
441 |
by simp_all |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
442 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
443 |
lemma minus_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
444 |
"k - 0 = (k::integer)" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
445 |
"0 - l = uminus (l::integer)" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
446 |
"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
|
447 |
"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
|
448 |
"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
|
449 |
"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
|
450 |
by (transfer, simp)+ |
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
451 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
452 |
lemma abs_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
453 |
"\<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
|
454 |
by simp |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46664
diff
changeset
|
455 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
456 |
lemma sgn_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
457 |
"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
|
458 |
by simp |
46028
9f113cdf3d66
attribute code_abbrev superseedes code_unfold_post
haftmann
parents:
45694
diff
changeset
|
459 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
460 |
lemma times_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
461 |
"k * 0 = (0::integer)" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
462 |
"0 * l = (0::integer)" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
463 |
"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
|
464 |
"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
|
465 |
"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
|
466 |
"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
|
467 |
by simp_all |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
468 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
469 |
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
|
470 |
where |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
471 |
"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
|
472 |
|
66801 | 473 |
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
|
474 |
"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
|
475 |
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
|
476 |
|
66801 | 477 |
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
|
478 |
"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
|
479 |
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
|
480 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
481 |
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
|
482 |
where |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
483 |
"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
|
484 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
485 |
lemma fst_divmod_abs [simp]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
486 |
"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
|
487 |
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
|
488 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
489 |
lemma snd_divmod_abs [simp]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
490 |
"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
|
491 |
by (simp add: divmod_abs_def) |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
492 |
|
53069
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents:
52435
diff
changeset
|
493 |
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
|
494 |
"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
|
495 |
"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
|
496 |
"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
|
497 |
"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
|
498 |
"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
|
499 |
"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
|
500 |
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
|
501 |
|
69946 | 502 |
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
|
503 |
"divmod_integer k l = |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
504 |
(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
|
505 |
(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
|
506 |
then divmod_abs k l |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
507 |
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
|
508 |
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
|
509 |
proof - |
69946 | 510 |
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
|
511 |
by (auto simp add: sgn_if) |
69946 | 512 |
have **: "- k = l * q \<longleftrightarrow> k = - (l * q)" for k l q :: int |
513 |
by auto |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
514 |
show ?thesis |
69946 | 515 |
by (simp add: divmod_integer_def divmod_abs_def) |
516 |
(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
|
517 |
qed |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
518 |
|
69946 | 519 |
lemma divmod_integer_code [code]: \<^marker>\<open>contributor \<open>René Thiemann\<close>\<close> \<^marker>\<open>contributor \<open>Akihisa Yamada\<close>\<close> |
520 |
"divmod_integer k l = |
|
521 |
(if k = 0 then (0, 0) |
|
522 |
else if l > 0 then |
|
523 |
(if k > 0 then Code_Numeral.divmod_abs k l |
|
524 |
else case Code_Numeral.divmod_abs k l of (r, s) \<Rightarrow> |
|
525 |
if s = 0 then (- r, 0) else (- r - 1, l - s)) |
|
526 |
else if l = 0 then (0, k) |
|
527 |
else apsnd uminus |
|
528 |
(if k < 0 then Code_Numeral.divmod_abs k l |
|
529 |
else case Code_Numeral.divmod_abs k l of (r, s) \<Rightarrow> |
|
530 |
if s = 0 then (- r, 0) else (- r - 1, - l - s)))" |
|
531 |
by (cases l "0 :: integer" rule: linorder_cases) |
|
532 |
(auto split: prod.splits simp add: divmod_integer_eq_cases) |
|
533 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
534 |
lemma div_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
535 |
"k div l = fst (divmod_integer k l)" |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
536 |
by simp |
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
537 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
538 |
lemma mod_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
539 |
"k mod l = snd (divmod_integer k l)" |
25767 | 540 |
by simp |
24999 | 541 |
|
68028 | 542 |
definition bit_cut_integer :: "integer \<Rightarrow> integer \<times> bool" |
543 |
where "bit_cut_integer k = (k div 2, odd k)" |
|
544 |
||
545 |
lemma bit_cut_integer_code [code]: |
|
546 |
"bit_cut_integer k = (if k = 0 then (0, False) |
|
547 |
else let (r, s) = Code_Numeral.divmod_abs k 2 |
|
548 |
in (if k > 0 then r else - r - s, s = 1))" |
|
549 |
proof - |
|
550 |
have "bit_cut_integer k = (let (r, s) = divmod_integer k 2 in (r, s = 1))" |
|
551 |
by (simp add: divmod_integer_def bit_cut_integer_def odd_iff_mod_2_eq_one) |
|
552 |
then show ?thesis |
|
553 |
by (simp add: divmod_integer_code) (auto simp add: split_def) |
|
554 |
qed |
|
555 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
556 |
lemma equal_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
557 |
"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
|
558 |
"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
|
559 |
"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
|
560 |
"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
|
561 |
"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
|
562 |
"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
|
563 |
"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
|
564 |
"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
|
565 |
"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
|
566 |
by (simp_all add: equal) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
567 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
568 |
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
|
569 |
"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
|
570 |
by (fact equal_refl) |
31266 | 571 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
572 |
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
|
573 |
"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
|
574 |
"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
|
575 |
"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
|
576 |
"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
|
577 |
"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
|
578 |
"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
|
579 |
"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
|
580 |
"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
|
581 |
"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
|
582 |
by simp_all |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
583 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
584 |
lemma less_integer_code [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
585 |
"0 < (0::integer) \<longleftrightarrow> False" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
586 |
"0 < Pos l \<longleftrightarrow> True" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
587 |
"0 < Neg l \<longleftrightarrow> False" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
588 |
"Pos k < 0 \<longleftrightarrow> False" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
589 |
"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
|
590 |
"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
|
591 |
"Neg k < 0 \<longleftrightarrow> True" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
592 |
"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
|
593 |
"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
|
594 |
by simp_all |
26140 | 595 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
596 |
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
|
597 |
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
|
598 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
599 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
600 |
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
|
601 |
"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
|
602 |
else let |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
603 |
(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
|
604 |
l' = num_of_integer l; |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
605 |
l'' = l' + l' |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
606 |
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
|
607 |
proof - |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
608 |
{ |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
609 |
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
|
610 |
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
|
611 |
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
|
612 |
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
|
613 |
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
|
614 |
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
|
615 |
by simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
616 |
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
|
617 |
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
|
618 |
by (simp add: mult_2) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
619 |
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
|
620 |
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
|
621 |
by simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
622 |
} |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
623 |
note aux = this |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
624 |
show ?thesis |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
54489
diff
changeset
|
625 |
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
|
626 |
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
|
627 |
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
|
628 |
mult_2 [where 'a=nat] aux add_One) |
25918 | 629 |
qed |
630 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
631 |
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
|
632 |
"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
|
633 |
else let |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
634 |
(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
|
635 |
l' = nat_of_integer l; |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
636 |
l'' = l' + l' |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
637 |
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
|
638 |
proof - |
66886 | 639 |
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
|
640 |
proof |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
641 |
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
|
642 |
qed |
66886 | 643 |
have *: "nat j mod 2 = nat_of_integer (of_int j mod 2)" if "j \<ge> 0" |
644 |
using that by transfer (simp add: nat_mod_distrib) |
|
645 |
from k show ?thesis |
|
646 |
by (auto simp add: split_def Let_def nat_of_integer_def nat_div_distrib mult_2 [symmetric] |
|
647 |
minus_mod_eq_mult_div [symmetric] *) |
|
33340
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents:
33318
diff
changeset
|
648 |
qed |
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
649 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
650 |
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
|
651 |
"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
|
652 |
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
|
653 |
else let |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
654 |
(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
|
655 |
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
|
656 |
in if j = 0 then l' else l' + 1)" |
64246 | 657 |
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
|
658 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
659 |
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
|
660 |
"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
|
661 |
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
|
662 |
else let |
60868
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
60758
diff
changeset
|
663 |
l = 2 * integer_of_int (k div 2); |
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
60758
diff
changeset
|
664 |
j = k mod 2 |
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
60758
diff
changeset
|
665 |
in if j = 0 then l else l + 1)" |
64246 | 666 |
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
|
667 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
668 |
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
|
669 |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
670 |
|
60758 | 671 |
subsection \<open>Serializer setup for target language integers\<close> |
24999 | 672 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
673 |
code_reserved Eval int Integer abs |
25767 | 674 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51375
diff
changeset
|
675 |
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
|
676 |
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
|
677 |
(SML) "IntInf.int" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
678 |
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
|
679 |
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
|
680 |
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
|
681 |
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
|
682 |
| 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
|
683 |
(Haskell) - |
24999 | 684 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51375
diff
changeset
|
685 |
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
|
686 |
constant "0::integer" \<rightharpoonup> |
58400
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
haftmann
parents:
58399
diff
changeset
|
687 |
(SML) "!(0/ :/ IntInf.int)" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
688 |
and (OCaml) "Z.zero" |
58400
d0d3c30806b4
always annotate potentially polymorphic Haskell numerals
haftmann
parents:
58399
diff
changeset
|
689 |
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
|
690 |
and (Scala) "BigInt(0)" |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46664
diff
changeset
|
691 |
|
60758 | 692 |
setup \<open> |
58399 | 693 |
fold (fn target => |
69593 | 694 |
Numeral.add_code \<^const_name>\<open>Code_Numeral.Pos\<close> I Code_Printer.literal_numeral target |
695 |
#> Numeral.add_code \<^const_name>\<open>Code_Numeral.Neg\<close> (~) Code_Printer.literal_numeral target) |
|
58399 | 696 |
["SML", "OCaml", "Haskell", "Scala"] |
60758 | 697 |
\<close> |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
698 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51375
diff
changeset
|
699 |
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
|
700 |
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
|
701 |
(SML) "IntInf.+ ((_), (_))" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
702 |
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
|
703 |
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
|
704 |
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
|
705 |
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
|
706 |
| 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
|
707 |
(SML) "IntInf.~" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
708 |
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
|
709 |
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
|
710 |
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
|
711 |
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
|
712 |
| 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
|
713 |
(SML) "IntInf.- ((_), (_))" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
714 |
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
|
715 |
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
|
716 |
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
|
717 |
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
|
718 |
| 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
|
719 |
(SML) "IntInf.*/ (2,/ (_))" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
720 |
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
|
721 |
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
|
722 |
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
|
723 |
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
|
724 |
| 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
|
725 |
(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
|
726 |
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
|
727 |
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
|
728 |
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
|
729 |
| 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
|
730 |
(SML) "IntInf.* ((_), (_))" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
731 |
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
|
732 |
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
|
733 |
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
|
734 |
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
|
735 |
| 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
|
736 |
(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
|
737 |
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
|
738 |
and (Haskell) "divMod/ (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
|
739 |
and (Scala) "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51375
diff
changeset
|
740 |
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
|
741 |
| 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
|
742 |
(SML) "!((_ : IntInf.int) = _)" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
743 |
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
|
744 |
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
|
745 |
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
|
746 |
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
|
747 |
| 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
|
748 |
(SML) "IntInf.<= ((_), (_))" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
749 |
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
|
750 |
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
|
751 |
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
|
752 |
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
|
753 |
| 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
|
754 |
(SML) "IntInf.< ((_), (_))" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
755 |
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
|
756 |
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
|
757 |
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
|
758 |
and (Eval) infixl 6 "<" |
61857
542f2c6da692
add serialisation for abs on integer to target language operation
Andreas Lochbihler
parents:
61275
diff
changeset
|
759 |
| constant "abs :: integer \<Rightarrow> _" \<rightharpoonup> |
542f2c6da692
add serialisation for abs on integer to target language operation
Andreas Lochbihler
parents:
61275
diff
changeset
|
760 |
(SML) "IntInf.abs" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
69593
diff
changeset
|
761 |
and (OCaml) "Z.abs" |
61857
542f2c6da692
add serialisation for abs on integer to target language operation
Andreas Lochbihler
parents:
61275
diff
changeset
|
762 |
and (Haskell) "Prelude.abs" |
542f2c6da692
add serialisation for abs on integer to target language operation
Andreas Lochbihler
parents:
61275
diff
changeset
|
763 |
and (Scala) "_.abs" |
542f2c6da692
add serialisation for abs on integer to target language operation
Andreas Lochbihler
parents:
61275
diff
changeset
|
764 |
and (Eval) "abs" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
765 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51375
diff
changeset
|
766 |
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
|
767 |
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
|
768 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
769 |
|
60758 | 770 |
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
|
771 |
|
61076 | 772 |
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
|
773 |
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
|
774 |
|
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
58889
diff
changeset
|
775 |
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
|
776 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
777 |
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
|
778 |
"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
|
779 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
780 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
781 |
lemma natural_eqI: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
782 |
"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
|
783 |
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
|
784 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
785 |
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
|
786 |
"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
|
787 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
788 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
789 |
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
|
790 |
"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
|
791 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
792 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
793 |
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
|
794 |
begin |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
795 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
796 |
lift_definition zero_natural :: natural |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
797 |
is "0 :: nat" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
798 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
799 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
800 |
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
|
801 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
802 |
lift_definition one_natural :: natural |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
803 |
is "1 :: nat" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
804 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
805 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
806 |
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
|
807 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
808 |
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
|
809 |
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
|
810 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
811 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
812 |
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
|
813 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
814 |
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
|
815 |
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
|
816 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
817 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
818 |
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
|
819 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
820 |
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
|
821 |
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
|
822 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
823 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
824 |
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
|
825 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
826 |
instance proof |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
827 |
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
|
828 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
829 |
end |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
830 |
|
64241
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
831 |
instance natural :: Rings.dvd .. |
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
832 |
|
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
833 |
lemma [transfer_rule]: |
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
834 |
"rel_fun pcr_natural (rel_fun pcr_natural HOL.iff) Rings.dvd Rings.dvd" |
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
835 |
unfolding dvd_def by transfer_prover |
430d74089d4d
transfer rules for divides relation on integer and natural
haftmann
parents:
64178
diff
changeset
|
836 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
837 |
lemma [transfer_rule]: |
68010 | 838 |
"rel_fun (=) pcr_natural (of_bool :: bool \<Rightarrow> nat) (of_bool :: bool \<Rightarrow> natural)" |
839 |
by (unfold of_bool_def [abs_def]) transfer_prover |
|
840 |
||
841 |
lemma [transfer_rule]: |
|
55945 | 842 |
"rel_fun HOL.eq pcr_natural (\<lambda>n::nat. n) (of_nat :: nat \<Rightarrow> natural)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
843 |
proof - |
55945 | 844 |
have "rel_fun HOL.eq pcr_natural (of_nat :: nat \<Rightarrow> nat) (of_nat :: nat \<Rightarrow> natural)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
845 |
by (unfold of_nat_def [abs_def]) transfer_prover |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
846 |
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
|
847 |
qed |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
848 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
849 |
lemma [transfer_rule]: |
55945 | 850 |
"rel_fun HOL.eq pcr_natural (numeral :: num \<Rightarrow> nat) (numeral :: num \<Rightarrow> natural)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
851 |
proof - |
55945 | 852 |
have "rel_fun HOL.eq pcr_natural (numeral :: num \<Rightarrow> nat) (\<lambda>n. of_nat (numeral n))" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
853 |
by transfer_prover |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
854 |
then show ?thesis by simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
855 |
qed |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
856 |
|
68010 | 857 |
lemma [transfer_rule]: |
858 |
"rel_fun pcr_natural (rel_fun (=) pcr_natural) (power :: _ \<Rightarrow> _ \<Rightarrow> nat) (power :: _ \<Rightarrow> _ \<Rightarrow> natural)" |
|
859 |
by (unfold power_def [abs_def]) transfer_prover |
|
860 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
861 |
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
|
862 |
"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
|
863 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
864 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
865 |
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
|
866 |
"natural_of_nat = of_nat" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
867 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
868 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
869 |
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
|
870 |
"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
|
871 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
872 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
873 |
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
|
874 |
"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
|
875 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
876 |
|
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
877 |
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
|
878 |
begin |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
879 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
880 |
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
|
881 |
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
|
882 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
883 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
884 |
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
|
885 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
886 |
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
|
887 |
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
|
888 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
889 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
890 |
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
|
891 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
892 |
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
|
893 |
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
|
894 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
895 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
896 |
instance proof |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
897 |
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
|
898 |
|
24999 | 899 |
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
|
900 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
901 |
lemma [transfer_rule]: |
55945 | 902 |
"rel_fun pcr_natural (rel_fun pcr_natural pcr_natural) (min :: _ \<Rightarrow> _ \<Rightarrow> nat) (min :: _ \<Rightarrow> _ \<Rightarrow> natural)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
903 |
by (unfold min_def [abs_def]) transfer_prover |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
904 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
905 |
lemma [transfer_rule]: |
55945 | 906 |
"rel_fun pcr_natural (rel_fun pcr_natural pcr_natural) (max :: _ \<Rightarrow> _ \<Rightarrow> nat) (max :: _ \<Rightarrow> _ \<Rightarrow> natural)" |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
907 |
by (unfold max_def [abs_def]) transfer_prover |
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 |
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
|
910 |
"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
|
911 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
912 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
913 |
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
|
914 |
"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
|
915 |
by transfer rule |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
916 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
917 |
instantiation natural :: unique_euclidean_semiring |
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
918 |
begin |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
919 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
920 |
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
|
921 |
is "divide :: nat \<Rightarrow> nat \<Rightarrow> nat" |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
922 |
. |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
923 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
924 |
declare divide_natural.rep_eq [simp] |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
925 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
926 |
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
|
927 |
is "modulo :: nat \<Rightarrow> nat \<Rightarrow> nat" |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
928 |
. |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
929 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
930 |
declare modulo_natural.rep_eq [simp] |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
931 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
932 |
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
|
933 |
is "euclidean_size :: nat \<Rightarrow> nat" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
934 |
. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
935 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
936 |
declare euclidean_size_natural.rep_eq [simp] |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
937 |
|
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
938 |
lift_definition division_segment_natural :: "natural \<Rightarrow> natural" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
939 |
is "division_segment :: nat \<Rightarrow> nat" |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
940 |
. |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
941 |
|
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
942 |
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
|
943 |
|
64592
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
944 |
instance |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
945 |
by (standard; transfer) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
946 |
(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
|
947 |
|
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
948 |
end |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
949 |
|
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
950 |
lemma [code]: |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
951 |
"euclidean_size = nat_of_natural" |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
952 |
by (simp add: fun_eq_iff) |
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
953 |
|
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
954 |
lemma [code]: |
66838
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
955 |
"division_segment (n::natural) = 1" |
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
haftmann
parents:
66836
diff
changeset
|
956 |
by (simp add: natural_eq_iff) |
66806
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents:
66801
diff
changeset
|
957 |
|
67905 | 958 |
instance natural :: linordered_semidom |
959 |
by (standard; transfer) simp_all |
|
960 |
||
70340 | 961 |
instance natural :: unique_euclidean_semiring_with_nat |
66839 | 962 |
by (standard; transfer) simp_all |
66815 | 963 |
|
68010 | 964 |
lemma [transfer_rule]: |
965 |
"rel_fun (=) (rel_fun pcr_natural pcr_natural) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> nat) (push_bit :: _ \<Rightarrow> _ \<Rightarrow> natural)" |
|
966 |
by (unfold push_bit_eq_mult [abs_def]) transfer_prover |
|
967 |
||
968 |
lemma [transfer_rule]: |
|
969 |
"rel_fun (=) (rel_fun pcr_natural pcr_natural) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> nat) (take_bit :: _ \<Rightarrow> _ \<Rightarrow> natural)" |
|
970 |
by (unfold take_bit_eq_mod [abs_def]) transfer_prover |
|
971 |
||
972 |
lemma [transfer_rule]: |
|
973 |
"rel_fun (=) (rel_fun pcr_natural pcr_natural) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> nat) (drop_bit :: _ \<Rightarrow> _ \<Rightarrow> natural)" |
|
974 |
by (unfold drop_bit_eq_div [abs_def]) transfer_prover |
|
975 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
976 |
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
|
977 |
is "nat :: int \<Rightarrow> nat" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
978 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
979 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
980 |
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
|
981 |
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
|
982 |
. |
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 |
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
|
985 |
"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
|
986 |
by transfer simp |
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 |
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
|
989 |
"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
|
990 |
by transfer auto |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
991 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
992 |
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
|
993 |
"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
|
994 |
by transfer rule |
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 |
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
|
997 |
"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
|
998 |
by transfer rule |
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 |
lemma [measure_function]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1001 |
"is_measure nat_of_natural" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1002 |
by (rule is_measure_trivial) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1003 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1004 |
|
60758 | 1005 |
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
|
1006 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1007 |
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
|
1008 |
is Nat.Suc |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1009 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1010 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1011 |
declare Suc.rep_eq [simp] |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1012 |
|
58306
117ba6cbe414
renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents:
57512
diff
changeset
|
1013 |
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
|
1014 |
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
|
1015 |
|
55416 | 1016 |
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
|
1017 |
fixes m :: natural |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1018 |
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
|
1019 |
shows P |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1020 |
using assms by transfer blast |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1021 |
|
67332 | 1022 |
instantiation natural :: size |
1023 |
begin |
|
58379
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1024 |
|
67332 | 1025 |
definition size_nat where [simp, code]: "size_nat = nat_of_natural" |
1026 |
||
1027 |
instance .. |
|
1028 |
||
1029 |
end |
|
58379
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1030 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1031 |
lemma natural_decr [termination_simp]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1032 |
"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
|
1033 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1034 |
|
58379
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1035 |
lemma natural_zero_minus_one: "(0::natural) - 1 = 0" |
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1036 |
by (rule zero_diff) |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1037 |
|
58379
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1038 |
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
|
1039 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1040 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1041 |
hide_const (open) Suc |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1042 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1043 |
|
60758 | 1044 |
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
|
1045 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1046 |
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
|
1047 |
is nat |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1048 |
. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1049 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1050 |
lemma [code_post]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1051 |
"Nat 0 = 0" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1052 |
"Nat 1 = 1" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1053 |
"Nat (numeral k) = numeral k" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1054 |
by (transfer, simp)+ |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1055 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1056 |
lemma [code abstype]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1057 |
"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
|
1058 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1059 |
|
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
61857
diff
changeset
|
1060 |
lemma [code]: |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
61857
diff
changeset
|
1061 |
"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
|
1062 |
by transfer simp |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1063 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1064 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1065 |
"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
|
1066 |
by simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1067 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1068 |
lemma [code_abbrev]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1069 |
"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
|
1070 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1071 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1072 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1073 |
"integer_of_natural 0 = 0" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1074 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1075 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1076 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1077 |
"integer_of_natural 1 = 1" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1078 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1079 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1080 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1081 |
"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
|
1082 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1083 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1084 |
lemma [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1085 |
"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
|
1086 |
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
|
1087 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1088 |
lemma [code, code_unfold]: |
55416 | 1089 |
"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
|
1090 |
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
|
1091 |
|
55642
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents:
55428
diff
changeset
|
1092 |
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
|
1093 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1094 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1095 |
"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
|
1096 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1097 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1098 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1099 |
"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
|
1100 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1101 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1102 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1103 |
"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
|
1104 |
by transfer simp |
7759f1766189
more fine-grained type class hierarchy for div and mod
haftmann
parents:
64246
diff
changeset
|
1105 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1106 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1107 |
"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
|
1108 |
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
|
1109 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1110 |
lemma [code abstract]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1111 |
"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
|
1112 |
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
|
1113 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1114 |
lemma [code]: |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1115 |
"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
|
1116 |
by transfer (simp add: equal) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1117 |
|
58379
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1118 |
lemma [code nbe]: "HOL.equal n (n::natural) \<longleftrightarrow> True" |
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1119 |
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
|
1120 |
|
58379
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1121 |
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
|
1122 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1123 |
|
58379
c044539a2bda
reintroduced an instantiation of 'size' for 'numerals'
blanchet
parents:
58377
diff
changeset
|
1124 |
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
|
1125 |
by transfer simp |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1126 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1127 |
hide_const (open) Nat |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1128 |
|
55736
f1ed1e9cd080
unregister lifting setup following the best practice of Lifting
kuncar
parents:
55642
diff
changeset
|
1129 |
lifting_update integer.lifting |
f1ed1e9cd080
unregister lifting setup following the best practice of Lifting
kuncar
parents:
55642
diff
changeset
|
1130 |
lifting_forget integer.lifting |
f1ed1e9cd080
unregister lifting setup following the best practice of Lifting
kuncar
parents:
55642
diff
changeset
|
1131 |
|
f1ed1e9cd080
unregister lifting setup following the best practice of Lifting
kuncar
parents:
55642
diff
changeset
|
1132 |
lifting_update natural.lifting |
f1ed1e9cd080
unregister lifting setup following the best practice of Lifting
kuncar
parents:
55642
diff
changeset
|
1133 |
lifting_forget natural.lifting |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1134 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1135 |
code_reflect Code_Numeral |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
61857
diff
changeset
|
1136 |
datatypes natural |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
61857
diff
changeset
|
1137 |
functions "Code_Numeral.Suc" "0 :: natural" "1 :: natural" |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
61857
diff
changeset
|
1138 |
"plus :: natural \<Rightarrow> _" "minus :: natural \<Rightarrow> _" |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
61857
diff
changeset
|
1139 |
"times :: natural \<Rightarrow> _" "divide :: natural \<Rightarrow> _" |
63950
cdc1e59aa513
syntactic type class for operation mod named after mod;
haftmann
parents:
63174
diff
changeset
|
1140 |
"modulo :: natural \<Rightarrow> _" |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
61857
diff
changeset
|
1141 |
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
|
1142 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1143 |
end |