src/HOL/Parity.thy
author wenzelm
Tue, 03 Sep 2013 01:12:40 +0200
changeset 53374 a14d2a854c02
parent 47225 650318981557
child 54227 63b441f49645
permissions -rw-r--r--
tuned proofs -- clarified flow of facts wrt. calculation;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 36840
diff changeset
     1
(*  Title:      HOL/Parity.thy
b460124855b8 tuned headers;
wenzelm
parents: 36840
diff changeset
     2
    Author:     Jeremy Avigad
b460124855b8 tuned headers;
wenzelm
parents: 36840
diff changeset
     3
    Author:     Jacques D. Fleuriot
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     4
*)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     5
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     6
header {* Even and Odd for int and nat *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     7
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     8
theory Parity
30738
0842e906300c normalized imports
haftmann
parents: 30056
diff changeset
     9
imports Main
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    10
begin
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    11
29608
564ea783ace8 no base sort in class import
haftmann
parents: 28952
diff changeset
    12
class even_odd = 
22390
378f34b1e380 now using "class"
haftmann
parents: 21404
diff changeset
    13
  fixes even :: "'a \<Rightarrow> bool"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    14
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    15
abbreviation
22390
378f34b1e380 now using "class"
haftmann
parents: 21404
diff changeset
    16
  odd :: "'a\<Colon>even_odd \<Rightarrow> bool" where
378f34b1e380 now using "class"
haftmann
parents: 21404
diff changeset
    17
  "odd x \<equiv> \<not> even x"
378f34b1e380 now using "class"
haftmann
parents: 21404
diff changeset
    18
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26236
diff changeset
    19
instantiation nat and int  :: even_odd
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    20
begin
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    21
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    22
definition
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    23
  even_def [presburger]: "even x \<longleftrightarrow> (x\<Colon>int) mod 2 = 0"
22390
378f34b1e380 now using "class"
haftmann
parents: 21404
diff changeset
    24
25571
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    25
definition
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    26
  even_nat_def [presburger]: "even x \<longleftrightarrow> even (int x)"
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    27
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    28
instance ..
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    29
c9e39eafc7a0 instantiation target rather than legacy instance
haftmann
parents: 25502
diff changeset
    30
end
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    31
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
    32
lemma transfer_int_nat_relations:
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
    33
  "even (int x) \<longleftrightarrow> even x"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
    34
  by (simp add: even_nat_def)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
    35
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35631
diff changeset
    36
declare transfer_morphism_int_nat[transfer add return:
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
    37
  transfer_int_nat_relations
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 31718
diff changeset
    38
]
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    39
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    40
lemma even_zero_int[simp]: "even (0::int)" by presburger
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    41
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    42
lemma odd_one_int[simp]: "odd (1::int)" by presburger
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    43
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    44
lemma even_zero_nat[simp]: "even (0::nat)" by presburger
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    45
31718
7715d4d3586f fixed thm name
nipkow
parents: 31148
diff changeset
    46
lemma odd_1_nat [simp]: "odd (1::nat)" by presburger
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    47
47224
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    48
lemma even_numeral_int [simp]: "even (numeral (Num.Bit0 k) :: int)"
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    49
  unfolding even_def by simp
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    50
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    51
lemma odd_numeral_int [simp]: "odd (numeral (Num.Bit1 k) :: int)"
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    52
  unfolding even_def by simp
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    53
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
    54
(* TODO: proper simp rules for Num.Bit0, Num.Bit1 *)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
    55
declare even_def[of "neg_numeral v", simp] for v
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    56
47224
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    57
lemma even_numeral_nat [simp]: "even (numeral (Num.Bit0 k) :: nat)"
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    58
  unfolding even_nat_def by simp
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    59
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    60
lemma odd_numeral_nat [simp]: "odd (numeral (Num.Bit1 k) :: nat)"
773fe2754b8c add simp rules for eve/odd on numerals
huffman
parents: 47163
diff changeset
    61
  unfolding even_nat_def by simp
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    62
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    63
subsection {* Even and odd are mutually exclusive *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    64
21263
wenzelm
parents: 21256
diff changeset
    65
lemma int_pos_lt_two_imp_zero_or_one:
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    66
    "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    67
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    68
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    69
lemma neq_one_mod_two [simp, presburger]: 
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
    70
  "((x::int) mod 2 ~= 0) = (x mod 2 = 1)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    71
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
    72
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    73
subsection {* Behavior under integer arithmetic operations *}
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
    74
declare dvd_def[algebra]
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
    75
lemma nat_even_iff_2_dvd[algebra]: "even (x::nat) \<longleftrightarrow> 2 dvd x"
36840
1e020f445846 fix duplicate simp rule warning
huffman
parents: 36722
diff changeset
    76
  by presburger
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
    77
lemma int_even_iff_2_dvd[algebra]: "even (x::int) \<longleftrightarrow> 2 dvd x"
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
    78
  by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    79
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    80
lemma even_times_anything: "even (x::int) ==> even (x * y)"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
    81
  by algebra
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    82
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
    83
lemma anything_times_even: "even (y::int) ==> even (x * y)" by algebra
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    84
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
    85
lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)" 
47163
248376f8881d remove redundant lemma
huffman
parents: 47108
diff changeset
    86
  by (simp add: even_def mod_mult_right_eq)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    87
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    88
lemma even_product[simp,presburger]: "even((x::int) * y) = (even x | even y)"
21263
wenzelm
parents: 21256
diff changeset
    89
  apply (auto simp add: even_times_anything anything_times_even)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    90
  apply (rule ccontr)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    91
  apply (auto simp add: odd_times_odd)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    92
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    93
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    94
lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    95
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    96
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    97
lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
    98
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    99
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   100
lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   101
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   102
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   103
lemma odd_plus_odd: "odd (x::int) ==> odd y ==> even (x + y)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   104
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   105
lemma even_sum[simp,presburger]:
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   106
  "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   107
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   108
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   109
lemma even_neg[simp,presburger,algebra]: "even (-(x::int)) = even x"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   110
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   111
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   112
lemma even_difference[simp]:
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   113
    "even ((x::int) - y) = ((even x & even y) | (odd x & odd y))" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   114
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   115
lemma even_power[simp,presburger]: "even ((x::int)^n) = (even x & n \<noteq> 0)"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   116
by (induct n) auto
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   117
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   118
lemma odd_pow: "odd x ==> odd((x::int)^n)" by simp
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   119
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   120
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   121
subsection {* Equivalent definitions *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   122
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   123
lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x" 
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   124
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   125
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   126
lemma two_times_odd_div_two_plus_one:
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   127
  "odd (x::int) ==> 2 * (x div 2) + 1 = x"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   128
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   129
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   130
lemma even_equiv_def: "even (x::int) = (EX y. x = 2 * y)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   131
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   132
lemma odd_equiv_def: "odd (x::int) = (EX y. x = 2 * y + 1)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   133
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   134
subsection {* even and odd for nats *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   135
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   136
lemma pos_int_even_equiv_nat_even: "0 \<le> x ==> even x = even (nat x)"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   137
by (simp add: even_nat_def)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   138
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   139
lemma even_product_nat[simp,presburger,algebra]:
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   140
  "even((x::nat) * y) = (even x | even y)"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   141
by (simp add: even_nat_def int_mult)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   142
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   143
lemma even_sum_nat[simp,presburger,algebra]:
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   144
  "even ((x::nat) + y) = ((even x & even y) | (odd x & odd y))"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   145
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   146
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   147
lemma even_difference_nat[simp,presburger,algebra]:
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   148
  "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   149
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   150
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   151
lemma even_Suc[simp,presburger,algebra]: "even (Suc x) = odd x"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   152
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   153
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   154
lemma even_power_nat[simp,presburger,algebra]:
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   155
  "even ((x::nat)^y) = (even x & 0 < y)"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   156
by (simp add: even_nat_def int_power)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   157
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   158
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   159
subsection {* Equivalent definitions *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   160
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   161
lemma nat_lt_two_imp_zero_or_one:
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   162
  "(x::nat) < Suc (Suc 0) ==> x = 0 | x = Suc 0"
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   163
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   164
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   165
lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   166
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   167
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   168
lemma odd_nat_mod_two_eq_one: "odd (x::nat) ==> x mod (Suc (Suc 0)) = Suc 0"
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   169
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   170
21263
wenzelm
parents: 21256
diff changeset
   171
lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   172
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   173
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   174
lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   175
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   176
21263
wenzelm
parents: 21256
diff changeset
   177
lemma even_nat_div_two_times_two: "even (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   178
    Suc (Suc 0) * (x div Suc (Suc 0)) = x" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   179
21263
wenzelm
parents: 21256
diff changeset
   180
lemma odd_nat_div_two_times_two_plus_one: "odd (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   181
    Suc( Suc (Suc 0) * (x div Suc (Suc 0))) = x" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   182
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   183
lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   184
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   185
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   186
lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))"
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   187
by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   188
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   189
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   190
subsection {* Parity and powers *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   191
21263
wenzelm
parents: 21256
diff changeset
   192
lemma  minus_one_even_odd_power:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   193
     "(even x --> (- 1::'a::{comm_ring_1})^x = 1) &
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   194
      (odd x --> (- 1::'a)^x = - 1)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   195
  apply (induct x)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   196
  apply (rule conjI)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   197
  apply simp
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   198
  apply (insert even_zero_nat, blast)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35043
diff changeset
   199
  apply simp
21263
wenzelm
parents: 21256
diff changeset
   200
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   201
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   202
lemma minus_one_even_power [simp]:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   203
    "even x ==> (- 1::'a::{comm_ring_1})^x = 1"
21263
wenzelm
parents: 21256
diff changeset
   204
  using minus_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   205
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   206
lemma minus_one_odd_power [simp]:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   207
    "odd x ==> (- 1::'a::{comm_ring_1})^x = - 1"
21263
wenzelm
parents: 21256
diff changeset
   208
  using minus_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   209
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   210
lemma neg_one_even_odd_power:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   211
     "(even x --> (-1::'a::{comm_ring_1})^x = 1) &
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   212
      (odd x --> (-1::'a)^x = -1)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   213
  apply (induct x)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35043
diff changeset
   214
  apply (simp, simp)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   215
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   216
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   217
lemma neg_one_even_power [simp]:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   218
    "even x ==> (-1::'a::{comm_ring_1})^x = 1"
21263
wenzelm
parents: 21256
diff changeset
   219
  using neg_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   220
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   221
lemma neg_one_odd_power [simp]:
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   222
    "odd x ==> (-1::'a::{comm_ring_1})^x = -1"
21263
wenzelm
parents: 21256
diff changeset
   223
  using neg_one_even_odd_power by blast
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   224
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   225
lemma neg_power_if:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   226
     "(-x::'a::{comm_ring_1}) ^ n =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   227
      (if even n then (x ^ n) else -(x ^ n))"
21263
wenzelm
parents: 21256
diff changeset
   228
  apply (induct n)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35043
diff changeset
   229
  apply simp_all
21263
wenzelm
parents: 21256
diff changeset
   230
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   231
21263
wenzelm
parents: 21256
diff changeset
   232
lemma zero_le_even_power: "even n ==>
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35216
diff changeset
   233
    0 <= (x::'a::{linordered_ring,monoid_mult}) ^ n"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   234
  apply (simp add: even_nat_equiv_def2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   235
  apply (erule exE)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   236
  apply (erule ssubst)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   237
  apply (subst power_add)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   238
  apply (rule zero_le_square)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   239
  done
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   240
21263
wenzelm
parents: 21256
diff changeset
   241
lemma zero_le_odd_power: "odd n ==>
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   242
    (0 <= (x::'a::{linordered_idom}) ^ n) = (0 <= x)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35043
diff changeset
   243
apply (auto simp: odd_nat_equiv_def2 power_add zero_le_mult_iff)
36722
c8ea75ea4a29 tuned proof
haftmann
parents: 35644
diff changeset
   244
apply (metis field_power_not_zero divisors_zero order_antisym_conv zero_le_square)
30056
0a35bee25c20 added lemmas
nipkow
parents: 29803
diff changeset
   245
done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   246
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   247
lemma zero_le_power_eq[presburger]: "(0 <= (x::'a::{linordered_idom}) ^ n) =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   248
    (even n | (odd n & 0 <= x))"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   249
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   250
  apply (subst zero_le_odd_power [symmetric])
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   251
  apply assumption+
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   252
  apply (erule zero_le_even_power)
21263
wenzelm
parents: 21256
diff changeset
   253
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   254
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   255
lemma zero_less_power_eq[presburger]: "(0 < (x::'a::{linordered_idom}) ^ n) =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   256
    (n = 0 | (even n & x ~= 0) | (odd n & 0 < x))"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   257
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   258
  unfolding order_less_le zero_le_power_eq by auto
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   259
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   260
lemma power_less_zero_eq[presburger]: "((x::'a::{linordered_idom}) ^ n < 0) =
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   261
    (odd n & x < 0)"
21263
wenzelm
parents: 21256
diff changeset
   262
  apply (subst linorder_not_le [symmetric])+
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   263
  apply (subst zero_le_power_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   264
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   265
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   266
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   267
lemma power_le_zero_eq[presburger]: "((x::'a::{linordered_idom}) ^ n <= 0) =
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   268
    (n ~= 0 & ((odd n & x <= 0) | (even n & x = 0)))"
21263
wenzelm
parents: 21256
diff changeset
   269
  apply (subst linorder_not_less [symmetric])+
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   270
  apply (subst zero_less_power_eq)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   271
  apply auto
21263
wenzelm
parents: 21256
diff changeset
   272
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   273
21263
wenzelm
parents: 21256
diff changeset
   274
lemma power_even_abs: "even n ==>
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   275
    (abs (x::'a::{linordered_idom}))^n = x^n"
21263
wenzelm
parents: 21256
diff changeset
   276
  apply (subst power_abs [symmetric])
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   277
  apply (simp add: zero_le_even_power)
21263
wenzelm
parents: 21256
diff changeset
   278
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   279
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   280
lemma zero_less_power_nat_eq[presburger]: "(0 < (x::nat) ^ n) = (n = 0 | 0 < x)"
21263
wenzelm
parents: 21256
diff changeset
   281
  by (induct n) auto
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   282
21263
wenzelm
parents: 21256
diff changeset
   283
lemma power_minus_even [simp]: "even n ==>
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   284
    (- x)^n = (x^n::'a::{comm_ring_1})"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   285
  apply (subst power_minus)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   286
  apply simp
21263
wenzelm
parents: 21256
diff changeset
   287
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   288
21263
wenzelm
parents: 21256
diff changeset
   289
lemma power_minus_odd [simp]: "odd n ==>
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30738
diff changeset
   290
    (- x)^n = - (x^n::'a::{comm_ring_1})"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   291
  apply (subst power_minus)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   292
  apply simp
21263
wenzelm
parents: 21256
diff changeset
   293
  done
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   294
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   295
lemma power_mono_even: fixes x y :: "'a :: {linordered_idom}"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   296
  assumes "even n" and "\<bar>x\<bar> \<le> \<bar>y\<bar>"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   297
  shows "x^n \<le> y^n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   298
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   299
  have "0 \<le> \<bar>x\<bar>" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   300
  with `\<bar>x\<bar> \<le> \<bar>y\<bar>`
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   301
  have "\<bar>x\<bar>^n \<le> \<bar>y\<bar>^n" by (rule power_mono)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   302
  thus ?thesis unfolding power_even_abs[OF `even n`] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   303
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   304
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   305
lemma odd_pos: "odd (n::nat) \<Longrightarrow> 0 < n" by presburger
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   306
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   307
lemma power_mono_odd: fixes x y :: "'a :: {linordered_idom}"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   308
  assumes "odd n" and "x \<le> y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   309
  shows "x^n \<le> y^n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   310
proof (cases "y < 0")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   311
  case True with `x \<le> y` have "-y \<le> -x" and "0 \<le> -y" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   312
  hence "(-y)^n \<le> (-x)^n" by (rule power_mono)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   313
  thus ?thesis unfolding power_minus_odd[OF `odd n`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   314
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   315
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   316
  show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   317
  proof (cases "x < 0")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   318
    case True hence "n \<noteq> 0" and "x \<le> 0" using `odd n`[THEN odd_pos] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   319
    hence "x^n \<le> 0" unfolding power_le_zero_eq using `odd n` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   320
    moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   321
    from `\<not> y < 0` have "0 \<le> y" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   322
    hence "0 \<le> y^n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   323
    ultimately show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   324
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   325
    case False hence "0 \<le> x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   326
    with `x \<le> y` show ?thesis using power_mono by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   327
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29654
diff changeset
   328
qed
21263
wenzelm
parents: 21256
diff changeset
   329
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   330
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   331
subsection {* More Even/Odd Results *}
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   332
 
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   333
lemma even_mult_two_ex: "even(n) = (\<exists>m::nat. n = 2*m)" by presburger
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   334
lemma odd_Suc_mult_two_ex: "odd(n) = (\<exists>m. n = Suc (2*m))" by presburger
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   335
lemma even_add [simp]: "even(m + n::nat) = (even m = even n)"  by presburger
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   336
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   337
lemma odd_add [simp]: "odd(m + n::nat) = (odd m \<noteq> odd n)" by presburger
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   338
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   339
lemma div_Suc: "Suc a div c = a div c + Suc 0 div c +
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   340
    (a mod c + Suc 0 mod c) div c" 
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   341
  apply (subgoal_tac "Suc a = a + Suc 0")
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   342
  apply (erule ssubst)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   343
  apply (rule div_add1_eq, simp)
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   344
  done
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   345
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   346
lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2" by presburger
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   347
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   348
lemma lemma_not_even_div2 [simp]: "~even n ==> (n + 1) div 2 = Suc (n div 2)"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   349
by presburger
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   350
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   351
lemma even_num_iff: "0 < n ==> even n = (~ even(n - 1 :: nat))"  by presburger
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   352
lemma even_even_mod_4_iff: "even (n::nat) = even (n mod 4)" by presburger
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   353
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   354
lemma lemma_odd_mod_4_div_2: "n mod 4 = (3::nat) ==> odd((n - 1) div 2)" by presburger
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   355
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   356
lemma lemma_even_mod_4_div_2: "n mod 4 = (1::nat) ==> even ((n - 1) div 2)"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 27651
diff changeset
   357
  by presburger
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25594
diff changeset
   358
21263
wenzelm
parents: 21256
diff changeset
   359
text {* Simplify, when the exponent is a numeral *}
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   360
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   361
lemmas zero_le_power_eq_numeral [simp] =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   362
    zero_le_power_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   363
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   364
lemmas zero_less_power_eq_numeral [simp] =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   365
    zero_less_power_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   366
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   367
lemmas power_le_zero_eq_numeral [simp] =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   368
    power_le_zero_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   369
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   370
lemmas power_less_zero_eq_numeral [simp] =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   371
    power_less_zero_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   372
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   373
lemmas zero_less_power_nat_eq_numeral [simp] =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   374
    zero_less_power_nat_eq [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   375
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   376
lemmas power_eq_0_iff_numeral [simp] = power_eq_0_iff [of _ "numeral w"] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   377
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45607
diff changeset
   378
lemmas power_even_abs_numeral [simp] = power_even_abs [of "numeral w" _] for w
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   379
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   380
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   381
subsection {* An Equivalence for @{term [source] "0 \<le> a^n"} *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   382
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   383
lemma even_power_le_0_imp_0:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   384
    "a ^ (2*k) \<le> (0::'a::{linordered_idom}) ==> a=0"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35043
diff changeset
   385
  by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   386
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   387
lemma zero_le_power_iff[presburger]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33358
diff changeset
   388
  "(0 \<le> a^n) = (0 \<le> (a::'a::{linordered_idom}) | even n)"
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   389
proof cases
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   390
  assume even: "even n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   391
  then obtain k where "n = 2*k"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   392
    by (auto simp add: even_nat_equiv_def2 numeral_2_eq_2)
21263
wenzelm
parents: 21256
diff changeset
   393
  thus ?thesis by (simp add: zero_le_even_power even)
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   394
next
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   395
  assume odd: "odd n"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   396
  then obtain k where "n = Suc(2*k)"
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   397
    by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   398
  thus ?thesis
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35043
diff changeset
   399
    by (auto simp add: zero_le_mult_iff zero_le_even_power
21263
wenzelm
parents: 21256
diff changeset
   400
             dest!: even_power_le_0_imp_0)
wenzelm
parents: 21256
diff changeset
   401
qed
wenzelm
parents: 21256
diff changeset
   402
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   403
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   404
subsection {* Miscellaneous *}
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   405
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   406
lemma [presburger]:"(x + 1) div 2 = x div 2 \<longleftrightarrow> even (x::int)" by presburger
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   407
lemma [presburger]: "(x + 1) div 2 = x div 2 + 1 \<longleftrightarrow> odd (x::int)" by presburger
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   408
lemma even_plus_one_div_two: "even (x::int) ==> (x + 1) div 2 = x div 2"  by presburger
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   409
lemma odd_plus_one_div_two: "odd (x::int) ==> (x + 1) div 2 = x div 2 + 1" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   410
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   411
lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \<longleftrightarrow> even x" by presburger
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   412
lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \<longleftrightarrow> even x" by presburger
21263
wenzelm
parents: 21256
diff changeset
   413
lemma even_nat_plus_one_div_two: "even (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   414
    (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   415
21263
wenzelm
parents: 21256
diff changeset
   416
lemma odd_nat_plus_one_div_two: "odd (x::nat) ==>
23522
7e8255828502 Tuned proofs
chaieb
parents: 23438
diff changeset
   417
    (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))" by presburger
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   418
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
   419
end