src/HOL/Library/Code_Binary_Nat.thy
author wenzelm
Fri, 21 Mar 2025 22:26:18 +0100
changeset 82317 231b6d8231c6
parent 81974 f30022be9213
permissions -rw-r--r--
more uniform Proof_Display.print_results for theory and proof output --- avoid loss of information seen in src/Doc/JEdit/document/output-and-state.png (the first bad changeset is f8c412a45af8, see also 53b59fa42696);
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: 47108
diff changeset
     1
(*  Title:      HOL/Library/Code_Binary_Nat.thy
51113
222fb6cb2c3e factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
haftmann
parents: 50023
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
     3
*)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
     4
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     5
section \<open>Implementation of natural numbers as binary numerals\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
     6
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents: 47108
diff changeset
     7
theory Code_Binary_Nat
51113
222fb6cb2c3e factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
haftmann
parents: 50023
diff changeset
     8
imports Code_Abstract_Nat
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
     9
begin
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    10
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    11
text \<open>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    12
  When generating code for functions on natural numbers, the
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66148
diff changeset
    13
  canonical representation using \<^term>\<open>0::nat\<close> and
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66148
diff changeset
    14
  \<^term>\<open>Suc\<close> is unsuitable for computations involving large
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    15
  numbers.  This theory refines the representation of
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    16
  natural numbers for code generation to use binary
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    17
  numerals, which do not grow linear in size but logarithmic.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    18
\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    19
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    20
subsection \<open>Representation\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    21
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents: 47108
diff changeset
    22
