src/HOL/Nat_Numeral.thy
author bulwahn
Mon, 12 Dec 2011 13:45:54 +0100
changeset 45818 53a697f5454a
parent 45607 16b4f5774621
child 46026 83caa4f4bd56
permissions -rw-r--r--
hiding constants and facts in the Quickcheck_Exhaustive and Quickcheck_Narrowing theory;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30925
c38cbc0ac8d1 theory NatBin now named Nat_Numeral
haftmann
parents: 30685
diff changeset
     1
(*  Title:      HOL/Nat_Numeral.thy
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     3
    Copyright   1999  University of Cambridge
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     4
*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     5
30925
c38cbc0ac8d1 theory NatBin now named Nat_Numeral
haftmann
parents: 30685
diff changeset
     6
header {* Binary numerals for the natural numbers *}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     7
30925
c38cbc0ac8d1 theory NatBin now named Nat_Numeral
haftmann
parents: 30685
diff changeset
     8
theory Nat_Numeral
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
     9
imports Int
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    10
begin
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    11
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    12
subsection {* Numerals for natural numbers *}
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    13
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    14
text {*
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    15
  Arithmetic for naturals is reduced to that for the non-negative integers.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    16
*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    17
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    18
instantiation nat :: number_semiring
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    19
begin
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    20
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    21
definition
32069
6d28bbd33e2c prefer code_inline over code_unfold; use code_unfold_post where appropriate
haftmann
parents: 31998
diff changeset
    22
  nat_number_of_def [code_unfold, code del]: "number_of v = nat (number_of v)"
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    23
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    24
instance proof
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    25
  fix n show "number_of (int n) = (of_nat n :: nat)"
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    26
    unfolding nat_number_of_def number_of_eq by simp
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    27
qed
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
    28
 
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25481
diff changeset
    29
end
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    30
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31790
diff changeset
    31
lemma [code_post]:
25965
05df64f786a4 improved code theorem setup
haftmann
parents: 25919
diff changeset
    32
  "nat (number_of v) = number_of v"
05df64f786a4 improved code theorem setup
haftmann
parents: 25919
diff changeset
    33
  unfolding nat_number_of_def ..
05df64f786a4 improved code theorem setup
haftmann
parents: 25919
diff changeset
    34
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    35
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    36
subsection {* Special case: squares and cubes *}
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    37
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    38
lemma numeral_2_eq_2: "2 = Suc (Suc 0)"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    39
  by (simp add: nat_number_of_def)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    40
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    41
lemma numeral_3_eq_3: "3 = Suc (Suc (Suc 0))"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    42
  by (simp add: nat_number_of_def)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    43
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    44
context power
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    45
begin
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    46
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    47
abbreviation (xsymbols)
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    48
  power2 :: "'a \<Rightarrow> 'a"  ("(_\<twosuperior>)" [1000] 999) where
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    49
  "x\<twosuperior> \<equiv> x ^ 2"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    50
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    51
notation (latex output)
29401
94fd5dd918f5 rename abbreviation square -> power2, to match theorem names
huffman
parents: 29045
diff changeset
    52
  power2  ("(_\<twosuperior>)" [1000] 999)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    53
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    54
notation (HTML output)
29401
94fd5dd918f5 rename abbreviation square -> power2, to match theorem names
huffman
parents: 29045
diff changeset
    55
  power2  ("(_\<twosuperior>)" [1000] 999)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    56
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    57
end
fec1a04b7220 power operation defined generic
haftmann
parents: 30925
diff changeset
    58
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    59
context monoid_mult
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    60
begin
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    61
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    62
lemma power2_eq_square: "a\<twosuperior> = a * a"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    63
  by (simp add: numeral_2_eq_2)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    64
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    65
lemma power3_eq_cube: "a ^ 3 = a * a * a"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    66
  by (simp add: numeral_3_eq_3 mult_assoc)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    67
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    68
lemma power_even_eq:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    69
  "a ^ (2*n) = (a ^ n) ^ 2"
35047
1b2bae06c796 hide fact Nat.add_0_right; make add_0_right from Groups priority
haftmann
parents: 35043
diff changeset
    70
  by (subst mult_commute) (simp add: power_mult)
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    71
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    72
lemma power_odd_eq:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    73
  "a ^ Suc (2*n) = a * (a ^ n) ^ 2"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    74
  by (simp add: power_even_eq)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    75
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    76
end
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    77
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    78
context semiring_1
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    79
begin
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    80
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    81
lemma zero_power2 [simp]: "0\<twosuperior> = 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    82
  by (simp add: power2_eq_square)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    83
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    84
lemma one_power2 [simp]: "1\<twosuperior> = 1"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    85
  by (simp add: power2_eq_square)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    86
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    87
end
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    88
36823
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
    89
context ring_1
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    90
begin
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    91
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    92
lemma power2_minus [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    93
  "(- a)\<twosuperior> = a\<twosuperior>"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    94
  by (simp add: power2_eq_square)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    95
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    96
text{*
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    97
  We cannot prove general results about the numeral @{term "-1"},
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    98
  so we have to use @{term "- 1"} instead.
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
    99
*}
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   100
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   101
lemma power_minus1_even [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   102
  "(- 1) ^ (2*n) = 1"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   103
proof (induct n)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   104
  case 0 show ?case by simp
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   105
next
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   106
  case (Suc n) then show ?case by (simp add: power_add)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   107
qed
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   108
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   109
lemma power_minus1_odd:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   110
  "(- 1) ^ Suc (2*n) = - 1"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   111
  by simp
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   112
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   113
lemma power_minus_even [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   114
  "(-a) ^ (2*n) = a ^ (2*n)"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   115
  by (simp add: power_minus [of a]) 
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   116
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   117
end
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   118
36823
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   119
context ring_1_no_zero_divisors
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   120
begin
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   121
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   122
lemma zero_eq_power2 [simp]:
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   123
  "a\<twosuperior> = 0 \<longleftrightarrow> a = 0"
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   124
  unfolding power2_eq_square by simp
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   125
36964
a354605f03dc remove simp attribute from power2_eq_1_iff
huffman
parents: 36841
diff changeset
   126
lemma power2_eq_1_iff:
36823
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   127
  "a\<twosuperior> = 1 \<longleftrightarrow> a = 1 \<or> a = - 1"
36964
a354605f03dc remove simp attribute from power2_eq_1_iff
huffman
parents: 36841
diff changeset
   128
  unfolding power2_eq_square by (rule square_eq_1_iff)
36823
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   129
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   130
end
001d1aad99de add lemma power2_eq_1_iff; generalize some other lemmas
huffman
parents: 36719
diff changeset
   131
44345
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   132
context idom
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   133
begin
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   134
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   135
lemma power2_eq_iff: "x\<twosuperior> = y\<twosuperior> \<longleftrightarrow> x = y \<or> x = - y"
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   136
  unfolding power2_eq_square by (rule square_eq_iff)
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   137
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   138
end
881c324470a2 add lemma power2_eq_iff
huffman
parents: 43531
diff changeset
   139
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   140
context linordered_ring
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   141
begin
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   142
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   143
lemma sum_squares_ge_zero:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   144
  "0 \<le> x * x + y * y"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   145
  by (intro add_nonneg_nonneg zero_le_square)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   146
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   147
lemma not_sum_squares_lt_zero:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   148
  "\<not> x * x + y * y < 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   149
  by (simp add: not_less sum_squares_ge_zero)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   150
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   151
end
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   152
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   153
context linordered_ring_strict
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   154
begin
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   155
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   156
lemma sum_squares_eq_zero_iff:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   157
  "x * x + y * y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
31034
736f521ad036 dropped duplicate lemma sum_nonneg_eq_zero_iff
haftmann
parents: 31014
diff changeset
   158
  by (simp add: add_nonneg_eq_0_iff)
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   159
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   160
lemma sum_squares_le_zero_iff:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   161
  "x * x + y * y \<le> 0 \<longleftrightarrow> x = 0 \<and> y = 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   162
  by (simp add: le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   163
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   164
lemma sum_squares_gt_zero_iff:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   165
  "0 < x * x + y * y \<longleftrightarrow> x \<noteq> 0 \<or> y \<noteq> 0"
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   166
  by (simp add: not_le [symmetric] sum_squares_le_zero_iff)
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   167
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   168
end
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   169
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33342
diff changeset
   170
context linordered_semidom
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   171
begin
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   172
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   173
lemma power2_le_imp_le:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   174
  "x\<twosuperior> \<le> y\<twosuperior> \<Longrightarrow> 0 \<le> y \<Longrightarrow> x \<le> y"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   175
  unfolding numeral_2_eq_2 by (rule power_le_imp_le_base)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   176
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   177
lemma power2_less_imp_less:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   178
  "x\<twosuperior> < y\<twosuperior> \<Longrightarrow> 0 \<le> y \<Longrightarrow> x < y"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   179
  by (rule power_less_imp_less_base)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   180
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   181
lemma power2_eq_imp_eq:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   182
  "x\<twosuperior> = y\<twosuperior> \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> x = y"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   183
  unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   184
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   185
end
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   186
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33342
diff changeset
   187
context linordered_idom
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   188
begin
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   189
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   190
lemma zero_le_power2 [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   191
  "0 \<le> a\<twosuperior>"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   192
  by (simp add: power2_eq_square)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   193
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   194
lemma zero_less_power2 [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   195
  "0 < a\<twosuperior> \<longleftrightarrow> a \<noteq> 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   196
  by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   197
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   198
lemma power2_less_0 [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   199
  "\<not> a\<twosuperior> < 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   200
  by (force simp add: power2_eq_square mult_less_0_iff) 
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   201
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   202
lemma abs_power2 [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   203
  "abs (a\<twosuperior>) = a\<twosuperior>"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   204
  by (simp add: power2_eq_square abs_mult abs_mult_self)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   205
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   206
lemma power2_abs [simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   207
  "(abs a)\<twosuperior> = a\<twosuperior>"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   208
  by (simp add: power2_eq_square abs_mult_self)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   209
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   210
lemma odd_power_less_zero:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   211
  "a < 0 \<Longrightarrow> a ^ Suc (2*n) < 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   212
proof (induct n)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   213
  case 0
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   214
  then show ?case by simp
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   215
next
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   216
  case (Suc n)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   217
  have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   218
    by (simp add: mult_ac power_add power2_eq_square)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   219
  thus ?case
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   220
    by (simp del: power_Suc add: Suc mult_less_0_iff mult_neg_neg)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   221
qed
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   222
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   223
lemma odd_0_le_power_imp_0_le:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   224
  "0 \<le> a ^ Suc (2*n) \<Longrightarrow> 0 \<le> a"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   225
  using odd_power_less_zero [of a n]
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   226
    by (force simp add: linorder_not_less [symmetric]) 
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   227
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   228
lemma zero_le_even_power'[simp]:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   229
  "0 \<le> a ^ (2*n)"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   230
proof (induct n)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   231
  case 0
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   232
    show ?case by simp
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   233
next
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   234
  case (Suc n)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   235
    have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" 
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   236
      by (simp add: mult_ac power_add power2_eq_square)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   237
    thus ?case
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   238
      by (simp add: Suc zero_le_mult_iff)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   239
qed
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   240
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   241
lemma sum_power2_ge_zero:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   242
  "0 \<le> x\<twosuperior> + y\<twosuperior>"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   243
  unfolding power2_eq_square by (rule sum_squares_ge_zero)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   244
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   245
lemma not_sum_power2_lt_zero:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   246
  "\<not> x\<twosuperior> + y\<twosuperior> < 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   247
  unfolding power2_eq_square by (rule not_sum_squares_lt_zero)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   248
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   249
lemma sum_power2_eq_zero_iff:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   250
  "x\<twosuperior> + y\<twosuperior> = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   251
  unfolding power2_eq_square by (rule sum_squares_eq_zero_iff)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   252
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   253
lemma sum_power2_le_zero_iff:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   254
  "x\<twosuperior> + y\<twosuperior> \<le> 0 \<longleftrightarrow> x = 0 \<and> y = 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   255
  unfolding power2_eq_square by (rule sum_squares_le_zero_iff)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   256
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   257
lemma sum_power2_gt_zero_iff:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   258
  "0 < x\<twosuperior> + y\<twosuperior> \<longleftrightarrow> x \<noteq> 0 \<or> y \<noteq> 0"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   259
  unfolding power2_eq_square by (rule sum_squares_gt_zero_iff)
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   260
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   261
end
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   262
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   263
lemma power2_sum:
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   264
  fixes x y :: "'a::number_semiring"
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   265
  shows "(x + y)\<twosuperior> = x\<twosuperior> + y\<twosuperior> + 2 * x * y"
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   266
  by (simp add: algebra_simps power2_eq_square semiring_mult_2_right)
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   267
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   268
lemma power2_diff:
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   269
  fixes x y :: "'a::number_ring"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   270
  shows "(x - y)\<twosuperior> = x\<twosuperior> + y\<twosuperior> - 2 * x * y"
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   271
  by (simp add: ring_distribs power2_eq_square mult_2) (rule mult_commute)
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   272
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   273
29040
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   274
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
   275
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   276
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
   277
  "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
   278
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   279
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
   280
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
   281
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   282
lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   283
by (simp add: neg_def del: of_nat_Suc)
29040
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   284
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   285
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
   286
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   287
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
   288
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
   289
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   290
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
   291
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   292
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
   293
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
   294
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   295
lemma not_neg_1: "~ neg 1"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   296
by (simp add: neg_def linorder_not_less)
29040
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   297
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   298
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
   299
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
   300
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   301
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
   302
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
   303
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   304
text {*
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   305
  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
   306
  @{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
   307
*}
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   308
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   309
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
   310
  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
   311
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   312
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
   313
  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
   314
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   315
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
   316
  "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
   317
  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
   318
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   319
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
   320
  "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
   321
  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
   322
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   323
lemmas neg_simps [simp] =
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   324
  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
   325
  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
   326
  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
   327
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 29039
diff changeset
   328
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   329
subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   330
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   331
declare nat_1 [simp]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   332
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   333
lemma nat_number_of [simp]: "nat (number_of w) = number_of w"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   334
by (simp add: nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   335
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31790
diff changeset
   336
lemma nat_numeral_0_eq_0 [simp, code_post]: "Numeral0 = (0::nat)"
44884
02efd5a6b6e5 tuned proofs
huffman
parents: 44857
diff changeset
   337
  by (rule semiring_numeral_0_eq_0)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   338
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   339
lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)"
44884
02efd5a6b6e5 tuned proofs
huffman
parents: 44857
diff changeset
   340
  by (rule semiring_numeral_1_eq_1)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   341
36719
d396f6f63d94 moved some lemmas from Groebner_Basis here
haftmann
parents: 36716
diff changeset
   342
lemma Numeral1_eq1_nat:
d396f6f63d94 moved some lemmas from Groebner_Basis here
haftmann
parents: 36716
diff changeset
   343
  "(1::nat) = Numeral1"
d396f6f63d94 moved some lemmas from Groebner_Basis here
haftmann
parents: 36716
diff changeset
   344
  by simp
d396f6f63d94 moved some lemmas from Groebner_Basis here
haftmann
parents: 36716
diff changeset
   345
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31790
diff changeset
   346
lemma numeral_1_eq_Suc_0 [code_post]: "Numeral1 = Suc 0"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   347
by (simp only: nat_numeral_1_eq_1 One_nat_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   348
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   349
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   350
subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*}
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 int_nat_number_of [simp]:
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23307
diff changeset
   353
     "int (number_of v) =  
23307
2fe3345035c7 modify proofs to avoid referring to int::nat=>int
huffman
parents: 23294
diff changeset
   354
         (if neg (number_of v :: int) then 0  
2fe3345035c7 modify proofs to avoid referring to int::nat=>int
huffman
parents: 23294
diff changeset
   355
          else (number_of v :: int))"
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   356
  unfolding nat_number_of_def number_of_is_id neg_def
44884
02efd5a6b6e5 tuned proofs
huffman
parents: 44857
diff changeset
   357
  by simp (* FIXME: redundant with of_nat_number_of_eq *)
23307
2fe3345035c7 modify proofs to avoid referring to int::nat=>int
huffman
parents: 23294
diff changeset
   358
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   359
lemma nonneg_int_cases:
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   360
  fixes k :: int assumes "0 \<le> k" obtains n where "k = of_nat n"
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   361
  using assms by (cases k, simp, simp)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   362
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   363
subsubsection{*Successor *}
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 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
   366
apply (rule sym)
44766
d4d33a4d7548 avoid using legacy theorem names
huffman
parents: 44345
diff changeset
   367
apply (simp add: nat_eq_iff)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   368
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   369
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   370
lemma Suc_nat_number_of_add:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   371
     "Suc (number_of v + n) =  
28984
060832a1f087 change arith_special simps to avoid using neg
huffman
parents: 28969
diff changeset
   372
        (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
   373
  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
   374
  by (simp add: Suc_nat_eq_nat_zadd1 add_ac)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   375
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   376
lemma Suc_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   377
     "Suc (number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   378
        (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
   379
apply (cut_tac n = 0 in Suc_nat_number_of_add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   380
apply (simp cong del: if_weak_cong)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   381
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   382
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   383
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   384
subsubsection{*Addition *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   385
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   386
lemma add_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   387
     "(number_of v :: nat) + number_of v' =  
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   388
         (if v < Int.Pls then number_of v'  
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   389
          else if v' < Int.Pls then number_of v  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   390
          else number_of (v + v'))"
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   391
  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
   392
  by (simp add: nat_add_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   393
30081
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   394
lemma nat_number_of_add_1 [simp]:
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   395
  "number_of v + (1::nat) =
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   396
    (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
   397
  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
   398
  by (simp add: nat_add_distrib)
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   399
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   400
lemma nat_1_add_number_of [simp]:
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   401
  "(1::nat) + number_of v =
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   402
    (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
   403
  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
   404
  by (simp add: nat_add_distrib)
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   405
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   406
lemma nat_1_add_1 [simp]: "1 + 1 = (2::nat)"
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   407
  by (rule semiring_one_add_one_is_two)
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   408
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   409
text {* TODO: replace simp rules above with these generic ones: *}
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   410
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   411
lemma semiring_add_number_of:
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   412
  "\<lbrakk>Int.Pls \<le> v; Int.Pls \<le> v'\<rbrakk> \<Longrightarrow>
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   413
    (number_of v :: 'a::number_semiring) + number_of v' = number_of (v + v')"
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   414
  unfolding Int.Pls_def
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   415
  by (elim nonneg_int_cases,
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   416
    simp only: number_of_int of_nat_add [symmetric])
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   417
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   418
lemma semiring_number_of_add_1:
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   419
  "Int.Pls \<le> v \<Longrightarrow>
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   420
    number_of v + (1::'a::number_semiring) = number_of (Int.succ v)"
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   421
  unfolding Int.Pls_def Int.succ_def
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   422
  by (elim nonneg_int_cases,
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   423
    simp only: number_of_int add_commute [where b=1] of_nat_Suc [symmetric])
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   424
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   425
lemma semiring_1_add_number_of:
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   426
  "Int.Pls \<le> v \<Longrightarrow>
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   427
    (1::'a::number_semiring) + number_of v = number_of (Int.succ v)"
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   428
  unfolding Int.Pls_def Int.succ_def
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   429
  by (elim nonneg_int_cases,
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   430
    simp only: number_of_int add_commute [where b=1] of_nat_Suc [symmetric])
30081
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   431
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   432
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   433
subsubsection{*Subtraction *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   434
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   435
lemma diff_nat_eq_if:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   436
     "nat z - nat z' =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   437
        (if neg z' then nat z   
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   438
         else let d = z-z' in     
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   439
              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
   440
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
   441
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   442
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   443
lemma diff_nat_number_of [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   444
     "(number_of v :: nat) - number_of v' =  
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   445
        (if v' < Int.Pls then number_of v  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   446
         else let d = number_of (v + uminus v') in     
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   447
              if neg d then 0 else nat d)"
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   448
  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
   449
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   450
30081
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   451
lemma nat_number_of_diff_1 [simp]:
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   452
  "number_of v - (1::nat) =
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   453
    (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
   454
  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
   455
  by auto
46b9c8ae3897 add simp rules for numerals with 1::nat
huffman
parents: 30079
diff changeset
   456
23164
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{*Multiplication *}
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 mult_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   461
     "(number_of v :: nat) * number_of v' =  
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   462
       (if v < Int.Pls then 0 else number_of (v * v'))"
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   463
  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
   464
  by (simp add: nat_mult_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   465
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   466
(* TODO: replace mult_nat_number_of with this next rule *)
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   467
lemma semiring_mult_number_of:
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   468
  "\<lbrakk>Int.Pls \<le> v; Int.Pls \<le> v'\<rbrakk> \<Longrightarrow>
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   469
    (number_of v :: 'a::number_semiring) * number_of v' = number_of (v * v')"
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   470
  unfolding Int.Pls_def
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   471
  by (elim nonneg_int_cases,
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   472
    simp only: number_of_int of_nat_mult [symmetric])
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   473
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   474
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   475
subsection{*Comparisons*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   476
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   477
subsubsection{*Equals (=) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   478
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   479
lemma eq_nat_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   480
     "((number_of v :: nat) = number_of v') =  
28969
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   481
      (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
   482
       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
   483
       else v = v')"
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
diff changeset
   484
  unfolding nat_number_of_def number_of_is_id neg_def
4ed63cdda799 change more lemmas to avoid using iszero
huffman
parents: 28968
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
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   488
subsubsection{*Less-than (<) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   489
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   490
lemma less_nat_number_of [simp]:
29011
a47003001699 simplify less_nat_number_of
huffman
parents: 29010
diff changeset
   491
  "(number_of v :: nat) < number_of v' \<longleftrightarrow>
a47003001699 simplify less_nat_number_of
huffman
parents: 29010
diff changeset
   492
    (if v < v' then Int.Pls < v' else False)"
a47003001699 simplify less_nat_number_of
huffman
parents: 29010
diff changeset
   493
  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
   494
  by auto
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
29010
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   497
subsubsection{*Less-than-or-equal *}
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   498
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   499
lemma le_nat_number_of [simp]:
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   500
  "(number_of v :: nat) \<le> number_of v' \<longleftrightarrow>
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   501
    (if v \<le> v' then True else v \<le> Int.Pls)"
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   502
  unfolding nat_number_of_def number_of_is_id numeral_simps
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   503
  by auto
5cd646abf6bc add lemma le_nat_number_of
huffman
parents: 28984
diff changeset
   504
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   505
(*Maps #n to n for n = 0, 1, 2*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   506
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
   507
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   508
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   509
subsection{*Powers with Numeric Exponents*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   510
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   511
text{*Squares of literal numerals will be evaluated.*}
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   512
lemmas power2_eq_square_number_of [simp] =
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 44884
diff changeset
   513
  power2_eq_square [of "number_of w"] for w
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   514
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   515
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   516
text{*Simprules for comparisons where common factors can be cancelled.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   517
lemmas zero_compare_simps =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   518
    add_strict_increasing add_strict_increasing2 add_increasing
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   519
    zero_le_mult_iff zero_le_divide_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   520
    zero_less_mult_iff zero_less_divide_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   521
    mult_le_0_iff divide_le_0_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   522
    mult_less_0_iff divide_less_0_iff 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   523
    zero_le_power2 power2_less_0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   524
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   525
subsubsection{*Nat *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   526
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   527
lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   528
by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   529
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   530
(*Expresses a natural number constant as the Suc of another one.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   531
  NOT suitable for rewriting because n recurs in the condition.*)
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 44884
diff changeset
   532
lemmas expand_Suc = Suc_pred' [of "number_of v"] for v
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   533
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   534
subsubsection{*Arith *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   535
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31182
diff changeset
   536
lemma Suc_eq_plus1: "Suc n = n + 1"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   537
  unfolding One_nat_def by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   538
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31182
diff changeset
   539
lemma Suc_eq_plus1_left: "Suc n = 1 + n"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   540
  unfolding One_nat_def by simp
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   541
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   542
(* These two can be useful when m = number_of... *)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   543
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   544
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
   545
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   546
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   547
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
   548
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   549
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   550
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
   551
  unfolding One_nat_def by (cases m) simp_all
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   552
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   553
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   554
subsection{*Comparisons involving (0::nat) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   555
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   556
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
   557
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   558
lemma eq_number_of_0 [simp]:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   559
  "number_of v = (0::nat) \<longleftrightarrow> v \<le> Int.Pls"
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   560
  unfolding nat_number_of_def number_of_is_id numeral_simps
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   561
  by auto
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   562
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   563
lemma eq_0_number_of [simp]:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   564
  "(0::nat) = number_of v \<longleftrightarrow> v \<le> Int.Pls"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   565
by (rule trans [OF eq_sym_conv eq_number_of_0])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   566
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   567
lemma less_0_number_of [simp]:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   568
   "(0::nat) < number_of v \<longleftrightarrow> Int.Pls < v"
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   569
  unfolding nat_number_of_def number_of_is_id numeral_simps
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 29011
diff changeset
   570
  by simp
23164
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 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
   573
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
   574
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   575
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   576
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   577
subsection{*Comparisons involving  @{term Suc} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   578
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   579
lemma eq_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   580
     "(number_of v = Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   581
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   582
         if neg pv then False else nat pv = n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   583
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
   584
                  number_of_pred nat_number_of_def 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   585
            split add: split_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   586
apply (rule_tac x = "number_of v" in spec)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   587
apply (auto simp add: nat_eq_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   588
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   589
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   590
lemma Suc_eq_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   591
     "(Suc n = number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   592
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   593
         if neg pv then False else nat pv = n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   594
by (rule trans [OF eq_sym_conv eq_number_of_Suc])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   595
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   596
lemma less_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   597
     "(number_of v < Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   598
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   599
         if neg pv then True else nat pv < n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   600
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
   601
                  number_of_pred nat_number_of_def  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   602
            split add: split_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   603
apply (rule_tac x = "number_of v" in spec)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   604
apply (auto simp add: nat_less_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   605
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   606
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   607
lemma less_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   608
     "(Suc n < number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   609
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   610
         if neg pv then False else n < nat pv)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   611
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
   612
                  number_of_pred nat_number_of_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   613
            split add: split_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   614
apply (rule_tac x = "number_of v" in spec)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   615
apply (auto simp add: zless_nat_eq_int_zless)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   616
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   617
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   618
lemma le_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   619
     "(number_of v <= Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   620
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   621
         if neg pv then True else nat pv <= n)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   622
by (simp add: Let_def linorder_not_less [symmetric])
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   623
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   624
lemma le_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   625
     "(Suc n <= number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   626
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   627
         if neg pv then False else n <= nat pv)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   628
by (simp add: Let_def linorder_not_less [symmetric])
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   629
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   630
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   631
lemma eq_number_of_Pls_Min: "(Numeral0 ::int) ~= number_of Int.Min"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   632
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   633
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   634
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   635
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   636
subsection{*Max and Min Combined with @{term Suc} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   637
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   638
lemma max_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   639
     "max (Suc n) (number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   640
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   641
         if neg pv then Suc n else Suc(max n (nat pv)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   642
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
   643
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   644
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   645
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   646
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   647
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   648
lemma max_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   649
     "max (number_of v) (Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   650
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   651
         if neg pv then Suc n else Suc(max (nat pv) n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   652
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
   653
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   654
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   655
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   656
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   657
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   658
lemma min_number_of_Suc [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   659
     "min (Suc n) (number_of v) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   660
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   661
         if neg pv then 0 else Suc(min n (nat pv)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   662
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
   663
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   664
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   665
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   666
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   667
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   668
lemma min_Suc_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   669
     "min (number_of v) (Suc n) =  
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25571
diff changeset
   670
        (let pv = number_of (Int.pred v) in  
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   671
         if neg pv then 0 else Suc(min (nat pv) n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   672
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
   673
            split add: split_if nat.split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   674
apply (rule_tac x = "number_of v" in spec) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   675
apply auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   676
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   677
 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   678
subsection{*Literal arithmetic involving powers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   679
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   680
lemma power_nat_number_of:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   681
     "(number_of v :: nat) ^ n =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   682
       (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
   683
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
   684
         split add: split_if cong: imp_cong)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   685
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   686
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 44884
diff changeset
   687
lemmas power_nat_number_of_number_of = power_nat_number_of [of _ "number_of w"] for w
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   688
declare power_nat_number_of_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   689
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   690
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   691
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   692
text{*For arbitrary rings*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   693
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   694
lemma power_number_of_even:
43526
2b92a6943915 generalize lemmas power_number_of_even and power_number_of_odd
huffman
parents: 40690
diff changeset
   695
  fixes z :: "'a::monoid_mult"
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   696
  shows "z ^ number_of (Int.Bit0 w) = (let w = z ^ (number_of w) in w * w)"
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   697
by (cases "w \<ge> 0") (auto simp add: Let_def Bit0_def nat_number_of_def number_of_is_id
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   698
  nat_add_distrib power_add simp del: nat_number_of)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   699
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   700
lemma power_number_of_odd:
43526
2b92a6943915 generalize lemmas power_number_of_even and power_number_of_odd
huffman
parents: 40690
diff changeset
   701
  fixes z :: "'a::monoid_mult"
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   702
  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
   703
     then (let w = z ^ (number_of w) in z * w * w) else 1)"
35815
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35631
diff changeset
   704
unfolding Let_def Bit1_def nat_number_of_def number_of_is_id
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35631
diff changeset
   705
apply (cases "0 <= w")
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35631
diff changeset
   706
apply (simp only: mult_assoc nat_add_distrib power_add, simp)
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   707
apply (simp add: not_le mult_2 [symmetric] add_assoc)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   708
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   709
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   710
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
   711
lemmas zpower_number_of_odd = power_number_of_odd [where 'a=int]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   712
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   713
lemmas power_number_of_even_number_of [simp] =
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 44884
diff changeset
   714
    power_number_of_even [of "number_of v"] for v
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   715
23294
9302a50a5bc9 generalize zpower_number_of_{even,odd} lemmas
huffman
parents: 23277
diff changeset
   716
lemmas power_number_of_odd_number_of [simp] =
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 44884
diff changeset
   717
    power_number_of_odd [of "number_of v"] for v
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   718
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   719
lemma nat_number_of_Pls: "Numeral0 = (0::nat)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   720
  by (simp add: nat_number_of_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   721
40690
3f472e57446a added "no_atp" for fact that confuses the SMT normalizer and that doesn't help ATPs anyway
blanchet
parents: 40077
diff changeset
   722
lemma nat_number_of_Min [no_atp]: "number_of Int.Min = (0::nat)"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   723
  apply (simp only: number_of_Min nat_number_of_def nat_zminus_int)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   724
  done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   725
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   726
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
   727
    "number_of (Int.Bit0 w) = (let n::nat = number_of w in n + n)"
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   728
by (cases "w \<ge> 0") (auto simp add: Let_def Bit0_def nat_number_of_def number_of_is_id
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   729
  nat_add_distrib simp del: nat_number_of)
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   730
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25965
diff changeset
   731
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
   732
  "number_of (Int.Bit1 w) =
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   733
    (if neg (number_of w :: int) then 0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   734
     else let n = number_of w in Suc (n + n))"
35815
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35631
diff changeset
   735
unfolding Let_def Bit1_def nat_number_of_def number_of_is_id neg_def
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35631
diff changeset
   736
apply (cases "w < 0")
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   737
apply (simp add: mult_2 [symmetric] add_assoc)
35815
10e723e54076 tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
boehmes
parents: 35631
diff changeset
   738
apply (simp only: nat_add_distrib, simp)
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   739
done
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   740
40077
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 36964
diff changeset
   741
lemmas eval_nat_numeral =
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   742
  nat_number_of_Bit0 nat_number_of_Bit1
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   743
36699
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   744
lemmas nat_arith =
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   745
  add_nat_number_of
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   746
  diff_nat_number_of
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   747
  mult_nat_number_of
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   748
  eq_nat_number_of
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   749
  less_nat_number_of
816da1023508 moved nat_arith ot Nat_Numeral: clarified normalizer setup
haftmann
parents: 35815
diff changeset
   750
36716
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   751
lemmas semiring_norm =
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   752
  Let_def arith_simps nat_arith rel_simps neg_simps if_False
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   753
  if_True add_0 add_Suc add_number_of_left mult_number_of_left
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   754
  numeral_1_eq_1 [symmetric] Suc_eq_plus1
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   755
  numeral_0_eq_0 [symmetric] numerals [symmetric]
36841
02df88789641 include iszero_simps in semiring_norm just once (they are already included in rel_simps)
huffman
parents: 36823
diff changeset
   756
  not_iszero_Numeral1
36716
b09f3ad3208f moved generic lemmas to appropriate places
haftmann
parents: 36699
diff changeset
   757
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   758
lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)"
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   759
  by (fact Let_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   760
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   761
lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring})"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   762
  by (simp only: number_of_Min power_minus1_even)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   763
31014
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   764
lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring})"
79f0858d9d49 collected square lemmas in Nat_Numeral
haftmann
parents: 31002
diff changeset
   765
  by (simp only: number_of_Min power_minus1_odd)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   766
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   767
lemma nat_number_of_add_left:
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   768
     "number_of v + (number_of v' + (k::nat)) =  
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   769
         (if neg (number_of v :: int) then number_of v' + k  
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   770
          else if neg (number_of v' :: int) then number_of v + k  
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   771
          else number_of (v + v') + k)"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   772
by (auto simp add: neg_def)
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   773
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   774
lemma nat_number_of_mult_left:
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   775
     "number_of v * (number_of v' * (k::nat)) =  
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   776
         (if v < Int.Pls then 0
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   777
          else number_of (v * v') * k)"
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   778
by (auto simp add: not_less Pls_def nat_number_of_def number_of_is_id
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   779
  nat_mult_distrib simp del: nat_number_of)
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   780
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   781
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   782
subsection{*Literal arithmetic and @{term of_nat}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   783
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   784
lemma of_nat_double:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   785
     "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
   786
by (simp only: mult_2 nat_add_distrib of_nat_add) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   787
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   788
lemma nat_numeral_m1_eq_0: "-1 = (0::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   789
by (simp only: nat_number_of_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   790
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   791
lemma of_nat_number_of_lemma:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   792
     "of_nat (number_of v :: nat) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   793
         (if 0 \<le> (number_of v :: int) 
44857
73d5b722c4b4 generalize lemma of_nat_number_of_eq to class number_semiring
huffman
parents: 44766
diff changeset
   794
          then (number_of v :: 'a :: number_semiring)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   795
          else 0)"
44857
73d5b722c4b4 generalize lemma of_nat_number_of_eq to class number_semiring
huffman
parents: 44766
diff changeset
   796
  by (auto simp add: int_number_of_def nat_number_of_def number_of_int
73d5b722c4b4 generalize lemma of_nat_number_of_eq to class number_semiring
huffman
parents: 44766
diff changeset
   797
    elim!: nonneg_int_cases)
23164
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 of_nat_number_of_eq [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   800
     "of_nat (number_of v :: nat) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   801
         (if neg (number_of v :: int) then 0  
44857
73d5b722c4b4 generalize lemma of_nat_number_of_eq to class number_semiring
huffman
parents: 44766
diff changeset
   802
          else (number_of v :: 'a :: number_semiring))"
73d5b722c4b4 generalize lemma of_nat_number_of_eq to class number_semiring
huffman
parents: 44766
diff changeset
   803
  by (simp only: of_nat_number_of_lemma neg_def, simp)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   804
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   805
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   806
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
   807
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   808
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
   809
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   810
lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   811
by (simp split: nat_diff_split)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   812
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   813
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
   814
  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
   815
  tests.*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   816
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
   817
  "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
   818
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
   819
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
   820
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
   821
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   822
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   823
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
   824
   simproc*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   825
lemma Suc_diff_number_of:
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   826
     "Int.Pls < v ==>
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   827
      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
   828
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
   829
apply simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   830
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
   831
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
   832
                        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
   833
done
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 diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35047
diff changeset
   836
by (simp split: nat_diff_split)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   837
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   838
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   839
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
   840
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   841
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
   842
     "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
   843
        (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
   844
         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
   845
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
   846
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   847
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
   848
     "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
   849
       (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
   850
         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
   851
apply (subst add_eq_if)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   852
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
   853
            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
   854
            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
   855
                 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
   856
                 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
   857
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   858
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   859
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
   860
     "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
   861
        (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
   862
         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
   863
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
   864
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
   865
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
   866
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   867
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   868
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
   869
     "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
   870
        (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
   871
         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
   872
                   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
   873
apply (subst add_eq_if)
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   874
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
   875
            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
   876
            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
   877
                 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
   878
                 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
   879
done
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   880
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   881
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   882
subsubsection{*Various Other Lemmas*}
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   883
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31068
diff changeset
   884
lemma card_UNIV_bool[simp]: "card (UNIV :: bool set) = 2"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31068
diff changeset
   885
by(simp add: UNIV_bool)
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31068
diff changeset
   886
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   887
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
   888
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   889
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
   890
lemma nat_mult_2: "2 * z = (z+z::nat)"
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   891
by (rule semiring_mult_2)
30652
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_mult_2_right: "z * 2 = (z+z::nat)"
43531
cc46a678faaf added number_semiring class, plus a few new lemmas;
huffman
parents: 43526
diff changeset
   894
by (rule semiring_mult_2_right)
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   895
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   896
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
   897
lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32069
diff changeset
   898
by (auto simp add: nat_1_add_1 [symmetric])
30652
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   899
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   900
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
   901
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   902
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
   903
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   904
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   905
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
   906
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   907
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   908
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
   909
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
   910
by simp
752329615264 distributed contents of theory Arith_Tools to theories Int, IntDiv and NatBin accordingly
haftmann
parents: 30497
diff changeset
   911
31096
e546e15089ef newline at end of file
huffman
parents: 31080
diff changeset
   912
end