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