src/HOL/Nat_Numeral.thy
author bulwahn
Thu, 29 Mar 2012 17:40:44 +0200
changeset 47197 ed681ca1188a
parent 47196 6012241abe93
child 47207 9368aa814518
permissions -rw-r--r--
announcing NEWS (cf. 446cfc760ccf)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30925
c38cbc0ac8d1 theory NatBin now named Nat_Numeral
haftmann
parents: 30685
diff changeset
     1
(*  Title:      HOL/Nat_Numeral.thy
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     3
    Copyright   1999  University of Cambridge
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     4
*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     5
30925
c38cbc0ac8d1 theory NatBin now named Nat_Numeral
haftmann
parents: 30685
diff changeset
     6
header {* Binary numerals for the natural numbers *}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     7
30925
c38cbc0ac8d1 theory NatBin now named Nat_Numeral
haftmann
parents: 30685
diff changeset
     8
theory Nat_Numeral
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
     9
imports Int
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    10
begin
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    11
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    12
subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    13
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
    14
declare nat_1 [simp]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    15
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    16
lemma nat_neg_numeral [simp]: "nat (neg_numeral w) = 0"
36719
d396f6f63d94 moved some lemmas from Groebner_Basis here
haftmann
parents: 36716
diff changeset
    17
  by simp
d396f6f63d94 moved some lemmas from Groebner_Basis here
haftmann
parents: 36716
diff changeset
    18
46026
83caa4f4bd56 semiring_numeral_0_eq_0, semiring_numeral_1_eq_1 now [simp], superseeding corresponding simp rules on type nat
haftmann
parents: 45607
diff changeset
    19
lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    20
  by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    21
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    22
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    23
subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    24
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    25
lemma int_numeral: "int (numeral v) = numeral v"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    26
  by (rule of_nat_numeral) (* already simp *)
23307
2fe3345035c7 modify proofs to avoid referring to int::nat=>int
huffman
parents: 23294
diff changeset
    27
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    28
lemma nonneg_int_cases:
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    29
  fixes k :: int assumes "0 \<le> k" obtains n where "k = of_nat n"
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    30
  using assms by (cases k, simp, simp)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    31
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    32
subsubsection{*Successor *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    33
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    34
lemma Suc_nat_eq_nat_zadd1: "(0::int) <= z ==> Suc (nat z) = nat (1 + z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    35
apply (rule sym)
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 44345
diff changeset
    36
apply (simp add: nat_eq_iff)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    37
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    38
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    39
lemma Suc_nat_number_of_add:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    40
  "Suc (numeral v + n) = numeral (v + Num.One) + n"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    41
  by simp
30081
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
    42
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    43
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    44
subsubsection{*Subtraction *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    45
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    46
lemma diff_nat_eq_if:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    47
     "nat z - nat z' =  
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    48
        (if z' < 0 then nat z   
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    49
         else let d = z-z' in     
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    50
              if d < 0 then 0 else nat d)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    51
by (simp add: Let_def nat_diff_distrib [symmetric])
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    52
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    53
(* Int.nat_diff_distrib has too-strong premises *)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    54
lemma nat_diff_distrib': "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x - y) = nat x - nat y"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    55
apply (rule int_int_eq [THEN iffD1], clarsimp)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    56
apply (subst zdiff_int [symmetric])
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    57
apply (rule nat_mono, simp_all)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    58
done
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    59
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    60
lemma diff_nat_numeral [simp]: 
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    61
  "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    62
  by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    63
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    64
lemma nat_numeral_diff_1 [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    65
  "numeral v - (1::nat) = nat (numeral v - 1)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    66
  using diff_nat_numeral [of v Num.One] by simp
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    67
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    68
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    69
subsection{*Comparisons*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    70
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    71
(*Maps #n to n for n = 1, 2*)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    72
lemmas numerals = numeral_1_eq_1 [where 'a=nat] numeral_2_eq_2
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    73
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    74
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    75
text{*Simprules for comparisons where common factors can be cancelled.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    76
lemmas zero_compare_simps =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    77
    add_strict_increasing add_strict_increasing2 add_increasing
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    78
    zero_le_mult_iff zero_le_divide_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    79
    zero_less_mult_iff zero_less_divide_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    80
    mult_le_0_iff divide_le_0_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    81
    mult_less_0_iff divide_less_0_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    82
    zero_le_power2 power2_less_0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    83
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    84
subsubsection{*Nat *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    85
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    86
lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
    87
by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    88
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    89
(*Expresses a natural number constant as the Suc of another one.
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    90
  NOT suitable for rewriting because n recurs on the right-hand side.*)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
    91
lemmas expand_Suc = Suc_pred' [of "numeral v", OF zero_less_numeral] for v
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    92
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    93
subsubsection{*Arith *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    94
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31182
diff changeset
    95
lemma Suc_eq_plus1: "Suc n = n + 1"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
    96
  unfolding One_nat_def by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    97
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31182
diff changeset
    98
lemma Suc_eq_plus1_left: "Suc n = 1 + n"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
    99
  unfolding One_nat_def by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   100
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   101
(* These two can be useful when m = numeral... *)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   102
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   103
lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29958
diff changeset
   104
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   105
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   106
lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29958
diff changeset
   107
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   108
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   109
lemma power_eq_if: "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29958
diff changeset
   110
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   111
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   112
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   113
subsection{*Comparisons involving  @{term Suc} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   114
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   115
lemma eq_numeral_Suc [simp]: "numeral v = Suc n \<longleftrightarrow> nat (numeral v - 1) = n"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   116
  by (subst expand_Suc, simp only: nat.inject nat_numeral_diff_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   117
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   118
lemma Suc_eq_numeral [simp]: "Suc n = numeral v \<longleftrightarrow> n = nat (numeral v - 1)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   119
  by (subst expand_Suc, simp only: nat.inject nat_numeral_diff_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   120
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   121
lemma less_numeral_Suc [simp]: "numeral v < Suc n \<longleftrightarrow> nat (numeral v - 1) < n"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   122
  by (subst expand_Suc, simp only: Suc_less_eq nat_numeral_diff_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   123
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   124
lemma less_Suc_numeral [simp]: "Suc n < numeral v \<longleftrightarrow> n < nat (numeral v - 1)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   125
  by (subst expand_Suc, simp only: Suc_less_eq nat_numeral_diff_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   126
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   127
lemma le_numeral_Suc [simp]: "numeral v \<le> Suc n \<longleftrightarrow> nat (numeral v - 1) \<le> n"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   128
  by (subst expand_Suc, simp only: Suc_le_mono nat_numeral_diff_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   129
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   130
lemma le_Suc_numeral [simp]: "Suc n \<le> numeral v \<longleftrightarrow> n \<le> nat (numeral v - 1)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   131
  by (subst expand_Suc, simp only: Suc_le_mono nat_numeral_diff_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   132
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   133
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   134
subsection{*Max and Min Combined with @{term Suc} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   135
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   136
lemma max_Suc_numeral [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   137
  "max (Suc n) (numeral v) = Suc (max n (nat (numeral v - 1)))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   138
  by (subst expand_Suc, simp only: max_Suc_Suc nat_numeral_diff_1)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   139
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   140
lemma max_numeral_Suc [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   141
  "max (numeral v) (Suc n) = Suc (max (nat (numeral v - 1)) n)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   142
  by (subst expand_Suc, simp only: max_Suc_Suc nat_numeral_diff_1)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   143
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   144
lemma min_Suc_numeral [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   145
  "min (Suc n) (numeral v) = Suc (min n (nat (numeral v - 1)))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   146
  by (subst expand_Suc, simp only: min_Suc_Suc nat_numeral_diff_1)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   147
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   148
lemma min_numeral_Suc [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   149
  "min (numeral v) (Suc n) = Suc (min (nat (numeral v - 1)) n)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   150
  by (subst expand_Suc, simp only: min_Suc_Suc nat_numeral_diff_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   151
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   152
subsection{*Literal arithmetic involving powers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   153
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   154
text{*For arbitrary rings*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   155
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   156
lemma power_numeral_even:
43526
2b92a6943915 generalize lemmas power_number_of_even and power_number_of_odd
huffman
parents: 40690
diff changeset
   157
  fixes z :: "'a::monoid_mult"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   158
  shows "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   159
  unfolding numeral_Bit0 power_add Let_def ..
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   160
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   161
lemma power_numeral_odd:
43526
2b92a6943915 generalize lemmas power_number_of_even and power_number_of_odd
huffman
parents: 40690
diff changeset
   162
  fixes z :: "'a::monoid_mult"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   163
  shows "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   164
  unfolding numeral_Bit1 One_nat_def add_Suc_right add_0_right
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   165
  unfolding power_Suc power_add Let_def mult_assoc ..
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   166
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   167
lemmas zpower_numeral_even = power_numeral_even [where 'a=int]
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   168
lemmas zpower_numeral_odd = power_numeral_odd [where 'a=int]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   169
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   170
lemma nat_numeral_Bit0:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   171
  "numeral (Num.Bit0 w) = (let n::nat = numeral w in n + n)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   172
  unfolding numeral_Bit0 Let_def ..
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   173
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   174
lemma nat_numeral_Bit1:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   175
  "numeral (Num.Bit1 w) = (let n = numeral w in Suc (n + n))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   176
  unfolding numeral_Bit1 Let_def by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   177
40077
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 36964
diff changeset
   178
lemmas eval_nat_numeral =
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   179
  nat_numeral_Bit0 nat_numeral_Bit1
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   180
36699
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   181
lemmas nat_arith =
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   182
  diff_nat_numeral
36699
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   183
36716
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   184
lemmas semiring_norm =
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   185
  Let_def arith_simps nat_arith rel_simps
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   186
  if_False if_True
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   187
  add_0 add_Suc add_numeral_left
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   188
  add_neg_numeral_left mult_numeral_left
36716
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   189
  numeral_1_eq_1 [symmetric] Suc_eq_plus1
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   190
  eq_numeral_iff_iszero not_iszero_Numeral1
36716
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   191
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   192
lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)"
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   193
  by (fact Let_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   194
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   195
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   196
subsection{*Literal arithmetic and @{term of_nat}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   197
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   198
lemma of_nat_double:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   199
     "0 \<le> x ==> of_nat (nat (2 * x)) = of_nat (nat x) + of_nat (nat x)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   200
by (simp only: mult_2 nat_add_distrib of_nat_add) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   201
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   202
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   203
subsubsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   204
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   205
text{*Where K above is a literal*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   206
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   207
lemma Suc_diff_eq_diff_pred: "0 < n ==> Suc m - n = m - (n - Numeral1)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   208
by (simp split: nat_diff_split)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   209
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   210
text{*No longer required as a simprule because of the @{text inverse_fold}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   211
   simproc*}
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   212
lemma Suc_diff_numeral: "Suc m - (numeral v) = m - (numeral v - 1)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   213
  by (subst expand_Suc, simp only: diff_Suc_Suc)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   214
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   215
lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   216
by (simp split: nat_diff_split)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   217
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   218
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   219
subsubsection{*For @{term nat_case} and @{term nat_rec}*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   220
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   221
lemma nat_case_numeral [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   222
  "nat_case a f (numeral v) = (let pv = nat (numeral v - 1) in f pv)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   223
  by (subst expand_Suc, simp only: nat.cases nat_numeral_diff_1 Let_def)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   224
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   225
lemma nat_case_add_eq_if [simp]:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   226
  "nat_case a f ((numeral v) + n) = (let pv = nat (numeral v - 1) in f (pv + n))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   227
  by (subst expand_Suc, simp only: nat.cases nat_numeral_diff_1 Let_def add_Suc)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   228
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   229
lemma nat_rec_numeral [simp]:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   230
  "nat_rec a f (numeral v) = (let pv = nat (numeral v - 1) in f pv (nat_rec a f pv))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   231
  by (subst expand_Suc, simp only: nat_rec_Suc nat_numeral_diff_1 Let_def)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   232
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   233
lemma nat_rec_add_eq_if [simp]:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   234
  "nat_rec a f (numeral v + n) =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   235
    (let pv = nat (numeral v - 1) in f (pv + n) (nat_rec a f (pv + n)))"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   236
  by (subst expand_Suc, simp only: nat_rec_Suc nat_numeral_diff_1 Let_def add_Suc)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   237
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   238
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   239
subsubsection{*Various Other Lemmas*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   240
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31068
diff changeset
   241
lemma card_UNIV_bool[simp]: "card (UNIV :: bool set) = 2"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31068
diff changeset
   242
by(simp add: UNIV_bool)
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31068
diff changeset
   243
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   244
text {*Evens and Odds, for Mutilated Chess Board*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   245
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   246
text{*Lemmas for specialist use, NOT as default simprules*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   247
lemma nat_mult_2: "2 * z = (z+z::nat)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   248
by (rule mult_2) (* FIXME: duplicate *)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   249
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   250
lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   251
by (rule mult_2_right) (* FIXME: duplicate *)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   252
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   253
text{*Case analysis on @{term "n<2"}*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   254
lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   255
by (auto simp add: numeral_2_eq_2)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   256
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   257
text{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   258
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   259
lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   260
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   261
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   262
lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   263
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   264
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   265
text{*Can be used to eliminate long strings of Sucs, but not by default*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   266
lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   267
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   268
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   269
text{*Legacy theorems*}
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   270
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   271
lemmas nat_1_add_1 = one_add_one [where 'a=nat]
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46026
diff changeset
   272
31096
e546e15089ef newline at end of file
huffman
parents: 31080
diff changeset
   273
end