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