src/HOL/Library/Code_Target_Nat.thy
author wenzelm
Thu, 10 Apr 2025 13:15:57 +0200
changeset 82472 d4b3eea69371
parent 82452 8b575e1fef3b
child 82773 4ec8e654112f
permissions -rw-r--r--
more robust shell script;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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: 51095
diff 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
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     5
section \<open>Implementation of natural numbers by target-language integers\<close>
50023
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: 51114
diff 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
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    11
subsection \<open>Implementation for \<^typ>\<open>nat\<close>\<close>
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    12
55736
f1ed1e9cd080 unregister lifting setup following the best practice of Lifting
kuncar
parents: 54796
diff changeset
    13
context
82452
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
    14
includes integer.lifting
55736
f1ed1e9cd080 unregister lifting setup following the best practice of Lifting
kuncar
parents: 54796
diff changeset
    15
begin
f1ed1e9cd080 unregister lifting setup following the best practice of Lifting
kuncar
parents: 54796
diff changeset
    16
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    17
lift_definition Nat :: "integer \<Rightarrow> nat"
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    18
  is nat
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    19
  .
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    20
51095
7ae79f2e3cc7 explicit conversion integer_of_nat already in Code_Numeral_Types;
haftmann
parents: 50023
diff changeset
    21
lemma [code_post]:
7ae79f2e3cc7 explicit conversion integer_of_nat already in Code_Numeral_Types;
haftmann
parents: 50023
diff changeset
    22
  "Nat 0 = 0"
7ae79f2e3cc7 explicit conversion integer_of_nat already in Code_Numeral_Types;
haftmann
parents: 50023
diff changeset
    23
  "Nat 1 = 1"
7ae79f2e3cc7 explicit conversion integer_of_nat already in Code_Numeral_Types;
haftmann
parents: 50023
diff changeset
    24
  "Nat (numeral k) = numeral k"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    25
  by (transfer, simp)+
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    26
51095
7ae79f2e3cc7 explicit conversion integer_of_nat already in Code_Numeral_Types;
haftmann
parents: 50023
diff changeset
    27
lemma [code_abbrev]:
7ae79f2e3cc7 explicit conversion integer_of_nat already in Code_Numeral_Types;
haftmann
parents: 50023
diff changeset
    28
  "integer_of_nat = of_nat"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    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_unfold]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    32
  "Int.nat (int_of_integer k) = nat_of_integer k"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    33
  by transfer rule
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 abstype]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    36
  "Code_Target_Nat.Nat (integer_of_nat n) = n"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    37
  by transfer simp
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 abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    40
  "integer_of_nat (nat_of_integer k) = max 0 k"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    41
  by transfer auto
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_abbrev]:
51095
7ae79f2e3cc7 explicit conversion integer_of_nat already in Code_Numeral_Types;
haftmann
parents: 50023
diff changeset
    44
  "nat_of_integer (numeral k) = nat_of_num k"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    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
64997
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    47
context
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    48
begin  
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    49
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    50
qualified definition natural :: "num \<Rightarrow> nat"
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    51
  where [simp]: "natural = nat_of_num"
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    52
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    53
lemma [code_computation_unfold]:
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    54
  "numeral = natural"
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    55
  "nat_of_num = natural"
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    56
  by (simp_all add: nat_of_num_numeral)
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    57
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    58
end
067a6cca39f0 dedicated computation preprocessing rules for nat, int implemented by target language literals
haftmann
parents: 64242
diff changeset
    59
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    60
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    61
  "integer_of_nat (nat_of_num n) = integer_of_num n"
66190
a41435469559 more direct construction of integer_of_num;
haftmann
parents: 64997
diff changeset
    62
  by (simp add: nat_of_num_numeral integer_of_nat_numeral)
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    63
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    64
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    65
  "integer_of_nat 0 = 0"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    66
  by transfer simp
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    67
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    68
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    69
  "integer_of_nat 1 = 1"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    70
  by transfer simp
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    71
51113
222fb6cb2c3e factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
haftmann
parents: 51095
diff changeset
    72
lemma [code]:
222fb6cb2c3e factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
haftmann
parents: 51095
diff changeset
    73
  "Suc n = n + 1"
222fb6cb2c3e factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
haftmann
parents: 51095
diff changeset
    74
  by simp
222fb6cb2c3e factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
haftmann
parents: 51095
diff changeset
    75
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    76
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    77
  "integer_of_nat (m + n) = of_nat m + of_nat n"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    78
  by transfer simp
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    79
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    80
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    81
  "integer_of_nat (m - n) = max 0 (of_nat m - of_nat n)"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    82
  by transfer simp
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    83
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    84
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    85
  "integer_of_nat (m * n) = of_nat m * of_nat n"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    86
  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
    87
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    88
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    89
  "integer_of_nat (m div n) = of_nat m div of_nat n"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    90
  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
    91
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    92
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
    93
  "integer_of_nat (m mod n) = of_nat m mod of_nat n"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
    94
  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
    95
