src/HOL/Nat_Numeral.thy
author haftmann
Wed, 22 Apr 2009 19:09:21 +0200
changeset 30960 fec1a04b7220
parent 30925 c38cbc0ac8d1
child 31002 bc4117fe72ab
permissions -rw-r--r--
power operation defined generic
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
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     9
imports IntDiv
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
    10
uses ("Tools/nat_simprocs.ML")
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    11
begin
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    12
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    13
text {*
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    14
  Arithmetic for naturals is reduced to that for the non-negative integers.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    15
*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    16
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    17
instantiation nat :: number
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    18
begin
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    19
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    20
definition
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28229
diff changeset
    21
  nat_number_of_def [code inline, code del]: "number_of v = nat (number_of v)"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    22
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    23
instance ..
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    24
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    25
end
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    26
25965
05df64f786a4 improved code theorem setup
haftmann
parents: 25919
diff changeset
    27
lemma [code post]:
05df64f786a4 improved code theorem setup
haftmann
parents: 25919
diff changeset
    28
  "nat (number_of v) = number_of v"
05df64f786a4 improved code theorem setup
haftmann
parents: 25919
diff changeset
    29
  unfolding nat_number_of_def ..
05df64f786a4 improved code theorem setup
haftmann
parents: 25919
diff changeset
    30
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    31
context recpower
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    32
begin
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    33
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    34
abbreviation (xsymbols)
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    35
  power2 :: "'a \<Rightarrow> 'a"  ("(_\<twosuperior>)" [1000] 999) where
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    36
  "x\<twosuperior> \<equiv> x ^ 2"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    37
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    38
notation (latex output)
29401
94fd5dd918f5 rename abbreviation square -> power2, to match theorem names
huffman
parents: 29045
diff changeset
    39
  power2  ("(_\<twosuperior>)" [1000] 999)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    40
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    41
notation (HTML output)
29401
94fd5dd918f5 rename abbreviation square -> power2, to match theorem names
huffman
parents: 29045
diff changeset
    42
  power2  ("(_\<twosuperior>)" [1000] 999)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    43
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    44
end
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    45
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    46
29040
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    47
subsection {* Predicate for negative binary numbers *}
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    48
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
    49
definition neg  :: "int \<Rightarrow> bool" where
29040
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    50
  "neg Z \<longleftrightarrow> Z < 0"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    51
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    52
lemma not_neg_int [simp]: "~ neg (of_nat n)"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    53
by (simp add: neg_def)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    54
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    55
lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    56
by (simp add: neg_def neg_less_0_iff_less del: of_nat_Suc)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    57
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    58
lemmas neg_eq_less_0 = neg_def
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    59
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    60
lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    61
by (simp add: neg_def linorder_not_less)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    62
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    63
text{*To simplify inequalities when Numeral1 can get simplified to 1*}
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    64
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    65
lemma not_neg_0: "~ neg 0"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    66
by (simp add: One_int_def neg_def)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    67
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    68
lemma not_neg_1: "~ neg 1"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    69
by (simp add: neg_def linorder_not_less zero_le_one)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    70
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    71
lemma neg_nat: "neg z ==> nat z = 0"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    72
by (simp add: neg_def order_less_imp_le) 
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    73
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    74
lemma not_neg_nat: "~ neg z ==> of_nat (nat z) = z"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    75
by (simp add: linorder_not_less neg_def)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    76
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    77
text {*
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    78
  If @{term Numeral0} is rewritten to 0 then this rule can't be applied:
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    79
  @{term Numeral0} IS @{term "number_of Pls"}
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    80
*}
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    81
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    82
lemma not_neg_number_of_Pls: "~ neg (number_of Int.Pls)"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    83
  by (simp add: neg_def)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    84
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    85
lemma neg_number_of_Min: "neg (number_of Int.Min)"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    86
  by (simp add: neg_def)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    87
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    88
lemma neg_number_of_Bit0:
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    89
  "neg (number_of (Int.Bit0 w)) = neg (number_of w)"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    90
  by (simp add: neg_def)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    91
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    92
lemma neg_number_of_Bit1:
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    93
  "neg (number_of (Int.Bit1 w)) = neg (number_of w)"
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    94
  by (simp add: neg_def)
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    95
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    96
lemmas neg_simps [simp] =
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    97
  not_neg_0 not_neg_1
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    98
  not_neg_number_of_Pls neg_number_of_Min
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
    99
  neg_number_of_Bit0 neg_number_of_Bit1
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   100
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   101
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   102
subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   103
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   104
declare nat_0 [simp] nat_1 [simp]
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 nat_number_of [simp]: "nat (number_of w) = number_of w"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   107
by (simp add: nat_number_of_def)
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 nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   110
by (simp add: nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   111
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   112
lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   113
by (simp add: nat_1 nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   114
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   115
lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   116
by (simp add: nat_numeral_1_eq_1)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   117
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   118
lemma numeral_2_eq_2: "2 = Suc (Suc 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   119
apply (unfold nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   120
apply (rule nat_2)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   121
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   122
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   123
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   124
subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   125
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   126
lemma int_nat_number_of [simp]:
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23307
diff changeset
   127
     "int (number_of v) =  
23307
2fe3345035c7 modify proofs to avoid referring to int::nat=>int
huffman
parents: 23294
diff changeset
   128
         (if neg (number_of v :: int) then 0  
2fe3345035c7 modify proofs to avoid referring to int::nat=>int
huffman
parents: 23294
diff changeset
   129
          else (number_of v :: int))"
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   130
  unfolding nat_number_of_def number_of_is_id neg_def
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   131
  by simp
23307
2fe3345035c7 modify proofs to avoid referring to int::nat=>int
huffman
parents: 23294
diff changeset
   132
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   133
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   134
subsubsection{*Successor *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   135
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   136
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
   137
apply (rule sym)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   138
apply (simp add: nat_eq_iff int_Suc)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   139
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   140
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   141
lemma Suc_nat_number_of_add:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   142
     "Suc (number_of v + n) =  
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   143
        (if neg (number_of v :: int) then 1+n else number_of (Int.succ v) + n)"
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   144
  unfolding nat_number_of_def number_of_is_id neg_def numeral_simps
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   145
  by (simp add: Suc_nat_eq_nat_zadd1 add_ac)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   146
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   147
lemma Suc_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   148
     "Suc (number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   149
        (if neg (number_of v :: int) then 1 else number_of (Int.succ v))"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   150
apply (cut_tac n = 0 in Suc_nat_number_of_add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   151
apply (simp cong del: if_weak_cong)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   152
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   153
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   154
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   155
subsubsection{*Addition *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   156
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   157
lemma add_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   158
     "(number_of v :: nat) + number_of v' =  
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   159
         (if v < Int.Pls then number_of v'  
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   160
          else if v' < Int.Pls then number_of v  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   161
          else number_of (v + v'))"
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   162
  unfolding nat_number_of_def number_of_is_id numeral_simps
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   163
  by (simp add: nat_add_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   164
30081
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   165
lemma nat_number_of_add_1 [simp]:
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   166
  "number_of v + (1::nat) =
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   167
    (if v < Int.Pls then 1 else number_of (Int.succ v))"
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   168
  unfolding nat_number_of_def number_of_is_id numeral_simps
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   169
  by (simp add: nat_add_distrib)
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   170
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   171
lemma nat_1_add_number_of [simp]:
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   172
  "(1::nat) + number_of v =
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   173
    (if v < Int.Pls then 1 else number_of (Int.succ v))"
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   174
  unfolding nat_number_of_def number_of_is_id numeral_simps
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   175
  by (simp add: nat_add_distrib)
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   176
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   177
lemma nat_1_add_1 [simp]: "1 + 1 = (2::nat)"
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   178
  by (rule int_int_eq [THEN iffD1]) simp
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   179
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   180
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   181
subsubsection{*Subtraction *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   182
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   183
lemma diff_nat_eq_if:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   184
     "nat z - nat z' =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   185
        (if neg z' then nat z   
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   186
         else let d = z-z' in     
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   187
              if neg d then 0 else nat d)"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26342
diff changeset
   188
by (simp add: Let_def nat_diff_distrib [symmetric] neg_eq_less_0 not_neg_eq_ge_0)
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26342
diff changeset
   189
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   190
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   191
lemma diff_nat_number_of [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   192
     "(number_of v :: nat) - number_of v' =  
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   193
        (if v' < Int.Pls then number_of v  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   194
         else let d = number_of (v + uminus v') in     
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   195
              if neg d then 0 else nat d)"
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   196
  unfolding nat_number_of_def number_of_is_id numeral_simps neg_def
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   197
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   198
30081
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   199
lemma nat_number_of_diff_1 [simp]:
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   200
  "number_of v - (1::nat) =
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   201
    (if v \<le> Int.Pls then 0 else number_of (Int.pred v))"
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   202
  unfolding nat_number_of_def number_of_is_id numeral_simps
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   203
  by auto
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   204
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   205
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   206
subsubsection{*Multiplication *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   207
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   208
lemma mult_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   209
     "(number_of v :: nat) * number_of v' =  
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   210
       (if v < Int.Pls then 0 else number_of (v * v'))"
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   211
  unfolding nat_number_of_def number_of_is_id numeral_simps
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   212
  by (simp add: nat_mult_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   213
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   214
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   215
subsubsection{*Quotient *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   216
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   217
lemma div_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   218
     "(number_of v :: nat)  div  number_of v' =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   219
          (if neg (number_of v :: int) then 0  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   220
           else nat (number_of v div number_of v'))"
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   221
  unfolding nat_number_of_def number_of_is_id neg_def
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   222
  by (simp add: nat_div_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   223
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   224
lemma one_div_nat_number_of [simp]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26342
diff changeset
   225
     "Suc 0 div number_of v' = nat (1 div number_of v')" 
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   226
by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   227
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   228
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   229
subsubsection{*Remainder *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   230
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   231
lemma mod_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   232
     "(number_of v :: nat)  mod  number_of v' =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   233
        (if neg (number_of v :: int) then 0  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   234
         else if neg (number_of v' :: int) then number_of v  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   235
         else nat (number_of v mod number_of v'))"
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   236
  unfolding nat_number_of_def number_of_is_id neg_def
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   237
  by (simp add: nat_mod_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   238
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   239
lemma one_mod_nat_number_of [simp]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 26342
diff changeset
   240
     "Suc 0 mod number_of v' =  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   241
        (if neg (number_of v' :: int) then Suc 0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   242
         else nat (1 mod number_of v'))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   243
by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   244
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   245
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   246
subsubsection{* Divisibility *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   247
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   248
lemmas dvd_eq_mod_eq_0_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   249
  dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   250
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   251
declare dvd_eq_mod_eq_0_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   252
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   253
ML
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   254
{*
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   255
val nat_number_of_def = thm"nat_number_of_def";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   256
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   257
val nat_number_of = thm"nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   258
val nat_numeral_0_eq_0 = thm"nat_numeral_0_eq_0";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   259
val nat_numeral_1_eq_1 = thm"nat_numeral_1_eq_1";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   260
val numeral_1_eq_Suc_0 = thm"numeral_1_eq_Suc_0";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   261
val numeral_2_eq_2 = thm"numeral_2_eq_2";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   262
val nat_div_distrib = thm"nat_div_distrib";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   263
val nat_mod_distrib = thm"nat_mod_distrib";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   264
val int_nat_number_of = thm"int_nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   265
val Suc_nat_eq_nat_zadd1 = thm"Suc_nat_eq_nat_zadd1";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   266
val Suc_nat_number_of_add = thm"Suc_nat_number_of_add";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   267
val Suc_nat_number_of = thm"Suc_nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   268
val add_nat_number_of = thm"add_nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   269
val diff_nat_eq_if = thm"diff_nat_eq_if";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   270
val diff_nat_number_of = thm"diff_nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   271
val mult_nat_number_of = thm"mult_nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   272
val div_nat_number_of = thm"div_nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   273
val mod_nat_number_of = thm"mod_nat_number_of";
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   274
*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   275
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   276
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   277
subsection{*Comparisons*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   278
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   279
subsubsection{*Equals (=) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   280
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   281
lemma eq_nat_nat_iff:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   282
     "[| (0::int) <= z;  0 <= z' |] ==> (nat z = nat z') = (z=z')"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   283
by (auto elim!: nonneg_eq_int)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   284
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   285
lemma eq_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   286
     "((number_of v :: nat) = number_of v') =  
28969
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   287
      (if neg (number_of v :: int) then (number_of v' :: int) \<le> 0
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   288
       else if neg (number_of v' :: int) then (number_of v :: int) = 0
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   289
       else v = v')"
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   290
  unfolding nat_number_of_def number_of_is_id neg_def
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   291
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   292
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   293
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   294
subsubsection{*Less-than (<) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   295
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   296
lemma less_nat_number_of [simp]:
29011
a47003001699 simplify less_nat_number_of
huffman
parents: 29010
diff changeset
   297
  "(number_of v :: nat) < number_of v' \<longleftrightarrow>
a47003001699 simplify less_nat_number_of
huffman
parents: 29010
diff changeset
   298
    (if v < v' then Int.Pls < v' else False)"
a47003001699 simplify less_nat_number_of
huffman
parents: 29010
diff changeset
   299
  unfolding nat_number_of_def number_of_is_id numeral_simps
28961
9f33ab8e15db simplify proof of less_nat_number_of
huffman
parents: 28562
diff changeset
   300
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   301
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   302
29010
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   303
subsubsection{*Less-than-or-equal *}
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   304
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   305
lemma le_nat_number_of [simp]:
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   306
  "(number_of v :: nat) \<le> number_of v' \<longleftrightarrow>
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   307
    (if v \<le> v' then True else v \<le> Int.Pls)"
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   308
  unfolding nat_number_of_def number_of_is_id numeral_simps
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   309
  by auto
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   310
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   311
(*Maps #n to n for n = 0, 1, 2*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   312
lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   313
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   314
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   315
subsection{*Powers with Numeric Exponents*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   316
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   317
text{*We cannot refer to the number @{term 2} in @{text Ring_and_Field.thy}.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   318
We cannot prove general results about the numeral @{term "-1"}, so we have to
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   319
use @{term "- 1"} instead.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   320
23277
aa158e145ea3 generalize class constraints on some lemmas
huffman
parents: 23164
diff changeset
   321
lemma power2_eq_square: "(a::'a::recpower)\<twosuperior> = a * a"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   322
  by (simp add: numeral_2_eq_2 Power.power_Suc)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   323
23277
aa158e145ea3 generalize class constraints on some lemmas
huffman
parents: 23164
diff changeset
   324
lemma zero_power2 [simp]: "(0::'a::{semiring_1,recpower})\<twosuperior> = 0"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   325
  by (simp add: power2_eq_square)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   326
23277
aa158e145ea3 generalize class constraints on some lemmas
huffman
parents: 23164
diff changeset
   327
lemma one_power2 [simp]: "(1::'a::{semiring_1,recpower})\<twosuperior> = 1"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   328
  by (simp add: power2_eq_square)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   329
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   330
lemma power3_eq_cube: "(x::'a::recpower) ^ 3 = x * x * x"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   331
  apply (subgoal_tac "3 = Suc (Suc (Suc 0))")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   332
  apply (erule ssubst)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   333
  apply (simp add: power_Suc mult_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   334
  apply (unfold nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   335
  apply (subst nat_eq_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   336
  apply simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   337
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   338
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   339
text{*Squares of literal numerals will be evaluated.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   340
lemmas power2_eq_square_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   341
    power2_eq_square [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   342
declare power2_eq_square_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   343
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   344
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   345
lemma zero_le_power2[simp]: "0 \<le> (a\<twosuperior>::'a::{ordered_idom,recpower})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   346
  by (simp add: power2_eq_square)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   347
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   348
lemma zero_less_power2[simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   349
     "(0 < a\<twosuperior>) = (a \<noteq> (0::'a::{ordered_idom,recpower}))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   350
  by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   351
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   352
lemma power2_less_0[simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   353
  fixes a :: "'a::{ordered_idom,recpower}"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   354
  shows "~ (a\<twosuperior> < 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   355
by (force simp add: power2_eq_square mult_less_0_iff) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   356
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   357
lemma zero_eq_power2[simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   358
     "(a\<twosuperior> = 0) = (a = (0::'a::{ordered_idom,recpower}))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   359
  by (force simp add: power2_eq_square mult_eq_0_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   360
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   361
lemma abs_power2[simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   362
     "abs(a\<twosuperior>) = (a\<twosuperior>::'a::{ordered_idom,recpower})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   363
  by (simp add: power2_eq_square abs_mult abs_mult_self)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   364
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   365
lemma power2_abs[simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   366
     "(abs a)\<twosuperior> = (a\<twosuperior>::'a::{ordered_idom,recpower})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   367
  by (simp add: power2_eq_square abs_mult_self)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   368
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   369
lemma power2_minus[simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   370
     "(- a)\<twosuperior> = (a\<twosuperior>::'a::{comm_ring_1,recpower})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   371
  by (simp add: power2_eq_square)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   372
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   373
lemma power2_le_imp_le:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   374
  fixes x y :: "'a::{ordered_semidom,recpower}"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   375
  shows "\<lbrakk>x\<twosuperior> \<le> y\<twosuperior>; 0 \<le> y\<rbrakk> \<Longrightarrow> x \<le> y"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   376
unfolding numeral_2_eq_2 by (rule power_le_imp_le_base)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   377
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   378
lemma power2_less_imp_less:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   379
  fixes x y :: "'a::{ordered_semidom,recpower}"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   380
  shows "\<lbrakk>x\<twosuperior> < y\<twosuperior>; 0 \<le> y\<rbrakk> \<Longrightarrow> x < y"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   381
by (rule power_less_imp_less_base)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   382
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   383
lemma power2_eq_imp_eq:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   384
  fixes x y :: "'a::{ordered_semidom,recpower}"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   385
  shows "\<lbrakk>x\<twosuperior> = y\<twosuperior>; 0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> x = y"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   386
unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base, simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   387
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   388
lemma power_minus1_even[simp]: "(- 1) ^ (2*n) = (1::'a::{comm_ring_1,recpower})"
29958
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   389
proof (induct n)
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   390
  case 0 show ?case by simp
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   391
next
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   392
  case (Suc n) then show ?case by (simp add: power_Suc power_add)
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   393
qed
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   394
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   395
lemma power_minus1_odd: "(- 1) ^ Suc(2*n) = -(1::'a::{comm_ring_1,recpower})"
6d84e2f9f5cf Even and odd powers of -1
paulson
parents: 29401
diff changeset
   396
  by (simp add: power_Suc) 
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   397
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   398
lemma power_even_eq: "(a::'a::recpower) ^ (2*n) = (a^n)^2"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   399
by (subst mult_commute) (simp add: power_mult)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   400
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   401
lemma power_odd_eq: "(a::int) ^ Suc(2*n) = a * (a^n)^2"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   402
by (simp add: power_even_eq) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   403
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   404
lemma power_minus_even [simp]:
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
   405
  "(-a) ^ (2*n) = (a::'a::{comm_ring_1,recpower}) ^ (2*n)"
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
   406
  by (simp add: power_minus [of a]) 
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   407
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   408
lemma zero_le_even_power'[simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   409
     "0 \<le> (a::'a::{ordered_idom,recpower}) ^ (2*n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   410
proof (induct "n")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   411
  case 0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   412
    show ?case by (simp add: zero_le_one)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   413
next
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   414
  case (Suc n)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   415
    have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   416
      by (simp add: mult_ac power_add power2_eq_square)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   417
    thus ?case
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   418
      by (simp add: prems zero_le_mult_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   419
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   420
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   421
lemma odd_power_less_zero:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   422
     "(a::'a::{ordered_idom,recpower}) < 0 ==> a ^ Suc(2*n) < 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   423
proof (induct "n")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   424
  case 0
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   425
  then show ?case by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   426
next
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   427
  case (Suc n)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   428
  have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)"
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   429
    by (simp add: mult_ac power_add power2_eq_square)
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23365
diff changeset
   430
  thus ?case
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   431
    by (simp del: power_Suc add: prems mult_less_0_iff mult_neg_neg)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   432
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   433
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   434
lemma odd_0_le_power_imp_0_le:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   435
     "0 \<le> a  ^ Suc(2*n) ==> 0 \<le> (a::'a::{ordered_idom,recpower})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   436
apply (insert odd_power_less_zero [of a n]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   437
apply (force simp add: linorder_not_less [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   438
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   439
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   440
text{*Simprules for comparisons where common factors can be cancelled.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   441
lemmas zero_compare_simps =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   442
    add_strict_increasing add_strict_increasing2 add_increasing
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   443
    zero_le_mult_iff zero_le_divide_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   444
    zero_less_mult_iff zero_less_divide_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   445
    mult_le_0_iff divide_le_0_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   446
    mult_less_0_iff divide_less_0_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   447
    zero_le_power2 power2_less_0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   448
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   449
subsubsection{*Nat *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   450
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   451
lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   452
by (simp add: numerals)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   453
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   454
(*Expresses a natural number constant as the Suc of another one.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   455
  NOT suitable for rewriting because n recurs in the condition.*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   456
lemmas expand_Suc = Suc_pred' [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   457
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   458
subsubsection{*Arith *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   459
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   460
lemma Suc_eq_add_numeral_1: "Suc n = n + 1"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   461
by (simp add: numerals)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   462
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   463
lemma Suc_eq_add_numeral_1_left: "Suc n = 1 + n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   464
by (simp add: numerals)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   465
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   466
(* These two can be useful when m = number_of... *)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   467
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   468
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
   469
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   470
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   471
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
   472
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   473
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   474
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
   475
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   476
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   477
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   478
subsection{*Comparisons involving (0::nat) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   479
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   480
text{*Simplification already does @{term "n<0"}, @{term "n\<le>0"} and @{term "0\<le>n"}.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   481
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   482
lemma eq_number_of_0 [simp]:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   483
  "number_of v = (0::nat) \<longleftrightarrow> v \<le> Int.Pls"
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   484
  unfolding nat_number_of_def number_of_is_id numeral_simps
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   485
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   486
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   487
lemma eq_0_number_of [simp]:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   488
  "(0::nat) = number_of v \<longleftrightarrow> v \<le> Int.Pls"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   489
by (rule trans [OF eq_sym_conv eq_number_of_0])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   490
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   491
lemma less_0_number_of [simp]:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   492
   "(0::nat) < number_of v \<longleftrightarrow> Int.Pls < v"
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   493
  unfolding nat_number_of_def number_of_is_id numeral_simps
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   494
  by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   495
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   496
lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)"
28969
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   497
by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric])
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   498
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   499
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   500
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   501
subsection{*Comparisons involving  @{term Suc} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   502
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   503
lemma eq_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   504
     "(number_of v = Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   505
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   506
         if neg pv then False else nat pv = n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   507
apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   508
                  number_of_pred nat_number_of_def 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   509
            split add: split_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   510
apply (rule_tac x = "number_of v" in spec)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   511
apply (auto simp add: nat_eq_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   512
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   513
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   514
lemma Suc_eq_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   515
     "(Suc n = number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   516
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   517
         if neg pv then False else nat pv = n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   518
by (rule trans [OF eq_sym_conv eq_number_of_Suc])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   519
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   520
lemma less_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   521
     "(number_of v < Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   522
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   523
         if neg pv then True else nat pv < n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   524
apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   525
                  number_of_pred nat_number_of_def  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   526
            split add: split_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   527
apply (rule_tac x = "number_of v" in spec)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   528
apply (auto simp add: nat_less_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   529
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   530
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   531
lemma less_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   532
     "(Suc n < number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   533
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   534
         if neg pv then False else n < nat pv)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   535
apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   536
                  number_of_pred nat_number_of_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   537
            split add: split_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   538
apply (rule_tac x = "number_of v" in spec)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   539
apply (auto simp add: zless_nat_eq_int_zless)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   540
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   541
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   542
lemma le_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   543
     "(number_of v <= Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   544
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   545
         if neg pv then True else nat pv <= n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   546
by (simp add: Let_def less_Suc_number_of linorder_not_less [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   547
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   548
lemma le_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   549
     "(Suc n <= number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   550
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   551
         if neg pv then False else n <= nat pv)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   552
by (simp add: Let_def less_number_of_Suc linorder_not_less [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   553
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   554
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   555
lemma eq_number_of_Pls_Min: "(Numeral0 ::int) ~= number_of Int.Min"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   556
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   557
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   558
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   559
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   560
subsection{*Max and Min Combined with @{term Suc} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   561
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   562
lemma max_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   563
     "max (Suc n) (number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   564
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   565
         if neg pv then Suc n else Suc(max n (nat pv)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   566
apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   567
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   568
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   569
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   570
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   571
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   572
lemma max_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   573
     "max (number_of v) (Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   574
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   575
         if neg pv then Suc n else Suc(max (nat pv) n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   576
apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   577
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   578
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   579
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   580
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   581
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   582
lemma min_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   583
     "min (Suc n) (number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   584
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   585
         if neg pv then 0 else Suc(min n (nat pv)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   586
apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   587
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   588
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   589
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   590
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   591
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   592
lemma min_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   593
     "min (number_of v) (Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   594
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   595
         if neg pv then 0 else Suc(min (nat pv) n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   596
apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   597
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   598
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   599
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   600
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   601
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   602
subsection{*Literal arithmetic involving powers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   603
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   604
lemma nat_power_eq: "(0::int) <= z ==> nat (z^n) = nat z ^ n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   605
apply (induct "n")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   606
apply (simp_all (no_asm_simp) add: nat_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   607
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   608
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   609
lemma power_nat_number_of:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   610
     "(number_of v :: nat) ^ n =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   611
       (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   612
by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def nat_power_eq
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   613
         split add: split_if cong: imp_cong)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   614
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   615
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   616
lemmas power_nat_number_of_number_of = power_nat_number_of [of _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   617
declare power_nat_number_of_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   618
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   619
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   620
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   621
text{*For arbitrary rings*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   622
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   623
lemma power_number_of_even:
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   624
  fixes z :: "'a::{number_ring,recpower}"
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   625
  shows "z ^ number_of (Int.Bit0 w) = (let w = z ^ (number_of w) in w * w)"
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   626
unfolding Let_def nat_number_of_def number_of_Bit0
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   627
apply (rule_tac x = "number_of w" in spec, clarify)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   628
apply (case_tac " (0::int) <= x")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   629
apply (auto simp add: nat_mult_distrib power_even_eq power2_eq_square)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   630
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   631
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   632
lemma power_number_of_odd:
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   633
  fixes z :: "'a::{number_ring,recpower}"
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   634
  shows "z ^ number_of (Int.Bit1 w) = (if (0::int) <= number_of w
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   635
     then (let w = z ^ (number_of w) in z * w * w) else 1)"
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   636
unfolding Let_def nat_number_of_def number_of_Bit1
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   637
apply (rule_tac x = "number_of w" in spec, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   638
apply (simp only: nat_add_distrib nat_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   639
apply simp
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   640
apply (auto simp add: nat_add_distrib nat_mult_distrib power_even_eq power2_eq_square neg_nat power_Suc)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   641
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   642
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   643
lemmas zpower_number_of_even = power_number_of_even [where 'a=int]
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   644
lemmas zpower_number_of_odd = power_number_of_odd [where 'a=int]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   645
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   646
lemmas power_number_of_even_number_of [simp] =
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   647
    power_number_of_even [of "number_of v", standard]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   648
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   649
lemmas power_number_of_odd_number_of [simp] =
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   650
    power_number_of_odd [of "number_of v", standard]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   651
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   652
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   653
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   654
ML
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   655
{*
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 26086
diff changeset
   656
val numeral_ss = @{simpset} addsimps @{thms numerals};
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   657
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   658
val nat_bin_arith_setup =
30685
dd5fe091ff04 structure LinArith now named Lin_Arith
haftmann
parents: 30652
diff changeset
   659
 Lin_Arith.map_data
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   660
   (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   661
     {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   662
      inj_thms = inj_thms,
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   663
      lessD = lessD, neqE = neqE,
29039
8b9207f82a78 separate neg_simps from rel_simps
huffman
parents: 29012
diff changeset
   664
      simpset = simpset addsimps @{thms neg_simps} @
8b9207f82a78 separate neg_simps from rel_simps
huffman
parents: 29012
diff changeset
   665
        [@{thm Suc_nat_number_of}, @{thm int_nat_number_of}]})
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   666
*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   667
24075
366d4d234814 arith method setup: proper context;
wenzelm
parents: 23969
diff changeset
   668
declaration {* K nat_bin_arith_setup *}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   669
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   670
(* Enable arith to deal with div/mod k where k is a numeral: *)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   671
declare split_div[of _ _ "number_of k", standard, arith_split]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   672
declare split_mod[of _ _ "number_of k", standard, arith_split]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   673
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   674
lemma nat_number_of_Pls: "Numeral0 = (0::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   675
  by (simp add: number_of_Pls nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   676
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   677
lemma nat_number_of_Min: "number_of Int.Min = (0::nat)"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   678
  apply (simp only: number_of_Min nat_number_of_def nat_zminus_int)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   679
  done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   680
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   681
lemma nat_number_of_Bit0:
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   682
    "number_of (Int.Bit0 w) = (let n::nat = number_of w in n + n)"
28969
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   683
  unfolding nat_number_of_def number_of_is_id numeral_simps Let_def
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   684
  by auto
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   685
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   686
lemma nat_number_of_Bit1:
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   687
  "number_of (Int.Bit1 w) =
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   688
    (if neg (number_of w :: int) then 0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   689
     else let n = number_of w in Suc (n + n))"
28969
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   690
  unfolding nat_number_of_def number_of_is_id numeral_simps neg_def Let_def
28968
a4f3db5d1393 change some lemmas to avoid using iszero
huffman
parents: 28961
diff changeset
   691
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   692
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   693
lemmas nat_number =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   694
  nat_number_of_Pls nat_number_of_Min
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   695
  nat_number_of_Bit0 nat_number_of_Bit1
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   696
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   697
lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   698
  by (simp add: Let_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   699
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   700
lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring,recpower})"
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   701
by (simp add: power_mult power_Suc); 
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   702
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   703
lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring,recpower})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   704
by (simp add: power_mult power_Suc); 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   705
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   706
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   707
subsection{*Literal arithmetic and @{term of_nat}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   708
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   709
lemma of_nat_double:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   710
     "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
   711
by (simp only: mult_2 nat_add_distrib of_nat_add) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   712
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   713
lemma nat_numeral_m1_eq_0: "-1 = (0::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   714
by (simp only: nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   715
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   716
lemma of_nat_number_of_lemma:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   717
     "of_nat (number_of v :: nat) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   718
         (if 0 \<le> (number_of v :: int) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   719
          then (number_of v :: 'a :: number_ring)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   720
          else 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   721
by (simp add: int_number_of_def nat_number_of_def number_of_eq of_nat_nat);
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   722
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   723
lemma of_nat_number_of_eq [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   724
     "of_nat (number_of v :: nat) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   725
         (if neg (number_of v :: int) then 0  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   726
          else (number_of v :: 'a :: number_ring))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   727
by (simp only: of_nat_number_of_lemma neg_def, simp) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   728
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   729
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   730
subsection {*Lemmas for the Combination and Cancellation Simprocs*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   731
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   732
lemma nat_number_of_add_left:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   733
     "number_of v + (number_of v' + (k::nat)) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   734
         (if neg (number_of v :: int) then number_of v' + k  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   735
          else if neg (number_of v' :: int) then number_of v + k  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   736
          else number_of (v + v') + k)"
28968
a4f3db5d1393 change some lemmas to avoid using iszero
huffman
parents: 28961
diff changeset
   737
  unfolding nat_number_of_def number_of_is_id neg_def
a4f3db5d1393 change some lemmas to avoid using iszero
huffman
parents: 28961
diff changeset
   738
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   739
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   740
lemma nat_number_of_mult_left:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   741
     "number_of v * (number_of v' * (k::nat)) =  
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   742
         (if v < Int.Pls then 0
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   743
          else number_of (v * v') * k)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   744
by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   745
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   746
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   747
subsubsection{*For @{text combine_numerals}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   748
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   749
lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   750
by (simp add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   751
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   752
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   753
subsubsection{*For @{text cancel_numerals}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   754
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   755
lemma nat_diff_add_eq1:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   756
     "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   757
by (simp split add: nat_diff_split add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   758
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   759
lemma nat_diff_add_eq2:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   760
     "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   761
by (simp split add: nat_diff_split add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   762
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   763
lemma nat_eq_add_iff1:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   764
     "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   765
by (auto split add: nat_diff_split simp add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   766
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   767
lemma nat_eq_add_iff2:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   768
     "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   769
by (auto split add: nat_diff_split simp add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   770
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   771
lemma nat_less_add_iff1:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   772
     "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   773
by (auto split add: nat_diff_split simp add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   774
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   775
lemma nat_less_add_iff2:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   776
     "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   777
by (auto split add: nat_diff_split simp add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   778
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   779
lemma nat_le_add_iff1:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   780
     "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   781
by (auto split add: nat_diff_split simp add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   782
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   783
lemma nat_le_add_iff2:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   784
     "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   785
by (auto split add: nat_diff_split simp add: add_mult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   786
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   787
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   788
subsubsection{*For @{text cancel_numeral_factors} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   789
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   790
lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   791
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   792
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   793
lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m<n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   794
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   795
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   796
lemma nat_mult_eq_cancel1: "(0::nat) < k ==> (k*m = k*n) = (m=n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   797
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   798
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   799
lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   800
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   801
23969
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   802
lemma nat_mult_dvd_cancel_disj[simp]:
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   803
  "(k*m) dvd (k*n) = (k=0 | m dvd (n::nat))"
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   804
by(auto simp: dvd_eq_mod_eq_0 mod_mult_distrib2[symmetric])
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   805
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   806
lemma nat_mult_dvd_cancel1: "0 < k \<Longrightarrow> (k*m) dvd (k*n::nat) = (m dvd n)"
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   807
by(auto)
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   808
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   809
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   810
subsubsection{*For @{text cancel_factor} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   811
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   812
lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   813
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   814
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   815
lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m<n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   816
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   817
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   818
lemma nat_mult_eq_cancel_disj: "(k*m = k*n) = (k = (0::nat) | m=n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   819
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   820
23969
ef782bbf2d09 Added cancel simprocs for dvd on nat and int
nipkow
parents: 23389
diff changeset
   821
lemma nat_mult_div_cancel_disj[simp]:
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   822
     "(k*m) div (k*n) = (if k = (0::nat) then 0 else m div n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   823
by (simp add: nat_mult_div_cancel1)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   824
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   825
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   826
subsection {* Simprocs for the Naturals *}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   827
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   828
use "Tools/nat_simprocs.ML"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   829
declaration {* K nat_simprocs_setup *}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   830
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   831
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
   832
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   833
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
   834
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   835
lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   836
by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   837
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   838
text {*Now just instantiating @{text n} to @{text "number_of v"} does
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   839
  the right simplification, but with some redundant inequality
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   840
  tests.*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   841
lemma neg_number_of_pred_iff_0:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   842
  "neg (number_of (Int.pred v)::int) = (number_of v = (0::nat))"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   843
apply (subgoal_tac "neg (number_of (Int.pred v)) = (number_of v < Suc 0) ")
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   844
apply (simp only: less_Suc_eq_le le_0_eq)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   845
apply (subst less_number_of_Suc, simp)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   846
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   847
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   848
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
   849
   simproc*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   850
lemma Suc_diff_number_of:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   851
     "Int.Pls < v ==>
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   852
      Suc m - (number_of v) = m - (number_of (Int.pred v))"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   853
apply (subst Suc_diff_eq_diff_pred)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   854
apply simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   855
apply (simp del: nat_numeral_1_eq_1)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   856
apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   857
                        neg_number_of_pred_iff_0)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   858
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   859
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   860
lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   861
by (simp add: numerals split add: nat_diff_split)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   862
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   863
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   864
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
   865
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   866
lemma nat_case_number_of [simp]:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   867
     "nat_case a f (number_of v) =
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   868
        (let pv = number_of (Int.pred v) in
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   869
         if neg pv then a else f (nat pv))"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   870
by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   871
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   872
lemma nat_case_add_eq_if [simp]:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   873
     "nat_case a f ((number_of v) + n) =
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   874
       (let pv = number_of (Int.pred v) in
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   875
         if neg pv then nat_case a f n else f (nat pv + n))"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   876
apply (subst add_eq_if)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   877
apply (simp split add: nat.split
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   878
            del: nat_numeral_1_eq_1
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   879
            add: nat_numeral_1_eq_1 [symmetric]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   880
                 numeral_1_eq_Suc_0 [symmetric]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   881
                 neg_number_of_pred_iff_0)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   882
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   883
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   884
lemma nat_rec_number_of [simp]:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   885
     "nat_rec a f (number_of v) =
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   886
        (let pv = number_of (Int.pred v) in
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   887
         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   888
apply (case_tac " (number_of v) ::nat")
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   889
apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   890
apply (simp split add: split_if_asm)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   891
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   892
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   893
lemma nat_rec_add_eq_if [simp]:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   894
     "nat_rec a f (number_of v + n) =
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   895
        (let pv = number_of (Int.pred v) in
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   896
         if neg pv then nat_rec a f n
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   897
                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   898
apply (subst add_eq_if)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   899
apply (simp split add: nat.split
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   900
            del: nat_numeral_1_eq_1
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   901
            add: nat_numeral_1_eq_1 [symmetric]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   902
                 numeral_1_eq_Suc_0 [symmetric]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   903
                 neg_number_of_pred_iff_0)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   904
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   905
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   906
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   907
subsubsection{*Various Other Lemmas*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   908
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   909
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
   910
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   911
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
   912
lemma nat_mult_2: "2 * z = (z+z::nat)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   913
proof -
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   914
  have "2*z = (1 + 1)*z" by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   915
  also have "... = z+z" by (simp add: left_distrib)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   916
  finally show ?thesis .
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   917
qed
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   918
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   919
lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   920
by (subst mult_commute, rule nat_mult_2)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   921
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   922
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
   923
lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   924
by arith
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   925
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   926
lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   927
by arith
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   928
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   929
lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   930
by (simp add: nat_mult_2 [symmetric])
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   931
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   932
lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   933
apply (subgoal_tac "m mod 2 < 2")
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   934
apply (erule less_2_cases [THEN disjE])
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   935
apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   936
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   937
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   938
lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   939
apply (subgoal_tac "m mod 2 < 2")
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   940
apply (force simp del: mod_less_divisor, simp)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   941
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   942
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   943
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
   944
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   945
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
   946
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   947
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   948
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
   949
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   950
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   951
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
   952
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
   953
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   954
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   955
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   956
text{*These lemmas collapse some needless occurrences of Suc:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   957
    at least three Sucs, since two and fewer are rewritten back to Suc again!
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   958
    We already have some rules to simplify operands smaller than 3.*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   959
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   960
lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   961
by (simp add: Suc3_eq_add_3)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   962
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   963
lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   964
by (simp add: Suc3_eq_add_3)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   965
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   966
lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   967
by (simp add: Suc3_eq_add_3)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   968
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   969
lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   970
by (simp add: Suc3_eq_add_3)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   971
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   972
lemmas Suc_div_eq_add3_div_number_of =
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   973
    Suc_div_eq_add3_div [of _ "number_of v", standard]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   974
declare Suc_div_eq_add3_div_number_of [simp]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   975
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   976
lemmas Suc_mod_eq_add3_mod_number_of =
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   977
    Suc_mod_eq_add3_mod [of _ "number_of v", standard]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   978
declare Suc_mod_eq_add3_mod_number_of [simp]
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   979
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
   980
end