src/HOL/Integ/NatBin.thy
author obua
Tue, 11 May 2004 20:11:08 +0200
changeset 14738 83f1a514dcb4
parent 14467 bbfa6b01a55f
child 15003 6145dd7538d7
permissions -rw-r--r--
changes made due to new Ring_and_Field theory
paulson@7032
     1
(*  Title:      HOL/NatBin.thy
paulson@7032
     2
    ID:         $Id$
paulson@7032
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
paulson@7032
     4
    Copyright   1999  University of Cambridge
wenzelm@12838
     5
*)
paulson@7032
     6
wenzelm@12838
     7
header {* Binary arithmetic for the natural numbers *}
paulson@7032
     8
paulson@14353
     9
theory NatBin = IntDiv:
paulson@7032
    10
wenzelm@12838
    11
text {*
paulson@14273
    12
  Arithmetic for naturals is reduced to that for the non-negative integers.
wenzelm@12838
    13
*}
paulson@7032
    14
wenzelm@12838
    15
instance nat :: number ..
wenzelm@12838
    16
wenzelm@12838
    17
defs (overloaded)
paulson@14273
    18
  nat_number_of_def:
paulson@14273
    19
     "(number_of::bin => nat) v == nat ((number_of :: bin => int) v)"
wenzelm@12838
    20
paulson@14272
    21
paulson@14353
    22
subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*}
paulson@14272
    23
paulson@14273
    24
declare nat_0 [simp] nat_1 [simp]
paulson@14273
    25
paulson@14273
    26
lemma nat_number_of [simp]: "nat (number_of w) = number_of w"
paulson@14273
    27
by (simp add: nat_number_of_def)
paulson@14272
    28
paulson@14387
    29
lemma nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)"
paulson@14273
    30
by (simp add: nat_number_of_def)
paulson@14272
    31
paulson@14387
    32
lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)"
paulson@14273
    33
by (simp add: nat_1 nat_number_of_def)
paulson@14272
    34
paulson@14272
    35
lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0"
paulson@14387
    36
by (simp add: nat_numeral_1_eq_1)
paulson@14272
    37
paulson@14272
    38
lemma numeral_2_eq_2: "2 = Suc (Suc 0)"
paulson@14272
    39
apply (unfold nat_number_of_def)
paulson@14272
    40
apply (rule nat_2)
paulson@14272
    41
done
paulson@14272
    42
paulson@14272
    43
paulson@14273
    44
text{*Distributive laws for type @{text nat}.  The others are in theory
paulson@14273
    45
   @{text IntArith}, but these require div and mod to be defined for type
paulson@14273
    46
   "int".  They also need some of the lemmas proved above.*}
paulson@14272
    47
paulson@14272
    48
lemma nat_div_distrib: "(0::int) <= z ==> nat (z div z') = nat z div nat z'"
paulson@14272
    49
apply (case_tac "0 <= z'")
paulson@14272
    50
apply (auto simp add: div_nonneg_neg_le0 DIVISION_BY_ZERO_DIV)
paulson@14273
    51
apply (case_tac "z' = 0", simp add: DIVISION_BY_ZERO)
paulson@14272
    52
apply (auto elim!: nonneg_eq_int)
paulson@14272
    53
apply (rename_tac m m')
paulson@14272
    54
apply (subgoal_tac "0 <= int m div int m'")
paulson@14387
    55
 prefer 2 apply (simp add: nat_numeral_0_eq_0 pos_imp_zdiv_nonneg_iff) 
paulson@14273
    56
apply (rule inj_int [THEN injD], simp)
paulson@14272
    57
apply (rule_tac r = "int (m mod m') " in quorem_div)
paulson@14273
    58
 prefer 2 apply force
paulson@14387
    59
apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int 
paulson@14273
    60
                 zmult_int)
paulson@14272
    61
done
paulson@14272
    62
paulson@14272
    63
(*Fails if z'<0: the LHS collapses to (nat z) but the RHS doesn't*)
paulson@14273
    64
lemma nat_mod_distrib:
paulson@14273
    65
     "[| (0::int) <= z;  0 <= z' |] ==> nat (z mod z') = nat z mod nat z'"
paulson@14273
    66
apply (case_tac "z' = 0", simp add: DIVISION_BY_ZERO)
paulson@14272
    67
apply (auto elim!: nonneg_eq_int)
paulson@14272
    68
apply (rename_tac m m')
paulson@14272
    69
apply (subgoal_tac "0 <= int m mod int m'")
paulson@14387
    70
 prefer 2 apply (simp add: nat_less_iff nat_numeral_0_eq_0 pos_mod_sign) 
paulson@14273
    71
apply (rule inj_int [THEN injD], simp)
paulson@14272
    72
apply (rule_tac q = "int (m div m') " in quorem_mod)
paulson@14273
    73
 prefer 2 apply force
paulson@14387
    74
apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int zmult_int)
paulson@14272
    75
done
paulson@14272
    76
paulson@14272
    77
paulson@14353
    78
subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*}
paulson@14272
    79
paulson@14272
    80
(*"neg" is used in rewrite rules for binary comparisons*)
paulson@14387
    81
lemma int_nat_number_of [simp]:
paulson@14273
    82
     "int (number_of v :: nat) =  
paulson@14378
    83
         (if neg (number_of v :: int) then 0  
paulson@14272
    84
          else (number_of v :: int))"
paulson@14272
    85
by (simp del: nat_number_of
paulson@14272
    86
	 add: neg_nat nat_number_of_def not_neg_nat add_assoc)
paulson@14272
    87
paulson@14272
    88
paulson@14390
    89
subsubsection{*Successor *}
paulson@14272
    90
paulson@14272
    91
lemma Suc_nat_eq_nat_zadd1: "(0::int) <= z ==> Suc (nat z) = nat (1 + z)"
paulson@14272
    92
apply (rule sym)
paulson@14273
    93
apply (simp add: nat_eq_iff int_Suc)
paulson@14272
    94
done
paulson@14272
    95
paulson@14273
    96
lemma Suc_nat_number_of_add:
paulson@14273
    97
     "Suc (number_of v + n) =  
paulson@14378
    98
        (if neg (number_of v :: int) then 1+n else number_of (bin_succ v) + n)" 
paulson@14272
    99
by (simp del: nat_number_of 
paulson@14272
   100
         add: nat_number_of_def neg_nat
paulson@14272
   101
              Suc_nat_eq_nat_zadd1 number_of_succ) 
paulson@14272
   102
paulson@14387
   103
lemma Suc_nat_number_of [simp]:
paulson@14273
   104
     "Suc (number_of v) =  
paulson@14378
   105
        (if neg (number_of v :: int) then 1 else number_of (bin_succ v))"
paulson@14273
   106
apply (cut_tac n = 0 in Suc_nat_number_of_add)
paulson@14272
   107
apply (simp cong del: if_weak_cong)
paulson@14272
   108
done
paulson@14272
   109
paulson@14272
   110
paulson@14390
   111
subsubsection{*Addition *}
paulson@14272
   112