69946
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
    96
context
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
    97
  includes integer.lifting
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
    98
begin
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
    99
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   100
lemma divmod_nat_code [code]: \<^marker>\<open>contributor \<open>René Thiemann\<close>\<close> \<^marker>\<open>contributor \<open>Akihisa Yamada\<close>\<close>
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   101
  "Euclidean_Rings.divmod_nat m n = (
69946
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   102
     let k = integer_of_nat m; l = integer_of_nat n
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   103
     in map_prod nat_of_integer nat_of_integer
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   104
       (if k = 0 then (0, 0)
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   105
        else if l = 0 then (0, k) else
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   106
          Code_Numeral.divmod_abs k l))"
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   107
  by (simp add: prod_eq_iff Let_def Euclidean_Rings.divmod_nat_def; transfer)
69946
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   108
    (simp add: nat_div_distrib nat_mod_distrib)
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   109
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   110
end
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   111
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   112
lemma [code]:
61275
053ec04ea866 monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents: 60500
diff changeset
   113
  "divmod m n = map_prod nat_of_integer nat_of_integer (divmod m n)"
69946
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   114
  by (simp only: prod_eq_iff divmod_def map_prod_def case_prod_beta fst_conv snd_conv; transfer)
494934c30f38 improved code equations taken over from AFP
haftmann
parents: 69593
diff changeset
   115
    (simp_all only: nat_div_distrib nat_mod_distrib
61275
053ec04ea866 monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents: 60500
diff changeset
   116
        zero_le_numeral nat_numeral)
053ec04ea866 monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents: 60500
diff changeset
   117
  
053ec04ea866 monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
haftmann
parents: 60500
diff changeset
   118
lemma [code]:
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   119
  "HOL.equal m n = HOL.equal (of_nat m :: integer) (of_nat n)"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
   120
  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
   121
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   122
lemma [code]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   123
  "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
   124
  by simp
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
lemma [code]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   127
  "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
   128
  by simp
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   129
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   130
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
   131
  "num_of_nat = num_of_integer \<circ> of_nat"
51114
3e913a575dc6 type lifting setup for code numeral types
haftmann
parents: 51113
diff changeset
   132
  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
   133
55736
f1ed1e9cd080 unregister lifting setup following the best practice of Lifting
kuncar
parents: 54796
diff changeset
   134
end
f1ed1e9cd080 unregister lifting setup following the best practice of Lifting
kuncar
parents: 54796
diff changeset
   135
54796
cdc6d8cbf770 avoid clashes of fact names
haftmann
parents: 53027
diff changeset
   136
lemma (in semiring_1) of_nat_code_if:
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   137
  "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
   138
     else let
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   139
       (m, q) = Euclidean_Rings.divmod_nat n 2;
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   140
       m' = 2 * of_nat m
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   141
     in if q = 0 then m' else m' + 1)"
75937
02b18f59f903 streamlined
haftmann
parents: 69946
diff changeset
   142
  by (cases n)
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   143
    (simp_all add: Let_def Euclidean_Rings.divmod_nat_def ac_simps
75937
02b18f59f903 streamlined
haftmann
parents: 69946
diff changeset
   144
      flip: of_nat_numeral of_nat_mult minus_mod_eq_mult_div)
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   145
54796
cdc6d8cbf770 avoid clashes of fact names
haftmann
parents: 53027
diff changeset
   146