code_datatype "0::nat" nat_of_num
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents: 47108
diff changeset
    23
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    24
lemma [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    25
  "num_of_nat 0 = Num.One"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    26
  "num_of_nat (nat_of_num k) = k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    27
  by (simp_all add: nat_of_num_inverse)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    28
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    29
lemma [code]:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60868
diff changeset
    30
  "(1::nat) = Numeral1"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    31
  by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    32
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60868
diff changeset
    33
lemma [code_abbrev]: "Numeral1 = (1::nat)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    34
  by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    35
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    36
lemma [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    37
  "Suc n = n + 1"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    38
  by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    39
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    40
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    41
subsection \<open>Basic arithmetic\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    42
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
    43
context
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
    44
begin
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
    45
66148
5e60c2d0a1f1 avoid ancient [code, code del] antipattern
haftmann
parents: 61433
diff changeset
    46
declare [[code drop: "plus :: nat \<Rightarrow> _"]]  
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    47
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    48
lemma plus_nat_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    49
  "nat_of_num k + nat_of_num l = nat_of_num (k + l)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    50
  "m + 0 = (m::nat)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    51
  "0 + n = (n::nat)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    52
  by (simp_all add: nat_of_num_numeral)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    53
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    54
text \<open>Bounded subtraction needs some auxiliary\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    55
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
    56
qualified definition dup :: "nat \<Rightarrow> nat" where
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    57
  "dup n = n + n"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    58
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    59
lemma dup_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    60
  "dup 0 = 0"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    61
  "dup (nat_of_num k) = nat_of_num (Num.Bit0 k)"
50023
28f3263d4d1b refined stack of library theories implementing int and/or nat by target language numerals
haftmann
parents: 47108
diff changeset
    62
  by (simp_all add: dup_def numeral_Bit0)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    63
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
    64
qualified definition sub :: "num \<Rightarrow> num \<Rightarrow> nat option" where
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    65
  "sub k l = (if k \<ge> l then Some (numeral k - numeral l) else None)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    66
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    67
lemma sub_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    68
  "sub Num.One Num.One = Some 0"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    69
  "sub (Num.Bit0 m) Num.One = Some (nat_of_num (Num.BitM m))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    70
  "sub (Num.Bit1 m) Num.One = Some (nat_of_num (Num.Bit0 m))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    71
  "sub Num.One (Num.Bit0 n) = None"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    72
  "sub Num.One (Num.Bit1 n) = None"
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 53069
diff changeset
    73
  "sub (Num.Bit0 m) (Num.Bit0 n) = map_option dup (sub m n)"
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 53069
diff changeset
    74
  "sub (Num.Bit1 m) (Num.Bit1 n) = map_option dup (sub m n)"
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 53069
diff changeset
    75
  "sub (Num.Bit1 m) (Num.Bit0 n) = map_option (\<lambda>q. dup q + 1) (sub m n)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    76
  "sub (Num.Bit0 m) (Num.Bit1 n) = (case sub m n of None \<Rightarrow> None
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    77
     | Some q \<Rightarrow> if q = 0 then None else Some (dup q - 1))"
81974
f30022be9213 Tidying more old proofs
paulson <lp15@cam.ac.uk>
parents: 77061
diff changeset
    78
  by (auto simp: nat_of_num_numeral Num.dbl_def Num.dbl_inc_def Num.dbl_dec_def
f30022be9213 Tidying more old proofs
paulson <lp15@cam.ac.uk>
parents: 77061
diff changeset
    79
    Let_def le_imp_diff_is_add BitM_plus_one sub_def dup_def 
f30022be9213 Tidying more old proofs
paulson <lp15@cam.ac.uk>
parents: 77061
diff changeset
    80
    sub_non_positive nat_add_distrib sub_non_negative)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    81
66148
5e60c2d0a1f1 avoid ancient [code, code del] antipattern
haftmann
parents: 61433
diff changeset
    82
declare [[code drop: "minus :: nat \<Rightarrow> _"]]
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    83
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    84
lemma minus_nat_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    85
  "nat_of_num k - nat_of_num l = (case sub k l of None \<Rightarrow> 0 | Some j \<Rightarrow> j)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    86
  "m - 0 = (m::nat)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    87
  "0 - n = (0::nat)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    88
  by (simp_all add: nat_of_num_numeral sub_non_positive sub_def)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    89
66148
5e60c2d0a1f1 avoid ancient [code, code del] antipattern
haftmann
parents: 61433
diff changeset
    90
declare [[code drop: "times :: nat \<Rightarrow> _"]]
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    91
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    92
lemma times_nat_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    93
  "nat_of_num k * nat_of_num l = nat_of_num (k * l)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    94
  "m * 0 = (0::nat)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    95
  "0 * n = (0::nat)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    96
  by (simp_all add: nat_of_num_numeral)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    97
66148
5e60c2d0a1f1 avoid ancient [code, code del] antipattern
haftmann
parents: 61433
diff changeset
    98
declare [[code drop: "HOL.equal :: nat \<Rightarrow> _"]]
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
    99
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   100
lemma equal_nat_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   101
  "HOL.equal 0 (0::nat) \<longleftrightarrow> True"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   102
  "HOL.equal 0 (nat_of_num l) \<longleftrightarrow> False"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   103
  "HOL.equal (nat_of_num k) 0 \<longleftrightarrow> False"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   104
  "HOL.equal (nat_of_num k) (nat_of_num l) \<longleftrightarrow> HOL.equal k l"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   105
  by (simp_all add: nat_of_num_numeral equal)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   106
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   107
lemma equal_nat_refl [code nbe]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   108
  "HOL.equal (n::nat) n \<longleftrightarrow> True"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   109
  by (rule equal_refl)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   110
66148
5e60c2d0a1f1 avoid ancient [code, code del] antipattern
haftmann
parents: 61433
diff changeset
   111
declare [[code drop: "less_eq :: nat \<Rightarrow> _"]]
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   112
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   113
lemma less_eq_nat_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   114
  "0 \<le> (n::nat) \<longleftrightarrow> True"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   115
  "nat_of_num k \<le> 0 \<longleftrightarrow> False"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   116
  "nat_of_num k \<le> nat_of_num l \<longleftrightarrow> k \<le> l"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   117
  by (simp_all add: nat_of_num_numeral)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   118
66148
5e60c2d0a1f1 avoid ancient [code, code del] antipattern
haftmann
parents: 61433
diff changeset
   119
declare [[code drop: "less :: nat \<Rightarrow> _"]]
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   120
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   121
lemma less_nat_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   122
  "(m::nat) < 0 \<longleftrightarrow> False"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   123
  "0 < nat_of_num l \<longleftrightarrow> True"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   124
  "nat_of_num k < nat_of_num l \<longleftrightarrow> k < l"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   125
  by (simp_all add: nat_of_num_numeral)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   126
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   127
declare [[code drop: Euclidean_Rings.divmod_nat]]
53069
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   128
  
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   129
lemma divmod_nat_code [code]:
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   130
  "Euclidean_Rings.divmod_nat (nat_of_num k) (nat_of_num l) = divmod k l"
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   131
  "Euclidean_Rings.divmod_nat m 0 = (0, m)"
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   132
  "Euclidean_Rings.divmod_nat 0 n = (0, 0)"
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 75937
diff changeset
   133
  by (simp_all add: Euclidean_Rings.divmod_nat_def nat_of_num_numeral)
53069
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   134
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   135
end
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   136
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   137
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   138
subsection \<open>Conversions\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   139
66148
5e60c2d0a1f1 avoid ancient [code, code del] antipattern
haftmann
parents: 61433
diff changeset
   140
declare [[code drop: of_nat]]
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   141
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   142
lemma of_nat_code [code]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   143
  "of_nat 0 = 0"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   144
  "of_nat (nat_of_num k) = numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   145
  by (simp_all add: nat_of_num_numeral)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   146
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   147
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51143
diff changeset
   148
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
   149
  code_module Code_Binary_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
   150
    (SML) Arith and (OCaml) Arith and (Haskell) Arith
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   151
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents:
diff changeset
   152
end