paulson@14272
   113
(*"neg" is used in rewrite rules for binary comparisons*)
paulson@14387
   114
lemma add_nat_number_of [simp]:
paulson@14273
   115
     "(number_of v :: nat) + number_of v' =  
paulson@14378
   116
         (if neg (number_of v :: int) then number_of v'  
paulson@14378
   117
          else if neg (number_of v' :: int) then number_of v  
paulson@14272
   118
          else number_of (bin_add v v'))"
paulson@14272
   119
by (force dest!: neg_nat
paulson@14272
   120
          simp del: nat_number_of
paulson@14272
   121
          simp add: nat_number_of_def nat_add_distrib [symmetric]) 
paulson@14272
   122
paulson@14272
   123
paulson@14390
   124
subsubsection{*Subtraction *}
paulson@14272
   125
paulson@14273
   126
lemma diff_nat_eq_if:
paulson@14273
   127
     "nat z - nat z' =  
paulson@14272
   128
        (if neg z' then nat z   
paulson@14272
   129
         else let d = z-z' in     
paulson@14272
   130
              if neg d then 0 else nat d)"
paulson@14273
   131
apply (simp add: Let_def nat_diff_distrib [symmetric] neg_eq_less_0 not_neg_eq_ge_0)
paulson@14273
   132
apply (simp add: diff_is_0_eq nat_le_eq_zle)
paulson@14272
   133
done
paulson@14272
   134
paulson@14387
   135
lemma diff_nat_number_of [simp]: 
paulson@14272
   136
     "(number_of v :: nat) - number_of v' =  
paulson@14378
   137
        (if neg (number_of v' :: int) then number_of v  
paulson@14272
   138
         else let d = number_of (bin_add v (bin_minus v')) in     
paulson@14272
   139
              if neg d then 0 else nat d)"
paulson@14272
   140
by (simp del: nat_number_of add: diff_nat_eq_if nat_number_of_def) 
paulson@14272
   141
paulson@14272
   142
paulson@14272
   143
paulson@14390
   144
subsubsection{*Multiplication *}
paulson@14272
   145
paulson@14387
   146
lemma mult_nat_number_of [simp]:
paulson@14273
   147
     "(number_of v :: nat) * number_of v' =  
paulson@14378
   148
       (if neg (number_of v :: int) then 0 else number_of (bin_mult v v'))"
paulson@14272
   149
by (force dest!: neg_nat
paulson@14272
   150
          simp del: nat_number_of
paulson@14272
   151
          simp add: nat_number_of_def nat_mult_distrib [symmetric]) 
paulson@14272
   152
paulson@14272
   153
paulson@14272
   154
paulson@14390
   155
subsubsection{*Quotient *}
paulson@14272
   156
paulson@14387
   157
lemma div_nat_number_of [simp]:
paulson@14273
   158
     "(number_of v :: nat)  div  number_of v' =  
paulson@14378
   159
          (if neg (number_of v :: int) then 0  
paulson@14272
   160
           else nat (number_of v div number_of v'))"
paulson@14272
   161
by (force dest!: neg_nat
paulson@14272
   162
          simp del: nat_number_of
paulson@14272
   163
          simp add: nat_number_of_def nat_div_distrib [symmetric]) 
paulson@14272
   164
paulson@14387
   165
lemma one_div_nat_number_of [simp]:
paulson@14387
   166
     "(Suc 0)  div  number_of v' = (nat (1 div number_of v'))" 
paulson@14387
   167
by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
paulson@14272
   168
paulson@14272
   169
paulson@14390
   170
subsubsection{*Remainder *}
paulson@14272
   171
paulson@14387
   172
lemma mod_nat_number_of [simp]:
paulson@14273
   173
     "(number_of v :: nat)  mod  number_of v' =  
paulson@14378
   174
        (if neg (number_of v :: int) then 0  
paulson@14378
   175
         else if neg (number_of v' :: int) then number_of v  
paulson@14272
   176
         else nat (number_of v mod number_of v'))"
paulson@14272
   177
by (force dest!: neg_nat
paulson@14272
   178
          simp del: nat_number_of
paulson@14272
   179
          simp add: nat_number_of_def nat_mod_distrib [symmetric]) 
paulson@14272
   180
paulson@14387
   181
lemma one_mod_nat_number_of [simp]:
paulson@14387
   182
     "(Suc 0)  mod  number_of v' =  
paulson@14387
   183
        (if neg (number_of v' :: int) then Suc 0
paulson@14387
   184
         else nat (1 mod number_of v'))"
paulson@14387
   185
by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
paulson@14387
   186
paulson@14387
   187
paulson@14272
   188
paulson@14272
   189
ML
paulson@14272
   190
{*
paulson@14272
   191
val nat_number_of_def = thm"nat_number_of_def";
paulson@14272
   192
paulson@14272
   193
val nat_number_of = thm"nat_number_of";
paulson@14387
   194
val nat_numeral_0_eq_0 = thm"nat_numeral_0_eq_0";
paulson@14387
   195
val nat_numeral_1_eq_1 = thm"nat_numeral_1_eq_1";
paulson@14272
   196
val numeral_1_eq_Suc_0 = thm"numeral_1_eq_Suc_0";
paulson@14272
   197
val numeral_2_eq_2 = thm"numeral_2_eq_2";
paulson@14272
   198
val nat_div_distrib = thm"nat_div_distrib";
paulson@14272
   199
val nat_mod_distrib = thm"nat_mod_distrib";
paulson@14272
   200
val int_nat_number_of = thm"int_nat_number_of";
paulson@14272
   201
val Suc_nat_eq_nat_zadd1 = thm"Suc_nat_eq_nat_zadd1";
paulson@14272
   202
val Suc_nat_number_of_add = thm"Suc_nat_number_of_add";
paulson@14272
   203
val Suc_nat_number_of = thm"Suc_nat_number_of";
paulson@14272
   204
val add_nat_number_of = thm"add_nat_number_of";
paulson@14272
   205
val diff_nat_eq_if = thm"diff_nat_eq_if";
paulson@14272
   206
val diff_nat_number_of = thm"diff_nat_number_of";
paulson@14272
   207
val mult_nat_number_of = thm"mult_nat_number_of";
paulson@14272
   208
val div_nat_number_of = thm"div_nat_number_of";
paulson@14272
   209
val mod_nat_number_of = thm"mod_nat_number_of";
paulson@14272
   210
*}
paulson@14272
   211
paulson@14272
   212
paulson@14390
   213
subsection{*Comparisons*}
paulson@14272
   214
paulson@14390
   215
subsubsection{*Equals (=) *}
paulson@14272
   216
paulson@14273
   217
lemma eq_nat_nat_iff:
paulson@14273
   218
     "[| (0::int) <= z;  0 <= z' |] ==> (nat z = nat z') = (z=z')"
paulson@14273
   219
by (auto elim!: nonneg_eq_int)
paulson@14272
   220
paulson@14272
   221
(*"neg" is used in rewrite rules for binary comparisons*)
paulson@14390
   222
lemma eq_nat_number_of [simp]:
paulson@14273
   223
     "((number_of v :: nat) = number_of v') =  
paulson@14378
   224
      (if neg (number_of v :: int) then (iszero (number_of v' :: int) | neg (number_of v' :: int))  
paulson@14378
   225
       else if neg (number_of v' :: int) then iszero (number_of v :: int)  
paulson@14378
   226
       else iszero (number_of (bin_add v (bin_minus v')) :: int))"
paulson@14272
   227
apply (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def
paulson@14272
   228
                  eq_nat_nat_iff eq_number_of_eq nat_0 iszero_def
paulson@14273
   229
            split add: split_if cong add: imp_cong)
paulson@14272
   230
apply (simp only: nat_eq_iff nat_eq_iff2)
paulson@14272
   231
apply (simp add: not_neg_eq_ge_0 [symmetric])
paulson@14272
   232
done
paulson@14272
   233
paulson@14272
   234
paulson@14390
   235
subsubsection{*Less-than (<) *}
paulson@14272
   236
paulson@14272
   237
(*"neg" is used in rewrite rules for binary comparisons*)
paulson@14390
   238
lemma less_nat_number_of [simp]:
paulson@14273
   239
     "((number_of v :: nat) < number_of v') =  
paulson@14378
   240
         (if neg (number_of v :: int) then neg (number_of (bin_minus v') :: int)  
paulson@14378
   241
          else neg (number_of (bin_add v (bin_minus v')) :: int))"
paulson@14390
   242
by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def
paulson@14272
   243
                nat_less_eq_zless less_number_of_eq_neg zless_nat_eq_int_zless
paulson@14390
   244
         cong add: imp_cong, simp) 
paulson@14272
   245
paulson@14390
   246
paulson@14272
   247
paulson@14272
   248
paulson@14272
   249
(*Maps #n to n for n = 0, 1, 2*)
paulson@14387
   250
lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2
paulson@14272
   251
paulson@14272
   252
paulson@14353
   253
subsection{*General Theorems About Powers Involving Binary Numerals*}
paulson@14353
   254
paulson@14353
   255
text{*We cannot refer to the number @{term 2} in @{text Ring_and_Field.thy}.
paulson@14353
   256
We cannot prove general results about the numeral @{term "-1"}, so we have to
paulson@14353
   257
use @{term "- 1"} instead.*}
paulson@14353
   258
obua@14738
   259
lemma power2_eq_square: "(a::'a::{comm_semiring_1_cancel,ringpower})\<twosuperior> = a * a"
paulson@14353
   260
  by (simp add: numeral_2_eq_2 Power.power_Suc)
paulson@14353
   261
obua@14738
   262
lemma [simp]: "(0::'a::{comm_semiring_1_cancel,ringpower})\<twosuperior> = 0"
paulson@14353
   263
  by (simp add: power2_eq_square)
paulson@14353
   264
obua@14738
   265
lemma [simp]: "(1::'a::{comm_semiring_1_cancel,ringpower})\<twosuperior> = 1"
paulson@14353
   266
  by (simp add: power2_eq_square)
paulson@14353
   267
paulson@14353
   268
text{*Squares of literal numerals will be evaluated.*}
paulson@14353
   269
declare power2_eq_square [of "number_of w", standard, simp]
paulson@14353
   270
obua@14738
   271
lemma zero_le_power2 [simp]: "0 \<le> (a\<twosuperior>::'a::{ordered_idom,ringpower})"
paulson@14353
   272
  by (simp add: power2_eq_square zero_le_square)
paulson@14353
   273
paulson@14353
   274
lemma zero_less_power2 [simp]:
obua@14738
   275
     "(0 < a\<twosuperior>) = (a \<noteq> (0::'a::{ordered_idom,ringpower}))"
paulson@14353
   276
  by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
paulson@14353
   277
paulson@14353
   278
lemma zero_eq_power2 [simp]:
obua@14738
   279
     "(a\<twosuperior> = 0) = (a = (0::'a::{ordered_idom,ringpower}))"
paulson@14353
   280
  by (force simp add: power2_eq_square mult_eq_0_iff)
paulson@14353
   281
paulson@14353
   282
lemma abs_power2 [simp]:
obua@14738
   283
     "abs(a\<twosuperior>) = (a\<twosuperior>::'a::{ordered_idom,ringpower})"
paulson@14353
   284
  by (simp add: power2_eq_square abs_mult abs_mult_self)
paulson@14353
   285
paulson@14353
   286
lemma power2_abs [simp]:
obua@14738
   287
     "(abs a)\<twosuperior> = (a\<twosuperior>::'a::{ordered_idom,ringpower})"
paulson@14353
   288
  by (simp add: power2_eq_square abs_mult_self)
paulson@14353
   289
paulson@14353
   290
lemma power2_minus [simp]:
obua@14738
   291
     "(- a)\<twosuperior> = (a\<twosuperior>::'a::{comm_ring_1,ringpower})"
paulson@14353
   292
  by (simp add: power2_eq_square)
paulson@14353
   293
obua@14738
   294
lemma power_minus1_even: "(- 1) ^ (2*n) = (1::'a::{comm_ring_1,ringpower})"
paulson@14353
   295
apply (induct_tac "n")
paulson@14353
   296
apply (auto simp add: power_Suc power_add)
paulson@14353
   297
done
paulson@14353
   298
paulson@14443
   299
lemma power_even_eq: "(a::'a::ringpower) ^ (2*n) = (a^n)^2"
paulson@14443
   300
by (simp add: power_mult power_mult_distrib power2_eq_square)
paulson@14443
   301
paulson@14443
   302
lemma power_odd_eq: "(a::int) ^ Suc(2*n) = a * (a^n)^2"
paulson@14443
   303
by (simp add: power_even_eq) 
paulson@14443
   304
paulson@14353
   305
lemma power_minus_even [simp]:
obua@14738
   306
     "(-a) ^ (2*n) = (a::'a::{comm_ring_1,ringpower}) ^ (2*n)"
paulson@14353
   307
by (simp add: power_minus1_even power_minus [of a]) 
paulson@14353
   308
paulson@14353
   309
lemma zero_le_even_power:
obua@14738
   310
     "0 \<le> (a::'a::{ordered_idom,ringpower}) ^ (2*n)"
paulson@14353
   311
proof (induct "n")
paulson@14353
   312
  case 0
paulson@14353
   313
    show ?case by (simp add: zero_le_one)
paulson@14353
   314
next
paulson@14353
   315
  case (Suc n)
paulson@14353
   316
    have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" 
paulson@14353
   317
      by (simp add: mult_ac power_add power2_eq_square)
paulson@14353
   318
    thus ?case
paulson@14353
   319
      by (simp add: prems zero_le_square zero_le_mult_iff)
paulson@14353
   320
qed
paulson@14353
   321
paulson@14353
   322
lemma odd_power_less_zero:
obua@14738
   323
     "(a::'a::{ordered_idom,ringpower}) < 0 ==> a ^ Suc(2*n) < 0"
paulson@14353
   324
proof (induct "n")
paulson@14353
   325
  case 0
paulson@14353
   326
    show ?case by (simp add: Power.power_Suc)
paulson@14353
   327
next
paulson@14353
   328
  case (Suc n)
paulson@14353
   329
    have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" 
paulson@14353
   330
      by (simp add: mult_ac power_add power2_eq_square Power.power_Suc)
paulson@14353
   331
    thus ?case
paulson@14353
   332
      by (simp add: prems mult_less_0_iff mult_neg)
paulson@14353
   333
qed
paulson@14353
   334
paulson@14353
   335
lemma odd_0_le_power_imp_0_le:
obua@14738
   336
     "0 \<le> a  ^ Suc(2*n) ==> 0 \<le> (a::'a::{ordered_idom,ringpower})"
paulson@14353
   337
apply (insert odd_power_less_zero [of a n]) 
paulson@14353
   338
apply (force simp add: linorder_not_less [symmetric]) 
paulson@14353
   339
done
paulson@14353
   340
paulson@14353
   341
paulson@14390
   342
subsubsection{*Nat *}
paulson@14272
   343
paulson@14272
   344
lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
paulson@14273
   345
by (simp add: numerals)
paulson@14272
   346
paulson@14272
   347
(*Expresses a natural number constant as the Suc of another one.
paulson@14272
   348
  NOT suitable for rewriting because n recurs in the condition.*)
paulson@14272
   349
lemmas expand_Suc = Suc_pred' [of "number_of v", standard]
paulson@14272
   350
paulson@14390
   351
subsubsection{*Arith *}
paulson@14272
   352
paulson@14272
   353
lemma Suc_eq_add_numeral_1: "Suc n = n + 1"
paulson@14273
   354
by (simp add: numerals)
paulson@14272
   355
paulson@14467
   356
lemma Suc_eq_add_numeral_1_left: "Suc n = 1 + n"
paulson@14467
   357
by (simp add: numerals)
paulson@14467
   358
paulson@14272
   359
(* These two can be useful when m = number_of... *)
paulson@14272
   360
paulson@14272
   361
lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))"
paulson@14272
   362
apply (case_tac "m")
paulson@14272
   363
apply (simp_all add: numerals)
paulson@14272
   364
done
paulson@14272
   365
paulson@14272
   366
lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))"
paulson@14272
   367
apply (case_tac "m")
paulson@14272
   368
apply (simp_all add: numerals)
paulson@14272
   369
done
paulson@14272
   370
paulson@14272
   371
lemma power_eq_if: "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))"
paulson@14272
   372
apply (case_tac "m")
paulson@14272
   373
apply (simp_all add: numerals)
paulson@14272
   374
done
paulson@14272
   375
paulson@14272
   376
lemma diff_less': "[| 0<n; 0<m |] ==> m - n < (m::nat)"
paulson@14273
   377
by (simp add: diff_less numerals)
paulson@14272
   378
paulson@14272
   379
declare diff_less' [of "number_of v", standard, simp]
paulson@14272
   380
paulson@14272
   381
paulson@14390
   382
subsection{*Comparisons involving (0::nat) *}
paulson@14272
   383
paulson@14390
   384
text{*Simplification already does @{term "n<0"}, @{term "n\<le>0"} and @{term "0\<le>n"}.*}
paulson@14390
   385
paulson@14390
   386
lemma eq_number_of_0 [simp]:
paulson@14273
   387
     "(number_of v = (0::nat)) =  
paulson@14378
   388
      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
paulson@14390
   389
by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
paulson@14272
   390
paulson@14390
   391
lemma eq_0_number_of [simp]:
paulson@14273
   392
     "((0::nat) = number_of v) =  
paulson@14378
   393
      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
paulson@14390
   394
by (rule trans [OF eq_sym_conv eq_number_of_0])
paulson@14272
   395
paulson@14390
   396
lemma less_0_number_of [simp]:
paulson@14378
   397
     "((0::nat) < number_of v) = neg (number_of (bin_minus v) :: int)"
paulson@14387
   398
by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric])
paulson@14272
   399
paulson@14272
   400
paulson@14378
   401
lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)"
paulson@14387
   402
by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
paulson@14272
   403
paulson@14272
   404
paulson@14272
   405
paulson@14390
   406
subsection{*Comparisons involving Suc *}
paulson@14272
   407
paulson@14273
   408
lemma eq_number_of_Suc [simp]:
paulson@14273
   409
     "(number_of v = Suc n) =  
paulson@14272
   410
        (let pv = number_of (bin_pred v) in  
paulson@14272
   411
         if neg pv then False else nat pv = n)"
paulson@14272
   412
apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
paulson@14272
   413
                  number_of_pred nat_number_of_def 
paulson@14273
   414
            split add: split_if)
paulson@14272
   415
apply (rule_tac x = "number_of v" in spec)
paulson@14272
   416
apply (auto simp add: nat_eq_iff)
paulson@14272
   417
done
paulson@14272
   418
paulson@14273
   419
lemma Suc_eq_number_of [simp]:
paulson@14273
   420
     "(Suc n = number_of v) =  
paulson@14272
   421
        (let pv = number_of (bin_pred v) in  
paulson@14272
   422
         if neg pv then False else nat pv = n)"
paulson@14390
   423
by (rule trans [OF eq_sym_conv eq_number_of_Suc])
paulson@14272
   424
paulson@14273
   425
lemma less_number_of_Suc [simp]:
paulson@14273
   426
     "(number_of v < Suc n) =  
paulson@14272
   427
        (let pv = number_of (bin_pred v) in  
paulson@14272
   428
         if neg pv then True else nat pv < n)"
paulson@14272
   429
apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
paulson@14272
   430
                  number_of_pred nat_number_of_def  
paulson@14273
   431
            split add: split_if)
paulson@14272
   432
apply (rule_tac x = "number_of v" in spec)
paulson@14272
   433
apply (auto simp add: nat_less_iff)
paulson@14272
   434
done
paulson@14272
   435
paulson@14273
   436
lemma less_Suc_number_of [simp]:
paulson@14273
   437
     "(Suc n < number_of v) =  
paulson@14272
   438
        (let pv = number_of (bin_pred v) in  
paulson@14272
   439
         if neg pv then False else n < nat pv)"
paulson@14272
   440
apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
paulson@14272
   441
                  number_of_pred nat_number_of_def
paulson@14273
   442
            split add: split_if)
paulson@14272
   443
apply (rule_tac x = "number_of v" in spec)
paulson@14272
   444
apply (auto simp add: zless_nat_eq_int_zless)
paulson@14272
   445
done
paulson@14272
   446
paulson@14273
   447
lemma le_number_of_Suc [simp]:
paulson@14273
   448
     "(number_of v <= Suc n) =  
paulson@14272
   449
        (let pv = number_of (bin_pred v) in  
paulson@14272
   450
         if neg pv then True else nat pv <= n)"
paulson@14390
   451
by (simp add: Let_def less_Suc_number_of linorder_not_less [symmetric])
paulson@14272
   452
paulson@14273
   453
lemma le_Suc_number_of [simp]:
paulson@14273
   454
     "(Suc n <= number_of v) =  
paulson@14272
   455
        (let pv = number_of (bin_pred v) in  
paulson@14272
   456
         if neg pv then False else n <= nat pv)"
paulson@14390
   457
by (simp add: Let_def less_number_of_Suc linorder_not_less [symmetric])
paulson@14272
   458
paulson@14272
   459
paulson@14272
   460
(* Push int(.) inwards: *)
paulson@14272
   461
declare zadd_int [symmetric, simp]
paulson@14272
   462
paulson@14272
   463
lemma lemma1: "(m+m = n+n) = (m = (n::int))"
paulson@14273
   464
by auto
paulson@14272
   465
paulson@14272
   466
lemma lemma2: "m+m ~= (1::int) + (n + n)"
paulson@14272
   467
apply auto
paulson@14272
   468
apply (drule_tac f = "%x. x mod 2" in arg_cong)
paulson@14273
   469
apply (simp add: zmod_zadd1_eq)
paulson@14272
   470
done
paulson@14272
   471
paulson@14273
   472
lemma eq_number_of_BIT_BIT:
paulson@14273
   473
     "((number_of (v BIT x) ::int) = number_of (w BIT y)) =  
paulson@14272
   474
      (x=y & (((number_of v) ::int) = number_of w))"
paulson@14272
   475
by (simp only: simp_thms number_of_BIT lemma1 lemma2 eq_commute
obua@14738
   476
               OrderedGroup.add_left_cancel add_assoc OrderedGroup.add_0
paulson@14273
   477
            split add: split_if cong: imp_cong) 
paulson@14272
   478
paulson@14272
   479
paulson@14273
   480
lemma eq_number_of_BIT_Pls:
paulson@14273
   481
     "((number_of (v BIT x) ::int) = number_of bin.Pls) =  
paulson@14272
   482
      (x=False & (((number_of v) ::int) = number_of bin.Pls))"
paulson@14272
   483
apply (simp only: simp_thms  add: number_of_BIT number_of_Pls eq_commute
paulson@14272
   484
            split add: split_if cong: imp_cong)
paulson@14273
   485
apply (rule_tac x = "number_of v" in spec, safe)
paulson@14272
   486
apply (simp_all (no_asm_use))
paulson@14272
   487
apply (drule_tac f = "%x. x mod 2" in arg_cong)
paulson@14273
   488
apply (simp add: zmod_zadd1_eq)
paulson@14272
   489
done
paulson@14272
   490
paulson@14273
   491
lemma eq_number_of_BIT_Min:
paulson@14273
   492
     "((number_of (v BIT x) ::int) = number_of bin.Min) =  
paulson@14272
   493
      (x=True & (((number_of v) ::int) = number_of bin.Min))"
paulson@14272
   494
apply (simp only: simp_thms  add: number_of_BIT number_of_Min eq_commute
paulson@14272
   495
            split add: split_if cong: imp_cong)
paulson@14273
   496
apply (rule_tac x = "number_of v" in spec, auto)
paulson@14273
   497
apply (drule_tac f = "%x. x mod 2" in arg_cong, auto)
paulson@14272
   498
done
paulson@14272
   499
paulson@14272
   500
lemma eq_number_of_Pls_Min: "(number_of bin.Pls ::int) ~= number_of bin.Min"
paulson@14273
   501
by auto
paulson@14272
   502
paulson@14272
   503
paulson@14272
   504
paulson@14390
   505
subsection{*Literal arithmetic involving powers*}
paulson@14272
   506
paulson@14272
   507
lemma nat_power_eq: "(0::int) <= z ==> nat (z^n) = nat z ^ n"
paulson@14272
   508
apply (induct_tac "n")
paulson@14272
   509
apply (simp_all (no_asm_simp) add: nat_mult_distrib)
paulson@14272
   510
done
paulson@14272
   511
paulson@14273
   512
lemma power_nat_number_of:
paulson@14273
   513
     "(number_of v :: nat) ^ n =  
paulson@14378
   514
       (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))"
paulson@14272
   515
by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def nat_power_eq
paulson@14272
   516
         split add: split_if cong: imp_cong)
paulson@14272
   517
paulson@14272
   518
paulson@14272
   519
declare power_nat_number_of [of _ "number_of w", standard, simp]
paulson@14272
   520
paulson@14272
   521
paulson@14390
   522
text{*For the integers*}
paulson@14272
   523
paulson@14273
   524
lemma zpower_number_of_even:
paulson@14273
   525
     "(z::int) ^ number_of (w BIT False) =  
paulson@14272
   526
      (let w = z ^ (number_of w) in  w*w)"
paulson@14272
   527
apply (simp del: nat_number_of  add: nat_number_of_def number_of_BIT Let_def)
paulson@14272
   528
apply (simp only: number_of_add) 
paulson@14273
   529
apply (rule_tac x = "number_of w" in spec, clarify)
paulson@14272
   530
apply (case_tac " (0::int) <= x")
paulson@14443
   531
apply (auto simp add: nat_mult_distrib power_even_eq power2_eq_square)
paulson@14272
   532
done
paulson@14272
   533
paulson@14273
   534
lemma zpower_number_of_odd:
paulson@14273
   535
     "(z::int) ^ number_of (w BIT True) =  
paulson@14272
   536
          (if (0::int) <= number_of w                    
paulson@14272
   537
           then (let w = z ^ (number_of w) in  z*w*w)    
paulson@14272
   538
           else 1)"
paulson@14272
   539
apply (simp del: nat_number_of  add: nat_number_of_def number_of_BIT Let_def)
paulson@14387
   540
apply (simp only: number_of_add nat_numeral_1_eq_1 not_neg_eq_ge_0 neg_eq_less_0) 
paulson@14273
   541
apply (rule_tac x = "number_of w" in spec, clarify)
paulson@14443
   542
apply (auto simp add: nat_add_distrib nat_mult_distrib power_even_eq power2_eq_square neg_nat)
paulson@14272
   543
done
paulson@14272
   544
paulson@14272
   545
declare zpower_number_of_even [of "number_of v", standard, simp]
paulson@14272
   546
declare zpower_number_of_odd  [of "number_of v", standard, simp]
paulson@14272
   547
paulson@14272
   548
paulson@14272
   549
paulson@14272
   550
ML
paulson@14272
   551
{*
paulson@14272
   552
val numerals = thms"numerals";
paulson@14272
   553
val numeral_ss = simpset() addsimps numerals;
paulson@14272
   554
paulson@14272
   555
val nat_bin_arith_setup =
paulson@14272
   556
 [Fast_Arith.map_data 
paulson@14272
   557
   (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
paulson@14272
   558
     {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
paulson@14272
   559
      inj_thms = inj_thms,
paulson@14272
   560
      lessD = lessD,
paulson@14272
   561
      simpset = simpset addsimps [Suc_nat_number_of, int_nat_number_of,
paulson@14272
   562
                                  not_neg_number_of_Pls,
paulson@14272
   563
                                  neg_number_of_Min,neg_number_of_BIT]})]
paulson@14272
   564
*}
paulson@14272
   565
wenzelm@12838
   566
setup nat_bin_arith_setup
wenzelm@12838
   567
nipkow@13189
   568
(* Enable arith to deal with div/mod k where k is a numeral: *)
nipkow@13189
   569
declare split_div[of _ _ "number_of k", standard, arith_split]
nipkow@13189
   570
declare split_mod[of _ _ "number_of k", standard, arith_split]
nipkow@13154
   571
nipkow@13491
   572
lemma nat_number_of_Pls: "number_of bin.Pls = (0::nat)"
wenzelm@12838
   573
  by (simp add: number_of_Pls nat_number_of_def)
wenzelm@12838
   574
nipkow@13491
   575
lemma nat_number_of_Min: "number_of bin.Min = (0::nat)"
wenzelm@12838
   576
  apply (simp only: number_of_Min nat_number_of_def nat_zminus_int)
wenzelm@12838
   577
  apply (simp add: neg_nat)
wenzelm@12838
   578
  done
wenzelm@12838
   579
wenzelm@12838
   580
lemma nat_number_of_BIT_True:
wenzelm@12838
   581
  "number_of (w BIT True) =
paulson@14378
   582
    (if neg (number_of w :: int) then 0
wenzelm@12838
   583
     else let n = number_of w in Suc (n + n))"
wenzelm@12838
   584
  apply (simp only: nat_number_of_def Let_def split: split_if)
wenzelm@12838
   585
  apply (intro conjI impI)
wenzelm@12838
   586
   apply (simp add: neg_nat neg_number_of_BIT)
wenzelm@12838
   587
  apply (rule int_int_eq [THEN iffD1])
wenzelm@12838
   588
  apply (simp only: not_neg_nat neg_number_of_BIT int_Suc zadd_int [symmetric] simp_thms)
wenzelm@12838
   589
  apply (simp only: number_of_BIT if_True zadd_assoc)
wenzelm@12838
   590
  done
wenzelm@12838
   591
wenzelm@12838
   592
lemma nat_number_of_BIT_False:
wenzelm@12838
   593
    "number_of (w BIT False) = (let n::nat = number_of w in n + n)"
wenzelm@12838
   594
  apply (simp only: nat_number_of_def Let_def)
paulson@14378
   595
  apply (cases "neg (number_of w :: int)")
wenzelm@12838
   596
   apply (simp add: neg_nat neg_number_of_BIT)
wenzelm@12838
   597
  apply (rule int_int_eq [THEN iffD1])
wenzelm@12838
   598
  apply (simp only: not_neg_nat neg_number_of_BIT int_Suc zadd_int [symmetric] simp_thms)
wenzelm@12838
   599
  apply (simp only: number_of_BIT if_False zadd_0 zadd_assoc)
wenzelm@12838
   600
  done
wenzelm@12838
   601
wenzelm@13043
   602
lemmas nat_number =
wenzelm@12838
   603
  nat_number_of_Pls nat_number_of_Min
wenzelm@12838
   604
  nat_number_of_BIT_True nat_number_of_BIT_False
wenzelm@12838
   605
wenzelm@12838
   606
lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)"
wenzelm@12838
   607
  by (simp add: Let_def)
nipkow@10574
   608
paulson@14443
   609
lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring,ringpower})"
paulson@14443
   610
by (simp add: power_mult); 
paulson@14443
   611
paulson@14443
   612
lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring,ringpower})"
paulson@14443
   613
by (simp add: power_mult power_Suc); 
paulson@14443
   614
berghofe@12440
   615
paulson@14390
   616
subsection{*Literal arithmetic and @{term of_nat}*}
paulson@14390
   617
paulson@14390
   618
lemma of_nat_double:
paulson@14390
   619
     "0 \<le> x ==> of_nat (nat (2 * x)) = of_nat (nat x) + of_nat (nat x)"
paulson@14390
   620
by (simp only: mult_2 nat_add_distrib of_nat_add) 
paulson@14390
   621
paulson@14390
   622
lemma nat_numeral_m1_eq_0: "-1 = (0::nat)"
paulson@14390
   623
by (simp only:  nat_number_of_def, simp)
paulson@14390
   624
paulson@14390
   625
lemma int_double_iff: "(0::int) \<le> 2*x + 1 = (0 \<le> x)"
paulson@14390
   626
by arith
paulson@14390
   627
paulson@14390
   628
lemma of_nat_number_of_lemma:
paulson@14390
   629
     "of_nat (number_of v :: nat) =  
paulson@14390
   630
         (if 0 \<le> (number_of v :: int) 
paulson@14390
   631
          then (number_of v :: 'a :: number_ring)
paulson@14390
   632
          else 0)"
paulson@14390
   633
apply (induct v, simp, simp add: nat_numeral_m1_eq_0)
paulson@14390
   634
apply (simp only: number_of nat_number_of_def)
paulson@14390
   635
txt{*Generalize in order to eliminate the function @{term number_of} and
paulson@14390
   636
its annoying simprules*}
paulson@14390
   637
apply (erule rev_mp)
paulson@14390
   638
apply (rule_tac x="number_of bin :: int" in spec) 
paulson@14390
   639
apply (rule_tac x="number_of bin :: 'a" in spec) 
paulson@14390
   640
apply (simp add: nat_add_distrib of_nat_double int_double_iff)
paulson@14390
   641
done
paulson@14390
   642
paulson@14390
   643
lemma of_nat_number_of_eq [simp]:
paulson@14390
   644
     "of_nat (number_of v :: nat) =  
paulson@14390
   645
         (if neg (number_of v :: int) then 0  
paulson@14390
   646
          else (number_of v :: 'a :: number_ring))"
paulson@14390
   647
by (simp only: of_nat_number_of_lemma neg_def, simp) 
paulson@14390
   648
paulson@14390
   649
paulson@14273
   650
subsection {*Lemmas for the Combination and Cancellation Simprocs*}
paulson@14273
   651
paulson@14273
   652
lemma nat_number_of_add_left:
paulson@14273
   653
     "number_of v + (number_of v' + (k::nat)) =  
paulson@14378
   654
         (if neg (number_of v :: int) then number_of v' + k  
paulson@14378
   655
          else if neg (number_of v' :: int) then number_of v + k  
paulson@14273
   656
          else number_of (bin_add v v') + k)"
paulson@14390
   657
by simp
paulson@14273
   658
paulson@14430
   659
lemma nat_number_of_mult_left:
paulson@14430
   660
     "number_of v * (number_of v' * (k::nat)) =  
paulson@14430
   661
         (if neg (number_of v :: int) then 0
paulson@14430
   662
          else number_of (bin_mult v v') * k)"
paulson@14430
   663
by simp
paulson@14430
   664
paulson@14273
   665
paulson@14390
   666
subsubsection{*For @{text combine_numerals}*}
paulson@14273
   667
paulson@14273
   668
lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)"
paulson@14273
   669
by (simp add: add_mult_distrib)
paulson@14273
   670
paulson@14273
   671
paulson@14390
   672
subsubsection{*For @{text cancel_numerals}*}
paulson@14273
   673
paulson@14273
   674
lemma nat_diff_add_eq1:
paulson@14273
   675
     "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)"
paulson@14273
   676
by (simp split add: nat_diff_split add: add_mult_distrib)
paulson@14273
   677
paulson@14273
   678
lemma nat_diff_add_eq2:
paulson@14273
   679
     "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))"
paulson@14273
   680
by (simp split add: nat_diff_split add: add_mult_distrib)
paulson@14273
   681
paulson@14273
   682
lemma nat_eq_add_iff1:
paulson@14273
   683
     "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)"
paulson@14273
   684
by (auto split add: nat_diff_split simp add: add_mult_distrib)
paulson@14273
   685
paulson@14273
   686
lemma nat_eq_add_iff2:
paulson@14273
   687
     "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)"
paulson@14273
   688
by (auto split add: nat_diff_split simp add: add_mult_distrib)
paulson@14273
   689
paulson@14273
   690
lemma nat_less_add_iff1:
paulson@14273
   691
     "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)"
paulson@14273
   692
by (auto split add: nat_diff_split simp add: add_mult_distrib)
paulson@14273
   693
paulson@14273
   694
lemma nat_less_add_iff2:
paulson@14273
   695
     "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)"
paulson@14273
   696
by (auto split add: nat_diff_split simp add: add_mult_distrib)
paulson@14273
   697
paulson@14273
   698
lemma nat_le_add_iff1:
paulson@14273
   699
     "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)"
paulson@14273
   700
by (auto split add: nat_diff_split simp add: add_mult_distrib)
paulson@14273
   701
paulson@14273
   702
lemma nat_le_add_iff2:
paulson@14273
   703
     "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)"
paulson@14273
   704
by (auto split add: nat_diff_split simp add: add_mult_distrib)
paulson@14273
   705
paulson@14273
   706
paulson@14390
   707
subsubsection{*For @{text cancel_numeral_factors} *}
paulson@14273
   708
paulson@14273
   709
lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)"
paulson@14273
   710
by auto
paulson@14273
   711
paulson@14273
   712
lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m<n)"
paulson@14273
   713
by auto
paulson@14273
   714
paulson@14273
   715
lemma nat_mult_eq_cancel1: "(0::nat) < k ==> (k*m = k*n) = (m=n)"
paulson@14273
   716
by auto
paulson@14273
   717
paulson@14273
   718
lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)"
paulson@14273
   719
by auto
paulson@14273
   720
paulson@14273
   721
paulson@14390
   722
subsubsection{*For @{text cancel_factor} *}
paulson@14273
   723
paulson@14273
   724
lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)"
paulson@14273
   725
by auto
paulson@14273
   726
paulson@14273
   727
lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m<n)"
paulson@14273
   728
by auto
paulson@14273
   729
paulson@14273
   730
lemma nat_mult_eq_cancel_disj: "(k*m = k*n) = (k = (0::nat) | m=n)"
paulson@14273
   731
by auto
paulson@14273
   732
paulson@14273
   733
lemma nat_mult_div_cancel_disj:
paulson@14273
   734
     "(k*m) div (k*n) = (if k = (0::nat) then 0 else m div n)"
paulson@14273
   735
by (simp add: nat_mult_div_cancel1)
paulson@14273
   736
paulson@14353
   737
paulson@14273
   738
ML
paulson@14273
   739
{*
paulson@14353
   740
val eq_nat_nat_iff = thm"eq_nat_nat_iff";
paulson@14353
   741
val eq_nat_number_of = thm"eq_nat_number_of";
paulson@14353
   742
val less_nat_number_of = thm"less_nat_number_of";
paulson@14353
   743
val power2_eq_square = thm "power2_eq_square";
paulson@14353
   744
val zero_le_power2 = thm "zero_le_power2";
paulson@14353
   745
val zero_less_power2 = thm "zero_less_power2";
paulson@14353
   746
val zero_eq_power2 = thm "zero_eq_power2";
paulson@14353
   747
val abs_power2 = thm "abs_power2";
paulson@14353
   748
val power2_abs = thm "power2_abs";
paulson@14353
   749
val power2_minus = thm "power2_minus";
paulson@14353
   750
val power_minus1_even = thm "power_minus1_even";
paulson@14353
   751
val power_minus_even = thm "power_minus_even";
paulson@14353
   752
val zero_le_even_power = thm "zero_le_even_power";
paulson@14353
   753
val odd_power_less_zero = thm "odd_power_less_zero";
paulson@14353
   754
val odd_0_le_power_imp_0_le = thm "odd_0_le_power_imp_0_le";
paulson@14353
   755
paulson@14353
   756
val Suc_pred' = thm"Suc_pred'";
paulson@14353
   757
val expand_Suc = thm"expand_Suc";
paulson@14353
   758
val Suc_eq_add_numeral_1 = thm"Suc_eq_add_numeral_1";
paulson@14467
   759
val Suc_eq_add_numeral_1_left = thm"Suc_eq_add_numeral_1_left";
paulson@14353
   760
val add_eq_if = thm"add_eq_if";
paulson@14353
   761
val mult_eq_if = thm"mult_eq_if";
paulson@14353
   762
val power_eq_if = thm"power_eq_if";
paulson@14353
   763
val diff_less' = thm"diff_less'";
paulson@14353
   764
val eq_number_of_0 = thm"eq_number_of_0";
paulson@14353
   765
val eq_0_number_of = thm"eq_0_number_of";
paulson@14353
   766
val less_0_number_of = thm"less_0_number_of";
paulson@14353
   767
val neg_imp_number_of_eq_0 = thm"neg_imp_number_of_eq_0";
paulson@14353
   768
val eq_number_of_Suc = thm"eq_number_of_Suc";
paulson@14353
   769
val Suc_eq_number_of = thm"Suc_eq_number_of";
paulson@14353
   770
val less_number_of_Suc = thm"less_number_of_Suc";
paulson@14353
   771
val less_Suc_number_of = thm"less_Suc_number_of";
paulson@14353
   772
val le_number_of_Suc = thm"le_number_of_Suc";
paulson@14353
   773
val le_Suc_number_of = thm"le_Suc_number_of";
paulson@14353
   774
val eq_number_of_BIT_BIT = thm"eq_number_of_BIT_BIT";
paulson@14353
   775
val eq_number_of_BIT_Pls = thm"eq_number_of_BIT_Pls";
paulson@14353
   776
val eq_number_of_BIT_Min = thm"eq_number_of_BIT_Min";
paulson@14353
   777
val eq_number_of_Pls_Min = thm"eq_number_of_Pls_Min";
paulson@14390
   778
val of_nat_number_of_eq = thm"of_nat_number_of_eq";
paulson@14353
   779
val nat_power_eq = thm"nat_power_eq";
paulson@14353
   780
val power_nat_number_of = thm"power_nat_number_of";
paulson@14353
   781
val zpower_number_of_even = thm"zpower_number_of_even";
paulson@14353
   782
val zpower_number_of_odd = thm"zpower_number_of_odd";
paulson@14353
   783
val nat_number_of_Pls = thm"nat_number_of_Pls";
paulson@14353
   784
val nat_number_of_Min = thm"nat_number_of_Min";
paulson@14353
   785
val nat_number_of_BIT_True = thm"nat_number_of_BIT_True";
paulson@14353
   786
val nat_number_of_BIT_False = thm"nat_number_of_BIT_False";
paulson@14353
   787
val Let_Suc = thm"Let_Suc";
paulson@14353
   788
paulson@14353
   789
val nat_number = thms"nat_number";
paulson@14353
   790
paulson@14273
   791
val nat_number_of_add_left = thm"nat_number_of_add_left";
paulson@14430
   792
val nat_number_of_mult_left = thm"nat_number_of_mult_left";
paulson@14273
   793
val left_add_mult_distrib = thm"left_add_mult_distrib";
paulson@14273
   794
val nat_diff_add_eq1 = thm"nat_diff_add_eq1";
paulson@14273
   795
val nat_diff_add_eq2 = thm"nat_diff_add_eq2";
paulson@14273
   796
val nat_eq_add_iff1 = thm"nat_eq_add_iff1";
paulson@14273
   797
val nat_eq_add_iff2 = thm"nat_eq_add_iff2";
paulson@14273
   798
val nat_less_add_iff1 = thm"nat_less_add_iff1";
paulson@14273
   799
val nat_less_add_iff2 = thm"nat_less_add_iff2";
paulson@14273
   800
val nat_le_add_iff1 = thm"nat_le_add_iff1";
paulson@14273
   801
val nat_le_add_iff2 = thm"nat_le_add_iff2";
paulson@14273
   802
val nat_mult_le_cancel1 = thm"nat_mult_le_cancel1";
paulson@14273
   803
val nat_mult_less_cancel1 = thm"nat_mult_less_cancel1";
paulson@14273
   804
val nat_mult_eq_cancel1 = thm"nat_mult_eq_cancel1";
paulson@14273
   805
val nat_mult_div_cancel1 = thm"nat_mult_div_cancel1";
paulson@14273
   806
val nat_mult_le_cancel_disj = thm"nat_mult_le_cancel_disj";
paulson@14273
   807
val nat_mult_less_cancel_disj = thm"nat_mult_less_cancel_disj";
paulson@14273
   808
val nat_mult_eq_cancel_disj = thm"nat_mult_eq_cancel_disj";
paulson@14273
   809
val nat_mult_div_cancel_disj = thm"nat_mult_div_cancel_disj";
paulson@14353
   810
paulson@14353
   811
val power_minus_even = thm"power_minus_even";
paulson@14353
   812
val zero_le_even_power = thm"zero_le_even_power";
paulson@14273
   813
*}
paulson@14273
   814
paulson@14273
   815
berghofe@12440
   816
subsection {* Configuration of the code generator *}
berghofe@12440
   817
berghofe@12933
   818
ML {*
berghofe@12933
   819
infix 7 `*;
berghofe@12933
   820
infix 6 `+;
berghofe@12933
   821
berghofe@12933
   822
val op `* = op * : int * int -> int;
berghofe@12933
   823
val op `+ = op + : int * int -> int;
berghofe@12933
   824
val `~ = ~ : int -> int;
berghofe@12933
   825
*}
berghofe@12933
   826
berghofe@12440
   827
types_code
berghofe@12440
   828
  "int" ("int")
berghofe@12440
   829
berghofe@14194
   830
constdefs
berghofe@14194
   831
  int_aux :: "int \<Rightarrow> nat \<Rightarrow> int"
berghofe@14194
   832
  "int_aux i n == (i + int n)"
berghofe@14194
   833
  nat_aux :: "nat \<Rightarrow> int \<Rightarrow> nat"
berghofe@14194
   834
  "nat_aux n i == (n + nat i)"
berghofe@12440
   835
berghofe@14194
   836
lemma [code]:
berghofe@14194
   837
  "int_aux i 0 = i"
berghofe@14194
   838
  "int_aux i (Suc n) = int_aux (i + 1) n" -- {* tail recursive *}
berghofe@14194
   839
  "int n = int_aux 0 n"
berghofe@14194
   840
  by (simp add: int_aux_def)+
berghofe@14194
   841
berghofe@14194
   842
lemma [code]: "nat_aux n i = (if i <= 0 then n else nat_aux (Suc n) (i - 1))"
berghofe@14194
   843
  by (simp add: nat_aux_def Suc_nat_eq_nat_zadd1) -- {* tail recursive *}
berghofe@14194
   844
lemma [code]: "nat i = nat_aux 0 i"
berghofe@14194
   845
  by (simp add: nat_aux_def)
berghofe@12440
   846
berghofe@12440
   847
consts_code
berghofe@12440
   848
  "0" :: "int"                  ("0")
berghofe@12440
   849
  "1" :: "int"                  ("1")
berghofe@12933
   850
  "uminus" :: "int => int"      ("`~")
berghofe@12933
   851
  "op +" :: "int => int => int" ("(_ `+/ _)")
berghofe@12933
   852
  "op *" :: "int => int => int" ("(_ `*/ _)")
paulson@14378
   853
  "op <" :: "int => int => bool" ("(_ </ _)")
paulson@14378
   854
  "op <=" :: "int => int => bool" ("(_ <=/ _)")
berghofe@12440
   855
  "neg"                         ("(_ < 0)")
berghofe@12440
   856
berghofe@14417
   857
ML {*
berghofe@14417
   858
fun number_of_codegen thy gr s b (Const ("Numeral.number_of",
berghofe@14417
   859
      Type ("fun", [_, Type ("IntDef.int", [])])) $ bin) =
berghofe@14417
   860
        (Some (gr, Pretty.str (string_of_int (HOLogic.dest_binum bin)))
berghofe@14417
   861
        handle TERM _ => None)
berghofe@14417
   862
  | number_of_codegen thy gr s b (Const ("Numeral.number_of",
berghofe@14417
   863
      Type ("fun", [_, Type ("nat", [])])) $ bin) =
berghofe@14417
   864
        Some (Codegen.invoke_codegen thy s b (gr,
berghofe@14417
   865
          Const ("IntDef.nat", HOLogic.intT --> HOLogic.natT) $
berghofe@14417
   866
            (Const ("Numeral.number_of", HOLogic.binT --> HOLogic.intT) $ bin)))
berghofe@14417
   867
  | number_of_codegen _ _ _ _ _ = None;
berghofe@14417
   868
*}
berghofe@14417
   869
berghofe@14417
   870
setup {* [Codegen.add_codegen "number_of_codegen" number_of_codegen] *}
berghofe@14417
   871
paulson@7032
   872
end