declare of_nat_code_if [code]
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   147
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   148
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
   149
  [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
   150
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   151
lemma [code]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   152
  "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
   153
  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
   154
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   155
lemma [code abstract]:
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   156
  "integer_of_nat (nat k) = max 0 (integer_of_int k)"
55736
f1ed1e9cd080 unregister lifting setup following the best practice of Lifting
kuncar
parents: 54796
diff changeset
   157
  including integer.lifting by transfer auto
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   158
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   159
definition char_of_nat :: "nat \<Rightarrow> char"
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   160
  where [code_abbrev]: "char_of_nat = char_of"
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   161
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   162
definition nat_of_char :: "char \<Rightarrow> nat"
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   163
  where [code_abbrev]: "nat_of_char = of_char"
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   164
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   165
lemma [code]:
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   166
  "char_of_nat = char_of_integer \<circ> integer_of_nat"
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   167
  including integer.lifting unfolding char_of_integer_def char_of_nat_def
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   168
  by transfer (simp add: fun_eq_iff)
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   169
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   170
lemma [code abstract]:
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   171
  "integer_of_nat (nat_of_char c) = integer_of_char c"
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   172
  by (cases c) (simp add: nat_of_char_def integer_of_char_def integer_of_nat_eq_of_nat)
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 66808
diff changeset
   173
53027
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   174
lemma term_of_nat_code [code]:
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
   175
  \<comment> \<open>Use \<^term>\<open>Code_Numeral.nat_of_integer\<close> in term reconstruction
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
   176
        instead of \<^term>\<open>Code_Target_Nat.Nat\<close> such that reconstructed
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   177
        terms can be fed back to the code generator\<close>
53027
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   178
  "term_of_class.term_of n =
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   179
   Code_Evaluation.App
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   180
     (Code_Evaluation.Const (STR ''Code_Numeral.nat_of_integer'')
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   181
        (typerep.Typerep (STR ''fun'')
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   182
           [typerep.Typerep (STR ''Code_Numeral.integer'') [],
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   183
         typerep.Typerep (STR ''Nat.nat'') []]))
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   184
     (term_of_class.term_of (integer_of_nat n))"
54796
cdc6d8cbf770 avoid clashes of fact names
haftmann
parents: 53027
diff changeset
   185
  by (simp add: term_of_anything)
53027
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   186
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   187
lemma nat_of_integer_code_post [code_post]:
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   188
  "nat_of_integer 0 = 0"
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   189
  "nat_of_integer 1 = 1"
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   190
  "nat_of_integer (numeral k) = numeral k"
55736
f1ed1e9cd080 unregister lifting setup following the best practice of Lifting
kuncar
parents: 54796
diff changeset
   191
  including integer.lifting by (transfer, simp)+
53027
1774d569b604 use nat_of_integer for term reconstruction instead of abstract constructor to allow reconstructed terms being fed back to the code generator
Andreas Lochbihler
parents: 52435
diff changeset
   192
82452
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   193
context
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   194
  includes integer.lifting and bit_operations_syntax
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   195
begin
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   196
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   197
declare [[code drop: \<open>bit :: nat \<Rightarrow> _\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   198
  \<open>and :: nat \<Rightarrow> _\<close> \<open>or :: nat \<Rightarrow> _\<close> \<open>xor :: nat \<Rightarrow> _\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   199
  \<open>push_bit :: _ \<Rightarrow> _ \<Rightarrow> nat\<close> \<open>drop_bit :: _ \<Rightarrow> _ \<Rightarrow> nat\<close> \<open>take_bit :: _ \<Rightarrow> _ \<Rightarrow> nat\<close>]]
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   200
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   201
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   202
  \<open>bit m n \<longleftrightarrow> bit (integer_of_nat m) n\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   203
  by transfer (simp add: bit_simps)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   204
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   205
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   206
  \<open>integer_of_nat (m AND n) = integer_of_nat m AND integer_of_nat n\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   207
  by transfer (simp add: of_nat_and_eq)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   208
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   209
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   210
  \<open>integer_of_nat (m OR n) = integer_of_nat m OR integer_of_nat n\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   211
  by transfer (simp add: of_nat_or_eq)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   212
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   213
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   214
  \<open>integer_of_nat (m XOR n) = integer_of_nat m XOR integer_of_nat n\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   215
  by transfer (simp add: of_nat_xor_eq)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   216
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   217
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   218
  \<open>integer_of_nat (push_bit n m) = push_bit n (integer_of_nat m)\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   219
  by transfer (simp add: of_nat_push_bit)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   220
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   221
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   222
  \<open>integer_of_nat (drop_bit n m) = drop_bit n (integer_of_nat m)\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   223
  by transfer (simp add: of_nat_drop_bit)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   224
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   225
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   226
  \<open>integer_of_nat (take_bit n m) = take_bit n (integer_of_nat m)\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   227
  by transfer (simp add: of_nat_take_bit)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   228
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   229
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   230
  \<open>integer_of_nat (mask n) = mask n\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   231
  by transfer (simp add: of_nat_mask_eq)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   232
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   233
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   234
  \<open>integer_of_nat (set_bit n m) = set_bit n (integer_of_nat m)\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   235
  by transfer (simp add: of_nat_set_bit_eq)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   236
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   237
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   238
  \<open>integer_of_nat (unset_bit n m) = unset_bit n (integer_of_nat m)\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   239
  by transfer (simp add: of_nat_unset_bit_eq)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   240
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   241
lemma [code]:
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   242
  \<open>integer_of_nat (flip_bit n m) = flip_bit n (integer_of_nat m)\<close>
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   243
  by transfer (simp add: of_nat_flip_bit_eq)
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   244
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   245
end
8b575e1fef3b use existing implementations of bit operations if nat is implemented by target-language integer
haftmann
parents: 81113
diff changeset
   246
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51143
diff changeset
   247
code_identifier
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51143
diff changeset
   248
  code_module Code_Target_Nat \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51143
diff changeset
   249
    (SML) Arith and (OCaml) Arith and (Haskell) Arith
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   250
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents:
diff changeset
   251
end