| author | hoelzl | 
| Tue, 26 Mar 2013 12:20:59 +0100 | |
| changeset 51528 | 66c3a7589de7 | 
| parent 51143 | 0a2371e7ced3 | 
| child 52435 | 6646bb548c6b | 
| permissions | -rw-r--r-- | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 1 | (* Title: HOL/Library/Code_Target_Nat.thy | 
| 51113 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 haftmann parents: 
51095diff
changeset | 2 | Author: Florian Haftmann, TU Muenchen | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 3 | *) | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 4 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 5 | header {* Implementation of natural numbers by target-language integers *}
 | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 6 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 7 | theory Code_Target_Nat | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
51114diff
changeset | 8 | imports Code_Abstract_Nat | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 9 | begin | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 10 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 11 | subsection {* Implementation for @{typ nat} *}
 | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 12 | |
| 51114 | 13 | lift_definition Nat :: "integer \<Rightarrow> nat" | 
| 14 | is nat | |
| 15 | . | |
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 16 | |
| 51095 
7ae79f2e3cc7
explicit conversion integer_of_nat already in Code_Numeral_Types;
 haftmann parents: 
50023diff
changeset | 17 | lemma [code_post]: | 
| 
7ae79f2e3cc7
explicit conversion integer_of_nat already in Code_Numeral_Types;
 haftmann parents: 
50023diff
changeset | 18 | "Nat 0 = 0" | 
| 
7ae79f2e3cc7
explicit conversion integer_of_nat already in Code_Numeral_Types;
 haftmann parents: 
50023diff
changeset | 19 | "Nat 1 = 1" | 
| 
7ae79f2e3cc7
explicit conversion integer_of_nat already in Code_Numeral_Types;
 haftmann parents: 
50023diff
changeset | 20 | "Nat (numeral k) = numeral k" | 
| 51114 | 21 | by (transfer, simp)+ | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 22 | |
| 51095 
7ae79f2e3cc7
explicit conversion integer_of_nat already in Code_Numeral_Types;
 haftmann parents: 
50023diff
changeset | 23 | lemma [code_abbrev]: | 
| 
7ae79f2e3cc7
explicit conversion integer_of_nat already in Code_Numeral_Types;
 haftmann parents: 
50023diff
changeset | 24 | "integer_of_nat = of_nat" | 
| 51114 | 25 | by transfer rule | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 26 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 27 | lemma [code_unfold]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 28 | "Int.nat (int_of_integer k) = nat_of_integer k" | 
| 51114 | 29 | by transfer rule | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 30 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 31 | lemma [code abstype]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 32 | "Code_Target_Nat.Nat (integer_of_nat n) = n" | 
| 51114 | 33 | by transfer simp | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 34 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 35 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 36 | "integer_of_nat (nat_of_integer k) = max 0 k" | 
| 51114 | 37 | by transfer auto | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 38 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 39 | lemma [code_abbrev]: | 
| 51095 
7ae79f2e3cc7
explicit conversion integer_of_nat already in Code_Numeral_Types;
 haftmann parents: 
50023diff
changeset | 40 | "nat_of_integer (numeral k) = nat_of_num k" | 
| 51114 | 41 | by transfer (simp add: nat_of_num_numeral) | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 42 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 43 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 44 | "integer_of_nat (nat_of_num n) = integer_of_num n" | 
| 51114 | 45 | by transfer (simp add: nat_of_num_numeral) | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 46 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 47 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 48 | "integer_of_nat 0 = 0" | 
| 51114 | 49 | by transfer simp | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 50 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 51 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 52 | "integer_of_nat 1 = 1" | 
| 51114 | 53 | by transfer simp | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 54 | |
| 51113 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 haftmann parents: 
51095diff
changeset | 55 | lemma [code]: | 
| 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 haftmann parents: 
51095diff
changeset | 56 | "Suc n = n + 1" | 
| 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 haftmann parents: 
51095diff
changeset | 57 | by simp | 
| 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 haftmann parents: 
51095diff
changeset | 58 | |
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 59 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 60 | "integer_of_nat (m + n) = of_nat m + of_nat n" | 
| 51114 | 61 | by transfer simp | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 62 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 63 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 64 | "integer_of_nat (m - n) = max 0 (of_nat m - of_nat n)" | 
| 51114 | 65 | by transfer simp | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 66 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 67 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 68 | "integer_of_nat (m * n) = of_nat m * of_nat n" | 
| 51114 | 69 | by transfer (simp add: of_nat_mult) | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 70 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 71 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 72 | "integer_of_nat (m div n) = of_nat m div of_nat n" | 
| 51114 | 73 | by transfer (simp add: zdiv_int) | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 74 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 75 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 76 | "integer_of_nat (m mod n) = of_nat m mod of_nat n" | 
| 51114 | 77 | by transfer (simp add: zmod_int) | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 78 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 79 | lemma [code]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 80 | "Divides.divmod_nat m n = (m div n, m mod n)" | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 81 | by (simp add: prod_eq_iff) | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 82 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 83 | lemma [code]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 84 | "HOL.equal m n = HOL.equal (of_nat m :: integer) (of_nat n)" | 
| 51114 | 85 | by transfer (simp add: equal) | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 86 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 87 | lemma [code]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 88 | "m \<le> n \<longleftrightarrow> (of_nat m :: integer) \<le> of_nat n" | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 89 | by simp | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 90 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 91 | lemma [code]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 92 | "m < n \<longleftrightarrow> (of_nat m :: integer) < of_nat n" | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 93 | by simp | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 94 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 95 | lemma num_of_nat_code [code]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 96 | "num_of_nat = num_of_integer \<circ> of_nat" | 
| 51114 | 97 | by transfer (simp add: fun_eq_iff) | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 98 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 99 | lemma (in semiring_1) of_nat_code: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 100 | "of_nat n = (if n = 0 then 0 | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 101 | else let | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 102 | (m, q) = divmod_nat n 2; | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 103 | m' = 2 * of_nat m | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 104 | in if q = 0 then m' else m' + 1)" | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 105 | proof - | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 106 | from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 107 | show ?thesis | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 108 | by (simp add: Let_def divmod_nat_div_mod mod_2_not_eq_zero_eq_one_nat | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 109 | of_nat_add [symmetric]) | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 110 | (simp add: * mult_commute of_nat_mult add_commute) | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 111 | qed | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 112 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 113 | declare of_nat_code [code] | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 114 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 115 | definition int_of_nat :: "nat \<Rightarrow> int" where | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 116 | [code_abbrev]: "int_of_nat = of_nat" | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 117 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 118 | lemma [code]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 119 | "int_of_nat n = int_of_integer (of_nat n)" | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 120 | by (simp add: int_of_nat_def) | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 121 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 122 | lemma [code abstract]: | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 123 | "integer_of_nat (nat k) = max 0 (integer_of_int k)" | 
| 51114 | 124 | by transfer auto | 
| 50023 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 125 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 126 | code_modulename SML | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 127 | Code_Target_Nat Arith | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 128 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 129 | code_modulename OCaml | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 130 | Code_Target_Nat Arith | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 131 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 132 | code_modulename Haskell | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 133 | Code_Target_Nat Arith | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 134 | |
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 135 | end | 
| 
28f3263d4d1b
refined stack of library theories implementing int and/or nat by target language numerals
 haftmann parents: diff
changeset | 136 |