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