src/HOL/Code_Numeral.thy
author blanchet
Mon, 30 Sep 2013 15:10:18 +0200
changeset 53997 8ff43f638da2
parent 53069 d165213e3924
child 54489 03ff4d1e6784
permissions -rw-r--r--
more "primrec_new" documentation
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
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
     5
header {* Numeric types for code generation onto target language numerals only *}
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
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
     8
imports Nat_Transfer Divides Lifting
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
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    11
subsection {* Type of target language integers *}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    12
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    13
typedef integer = "UNIV \<Colon> int set"
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
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    16
setup_lifting (no_code) type_definition_integer
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
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    78
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
    79
  "fun_rel HOL.eq pcr_integer (of_nat :: nat \<Rightarrow> int) (of_nat :: nat \<Rightarrow> integer)"
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
    80
  by (unfold of_nat_def [abs_def]) transfer_prover
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    81
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    82
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
    83
  "fun_rel HOL.eq pcr_integer (\<lambda>k :: int. k :: int) (of_int :: int \<Rightarrow> integer)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    84
proof -
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
    85
  have "fun_rel HOL.eq pcr_integer (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> integer)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    86
    by (unfold of_int_of_nat [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    87
  then show ?thesis by (simp add: id_def)
24999
haftmann
parents:
diff changeset
    88
qed
haftmann
parents:
diff changeset
    89
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    90
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
    91
  "fun_rel HOL.eq pcr_integer (numeral :: num \<Rightarrow> int) (numeral :: num \<Rightarrow> integer)"
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
    92
proof -
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
    93
  have "fun_rel HOL.eq pcr_integer (numeral :: num \<Rightarrow> int) (\<lambda>n. of_int (numeral n))"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    94
    by transfer_prover
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
    95
  then show ?thesis by simp
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
    96
qed
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
    97
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
    98
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
    99
  "fun_rel HOL.eq pcr_integer (neg_numeral :: num \<Rightarrow> int) (neg_numeral :: num \<Rightarrow> integer)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   100
  by (unfold neg_numeral_def [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   101
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   102
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   103
  "fun_rel HOL.eq (fun_rel HOL.eq pcr_integer) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> int) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> integer)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   104
  by (unfold Num.sub_def [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   105
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   106
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
   107
  "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
   108
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   109
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   110
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
   111
  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
   112
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   113
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   114
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
   115
  "integer_of_nat = of_nat"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   116
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   117
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   118
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
   119
  "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
   120
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   121
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   122
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
   123
  is Int.nat
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   124
  .
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   125
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   126
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
   127
  "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
   128
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   129
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   130
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
   131
  "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
   132
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   133
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   134
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
   135
  "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
   136
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   137
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   138
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
   139
  "integer_of_int = of_int"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   140
  by transfer (simp add: fun_eq_iff)
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   141
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   142
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
   143
  "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
   144
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   145
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   146
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
   147
  "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
   148
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   149
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   150
lemma int_of_integer_neg_numeral [simp]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   151
  "int_of_integer (neg_numeral k) = neg_numeral k"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   152
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   153
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   154
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
   155
  "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
   156
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   157
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   158
instantiation integer :: "{ring_div, equal, linordered_idom}"
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   159
begin
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   160
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   161
lift_definition div_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
   162
  is "Divides.div :: int \<Rightarrow> int \<Rightarrow> int"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   163
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   164
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   165
declare div_integer.rep_eq [simp]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   166
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   167
lift_definition mod_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
   168
  is "Divides.mod :: int \<Rightarrow> int \<Rightarrow> int"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   169
  .
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
declare mod_integer.rep_eq [simp]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   172
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   173
lift_definition abs_integer :: "integer \<Rightarrow> integer"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   174
  is "abs :: int \<Rightarrow> int"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   175
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   176
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   177
declare abs_integer.rep_eq [simp]
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   178
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   179
lift_definition sgn_integer :: "integer \<Rightarrow> integer"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   180
  is "sgn :: int \<Rightarrow> int"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   181
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   182
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   183
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
   184
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   185
lift_definition less_eq_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   186
  is "less_eq :: int \<Rightarrow> int \<Rightarrow> bool"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   187
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   188
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   189
lift_definition less_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   190
  is "less :: int \<Rightarrow> int \<Rightarrow> bool"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   191
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   192
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   193
lift_definition equal_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   194
  is "HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   195
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   196
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   197
instance proof
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   198
qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   199
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   200
end
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   201
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   202
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   203
  "fun_rel pcr_integer (fun_rel pcr_integer pcr_integer) (min :: _ \<Rightarrow> _ \<Rightarrow> int) (min :: _ \<Rightarrow> _ \<Rightarrow> integer)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   204
  by (unfold min_def [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   205
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   206
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   207
  "fun_rel pcr_integer (fun_rel pcr_integer pcr_integer) (max :: _ \<Rightarrow> _ \<Rightarrow> int) (max :: _ \<Rightarrow> _ \<Rightarrow> integer)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   208
  by (unfold max_def [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   209
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   210
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
   211
  "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
   212
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   213
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   214
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
   215
  "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
   216
  by transfer rule
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   217
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   218
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
   219
  "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
   220
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   221
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   222
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
   223
  "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
   224
  by transfer auto
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   225
53069
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   226
instance integer :: semiring_numeral_div
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   227
  by intro_classes (transfer,
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   228
    fact semiring_numeral_div_class.diff_invert_add1
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   229
    semiring_numeral_div_class.le_add_diff_inverse2
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   230
    semiring_numeral_div_class.mult_div_cancel
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   231
    semiring_numeral_div_class.div_less
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   232
    semiring_numeral_div_class.mod_less
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   233
    semiring_numeral_div_class.div_positive
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   234
    semiring_numeral_div_class.mod_less_eq_dividend
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   235
    semiring_numeral_div_class.pos_mod_bound
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   236
    semiring_numeral_div_class.pos_mod_sign
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   237
    semiring_numeral_div_class.mod_mult2_eq
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   238
    semiring_numeral_div_class.div_mult2_eq
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   239
    semiring_numeral_div_class.discrete)+
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   240
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   241
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   242
subsection {* Code theorems for target language integers *}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   243
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   244
text {* Constructors *}
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   245
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   246
definition Pos :: "num \<Rightarrow> integer"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   247
where
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   248
  [simp, code_abbrev]: "Pos = numeral"
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]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   251
  "fun_rel HOL.eq pcr_integer numeral Pos"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   252
  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
   253
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   254
definition Neg :: "num \<Rightarrow> integer"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   255
where
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   256
  [simp, code_abbrev]: "Neg = neg_numeral"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   257
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   258
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   259
  "fun_rel HOL.eq pcr_integer neg_numeral Neg"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   260
  by simp transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   261
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   262
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
   263
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   264
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   265
text {* Auxiliary operations *}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   266
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   267
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
   268
  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
   269
  .
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   270
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   271
lemma dup_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   272
  "dup 0 = 0"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   273
  "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
   274
  "dup (Neg n) = Neg (Num.Bit0 n)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   275
  by (transfer, simp only: neg_numeral_def numeral_Bit0 minus_add_distrib)+
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   276
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   277
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
   278
  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
   279
  .
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   280
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   281
lemma sub_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   282
  "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
   283
  "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
   284
  "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
   285
  "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
   286
  "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
   287
  "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
   288
  "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
   289
  "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
   290
  "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   291
  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
   292
24999
haftmann
parents:
diff changeset
   293
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   294
text {* Implementations *}
24999
haftmann
parents:
diff changeset
   295
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   296
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
   297
  "1 = Pos Num.One"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   298
  by simp
24999
haftmann
parents:
diff changeset
   299
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   300
lemma plus_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   301
  "k + 0 = (k::integer)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   302
  "0 + l = (l::integer)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   303
  "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
   304
  "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
   305
  "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
   306
  "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
   307
  by (transfer, simp)+
24999
haftmann
parents:
diff changeset
   308
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   309
lemma uminus_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   310
  "uminus 0 = (0::integer)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   311
  "uminus (Pos m) = Neg m"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   312
  "uminus (Neg m) = Pos m"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   313
  by simp_all
28708
a1a436f09ec6 explicit check for pattern discipline before code translation
haftmann
parents: 28562
diff changeset
   314
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   315
lemma minus_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   316
  "k - 0 = (k::integer)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   317
  "0 - l = uminus (l::integer)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   318
  "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
   319
  "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
   320
  "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
   321
  "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
   322
  by (transfer, simp)+
46028
9f113cdf3d66 attribute code_abbrev superseedes code_unfold_post
haftmann
parents: 45694
diff changeset
   323
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   324
lemma abs_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   325
  "\<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
   326
  by simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46664
diff changeset
   327
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   328
lemma sgn_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   329
  "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
   330
  by simp
46028
9f113cdf3d66 attribute code_abbrev superseedes code_unfold_post
haftmann
parents: 45694
diff changeset
   331
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   332
lemma times_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   333
  "k * 0 = (0::integer)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   334
  "0 * l = (0::integer)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   335
  "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
   336
  "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
   337
  "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
   338
  "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
   339
  by simp_all
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   340
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   341
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
   342
where
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   343
  "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
   344
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   345
lemma fst_divmod [simp]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   346
  "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
   347
  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
   348
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   349
lemma snd_divmod [simp]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   350
  "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
   351
  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
   352
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   353
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
   354
where
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   355
  "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
   356
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   357
lemma fst_divmod_abs [simp]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   358
  "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
   359
  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
   360
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   361
lemma snd_divmod_abs [simp]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   362
  "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
   363
  by (simp add: divmod_abs_def)
28708
a1a436f09ec6 explicit check for pattern discipline before code translation
haftmann
parents: 28562
diff changeset
   364
53069
d165213e3924 execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
haftmann
parents: 52435
diff changeset
   365
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
   366
  "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
   367
  "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
   368
  "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
   369
  "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
   370
  "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
   371
  "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
   372
  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
   373
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   374
lemma divmod_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   375
  "divmod_integer k l =
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   376
    (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
   377
    (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
   378
      then divmod_abs k l
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   379
      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
   380
        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
   381
proof -
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   382
  have aux1: "\<And>k l::int. sgn k = sgn l \<longleftrightarrow> k = 0 \<and> l = 0 \<or> 0 < l \<and> 0 < k \<or> l < 0 \<and> k < 0"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   383
    by (auto simp add: sgn_if)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   384
  have aux2: "\<And>q::int. - int_of_integer k = int_of_integer l * q \<longleftrightarrow> int_of_integer k = int_of_integer l * - q" by auto
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   385
  show ?thesis
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   386
    by (simp add: prod_eq_iff integer_eq_iff prod_case_beta aux1)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   387
      (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right aux2)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   388
qed
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   389
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   390
lemma div_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   391
  "k div l = fst (divmod_integer k l)"
28708
a1a436f09ec6 explicit check for pattern discipline before code translation
haftmann
parents: 28562
diff changeset
   392
  by simp
a1a436f09ec6 explicit check for pattern discipline before code translation
haftmann
parents: 28562
diff changeset
   393
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   394
lemma mod_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   395
  "k mod l = snd (divmod_integer k l)"
25767
852bce03412a index now a copy of nat rather than int
haftmann
parents: 25691
diff changeset
   396
  by simp
24999
haftmann
parents:
diff changeset
   397
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   398
lemma equal_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   399
  "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
   400
  "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
   401
  "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
   402
  "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
   403
  "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
   404
  "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
   405
  "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
   406
  "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
   407
  "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
   408
  by (simp_all add: equal)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   409
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   410
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
   411
  "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
   412
  by (fact equal_refl)
31266
55e70b6d812e added lemma about 0 - 1
haftmann
parents: 31205
diff changeset
   413
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   414
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
   415
  "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
   416
  "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
   417
  "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
   418
  "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
   419
  "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
   420
  "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
   421
  "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
   422
  "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
   423
  "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
   424
  by simp_all
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   425
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   426
lemma less_integer_code [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   427
  "0 < (0::integer) \<longleftrightarrow> False"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   428
  "0 < Pos l \<longleftrightarrow> True"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   429
  "0 < Neg l \<longleftrightarrow> False"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   430
  "Pos k < 0 \<longleftrightarrow> False"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   431
  "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
   432
  "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
   433
  "Neg k < 0 \<longleftrightarrow> True"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   434
  "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
   435
  "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
   436
  by simp_all
26140
e45375135052 Zero/Suc recursion combinator for type index
haftmann
parents: 26086
diff changeset
   437
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   438
lift_definition integer_of_num :: "num \<Rightarrow> integer"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   439
  is "numeral :: num \<Rightarrow> int"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   440
  .
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
lemma integer_of_num [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   443
  "integer_of_num num.One = 1"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   444
  "integer_of_num (num.Bit0 n) = (let k = integer_of_num n in k + k)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   445
  "integer_of_num (num.Bit1 n) = (let k = integer_of_num n in k + k + 1)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   446
  by (transfer, simp only: numeral.simps Let_def)+
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   447
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   448
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
   449
  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
   450
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   451
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   452
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
   453
  "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
   454
     else let
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   455
       (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
   456
       l' = num_of_integer l;
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   457
       l'' = l' + l'
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   458
     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
   459
proof -
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   460
  {
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   461
    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
   462
    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
   463
    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
   464
    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
   465
    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
   466
      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
   467
      by simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   468
    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
   469
      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
   470
      by (simp add: mult_2)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   471
    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
   472
      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
   473
      by simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   474
  }
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   475
  note aux = this
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   476
  show ?thesis
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   477
    by (auto simp add: num_of_integer_def nat_of_integer_def Let_def prod_case_beta
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   478
      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
   479
      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
   480
       mult_2 [where 'a=nat] aux add_One)
25918
haftmann
parents: 25767
diff changeset
   481
qed
haftmann
parents: 25767
diff changeset
   482
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   483
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
   484
  "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
   485
     else let
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   486
       (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
   487
       l' = nat_of_integer l;
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   488
       l'' = l' + l'
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   489
     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
   490
proof -
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   491
  obtain j where "k = integer_of_int j"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   492
  proof
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   493
    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
   494
  qed
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   495
  moreover have "2 * (j div 2) = j - j mod 2"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   496
    by (simp add: zmult_div_cancel mult_commute)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   497
  ultimately show ?thesis
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   498
    by (auto simp add: split_def Let_def mod_integer_def nat_of_integer_def not_le
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   499
      nat_add_distrib [symmetric] Suc_nat_eq_nat_zadd1)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   500
      (auto simp add: mult_2 [symmetric])
33340
a165b97f3658 moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33318
diff changeset
   501
qed
28708
a1a436f09ec6 explicit check for pattern discipline before code translation
haftmann
parents: 28562
diff changeset
   502
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   503
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
   504
  "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
   505
     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
   506
     else let
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   507
       (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
   508
       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
   509
     in if j = 0 then l' else l' + 1)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   510
  by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
28708
a1a436f09ec6 explicit check for pattern discipline before code translation
haftmann
parents: 28562
diff changeset
   511
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   512
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
   513
  "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
   514
     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
   515
     else let
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   516
       (l, j) = divmod_int k 2;
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   517
       l' = 2 * integer_of_int l
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   518
     in if j = 0 then l' else l' + 1)"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   519
  by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   520
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   521
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
   522
28708
a1a436f09ec6 explicit check for pattern discipline before code translation
haftmann
parents: 28562
diff changeset
   523
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   524
subsection {* Serializer setup for target language integers *}
24999
haftmann
parents:
diff changeset
   525
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   526
code_reserved Eval int Integer abs
25767
852bce03412a index now a copy of nat rather than int
haftmann
parents: 25691
diff changeset
   527
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   528
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
   529
  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
   530
    (SML) "IntInf.int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   531
    and (OCaml) "Big'_int.big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   532
    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
   533
    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
   534
    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
   535
| 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
   536
    (Haskell) -
24999
haftmann
parents:
diff changeset
   537
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   538
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
   539
  constant "0::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
   540
    (SML) "0"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   541
    and (OCaml) "Big'_int.zero'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   542
    and (Haskell) "0"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   543
    and (Scala) "BigInt(0)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46664
diff changeset
   544
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   545
setup {*
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   546
  fold (Numeral.add_code @{const_name Code_Numeral.Pos}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   547
    false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   548
*}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   549
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   550
setup {*
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   551
  fold (Numeral.add_code @{const_name Code_Numeral.Neg}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   552
    true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   553
*}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   554
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   555
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
   556
  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
   557
    (SML) "IntInf.+ ((_), (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   558
    and (OCaml) "Big'_int.add'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   559
    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
   560
    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
   561
    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
   562
| 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
   563
    (SML) "IntInf.~"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   564
    and (OCaml) "Big'_int.minus'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   565
    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
   566
    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
   567
    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
   568
| 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
   569
    (SML) "IntInf.- ((_), (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   570
    and (OCaml) "Big'_int.sub'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   571
    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
   572
    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
   573
    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
   574
| 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
   575
    (SML) "IntInf.*/ (2,/ (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   576
    and (OCaml) "Big'_int.mult'_big'_int/ (Big'_int.big'_int'_of'_int/ 2)"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   577
    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
   578
    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
   579
    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
   580
| 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
   581
    (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
   582
    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
   583
    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
   584
    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
   585
| 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
   586
    (SML) "IntInf.* ((_), (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   587
    and (OCaml) "Big'_int.mult'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   588
    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
   589
    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
   590
    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
   591
| 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
   592
    (SML) "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   593
    and (OCaml) "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   594
    and (Haskell) "divMod/ (abs _)/ (abs _)"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   595
    and (Scala) "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   596
    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
   597
| 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
   598
    (SML) "!((_ : IntInf.int) = _)"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   599
    and (OCaml) "Big'_int.eq'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   600
    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
   601
    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
   602
    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
   603
| 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
   604
    (SML) "IntInf.<= ((_), (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   605
    and (OCaml) "Big'_int.le'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   606
    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
   607
    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
   608
    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
   609
| 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
   610
    (SML) "IntInf.< ((_), (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   611
    and (OCaml) "Big'_int.lt'_big'_int"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   612
    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
   613
    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
   614
    and (Eval) infixl 6 "<"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   615
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51375
diff changeset
   616
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
   617
  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
   618
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   619
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   620
subsection {* Type of target language naturals *}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   621
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   622
typedef natural = "UNIV \<Colon> nat set"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   623
  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
   624
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   625
setup_lifting (no_code) type_definition_natural
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   626
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   627
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
   628
  "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
   629
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   630
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   631
lemma natural_eqI:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   632
  "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
   633
  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
   634
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   635
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
   636
  "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
   637
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   638
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   639
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
   640
  "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
   641
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   642
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   643
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
   644
begin
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   645
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   646
lift_definition zero_natural :: natural
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   647
  is "0 :: nat"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   648
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   649
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   650
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
   651
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   652
lift_definition one_natural :: natural
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   653
  is "1 :: nat"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   654
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   655
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   656
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
   657
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   658
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
   659
  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
   660
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   661
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   662
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
   663
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   664
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
   665
  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
   666
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   667
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   668
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
   669
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   670
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
   671
  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
   672
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   673
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   674
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
   675
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   676
instance proof
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   677
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
   678
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   679
end
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   680
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   681
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   682
  "fun_rel HOL.eq pcr_natural (\<lambda>n::nat. n) (of_nat :: nat \<Rightarrow> natural)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   683
proof -
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   684
  have "fun_rel HOL.eq pcr_natural (of_nat :: nat \<Rightarrow> nat) (of_nat :: nat \<Rightarrow> natural)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   685
    by (unfold of_nat_def [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   686
  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
   687
qed
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   688
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   689
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   690
  "fun_rel HOL.eq pcr_natural (numeral :: num \<Rightarrow> nat) (numeral :: num \<Rightarrow> natural)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   691
proof -
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   692
  have "fun_rel HOL.eq pcr_natural (numeral :: num \<Rightarrow> nat) (\<lambda>n. of_nat (numeral n))"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   693
    by transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   694
  then show ?thesis by simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   695
qed
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   696
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   697
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
   698
  "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
   699
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   700
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   701
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
   702
  "natural_of_nat = of_nat"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   703
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   704
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   705
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
   706
  "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
   707
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   708
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   709
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
   710
  "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
   711
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   712
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   713
instantiation natural :: "{semiring_div, equal, linordered_semiring}"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   714
begin
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   715
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   716
lift_definition div_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
   717
  is "Divides.div :: nat \<Rightarrow> nat \<Rightarrow> nat"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   718
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   719
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   720
declare div_natural.rep_eq [simp]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   721
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   722
lift_definition mod_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
   723
  is "Divides.mod :: nat \<Rightarrow> nat \<Rightarrow> nat"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   724
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   725
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   726
declare mod_natural.rep_eq [simp]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   727
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   728
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
   729
  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
   730
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   731
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   732
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
   733
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   734
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
   735
  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
   736
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   737
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   738
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
   739
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   740
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
   741
  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
   742
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   743
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   744
instance proof
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   745
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
   746
24999
haftmann
parents:
diff changeset
   747
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
   748
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   749
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   750
  "fun_rel pcr_natural (fun_rel pcr_natural pcr_natural) (min :: _ \<Rightarrow> _ \<Rightarrow> nat) (min :: _ \<Rightarrow> _ \<Rightarrow> natural)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   751
  by (unfold min_def [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   752
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   753
lemma [transfer_rule]:
51375
d9e62d9c98de patch Isabelle ditribution to conform to changes regarding the parametricity
kuncar
parents: 51143
diff changeset
   754
  "fun_rel pcr_natural (fun_rel pcr_natural pcr_natural) (max :: _ \<Rightarrow> _ \<Rightarrow> nat) (max :: _ \<Rightarrow> _ \<Rightarrow> natural)"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   755
  by (unfold max_def [abs_def]) transfer_prover
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   756
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   757
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
   758
  "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
   759
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   760
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   761
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
   762
  "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
   763
  by transfer rule
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
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
   766
  is "nat :: int \<Rightarrow> nat"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   767
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   768
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   769
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
   770
  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
   771
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   772
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   773
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
   774
  "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
   775
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   776
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   777
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
   778
  "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
   779
  by transfer auto
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 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
   782
  "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
   783
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   784
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   785
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
   786
  "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
   787
  by transfer rule
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   788
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   789
lemma [measure_function]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   790
  "is_measure nat_of_natural"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   791
  by (rule is_measure_trivial)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   792
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   793
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   794
subsection {* Inductive represenation of target language naturals *}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   795
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   796
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
   797
  is Nat.Suc
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   798
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   799
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   800
declare Suc.rep_eq [simp]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   801
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   802
rep_datatype "0::natural" Suc
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   803
  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
   804
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   805
lemma natural_case [case_names nat, cases type: natural]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   806
  fixes m :: natural
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   807
  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
   808
  shows P
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   809
  using assms by transfer blast
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   810
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   811
lemma [simp, code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   812
  "natural_size = nat_of_natural"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   813
proof (rule ext)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   814
  fix n
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   815
  show "natural_size n = nat_of_natural n"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   816
    by (induct n) simp_all
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   817
qed
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   818
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   819
lemma [simp, code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   820
  "size = nat_of_natural"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   821
proof (rule ext)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   822
  fix n
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   823
  show "size n = nat_of_natural n"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   824
    by (induct n) simp_all
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   825
qed
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   826
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   827
lemma natural_decr [termination_simp]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   828
  "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
   829
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   830
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   831
lemma natural_zero_minus_one:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   832
  "(0::natural) - 1 = 0"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   833
  by simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   834
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   835
lemma Suc_natural_minus_one:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   836
  "Suc n - 1 = n"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   837
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   838
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   839
hide_const (open) Suc
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   840
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   841
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   842
subsection {* Code refinement for target language naturals *}
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   843
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   844
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
   845
  is nat
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   846
  .
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   847
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   848
lemma [code_post]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   849
  "Nat 0 = 0"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   850
  "Nat 1 = 1"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   851
  "Nat (numeral k) = numeral k"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   852
  by (transfer, simp)+
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   853
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   854
lemma [code abstype]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   855
  "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
   856
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   857
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   858
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   859
  "integer_of_natural (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
   860
  by simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   861
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   862
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   863
  "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
   864
  by simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   865
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   866
lemma [code_abbrev]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   867
  "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
   868
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   869
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   870
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   871
  "integer_of_natural 0 = 0"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   872
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   873
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   874
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   875
  "integer_of_natural 1 = 1"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   876
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   877
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   878
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   879
  "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
   880
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   881
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   882
lemma [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   883
  "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
   884
  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
   885
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   886
lemma [code, code_unfold]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   887
  "natural_case f g n = (if n = 0 then f else g (n - 1))"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   888
  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
   889
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   890
declare natural.recs [code del]
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   891
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   892
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   893
  "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
   894
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   895
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   896
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   897
  "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
   898
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   899
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   900
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   901
  "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
   902
  by transfer (simp add: of_nat_mult)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   903
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   904
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   905
  "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
   906
  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
   907
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   908
lemma [code abstract]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   909
  "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
   910
  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
   911
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   912
lemma [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   913
  "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
   914
  by transfer (simp add: equal)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   915
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   916
lemma [code nbe]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   917
  "HOL.equal n (n::natural) \<longleftrightarrow> True"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   918
  by (simp add: equal)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   919
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   920
lemma [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   921
  "m \<le> n \<longleftrightarrow> integer_of_natural m \<le> integer_of_natural n"
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   922
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   923
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   924
lemma [code]:
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   925
  "m < n \<longleftrightarrow> 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
   926
  by transfer simp
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   927
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   928
hide_const (open) Nat
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   929
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   930
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   931
code_reflect Code_Numeral
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   932
  datatypes natural = _
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   933
  functions integer_of_natural natural_of_integer
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   934
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   935
end
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 49962
diff changeset
   936